Idris2/tests/idris2/reflection002/power.idr
Edwin Brady 106235c165 Allow scripts to inspect goal
If available (sometimes, say a top level expression, it might need
inferring so there'll be no goal available). Also add the ability to log
the current goal, or indeed any term.
2020-05-31 14:33:34 +01:00

11 lines
202 B
Idris

import Language.Reflection
%language ElabReflection
powerFn : Nat -> TTImp
powerFn Z = `(const 1)
powerFn (S k) = `(\x => mult x (~(powerFn k) x))
cube : Nat -> Nat
cube = %runElab check (powerFn 3)