2020-05-18 18:57:43 +03:00
|
|
|
module Language.Reflection
|
|
|
|
|
|
|
|
import public Language.Reflection.TT
|
|
|
|
import public Language.Reflection.TTImp
|
|
|
|
|
2020-06-02 00:16:27 +03:00
|
|
|
-- Elaboration scripts. Where types/terms are returned, binders will have
|
|
|
|
-- unique, if not necessarily human readabe, names
|
2020-05-31 16:33:34 +03:00
|
|
|
export
|
2020-05-18 18:57:43 +03:00
|
|
|
data Elab : Type -> Type where
|
|
|
|
Pure : a -> Elab a
|
|
|
|
Bind : Elab a -> (a -> Elab b) -> Elab b
|
2020-06-01 16:26:37 +03:00
|
|
|
Fail : String -> Elab a
|
|
|
|
|
2020-05-31 16:33:34 +03:00
|
|
|
LogMsg : Nat -> String -> Elab ()
|
|
|
|
LogTerm : Nat -> String -> TTImp -> Elab ()
|
2020-05-18 18:57:43 +03:00
|
|
|
|
2020-06-03 00:41:37 +03:00
|
|
|
-- Elaborate a TTImp term to a concrete value
|
|
|
|
Check : {expected : Type} -> TTImp -> Elab expected
|
2020-06-03 01:36:20 +03:00
|
|
|
-- Quote a concrete expression back to a TTImp
|
|
|
|
Quote : val -> Elab TTImp
|
2020-06-03 11:17:37 +03:00
|
|
|
|
|
|
|
-- Elaborate under a lambda
|
|
|
|
Lambda : (0 x : Type) -> (x -> Elab ty) -> Elab (x -> ty)
|
|
|
|
-- Elaborate under a forall
|
|
|
|
ForAll : (0 x : Type) -> (x -> Elab ty) -> Elab (x -> ty)
|
|
|
|
|
2020-05-31 16:33:34 +03:00
|
|
|
-- Get the current goal type, if known
|
|
|
|
-- (it might need to be inferred from the solution)
|
|
|
|
Goal : Elab (Maybe TTImp)
|
2020-06-02 00:16:27 +03:00
|
|
|
-- Get the names of local variables in scope
|
|
|
|
LocalVars : Elab (List Name)
|
2020-06-01 17:13:42 +03:00
|
|
|
-- Generate a new unique name, based on the given string
|
|
|
|
GenSym : String -> Elab Name
|
|
|
|
-- Put a name in the current namespace
|
|
|
|
InCurrentNS : Name -> Elab Name
|
2020-06-01 16:26:37 +03:00
|
|
|
-- Get the types of every name which matches.
|
|
|
|
-- There may be ambiguities - returns a list of fully explicit names
|
|
|
|
-- and their types. If there's no results, the name is undefined.
|
|
|
|
GetType : Name -> Elab (List (Name, TTImp))
|
2020-06-02 00:16:27 +03:00
|
|
|
-- Get the type of a local variable
|
|
|
|
GetLocalType : Name -> Elab TTImp
|
2020-06-01 17:01:52 +03:00
|
|
|
-- Get the constructors of a data type. The name must be fully resolved.
|
2020-06-01 16:26:37 +03:00
|
|
|
GetCons : Name -> Elab (List Name)
|
2020-06-01 17:01:52 +03:00
|
|
|
-- Check a group of top level declarations
|
|
|
|
Declare : List Decl -> Elab ()
|
|
|
|
|
2020-05-18 18:57:43 +03:00
|
|
|
mutual
|
2020-05-31 16:33:34 +03:00
|
|
|
export
|
2020-05-18 18:57:43 +03:00
|
|
|
Functor Elab where
|
|
|
|
map f e = do e' <- e
|
|
|
|
pure (f e')
|
|
|
|
|
2020-05-31 16:33:34 +03:00
|
|
|
export
|
2020-05-18 18:57:43 +03:00
|
|
|
Applicative Elab where
|
|
|
|
pure = Pure
|
|
|
|
f <*> a = do f' <- f
|
|
|
|
a' <- a
|
|
|
|
pure (f' a')
|
|
|
|
|
2020-05-31 16:33:34 +03:00
|
|
|
export
|
2020-05-18 18:57:43 +03:00
|
|
|
Monad Elab where
|
|
|
|
(>>=) = Bind
|
2020-05-31 16:33:34 +03:00
|
|
|
|
2020-06-01 16:26:37 +03:00
|
|
|
export
|
|
|
|
fail : String -> Elab a
|
|
|
|
fail = Fail
|
|
|
|
|
2020-05-31 16:33:34 +03:00
|
|
|
export
|
|
|
|
logMsg : Nat -> String -> Elab ()
|
|
|
|
logMsg = LogMsg
|
|
|
|
|
|
|
|
export
|
|
|
|
logTerm : Nat -> String -> TTImp -> Elab ()
|
|
|
|
logTerm = LogTerm
|
|
|
|
|
|
|
|
export
|
|
|
|
logGoal : Nat -> String -> Elab ()
|
|
|
|
logGoal n msg
|
|
|
|
= do g <- Goal
|
|
|
|
case g of
|
|
|
|
Nothing => pure ()
|
|
|
|
Just t => logTerm n msg t
|
|
|
|
|
|
|
|
export
|
2020-06-03 00:41:37 +03:00
|
|
|
check : {expected : Type} -> TTImp -> Elab expected
|
2020-05-31 16:33:34 +03:00
|
|
|
check = Check
|
|
|
|
|
2020-06-03 01:36:20 +03:00
|
|
|
export
|
|
|
|
quote : val -> Elab TTImp
|
|
|
|
quote = Quote
|
|
|
|
|
2020-06-03 11:17:37 +03:00
|
|
|
export
|
|
|
|
lambda : (0 x : Type) -> (x -> Elab ty) -> Elab (x -> ty)
|
|
|
|
lambda = Lambda
|
|
|
|
|
|
|
|
export
|
|
|
|
forAll : (0 x : Type) -> (x -> Elab ty) -> Elab (x -> ty)
|
|
|
|
forAll = ForAll
|
|
|
|
|
2020-05-31 16:33:34 +03:00
|
|
|
export
|
|
|
|
goal : Elab (Maybe TTImp)
|
|
|
|
goal = Goal
|
2020-06-01 16:26:37 +03:00
|
|
|
|
2020-06-02 00:16:27 +03:00
|
|
|
export
|
|
|
|
localVars : Elab (List Name)
|
|
|
|
localVars = LocalVars
|
|
|
|
|
2020-06-01 17:13:42 +03:00
|
|
|
export
|
|
|
|
genSym : String -> Elab Name
|
|
|
|
genSym = GenSym
|
|
|
|
|
|
|
|
export
|
|
|
|
inCurrentNS : Name -> Elab Name
|
|
|
|
inCurrentNS = InCurrentNS
|
|
|
|
|
2020-06-01 16:26:37 +03:00
|
|
|
export
|
|
|
|
getType : Name -> Elab (List (Name, TTImp))
|
|
|
|
getType = GetType
|
|
|
|
|
2020-06-02 00:16:27 +03:00
|
|
|
export
|
|
|
|
getLocalType : Name -> Elab TTImp
|
|
|
|
getLocalType = GetLocalType
|
|
|
|
|
2020-06-01 16:26:37 +03:00
|
|
|
export
|
|
|
|
getCons : Name -> Elab (List Name)
|
|
|
|
getCons = GetCons
|
2020-06-01 17:01:52 +03:00
|
|
|
|
|
|
|
export
|
|
|
|
declare : List Decl -> Elab ()
|
|
|
|
declare = Declare
|