2020-05-18 18:57:43 +03:00
|
|
|
module Language.Reflection
|
|
|
|
|
|
|
|
import public Language.Reflection.TT
|
|
|
|
import public Language.Reflection.TTImp
|
|
|
|
|
2021-10-25 18:16:13 +03:00
|
|
|
import public Control.Monad.Trans
|
|
|
|
|
2021-05-13 14:58:56 +03:00
|
|
|
%default total
|
|
|
|
|
2021-10-25 18:16:13 +03:00
|
|
|
----------------------------------
|
|
|
|
--- Elaboration data structure ---
|
|
|
|
----------------------------------
|
|
|
|
|
2020-07-17 17:18:23 +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
|
2021-08-11 14:26:17 +03:00
|
|
|
Fail : FC -> String -> Elab a
|
2020-06-01 16:26:37 +03:00
|
|
|
|
2021-10-16 13:24:12 +03:00
|
|
|
Try : Elab a -> Elab a -> Elab a
|
|
|
|
|
2021-10-14 14:35:32 +03:00
|
|
|
||| Log a message. Takes a
|
|
|
|
||| * topic
|
|
|
|
||| * level
|
|
|
|
||| * message
|
2020-08-20 20:45:34 +03:00
|
|
|
LogMsg : String -> Nat -> String -> Elab ()
|
2021-10-14 14:35:32 +03:00
|
|
|
||| Print and log a term. Takes a
|
|
|
|
||| * topic
|
|
|
|
||| * level
|
|
|
|
||| * message
|
|
|
|
||| * term
|
2020-08-20 20:45:34 +03:00
|
|
|
LogTerm : String -> Nat -> String -> TTImp -> Elab ()
|
2021-10-14 14:35:32 +03:00
|
|
|
||| Resugar, print and log a term. Takes a
|
|
|
|
||| * topic
|
|
|
|
||| * level
|
|
|
|
||| * message
|
|
|
|
||| * term
|
|
|
|
LogSugaredTerm : String -> 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
|
2021-09-15 17:01:36 +03:00
|
|
|
Check : TTImp -> Elab expected
|
2020-06-03 01:36:20 +03:00
|
|
|
-- Quote a concrete expression back to a TTImp
|
2021-09-15 17:01:36 +03:00
|
|
|
Quote : (0 _ : val) -> Elab TTImp
|
2020-06-03 11:17:37 +03:00
|
|
|
|
|
|
|
-- Elaborate under a lambda
|
2020-06-03 22:23:36 +03:00
|
|
|
Lambda : (0 x : Type) ->
|
|
|
|
{0 ty : x -> Type} ->
|
|
|
|
((val : x) -> Elab (ty val)) -> Elab ((val : x) -> (ty val))
|
2020-06-03 11:17:37 +03:00
|
|
|
|
2020-08-20 20:45:34 +03:00
|
|
|
-- Get the current goal type, if known
|
2020-05-31 16:33:34 +03:00
|
|
|
-- (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))
|
2021-11-07 18:06:53 +03:00
|
|
|
-- Get the metadata associated with a name
|
|
|
|
GetInfo : Name -> Elab (List (Name, NameInfo))
|
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 ()
|
|
|
|
|
2021-05-13 14:58:56 +03:00
|
|
|
export
|
|
|
|
Functor Elab where
|
|
|
|
map f e = Bind e $ Pure . f
|
|
|
|
|
|
|
|
export
|
|
|
|
Applicative Elab where
|
|
|
|
pure = Pure
|
|
|
|
f <*> a = Bind f (<$> a)
|
|
|
|
|
2021-10-17 00:23:35 +03:00
|
|
|
export
|
|
|
|
Alternative Elab where
|
|
|
|
empty = Fail EmptyFC ""
|
|
|
|
l <|> r = Try l r
|
|
|
|
|
2021-05-13 14:58:56 +03:00
|
|
|
export
|
|
|
|
Monad Elab where
|
|
|
|
(>>=) = Bind
|
2020-05-31 16:33:34 +03:00
|
|
|
|
2021-10-25 18:16:13 +03:00
|
|
|
-----------------------------
|
|
|
|
--- Elaboration interface ---
|
|
|
|
-----------------------------
|
2021-08-11 14:26:17 +03:00
|
|
|
|
2021-10-25 18:16:13 +03:00
|
|
|
public export
|
|
|
|
interface Monad m => Elaboration m where
|
2020-06-01 16:26:37 +03:00
|
|
|
|
2021-10-25 18:16:13 +03:00
|
|
|
||| Report an error in elaboration at some location
|
|
|
|
failAt : FC -> String -> m a
|
2021-10-16 13:24:12 +03:00
|
|
|
|
2021-10-25 18:16:13 +03:00
|
|
|
||| Try the first elaborator. If it fails, reset the elaborator state and
|
|
|
|
||| run the second
|
|
|
|
try : Elab a -> Elab a -> m a
|
2020-05-31 16:33:34 +03:00
|
|
|
|
2021-10-25 18:16:13 +03:00
|
|
|
||| Write a log message, if the log level is >= the given level
|
|
|
|
logMsg : String -> Nat -> String -> m ()
|
2020-05-31 16:33:34 +03:00
|
|
|
|
2021-10-25 18:16:13 +03:00
|
|
|
||| Write a log message and a rendered term, if the log level is >= the given level
|
|
|
|
logTerm : String -> Nat -> String -> TTImp -> m ()
|
2021-10-14 14:35:32 +03:00
|
|
|
|
2021-10-25 18:16:13 +03:00
|
|
|
||| Write a log message and a resugared & rendered term, if the log level is >= the given level
|
|
|
|
logSugaredTerm : String -> Nat -> String -> TTImp -> m ()
|
2020-05-31 16:33:34 +03:00
|
|
|
|
2021-10-25 18:16:13 +03:00
|
|
|
||| Check that some TTImp syntax has the expected type
|
|
|
|
||| Returns the type checked value
|
|
|
|
check : TTImp -> m expected
|
2020-06-03 01:36:20 +03:00
|
|
|
|
2021-10-25 18:16:13 +03:00
|
|
|
||| Return TTImp syntax of a given value
|
|
|
|
quote : (0 _ : val) -> m TTImp
|
2020-06-03 11:17:37 +03:00
|
|
|
|
2021-10-25 18:16:13 +03:00
|
|
|
||| Build a lambda expression
|
|
|
|
lambda : (0 x : Type) ->
|
|
|
|
{0 ty : x -> Type} ->
|
|
|
|
((val : x) -> Elab (ty val)) -> m ((val : x) -> (ty val))
|
2020-06-01 16:26:37 +03:00
|
|
|
|
2021-10-25 18:16:13 +03:00
|
|
|
||| Get the goal type of the current elaboration
|
|
|
|
goal : m (Maybe TTImp)
|
2020-06-02 00:16:27 +03:00
|
|
|
|
2021-10-25 18:16:13 +03:00
|
|
|
||| Get the names of the local variables in scope
|
|
|
|
localVars : m (List Name)
|
2020-06-01 17:13:42 +03:00
|
|
|
|
2021-10-25 18:16:13 +03:00
|
|
|
||| Generate a new unique name
|
|
|
|
genSym : String -> m Name
|
2020-06-01 17:13:42 +03:00
|
|
|
|
2021-10-25 18:16:13 +03:00
|
|
|
||| Given a name, return the name decorated with the current namespace
|
|
|
|
inCurrentNS : Name -> m Name
|
2020-06-01 16:26:37 +03:00
|
|
|
|
2021-10-25 18:16:13 +03:00
|
|
|
||| Given a possibly ambiguous name, get all the matching names and their types
|
|
|
|
getType : Name -> m (List (Name, TTImp))
|
2020-06-02 00:16:27 +03:00
|
|
|
|
2021-11-07 18:06:53 +03:00
|
|
|
||| Get the metadata associated with a name. Returns all matching namea and their types
|
|
|
|
getInfo : Name -> m (List (Name, NameInfo))
|
|
|
|
|
2021-10-25 18:16:13 +03:00
|
|
|
||| Get the type of a local variable
|
|
|
|
getLocalType : Name -> m TTImp
|
2020-06-01 17:01:52 +03:00
|
|
|
|
2021-10-25 18:16:13 +03:00
|
|
|
||| Get the constructors of a fully qualified data type name
|
|
|
|
getCons : Name -> m (List Name)
|
|
|
|
|
|
|
|
||| Make some top level declarations
|
|
|
|
declare : List Decl -> m ()
|
|
|
|
|
|
|
|
export %inline
|
|
|
|
||| Report an error in elaboration
|
|
|
|
fail : Elaboration m => String -> m a
|
|
|
|
fail = failAt EmptyFC
|
|
|
|
|
|
|
|
||| Log the current goal type, if the log level is >= the given level
|
|
|
|
export %inline
|
|
|
|
logGoal : Elaboration m => String -> Nat -> String -> m ()
|
|
|
|
logGoal str n msg = whenJust !goal $ logTerm str n msg
|
|
|
|
|
|
|
|
export
|
|
|
|
Elaboration Elab where
|
|
|
|
failAt = Fail
|
|
|
|
try = Try
|
|
|
|
logMsg = LogMsg
|
|
|
|
logTerm = LogTerm
|
|
|
|
logSugaredTerm = LogSugaredTerm
|
|
|
|
check = Check
|
|
|
|
quote = Quote
|
|
|
|
lambda = Lambda
|
|
|
|
goal = Goal
|
|
|
|
localVars = LocalVars
|
|
|
|
genSym = GenSym
|
|
|
|
inCurrentNS = InCurrentNS
|
|
|
|
getType = GetType
|
2021-11-07 18:06:53 +03:00
|
|
|
getInfo = GetInfo
|
2021-10-25 18:16:13 +03:00
|
|
|
getLocalType = GetLocalType
|
|
|
|
getCons = GetCons
|
|
|
|
declare = Declare
|
|
|
|
|
|
|
|
public export
|
|
|
|
Elaboration m => MonadTrans t => Monad (t m) => Elaboration (t m) where
|
|
|
|
failAt = lift .: failAt
|
|
|
|
try = lift .: try
|
|
|
|
logMsg s = lift .: logMsg s
|
|
|
|
logTerm s n = lift .: logTerm s n
|
|
|
|
logSugaredTerm s n = lift .: logSugaredTerm s n
|
|
|
|
check = lift . check
|
|
|
|
quote v = lift $ quote v
|
|
|
|
lambda x = lift . lambda x
|
|
|
|
goal = lift goal
|
|
|
|
localVars = lift localVars
|
|
|
|
genSym = lift . genSym
|
|
|
|
inCurrentNS = lift . inCurrentNS
|
|
|
|
getType = lift . getType
|
2021-11-07 18:06:53 +03:00
|
|
|
getInfo = lift . getInfo
|
2021-10-25 18:16:13 +03:00
|
|
|
getLocalType = lift . getLocalType
|
|
|
|
getCons = lift . getCons
|
|
|
|
declare = lift . declare
|