mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-19 09:12:34 +03:00
35 lines
657 B
Idris
35 lines
657 B
Idris
import Language.Reflection
|
|
|
|
%language ElabReflection
|
|
%default total
|
|
|
|
magicScript : Elab a
|
|
magicScript = check $ IAlternative EmptyFC FirstSuccess
|
|
[ `(the Nat 5)
|
|
, `(the String "foo")
|
|
]
|
|
|
|
x : Nat
|
|
x = %runElab magicScript
|
|
|
|
y : String
|
|
y = %runElab magicScript
|
|
|
|
-- Check necessary parts of the expected error message
|
|
|
|
failing "Sorry, I can't find any elaboration which works. All errors:"
|
|
z : Bool
|
|
z = %runElab magicScript
|
|
|
|
failing "Mismatch between: Nat and Bool"
|
|
z : Bool
|
|
z = %runElab magicScript
|
|
|
|
failing "Mismatch between: String and Bool"
|
|
z : Bool
|
|
z = %runElab magicScript
|
|
|
|
-- For the whole error message
|
|
z : Bool
|
|
z = %runElab magicScript
|