Idris2/tests/idris2/reflection006/refleq.idr
Edwin Brady 2a75731916 In reflection, check now takes a concrete type
So the type of Elab now gives the expected type that's being elaborated
to, meaning that we can run 'check' in the middle of scripts and use the
result.
2020-06-02 22:41:37 +01:00

25 lines
587 B
Idris

import Language.Reflection
%language ElabReflection
solveReflected : TTImp -> Elab any
solveReflected `(Builtin.Equal {a=_} {b=_} ~(left) ~(right))
= do logTerm 0 "Left" left
logTerm 0 "Right" right
fail "Not done"
solveReflected g
= do logTerm 0 "Goal" g
fail "I don't know how to prove this"
%macro
prove : Elab any
prove
= do env <- localVars
Just g <- goal
| Nothing => fail "No goal to solve"
logMsg 0 (show env)
solveReflected g
commutes : (x, y : Nat) -> plus x y = plus y x
commutes x y = prove