Idris2/libs/base/Language/Reflection.idr
Edwin Brady 85c002c771 Progress on elaborator reflection
Add %runElab and start on scripts, although all they can do so far is
check a term. This does gives us, sort of, "template Idris" (as
demonstrated in test reflection002)
2020-05-31 01:36:54 +01:00

31 lines
614 B
Idris

module Language.Reflection
import public Language.Reflection.TT
import public Language.Reflection.TTImp
public export
data Elab : Type -> Type where
Pure : a -> Elab a
Bind : Elab a -> (a -> Elab b) -> Elab b
Log : Nat -> String -> Elab ()
-- Check a TTImp term against the current goal type
Check : TTImp -> Elab TT
mutual
export
Functor Elab where
map f e = do e' <- e
pure (f e')
export
Applicative Elab where
pure = Pure
f <*> a = do f' <- f
a' <- a
pure (f' a')
export
Monad Elab where
(>>=) = Bind