Skip to content

Commit

Permalink
[ elab ] Print script's FC in the bad elaboration script error
Browse files Browse the repository at this point in the history
  • Loading branch information
buzden authored and gallais committed Sep 22, 2023
1 parent f6c000e commit a643e3a
Show file tree
Hide file tree
Showing 4 changed files with 42 additions and 0 deletions.
4 changes: 4 additions & 0 deletions src/Idris/Error.idr
Original file line number Diff line number Diff line change
Expand Up @@ -603,6 +603,10 @@ perrorRaw (BadRunElab fc env script desc)
= pure $ errorDesc (reflow "Bad elaborator script" <++> code !(pshow env script)
<++> parens (pretty0 desc) <+> dot)
<+> line <+> !(ploc fc)
<+> !(let scriptFC = getLoc script in
if isJust (isNonEmptyFC scriptFC)
then pure $ line <+> reflow "Stuck place in the script:" <+> line <+> !(ploc scriptFC)
else pure emptyDoc)
perrorRaw (RunElabFail e)
= pure $ reflow "Error during reflection" <+> colon <++> !(perrorRaw e)
perrorRaw (GenericMsg fc str) = pure $ pretty0 str <+> line <+> !(ploc fc)
Expand Down
14 changes: 14 additions & 0 deletions tests/idris2/reflection/reflection022/BadElabScript.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,14 @@
module BadElabScript

import Language.Reflection

%language ElabReflection

x : Elab Nat
x = do
ignore $ check {expected=Type} `(Nat)
?someHole
check `(%search)

y : Nat
y = %runElab x
21 changes: 21 additions & 0 deletions tests/idris2/reflection/reflection022/expected
Original file line number Diff line number Diff line change
@@ -0,0 +1,21 @@
1/1: Building BadElabScript (BadElabScript.idr)
Error: While processing right hand side of y. Bad elaborator script ?someHole [no locals in scope] (script is not a data value).

BadElabScript:14:5--14:15
10 | ?someHole
11 | check `(%search)
12 |
13 | y : Nat
14 | y = %runElab x
^^^^^^^^^^

Stuck place in the script:

BadElabScript:10:3--10:12
06 |
07 | x : Elab Nat
08 | x = do
09 | ignore $ check {expected=Type} `(Nat)
10 | ?someHole
^^^^^^^^^

3 changes: 3 additions & 0 deletions tests/idris2/reflection/reflection022/run
Original file line number Diff line number Diff line change
@@ -0,0 +1,3 @@
. ../../../testutils.sh

check BadElabScript.idr

0 comments on commit a643e3a

Please sign in to comment.