Skip to content

Commit

Permalink
Merge pull request #3253 from dunhamsteve/issue3251
Browse files Browse the repository at this point in the history
[ parser ] Fix issue parsing unquote
  • Loading branch information
andrevidela committed Apr 27, 2024
2 parents 3489cc3 + d7867c0 commit 84ce3a6
Show file tree
Hide file tree
Showing 5 changed files with 27 additions and 1 deletion.
3 changes: 3 additions & 0 deletions CHANGELOG_NEXT.md
Original file line number Diff line number Diff line change
Expand Up @@ -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
Expand Down
2 changes: 1 addition & 1 deletion src/Idris/Parser.idr
Original file line number Diff line number Diff line change
Expand Up @@ -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
Expand Down
19 changes: 19 additions & 0 deletions tests/idris2/error/perror031/Issue3251.idr
Original file line number Diff line number Diff line change
@@ -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 change: 1 addition & 0 deletions tests/idris2/error/perror031/expected
Original file line number Diff line number Diff line change
@@ -0,0 +1 @@
1/1: Building Issue3251 (Issue3251.idr)
3 changes: 3 additions & 0 deletions tests/idris2/error/perror031/run
Original file line number Diff line number Diff line change
@@ -0,0 +1,3 @@
. ../../../testutils.sh

check Issue3251.idr

0 comments on commit 84ce3a6

Please sign in to comment.