Merge pull request #3253 from dunhamsteve/issue3251

[ parser ] Fix issue parsing unquote
This commit is contained in:
André Videla 2024-04-28 08:17:18 +09:00 committed by GitHub
commit 84ce3a6836
No known key found for this signature in database
GPG Key ID: B5690EEEBB952194
5 changed files with 27 additions and 1 deletions

View File

@ -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

View File

@ -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

View 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)

View File

@ -0,0 +1 @@
1/1: Building Issue3251 (Issue3251.idr)

View File

@ -0,0 +1,3 @@
. ../../../testutils.sh
check Issue3251.idr