Idris2/tests/chez/chez036/Crash.idr
Steve Dunham ecf4765c4b
[ fix ] Fix issue with eager evaluation of crashing functions (fixes #3003) (#3004)
* [ fix ] Fix issue with eager evaluation of crashing functions

* Mark functions that call unsafe builtins as non-constant

* Better detection of crash primop when deciding if functions can be constant
2023-06-28 08:32:48 +01:00

15 lines
221 B
Idris

-- no_clauses.idr
%default total
bad : (0 _ : Z = S Z) -> Void
bad _ impossible
data BadPrf = Ok | No (Z = S Z)
Show BadPrf where
show Ok = "Ok"
show (No prf) = absurd $ bad prf
main : IO ()
main = printLn Ok