Skip to content

Commit

Permalink
[ fix ] codegen issue when using partial case statements in prelude. (#…
Browse files Browse the repository at this point in the history
  • Loading branch information
dunhamsteve committed Apr 23, 2023
1 parent a423715 commit 3ad391d
Show file tree
Hide file tree
Showing 5 changed files with 8 additions and 2 deletions.
2 changes: 1 addition & 1 deletion src/TTImp/ProcessDef.idr
Original file line number Diff line number Diff line change
Expand Up @@ -844,7 +844,7 @@ mkRunTime fc n

mkCrash : {vars : _} -> String -> Term vars
mkCrash msg
= apply fc (Ref fc Func (NS builtinNS (UN $ Basic "idris_crash")))
= apply fc (Ref fc Func (UN $ Basic "prim__crash"))
[Erased fc Placeholder, PrimVal fc (Str msg)]

matchAny : Term vars -> Term vars
Expand Down
2 changes: 1 addition & 1 deletion tests/Main.idr
Original file line number Diff line number Diff line change
Expand Up @@ -89,7 +89,7 @@ idrisTestsError = MkTestPool "Error messages" [] Nothing
"error011", "error012", "error013", "error014", "error015",
"error016", "error017", "error018", "error019", "error020",
"error021", "error022", "error023", "error024", "error025",
"error026",
"error026", "error027",
-- Parse errors
"perror001", "perror002", "perror003", "perror004", "perror005",
"perror006", "perror007", "perror008", "perror009", "perror010",
Expand Down
2 changes: 2 additions & 0 deletions tests/idris2/error027/Issue2950.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,2 @@
main : IO ()
main = printLn $ mod 10 0
1 change: 1 addition & 0 deletions tests/idris2/error027/expected
Original file line number Diff line number Diff line change
@@ -0,0 +1 @@
ERROR: Unhandled input for Prelude.Num.case block in mod at Prelude.Num:94:3--96:44
3 changes: 3 additions & 0 deletions tests/idris2/error027/run
Original file line number Diff line number Diff line change
@@ -0,0 +1,3 @@
rm -rf build

$1 --no-color --console-width 0 --exec main Issue2950.idr

0 comments on commit 3ad391d

Please sign in to comment.