mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-19 09:12:34 +03:00
Merge pull request #3253 from dunhamsteve/issue3251
[ parser ] Fix issue parsing unquote
This commit is contained in:
commit
84ce3a6836
@ -46,6 +46,9 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO
|
||||
environment variable adds to the "Package Search Paths." Functionally this is
|
||||
not a breaking change.
|
||||
|
||||
* The compiler now parses `~x.fun` as unquoting `x` rather than `x.fun`
|
||||
and `~(f 5).fun` as unquoting `(f 5)` rather than `(f 5).fun`.
|
||||
|
||||
### Backend changes
|
||||
|
||||
#### RefC Backend
|
||||
|
@ -631,7 +631,7 @@ mutual
|
||||
decoratedSymbol fname "]"
|
||||
pure ts
|
||||
pure (PQuoteDecl (boundToFC fname b) (collectDefs (concat b.val)))
|
||||
<|> do b <- bounds (decoratedSymbol fname "~" *> simpleExpr fname indents)
|
||||
<|> do b <- bounds (decoratedSymbol fname "~" *> simplerExpr fname indents)
|
||||
pure (PUnquote (boundToFC fname b) b.val)
|
||||
<|> do start <- bounds (symbol "(")
|
||||
bracketedExpr fname start indents
|
||||
|
19
tests/idris2/error/perror031/Issue3251.idr
Normal file
19
tests/idris2/error/perror031/Issue3251.idr
Normal file
@ -0,0 +1,19 @@
|
||||
import Language.Reflection
|
||||
|
||||
(.fun) : Nat -> Nat
|
||||
|
||||
x : TTImp
|
||||
|
||||
f : Nat -> TTImp
|
||||
|
||||
useX : TTImp
|
||||
useX = `(g (~x).fun)
|
||||
|
||||
useX' : TTImp
|
||||
useX' = `(g ~x.fun)
|
||||
|
||||
useFX : TTImp
|
||||
useFX = `(g (~(f 5)).fun)
|
||||
|
||||
useFX' : TTImp
|
||||
useFX' = `(g ~(f 5).fun)
|
1
tests/idris2/error/perror031/expected
Normal file
1
tests/idris2/error/perror031/expected
Normal file
@ -0,0 +1 @@
|
||||
1/1: Building Issue3251 (Issue3251.idr)
|
3
tests/idris2/error/perror031/run
Normal file
3
tests/idris2/error/perror031/run
Normal file
@ -0,0 +1,3 @@
|
||||
. ../../../testutils.sh
|
||||
|
||||
check Issue3251.idr
|
Loading…
Reference in New Issue
Block a user