mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-25 20:51:43 +03:00
[ fix ] codegen issue when using partial case statements in prelude. (#2952)
This commit is contained in:
parent
a423715da1
commit
3ad391d597
@ -844,7 +844,7 @@ mkRunTime fc n
|
|||||||
|
|
||||||
mkCrash : {vars : _} -> String -> Term vars
|
mkCrash : {vars : _} -> String -> Term vars
|
||||||
mkCrash msg
|
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)]
|
[Erased fc Placeholder, PrimVal fc (Str msg)]
|
||||||
|
|
||||||
matchAny : Term vars -> Term vars
|
matchAny : Term vars -> Term vars
|
||||||
|
@ -89,7 +89,7 @@ idrisTestsError = MkTestPool "Error messages" [] Nothing
|
|||||||
"error011", "error012", "error013", "error014", "error015",
|
"error011", "error012", "error013", "error014", "error015",
|
||||||
"error016", "error017", "error018", "error019", "error020",
|
"error016", "error017", "error018", "error019", "error020",
|
||||||
"error021", "error022", "error023", "error024", "error025",
|
"error021", "error022", "error023", "error024", "error025",
|
||||||
"error026",
|
"error026", "error027",
|
||||||
-- Parse errors
|
-- Parse errors
|
||||||
"perror001", "perror002", "perror003", "perror004", "perror005",
|
"perror001", "perror002", "perror003", "perror004", "perror005",
|
||||||
"perror006", "perror007", "perror008", "perror009", "perror010",
|
"perror006", "perror007", "perror008", "perror009", "perror010",
|
||||||
|
2
tests/idris2/error027/Issue2950.idr
Normal file
2
tests/idris2/error027/Issue2950.idr
Normal file
@ -0,0 +1,2 @@
|
|||||||
|
main : IO ()
|
||||||
|
main = printLn $ mod 10 0
|
1
tests/idris2/error027/expected
Normal file
1
tests/idris2/error027/expected
Normal file
@ -0,0 +1 @@
|
|||||||
|
ERROR: Unhandled input for Prelude.Num.case block in mod at Prelude.Num:94:3--96:44
|
3
tests/idris2/error027/run
Executable file
3
tests/idris2/error027/run
Executable file
@ -0,0 +1,3 @@
|
|||||||
|
rm -rf build
|
||||||
|
|
||||||
|
$1 --no-color --console-width 0 --exec main Issue2950.idr
|
Loading…
Reference in New Issue
Block a user