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