Idris2/libs/base/Language/Reflection.idr
Edwin Brady f303e469fb Improve elaborator reflection performance
In a 'Bind', normalise the result of the first action, rather than
quoting the HNF. This improves performance since the HNF could be quite
big when quoted back.

Ideally, we wouldn't have to quote and unquote here, and we can probably
achieve this by tinkering with the evaluator.

This has an unfortunate effect on the reflection002 test, in that the
"typed template Idris" example now evaluates too much. But, I think the
overall performance is too important for the primary motivation
behind elaborator reflection. I will return to this!
2020-07-17 15:18:23 +01:00

146 lines
3.9 KiB
Idris

module Language.Reflection
import public Language.Reflection.TT
import public Language.Reflection.TTImp
||| Elaboration scripts
||| Where types/terms are returned, binders will have unique, if not
||| necessarily human readabe, names
export
data Elab : Type -> Type where
Pure : a -> Elab a
Bind : Elab a -> (a -> Elab b) -> Elab b
Fail : String -> Elab a
LogMsg : Nat -> String -> Elab ()
LogTerm : Nat -> String -> TTImp -> Elab ()
-- Elaborate a TTImp term to a concrete value
Check : {expected : Type} -> TTImp -> Elab expected
-- Quote a concrete expression back to a TTImp
Quote : val -> Elab TTImp
-- Elaborate under a lambda
Lambda : (0 x : Type) ->
{0 ty : x -> Type} ->
((val : x) -> Elab (ty val)) -> Elab ((val : x) -> (ty val))
-- Get the current goal type, if known
-- (it might need to be inferred from the solution)
Goal : Elab (Maybe TTImp)
-- Get the names of local variables in scope
LocalVars : Elab (List Name)
-- 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
-- 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))
-- Get the type of a local variable
GetLocalType : Name -> Elab TTImp
-- Get the constructors of a data type. The name must be fully resolved.
GetCons : Name -> Elab (List Name)
-- Check a group of top level declarations
Declare : List Decl -> Elab ()
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
||| Report an error in elaboration
export
fail : String -> Elab a
fail = Fail
||| Write a log message, if the log level is >= the given level
export
logMsg : Nat -> String -> Elab ()
logMsg = LogMsg
||| Write a log message and a rendered term, if the log level is >= the given level
export
logTerm : Nat -> String -> TTImp -> Elab ()
logTerm = LogTerm
||| Log the current goal type, if the log level is >= the given level
export
logGoal : Nat -> String -> Elab ()
logGoal n msg
= do g <- Goal
case g of
Nothing => pure ()
Just t => logTerm n msg t
||| Check that some TTImp syntax has the expected type
||| Returns the type checked value
export
check : {expected : Type} -> TTImp -> Elab expected
check = Check
||| Return TTImp syntax of a given value
export
quote : val -> Elab TTImp
quote = Quote
||| Build a lambda expression
export
lambda : (0 x : Type) ->
{0 ty : x -> Type} ->
((val : x) -> Elab (ty val)) -> Elab ((val : x) -> (ty val))
lambda = Lambda
||| Get the goal type of the current elaboration
export
goal : Elab (Maybe TTImp)
goal = Goal
||| Get the names of the local variables in scope
export
localVars : Elab (List Name)
localVars = LocalVars
||| Generate a new unique name
export
genSym : String -> Elab Name
genSym = GenSym
||| Given a name, return the name decorated with the current namespace
export
inCurrentNS : Name -> Elab Name
inCurrentNS = InCurrentNS
||| Given a possibly ambiguous name, get all the matching names and their types
export
getType : Name -> Elab (List (Name, TTImp))
getType = GetType
||| Get the type of a local variable
export
getLocalType : Name -> Elab TTImp
getLocalType = GetLocalType
||| Get the constructors of a fully qualified data type name
export
getCons : Name -> Elab (List Name)
getCons = GetCons
||| Make some top level declarations
export
declare : List Decl -> Elab ()
declare = Declare