quox/lib/Quox/Typing/Context.idr

464 lines
14 KiB
Idris
Raw Normal View History

2023-03-13 16:41:57 -04:00
module Quox.Typing.Context
import Quox.Syntax
import Quox.Context
2023-03-15 10:54:51 -04:00
import Quox.Pretty
2024-05-27 15:28:22 -04:00
import public Quox.SingletonExtra
2023-09-17 13:09:54 -04:00
import Derive.Prelude
2023-03-13 16:41:57 -04:00
%default total
2023-09-17 13:09:54 -04:00
%language ElabReflection
2023-03-13 16:41:57 -04:00
2023-03-15 10:54:51 -04:00
public export
2024-05-27 15:28:22 -04:00
QContext : (q, n : Nat) -> Type
QContext q n = Context' (Qty q) n
2023-03-15 10:54:51 -04:00
public export
2024-05-27 15:28:22 -04:00
record LocalVar q d n where
constructor MkLocal
2024-05-27 15:28:22 -04:00
type : Term q d n
term : Maybe (Term q d n) -- if from a `let`
-- %runElab deriveIndexed "LocalVar" [Show]
export
[LocateVarType] Located (LocalVar q d n) where x.loc = x.type.loc
namespace LocalVar
export %inline
2024-05-27 15:28:22 -04:00
letVar : (type, term : Term q d n) -> LocalVar q d n
letVar type term = MkLocal {type, term = Just term}
export %inline
2024-05-27 15:28:22 -04:00
lamVar : (type : Term q d n) -> LocalVar q d n
lamVar type = MkLocal {type, term = Nothing}
2024-04-14 09:48:10 -04:00
export %inline
2024-05-27 15:28:22 -04:00
mapVar : (Term q d n -> Term q' d' n') -> LocalVar q d n -> LocalVar q' d' n'
2024-04-14 09:48:10 -04:00
mapVar f = {type $= f, term $= map f}
export %inline
2024-05-27 15:28:22 -04:00
subQ : {q1, q2 : Nat} -> QSubst q1 q2 -> LocalVar q1 d n -> LocalVar q2 d n
subQ th = mapVar (// th)
export %inline
weakQ : {q : Nat} -> LocalVar q d n -> LocalVar (S q) d n
weakQ = subQ $ shift 1
export %inline
subD : DSubst d1 d2 -> LocalVar q d1 n -> LocalVar q d2 n
2024-04-14 09:48:10 -04:00
subD th = mapVar (// th)
2024-04-14 09:48:10 -04:00
export %inline
2024-05-27 15:28:22 -04:00
weakD : LocalVar q d n -> LocalVar q (S d) n
weakD = subD $ shift 1
2024-05-27 15:28:22 -04:00
export %inline CanShift (LocalVar q d) where l // by = mapVar (// by) l
export %inline CanQSubst LocalVar where l // th = mapVar (// th) l
export %inline CanDSubst LocalVar where l // th = mapVar (// th) l
export %inline CanTSubst LocalVar where l // th = mapVar (// th) l
2024-04-14 09:48:10 -04:00
2023-03-13 16:41:57 -04:00
public export
2023-04-01 13:16:43 -04:00
TContext : TermLike
2024-05-27 15:28:22 -04:00
TContext q d = Context (LocalVar q d)
2023-03-13 16:41:57 -04:00
public export
2024-05-27 15:28:22 -04:00
QOutput : (q, n : Nat) -> Type
QOutput = QContext
2023-03-13 16:41:57 -04:00
public export
DimAssign : Nat -> Type
DimAssign = Context' DimConst
2023-03-13 16:41:57 -04:00
public export
2024-05-27 15:28:22 -04:00
record TyContext q d n where
2023-03-13 16:41:57 -04:00
constructor MkTyContext
2024-05-27 15:28:22 -04:00
{auto qtyLen : Singleton q}
{auto dimLen : Singleton d}
{auto termLen : Singleton n}
2024-05-27 15:28:22 -04:00
qnames : BContext q -- only used for printing
dctx : DimEq d
2024-04-14 09:48:10 -04:00
dnames : BContext d -- only used for printing
2024-05-27 15:28:22 -04:00
tctx : TContext q d n
tnames : BContext n -- only used for printing
2024-05-27 15:28:22 -04:00
qtys : QContext q n -- only used for printing
2023-03-13 16:41:57 -04:00
%name TyContext ctx
2024-05-27 15:28:22 -04:00
-- %runElab deriveIndexed "TyContext" [Show]
public export
2024-05-27 15:28:22 -04:00
record EqContext q n where
constructor MkEqContext
2024-05-27 15:28:22 -04:00
{auto qtyLen : Singleton q}
2023-03-15 10:54:51 -04:00
{dimLen : Nat}
{auto termLen : Singleton n}
2024-05-27 15:28:22 -04:00
qnames : BContext q -- only used for printing
2023-03-15 10:54:51 -04:00
dassign : DimAssign dimLen -- only used for printing
2023-05-01 21:06:25 -04:00
dnames : BContext dimLen -- only used for printing
2024-05-27 15:28:22 -04:00
tctx : TContext q 0 n
tnames : BContext n -- only used for printing
2024-05-27 15:28:22 -04:00
qtys : QContext q n -- only used for printing
%name EqContext ctx
2024-05-27 15:28:22 -04:00
-- %runElab deriveIndexed "EqContext" [Show]
2023-04-15 09:13:01 -04:00
public export
2024-05-27 15:28:22 -04:00
record WhnfContext q d n where
2023-04-15 09:13:01 -04:00
constructor MkWhnfContext
2024-05-27 15:28:22 -04:00
{auto qtyLen : Singleton q}
2023-10-15 10:23:38 -04:00
{auto dimLen : Singleton d}
{auto termLen : Singleton n}
2024-05-27 15:28:22 -04:00
qnames : BContext q -- only used for printing
2023-05-01 21:06:25 -04:00
dnames : BContext d
tnames : BContext n
2024-05-27 15:28:22 -04:00
tctx : TContext q d n
2023-04-15 09:13:01 -04:00
%name WhnfContext ctx
2024-05-27 15:28:22 -04:00
-- %runElab deriveIndexed "WhnfContext" [Show]
2023-04-15 09:13:01 -04:00
namespace TContext
export %inline
2024-05-27 15:28:22 -04:00
zeroFor : {q : Nat} -> Located1 tm => Context tm n -> QOutput q n
zeroFor = map $ \x => zero x.loc
2023-09-30 12:37:44 -04:00
public export
2023-06-23 12:32:05 -04:00
extendLen : Telescope a n1 n2 -> Singleton n1 -> Singleton n2
extendLen [<] x = x
extendLen (tel :< _) x = [|S $ extendLen tel x|]
2023-04-03 11:46:23 -04:00
public export
2024-05-27 15:28:22 -04:00
CtxExtension : (q, d, n1, n2 : Nat) -> Type
CtxExtension q d = Telescope $ \n => (Qty q, BindName, Term q d n)
2023-04-03 11:46:23 -04:00
public export
2024-05-27 15:28:22 -04:00
CtxExtension0 : (q, d, n1, n2 : Nat) -> Type
CtxExtension0 q d = Telescope $ \n => (BindName, Term q d n)
public export
2024-05-27 15:28:22 -04:00
CtxExtensionLet : (q, d, n1, n2 : Nat) -> Type
CtxExtensionLet q d = Telescope $ \n => (Qty q, BindName, LocalVar q d n)
public export
2024-05-27 15:28:22 -04:00
CtxExtensionLet0 : (q, d, n1, n2 : Nat) -> Type
CtxExtensionLet0 q d = Telescope $ \n => (BindName, LocalVar q d n)
namespace TyContext
public export %inline
2024-05-27 15:28:22 -04:00
empty : TyContext 0 0 0
2024-04-14 09:48:10 -04:00
empty = MkTyContext {
2024-05-27 15:28:22 -04:00
qnames = [<], dctx = new, dnames = [<], tctx = [<], tnames = [<], qtys = [<]
2024-04-14 09:48:10 -04:00
}
2023-03-26 10:10:39 -04:00
public export %inline
2024-05-27 15:28:22 -04:00
null : TyContext q d n -> Bool
null ctx = null ctx.qnames && null ctx.dnames && null ctx.tnames
2023-03-26 10:10:39 -04:00
export %inline
2024-05-27 15:28:22 -04:00
extendTyLetN : CtxExtensionLet q d n1 n2 ->
TyContext q d n1 -> TyContext q d n2
extendTyLetN xss
(MkTyContext {termLen, qnames, dctx, dnames, tctx, tnames, qtys}) =
let (qs, xs, ls) = unzip3 xss in
MkTyContext {
2024-05-27 15:28:22 -04:00
qnames, dctx, dnames,
termLen = extendLen xss termLen,
tctx = tctx . ls,
tnames = tnames . xs,
qtys = qtys . qs
}
export %inline
2024-05-27 15:28:22 -04:00
extendTyN : CtxExtension q d n1 n2 -> TyContext q d n1 -> TyContext q d n2
extendTyN = extendTyLetN . map (\(q, x, s) => (q, x, lamVar s))
export %inline
2024-05-27 15:28:22 -04:00
extendTyLetN0 : CtxExtensionLet0 q d n1 n2 ->
TyContext q d n1 -> TyContext q d n2
extendTyLetN0 xss ctx =
let Val q = ctx.qtyLen in
extendTyLetN (map (\(x, t) => (zero x.loc, x, t)) xss) ctx
export %inline
2024-05-27 15:28:22 -04:00
extendTyN0 : CtxExtension0 q d n1 n2 ->
TyContext q d n1 -> TyContext q d n2
extendTyN0 xss ctx =
let Val q = ctx.qtyLen in
extendTyN (map (\(x, t) => (zero x.loc, x, t)) xss) ctx
export %inline
2024-05-27 15:28:22 -04:00
extendTyLet : Qty q -> BindName -> Term q d n -> Term q d n ->
TyContext q d n -> TyContext q d (S n)
extendTyLet q x s e = extendTyLetN [< (q, x, letVar s e)]
export %inline
2024-05-27 15:28:22 -04:00
extendTy : Qty q -> BindName -> Term q d n ->
TyContext q d n -> TyContext q d (S n)
2023-03-15 10:54:51 -04:00
extendTy q x s = extendTyN [< (q, x, s)]
export %inline
2024-05-27 15:28:22 -04:00
extendTy0 : BindName -> Term q d n ->
TyContext q d n -> TyContext q d (S n)
extendTy0 x t ctx = let Val q = ctx.qtyLen in extendTy (zero x.loc) x t ctx
export %inline
2024-05-27 15:28:22 -04:00
extendDim : BindName -> TyContext q d n -> TyContext q (S d) n
extendDim x (MkTyContext {dimLen, dctx, dnames, qnames, tctx, tnames, qtys}) =
MkTyContext {
dctx = dctx :<? Nothing,
dnames = dnames :< x,
dimLen = [|S dimLen|],
tctx = map weakD tctx,
2024-05-27 15:28:22 -04:00
qnames, tnames, qtys
}
export %inline
2024-05-27 15:28:22 -04:00
eqDim : Dim d -> Dim d -> TyContext q d n -> TyContext q d n
eqDim p q = {dctx $= set p q}
2023-04-15 09:13:01 -04:00
export
2024-05-27 15:28:22 -04:00
toWhnfContext : TyContext q d n -> WhnfContext q d n
toWhnfContext (MkTyContext {qnames, dnames, tnames, tctx, _}) =
MkWhnfContext {qnames, dnames, tnames, tctx}
public export
(.names) : TyContext q d n -> NameContexts q d n
(MkTyContext {qnames, dnames, tnames, _}).names =
MkNameContexts {qnames, dnames, tnames}
2023-04-15 09:13:01 -04:00
namespace QOutput
2023-04-01 13:16:43 -04:00
export %inline
2024-05-27 15:28:22 -04:00
(+) : QOutput q n -> QOutput q n -> QOutput q n
2023-04-01 13:16:43 -04:00
(+) = zipWith (+)
2023-04-01 13:16:43 -04:00
export %inline
2024-05-27 15:28:22 -04:00
(*) : Qty q -> QOutput q n -> QOutput q n
2023-04-01 13:16:43 -04:00
(*) pi = map (pi *)
2023-04-01 13:16:43 -04:00
export %inline
2024-05-27 15:28:22 -04:00
zeroFor : TyContext q _ n -> QOutput q n
zeroFor ctx =
let Val q = ctx.qtyLen in
zeroFor ctx.tctx @{LocateVarType}
export
makeDAssign : DSubst d 0 -> DimAssign d
2023-05-01 21:06:25 -04:00
makeDAssign (Shift SZ) = [<]
makeDAssign (K e _ ::: th) = makeDAssign th :< e
export
2024-05-27 15:28:22 -04:00
makeEqContext' : {d : Nat} -> TyContext q d n -> DSubst d 0 -> EqContext q n
2023-03-15 10:54:51 -04:00
makeEqContext' ctx th = MkEqContext {
2024-05-27 15:28:22 -04:00
qtyLen = ctx.qtyLen,
qnames = ctx.qnames,
termLen = ctx.termLen,
dassign = makeDAssign th,
dnames = ctx.dnames,
tctx = map (subD th) ctx.tctx,
2023-03-15 10:54:51 -04:00
tnames = ctx.tnames,
qtys = ctx.qtys
}
2023-03-15 10:54:51 -04:00
export
2024-05-27 15:28:22 -04:00
makeEqContext : TyContext q d n -> DSubst d 0 -> EqContext q n
2023-03-15 10:54:51 -04:00
makeEqContext ctx@(MkTyContext {dnames, _}) th =
let Val d = lengthPrf0 dnames in makeEqContext' ctx th
2023-03-15 10:54:51 -04:00
namespace EqContext
public export %inline
2024-05-27 15:28:22 -04:00
empty : EqContext 0 0
2023-03-15 10:54:51 -04:00
empty = MkEqContext {
2024-05-27 15:28:22 -04:00
qnames = [<], dassign = [<], dnames = [<],
tctx = [<], tnames = [<], qtys = [<]
2023-03-15 10:54:51 -04:00
}
2023-03-26 10:10:39 -04:00
public export %inline
2024-05-27 15:28:22 -04:00
null : EqContext q n -> Bool
null ctx = null ctx.qnames && null ctx.dnames && null ctx.tnames
2023-03-26 10:10:39 -04:00
export %inline
2024-05-27 15:28:22 -04:00
extendTyLetN : CtxExtensionLet q 0 n1 n2 -> EqContext q n1 -> EqContext q n2
extendTyLetN xss
(MkEqContext {termLen, qnames, dassign, dnames, tctx, tnames, qtys}) =
let (qs, xs, ls) = unzip3 xss in
MkEqContext {
termLen = extendLen xss termLen,
tctx = tctx . ls,
tnames = tnames . xs,
qtys = qtys . qs,
2024-05-27 15:28:22 -04:00
dassign, dnames, qnames
}
export %inline
2024-05-27 15:28:22 -04:00
extendTyN : CtxExtension q 0 n1 n2 -> EqContext q n1 -> EqContext q n2
extendTyN = extendTyLetN . map (\(q, x, s) => (q, x, lamVar s))
export %inline
2024-05-27 15:28:22 -04:00
extendTyLetN0 : CtxExtensionLet0 q 0 n1 n2 -> EqContext q n1 -> EqContext q n2
extendTyLetN0 xss ctx =
let Val q = ctx.qtyLen in
extendTyLetN (map (\(x, t) => (zero x.loc, x, t)) xss) ctx
export %inline
2024-05-27 15:28:22 -04:00
extendTyN0 : CtxExtension0 q 0 n1 n2 -> EqContext q n1 -> EqContext q n2
extendTyN0 xss ctx =
let Val q = ctx.qtyLen in
extendTyN (map (\(x, t) => (zero x.loc, x, t)) xss) ctx
export %inline
2024-05-27 15:28:22 -04:00
extendTyLet : Qty q -> BindName -> Term q 0 n -> Term q 0 n ->
EqContext q n -> EqContext q (S n)
extendTyLet q x s e = extendTyLetN [< (q, x, letVar s e)]
export %inline
2024-05-27 15:28:22 -04:00
extendTy : Qty q -> BindName -> Term q 0 n ->
EqContext q n -> EqContext q (S n)
2023-03-15 10:54:51 -04:00
extendTy q x s = extendTyN [< (q, x, s)]
export %inline
2024-05-27 15:28:22 -04:00
extendTy0 : BindName -> Term q 0 n -> EqContext q n -> EqContext q (S n)
extendTy0 x t ctx = let Val q = ctx.qtyLen in extendTy (zero x.loc) x t ctx
export %inline
2024-05-27 15:28:22 -04:00
extendDim : BindName -> DimConst -> EqContext q n -> EqContext q n
extendDim x e (MkEqContext {dassign, dnames, qnames, tctx, tnames, qtys}) =
MkEqContext {dassign = dassign :< e, dnames = dnames :< x,
2024-05-27 15:28:22 -04:00
qnames, tctx, tnames, qtys}
export
2024-05-27 15:28:22 -04:00
toTyContext : (e : EqContext q n) -> TyContext q e.dimLen n
toTyContext
(MkEqContext {dimLen, dassign, dnames, qnames, tctx, tnames, qtys}) =
2023-03-15 10:54:51 -04:00
MkTyContext {
2023-11-27 15:01:36 -05:00
dctx = fromGround dnames dassign,
tctx = map (subD $ shift0 dimLen) tctx,
2024-05-27 15:28:22 -04:00
dnames, qnames, tnames, qtys
2023-03-15 10:54:51 -04:00
}
2023-04-15 09:13:01 -04:00
export
2024-05-27 15:28:22 -04:00
toWhnfContext : (ectx : EqContext q n) -> WhnfContext q 0 n
toWhnfContext (MkEqContext {qnames, tnames, tctx, _}) =
MkWhnfContext {dnames = [<], qnames, tnames, tctx}
2023-04-15 09:13:01 -04:00
2023-10-15 10:23:38 -04:00
export
2024-05-27 15:28:22 -04:00
injElim : WhnfContext q d n -> Elim q 0 0 -> Elim q d n
2023-10-15 10:23:38 -04:00
injElim ctx e =
let Val d = ctx.dimLen; Val n = ctx.termLen in
e // shift0 d // shift0 n
2024-05-27 15:28:22 -04:00
public export
(.names) : (e : EqContext q n) -> NameContexts q e.dimLen n
(MkEqContext {qnames, dnames, tnames, _}).names =
MkNameContexts {qnames, dnames, tnames}
public export
(.names0) : EqContext q n -> NameContexts q 0 n
(MkEqContext {qnames, tnames, _}).names0 =
MkNameContexts {qnames, dnames = [<], tnames}
2023-04-15 09:13:01 -04:00
namespace WhnfContext
public export %inline
2024-05-27 15:28:22 -04:00
empty : WhnfContext 0 0 0
empty = MkWhnfContext [<] [<] [<] [<]
2023-04-15 09:13:01 -04:00
export
2024-05-27 15:28:22 -04:00
extendTy' : BindName -> LocalVar q d n ->
WhnfContext q d n -> WhnfContext q d (S n)
extendTy' x var (MkWhnfContext {termLen, qnames, dnames, tnames, tctx}) =
MkWhnfContext {
2024-05-27 15:28:22 -04:00
dnames, qnames,
termLen = [|S termLen|],
tnames = tnames :< x,
tctx = tctx :< var
}
export %inline
2024-05-27 15:28:22 -04:00
extendTy : BindName -> Term q d n ->
WhnfContext q d n -> WhnfContext q d (S n)
extendTy x ty ctx = extendTy' x (lamVar ty) ctx
export %inline
2024-05-27 15:28:22 -04:00
extendTyLet : BindName -> (type, term : Term q d n) ->
WhnfContext q d n -> WhnfContext q d (S n)
extendTyLet x type term ctx = extendTy' x (letVar {type, term}) ctx
2023-04-15 09:13:01 -04:00
export
2024-05-27 15:28:22 -04:00
extendDimN : {s : Nat} -> BContext s -> WhnfContext q d n ->
WhnfContext q (s + d) n
extendDimN ns (MkWhnfContext {qnames, dnames, tnames, tctx, dimLen}) =
2023-04-15 09:13:01 -04:00
MkWhnfContext {
2023-10-15 10:23:38 -04:00
dimLen = [|Val s + dimLen|],
2023-04-15 09:13:01 -04:00
dnames = dnames ++ toSnocVect' ns,
tctx = map (subD $ shift s) tctx,
2024-05-27 15:28:22 -04:00
qnames, tnames
2023-04-15 09:13:01 -04:00
}
export
2024-05-27 15:28:22 -04:00
extendDim : BindName -> WhnfContext q d n -> WhnfContext q (S d) n
2023-04-15 09:13:01 -04:00
extendDim i = extendDimN [< i]
2024-05-27 15:28:22 -04:00
public export
(.names) : WhnfContext q d n -> NameContexts q d n
(MkWhnfContext {qnames, dnames, tnames, _}).names =
MkNameContexts {qnames, dnames, tnames}
2023-03-15 10:54:51 -04:00
private
2023-05-14 13:58:46 -04:00
prettyTContextElt : {opts : _} ->
2024-05-27 15:28:22 -04:00
NameContexts q d n ->
Doc opts -> BindName -> LocalVar q d n ->
2024-04-04 13:23:08 -04:00
Eff Pretty (Doc opts)
2024-05-27 15:28:22 -04:00
prettyTContextElt names q x s = do
let Val _ = lengthPrf0 names.qnames
2024-04-04 13:23:08 -04:00
dot <- dotD
x <- prettyTBind x; colon <- colonD
2024-05-27 15:28:22 -04:00
ty <- withPrec Outer $ prettyTerm names s.type; eq <- cstD
tm <- traverse (withPrec Outer . prettyTerm names) s.term
d <- askAt INDENT
let qx = hcat [q, dot, x]
pure $ case tm of
Nothing =>
ifMultiline (hsep [qx, colon, ty]) (vsep [qx, indent d $ colon <++> ty])
Just tm =>
ifMultiline (hsep [qx, colon, ty, eq, tm])
(vsep [qx, indent d $ colon <++> ty, indent d $ eq <++> tm])
2023-05-14 13:58:46 -04:00
private
prettyTContext' : {opts : _} ->
2024-05-27 15:28:22 -04:00
NameContexts q d n -> Context' (Doc opts) n ->
TContext q d n -> Eff Pretty (SnocList (Doc opts))
prettyTContext' _ [<] [<] = pure [<]
prettyTContext' names (qtys :< q) (tys :< t) =
let names' = {tnames $= tail, termLen $= map pred} names in
[|prettyTContext' names' qtys tys :<
prettyTContextElt names' q (head names.tnames) t|]
2023-05-14 13:58:46 -04:00
export
prettyTContext : {opts : _} ->
2024-05-27 15:28:22 -04:00
NameContexts q d n -> QContext q n ->
TContext q d n -> Eff Pretty (Doc opts)
prettyTContext names qtys tys = do
qtys <- traverse (prettyQty names.qnames) qtys
sepSingleTight !commaD . toList <$> prettyTContext' names qtys tys
2023-05-14 13:58:46 -04:00
export
2024-05-27 15:28:22 -04:00
prettyTyContext : {opts : _} -> TyContext q d n -> Eff Pretty (Doc opts)
prettyTyContext ctx = case ctx.dctx of
C [<] => prettyTContext ctx.names ctx.qtys ctx.tctx
_ => pure $
sepSingle [!(prettyDimEq ctx.dnames ctx.dctx) <++> !pipeD,
!(prettyTContext ctx.names ctx.qtys ctx.tctx)]
2023-05-14 13:58:46 -04:00
export
2024-05-27 15:28:22 -04:00
prettyEqContext : {opts : _} -> EqContext q n -> Eff Pretty (Doc opts)
2023-05-14 13:58:46 -04:00
prettyEqContext ctx = prettyTyContext $ toTyContext ctx
2024-04-04 13:23:08 -04:00
export
2024-05-27 15:28:22 -04:00
prettyWhnfContext : {opts : _} -> WhnfContext q d n -> Eff Pretty (Doc opts)
2024-04-04 13:23:08 -04:00
prettyWhnfContext ctx =
let Val n = ctx.termLen in
2024-05-27 15:28:22 -04:00
sepSingleTight !commaD . toList <$>
prettyTContext' ctx.names (replicate n "_") ctx.tctx