coercions and compositions
This commit is contained in:
parent
468ae7e444
commit
a5ccf0215a
25 changed files with 1344 additions and 651 deletions
|
@ -22,7 +22,7 @@ data DimConst = Zero | One
|
|||
|
||||
||| `ends l r e` returns `l` if `e` is `Zero`, or `r` if it is `One`.
|
||||
public export
|
||||
ends : a -> a -> DimConst -> a
|
||||
ends : Lazy a -> Lazy a -> DimConst -> a
|
||||
ends l r Zero = l
|
||||
ends l r One = r
|
||||
|
||||
|
@ -130,5 +130,5 @@ BV i = B $ V i
|
|||
|
||||
|
||||
export
|
||||
weakD : {default 1 by : Nat} -> Dim d -> Dim (by + d)
|
||||
weakD p = p // shift by
|
||||
weakD : (by : Nat) -> Dim d -> Dim (by + d)
|
||||
weakD by p = p // shift by
|
||||
|
|
|
@ -63,6 +63,11 @@ ifConsistent : Applicative f => (eqs : DimEq d) -> f a -> f (IfConsistent eqs a)
|
|||
ifConsistent ZeroIsOne act = pure Nothing
|
||||
ifConsistent (C _) act = Just <$> act
|
||||
|
||||
public export
|
||||
toMaybe : IfConsistent eqs a -> Maybe a
|
||||
toMaybe Nothing = Nothing
|
||||
toMaybe (Just x) = Just x
|
||||
|
||||
|
||||
export
|
||||
fromGround' : Context' DimConst d -> DimEq' d
|
||||
|
@ -245,7 +250,7 @@ PrettyHL (DimEq' d) where
|
|||
go [<] = pure [<]
|
||||
go (eqs :< Nothing) = local {dnames $= tail} $ go eqs
|
||||
go (eqs :< Just p) = do
|
||||
eq <- prettyCst (BV {d = 1} 0) (weakD p)
|
||||
eq <- prettyCst (BV {d = 1} 0) (weakD 1 p)
|
||||
eqs <- local {dnames $= tail} $ go eqs
|
||||
pure $ eqs :< eq
|
||||
|
||||
|
|
|
@ -5,6 +5,7 @@ import public Quox.Syntax.Shift
|
|||
import public Quox.Syntax.Subst
|
||||
import public Quox.Syntax.Qty
|
||||
import public Quox.Syntax.Dim
|
||||
import public Quox.Syntax.Term.TyConKind
|
||||
import public Quox.Name
|
||||
import public Quox.Context
|
||||
|
||||
|
@ -17,33 +18,35 @@ import Data.Nat
|
|||
import public Data.So
|
||||
import Data.String
|
||||
import public Data.SortedMap
|
||||
import public Data.SortedMap.Dependent
|
||||
import public Data.SortedSet
|
||||
|
||||
%default total
|
||||
|
||||
|
||||
public export
|
||||
0 TermLike : Type
|
||||
TermLike : Type
|
||||
TermLike = Nat -> Nat -> Type
|
||||
|
||||
public export
|
||||
0 TSubstLike : Type
|
||||
TSubstLike : Type
|
||||
TSubstLike = Nat -> Nat -> Nat -> Type
|
||||
|
||||
public export
|
||||
0 Universe : Type
|
||||
Universe : Type
|
||||
Universe = Nat
|
||||
|
||||
public export
|
||||
0 TagVal : Type
|
||||
TagVal : Type
|
||||
TagVal = String
|
||||
|
||||
|
||||
infixl 8 :#
|
||||
infixl 9 :@, :%
|
||||
mutual
|
||||
public export
|
||||
0 TSubst : TSubstLike
|
||||
TSubst d = Subst $ Elim d
|
||||
TSubst : TSubstLike
|
||||
TSubst d = Subst $ \n => Elim d n
|
||||
|
||||
||| first argument `d` is dimension scope size;
|
||||
||| second `n` is term scope size
|
||||
|
@ -138,16 +141,19 @@ mutual
|
|||
||| type-annotated term
|
||||
(:#) : (tm, ty : Term d n) -> Elim d n
|
||||
|
||||
||| match on types (needed for coercions :mario_flop:)
|
||||
TypeCase : (ty : Elim d n) ->
|
||||
(ret : Term d n) ->
|
||||
(univ : Term d n) ->
|
||||
(pi : ScopeTermN 2 d n) ->
|
||||
(sig : ScopeTermN 2 d n) ->
|
||||
(enum : Term d n) ->
|
||||
(eq : ScopeTermN 5 d n) ->
|
||||
(nat : Term d n) ->
|
||||
(box : ScopeTerm d n) ->
|
||||
||| coerce a value along a type equality, or show its coherence
|
||||
||| [@xtt; §2.1.1]
|
||||
Coe : (ty : DScopeTerm d n) -> (p, q : Dim d) ->
|
||||
(val : Term d n) -> Elim d n
|
||||
|
||||
||| "generalised composition" [@xtt; §2.1.2]
|
||||
Comp : (ty : Term d n) -> (p, q : Dim d) ->
|
||||
(val : Term d n) -> (r : Dim d) ->
|
||||
(zero, one : DScopeTerm d n) -> Elim d n
|
||||
|
||||
||| match on types. needed for b.s. of coercions [@xtt; §2.2]
|
||||
TypeCase : (ty : Elim d n) -> (ret : Term d n) ->
|
||||
(arms : TypeCaseArms d n) -> (def : Term d n) ->
|
||||
Elim d n
|
||||
|
||||
||| term closure/suspended substitution
|
||||
|
@ -158,9 +164,22 @@ mutual
|
|||
Elim dto n
|
||||
|
||||
public export
|
||||
0 CaseEnumArms : TermLike
|
||||
CaseEnumArms : TermLike
|
||||
CaseEnumArms d n = SortedMap TagVal (Term d n)
|
||||
|
||||
public export
|
||||
TypeCaseArms : TermLike
|
||||
TypeCaseArms d n = SortedDMap TyConKind (\k => TypeCaseArmBody k d n)
|
||||
|
||||
public export
|
||||
TypeCaseArm : TermLike
|
||||
TypeCaseArm d n = (k ** TypeCaseArmBody k d n)
|
||||
|
||||
public export
|
||||
TypeCaseArmBody : TyConKind -> TermLike
|
||||
TypeCaseArmBody k = ScopeTermN (arity k)
|
||||
|
||||
|
||||
||| a scoped term with names
|
||||
public export
|
||||
record Scoped (s : Nat) (f : Nat -> Type) (n : Nat) where
|
||||
|
@ -174,15 +193,16 @@ mutual
|
|||
N : (body : f n) -> ScopedBody s f n
|
||||
|
||||
public export
|
||||
0 ScopeTermN, DScopeTermN : Nat -> TermLike
|
||||
ScopeTermN, DScopeTermN : Nat -> TermLike
|
||||
ScopeTermN s d n = Scoped s (Term d) n
|
||||
DScopeTermN s d n = Scoped s (\d => Term d n) d
|
||||
|
||||
public export
|
||||
0 ScopeTerm, DScopeTerm : TermLike
|
||||
ScopeTerm, DScopeTerm : TermLike
|
||||
ScopeTerm = ScopeTermN 1
|
||||
DScopeTerm = DScopeTermN 1
|
||||
|
||||
|
||||
%name Term s, t, r
|
||||
%name Elim e, f
|
||||
%name Scoped body
|
||||
|
@ -260,6 +280,19 @@ makeNat : Nat -> Term d n
|
|||
makeNat 0 = Zero
|
||||
makeNat (S k) = Succ $ makeNat k
|
||||
|
||||
public export
|
||||
public export %inline
|
||||
enum : List TagVal -> Term d n
|
||||
enum = Enum . SortedSet.fromList
|
||||
|
||||
public export %inline
|
||||
typeCase : Elim d n -> Term d n ->
|
||||
List (TypeCaseArm d n) -> Term d n -> Elim d n
|
||||
typeCase ty ret arms def = TypeCase ty ret (fromList arms) def
|
||||
|
||||
public export %inline
|
||||
typeCase1 : Elim d n -> Term d n ->
|
||||
(k : TyConKind) -> NContext (arity k) -> Term d (arity k + n) ->
|
||||
{default Nat def : Term d n} ->
|
||||
Elim d n
|
||||
typeCase1 ty ret k ns body {def} =
|
||||
typeCase ty ret [(k ** SY ns body)] def
|
||||
|
|
|
@ -15,22 +15,22 @@ export %inline
|
|||
typeD, arrowD, darrowD, timesD, lamD, eqndD, dlamD, annD, natD :
|
||||
Pretty.HasEnv m => m (Doc HL)
|
||||
typeD = hlF Syntax $ ifUnicode "★" "Type"
|
||||
arrowD = hlF Syntax $ ifUnicode "→" "->"
|
||||
darrowD = hlF Syntax $ ifUnicode "⇒" "=>"
|
||||
timesD = hlF Syntax $ ifUnicode "×" "**"
|
||||
arrowD = hlF Delim $ ifUnicode "→" "->"
|
||||
darrowD = hlF Delim $ ifUnicode "⇒" "=>"
|
||||
timesD = hlF Delim $ ifUnicode "×" "**"
|
||||
lamD = hlF Syntax $ ifUnicode "λ" "fun"
|
||||
eqndD = hlF Syntax $ ifUnicode "≡" "=="
|
||||
eqndD = hlF Delim $ ifUnicode "≡" "=="
|
||||
dlamD = hlF Syntax $ ifUnicode "δ" "dfun"
|
||||
annD = hlF Syntax $ ifUnicode "∷" "::"
|
||||
annD = hlF Delim $ ifUnicode "∷" "::"
|
||||
natD = hlF Syntax $ ifUnicode "ℕ" "Nat"
|
||||
|
||||
export %inline
|
||||
eqD, colonD, commaD, semiD, caseD, typecaseD, returnD,
|
||||
ofD, dotD, zeroD, succD : Doc HL
|
||||
ofD, dotD, zeroD, succD, coeD, compD : Doc HL
|
||||
eqD = hl Syntax "Eq"
|
||||
colonD = hl Syntax ":"
|
||||
commaD = hl Syntax ","
|
||||
semiD = hl Syntax ";"
|
||||
colonD = hl Delim ":"
|
||||
commaD = hl Delim ","
|
||||
semiD = hl Delim ";"
|
||||
caseD = hl Syntax "case"
|
||||
typecaseD = hl Syntax "type-case"
|
||||
ofD = hl Syntax "of"
|
||||
|
@ -38,6 +38,8 @@ returnD = hl Syntax "return"
|
|||
dotD = hl Delim "."
|
||||
zeroD = hl Syntax "zero"
|
||||
succD = hl Syntax "succ"
|
||||
coeD = hl Syntax "coe"
|
||||
compD = hl Syntax "compD"
|
||||
|
||||
|
||||
export
|
||||
|
@ -138,10 +140,10 @@ prettyTuple = map (parens . align . separate commaD) . traverse prettyM
|
|||
|
||||
export
|
||||
prettyArms : PrettyHL a => Pretty.HasEnv m =>
|
||||
List (SnocList BaseName, Doc HL, a) -> m (Doc HL)
|
||||
prettyArms =
|
||||
BinderSort -> List (SnocList BaseName, Doc HL, a) -> m (Doc HL)
|
||||
prettyArms s =
|
||||
map (braces . aseparate semiD) .
|
||||
traverse (\(xs, l, r) => prettyArm T xs l r)
|
||||
traverse (\(xs, l, r) => prettyArm s xs l r)
|
||||
|
||||
export
|
||||
prettyCase' : (PrettyHL a, PrettyHL b, PrettyHL c, Pretty.HasEnv m) =>
|
||||
|
@ -153,7 +155,7 @@ prettyCase' intro elim r ret arms = do
|
|||
ret <- case r of
|
||||
Unused => pretty0M ret
|
||||
_ => prettyLams Nothing T [< r] ret
|
||||
arms <- prettyArms arms
|
||||
arms <- prettyArms T arms
|
||||
pure $ asep [intro <++> elim, returnD <++> ret, ofD <++> arms]
|
||||
|
||||
export
|
||||
|
@ -192,6 +194,12 @@ export
|
|||
prettyBoxVal : PrettyHL a => Pretty.HasEnv m => a -> m (Doc HL)
|
||||
prettyBoxVal val = bracks <$> pretty0M val
|
||||
|
||||
export
|
||||
prettyCompPat : Pretty.HasEnv m => Dim d -> DimConst -> BaseName -> m (Doc HL)
|
||||
prettyCompPat s e j = pure $
|
||||
hsep [parens (hsep [!(pretty0M s), hl Syntax "=", !(pretty0M e)]),
|
||||
!(pretty0M $ DV j)]
|
||||
|
||||
export
|
||||
toNatLit : Term d n -> Maybe Nat
|
||||
toNatLit Zero = Just 0
|
||||
|
@ -210,58 +218,77 @@ parameters (showSubsts : Bool)
|
|||
where
|
||||
prettyM (TYPE l) =
|
||||
parensIfM App $ !typeD <+> hl Syntax !(prettyUnivSuffix l)
|
||||
|
||||
prettyM (Pi qty s (S _ (N t))) = do
|
||||
dom <- pretty0M $ MkWithQty qty s
|
||||
cod <- withPrec AnnR $ prettyM t
|
||||
parensIfM AnnR $ asep [dom <++> !arrowD, cod]
|
||||
|
||||
prettyM (Pi qty s (S [< x] (Y t))) =
|
||||
prettyBindType (Just qty) x s !arrowD t
|
||||
|
||||
prettyM (Lam (S x t)) =
|
||||
let GotLams {names, body, _} = getLams' x t.term Refl in
|
||||
prettyLams (Just !lamD) T (toSnocList' names) body
|
||||
|
||||
prettyM (Sig s (S _ (N t))) = do
|
||||
s <- withPrec InTimes $ prettyM s
|
||||
t <- withPrec Times $ prettyM t
|
||||
parensIfM Times $ asep [s <++> !timesD, t]
|
||||
|
||||
prettyM (Sig s (S [< x] (Y t))) =
|
||||
prettyBindType Nothing x s !timesD t
|
||||
|
||||
prettyM (Pair s t) =
|
||||
let GotPairs {init, last, _} = getPairs' [< s] t in
|
||||
prettyTuple $ toList $ init :< last
|
||||
|
||||
prettyM (Enum tags) =
|
||||
pure $ delims "{" "}" . aseparate comma $ map prettyTagBare $
|
||||
Prelude.toList tags
|
||||
|
||||
prettyM (Tag t) =
|
||||
pure $ prettyTag t
|
||||
|
||||
prettyM (Eq (S _ (N ty)) l r) = do
|
||||
l <- withPrec InEq $ prettyM l
|
||||
r <- withPrec InEq $ prettyM r
|
||||
ty <- withPrec InEq $ prettyM ty
|
||||
parensIfM Eq $ asep [l <++> !eqndD, r <++> colonD, ty]
|
||||
|
||||
prettyM (Eq (S [< i] (Y ty)) l r) = do
|
||||
prettyApps Nothing (L eqD)
|
||||
[epretty $ MkTypeLine i ty, epretty l, epretty r]
|
||||
|
||||
prettyM (DLam (S i t)) =
|
||||
let GotDLams {names, body, _} = getDLams' i t.term Refl in
|
||||
prettyLams (Just !dlamD) D (toSnocList' names) body
|
||||
|
||||
prettyM Nat = natD
|
||||
|
||||
prettyM Zero = pure $ hl Syntax "0"
|
||||
|
||||
prettyM (Succ n) =
|
||||
case toNatLit n of
|
||||
Just n => pure $ hl Syntax $ pretty $ S n
|
||||
Nothing => prettyApps Nothing (L succD) [n]
|
||||
|
||||
prettyM (BOX pi ty) = do
|
||||
pi <- pretty0M pi
|
||||
ty <- pretty0M ty
|
||||
pure $ bracks $ hcat [pi, dotD, align ty]
|
||||
|
||||
prettyM (Box val) = prettyBoxVal val
|
||||
|
||||
prettyM (E e) = prettyM e
|
||||
|
||||
prettyM (CloT s th) =
|
||||
if showSubsts then
|
||||
parensIfM SApp . hang 2 =<<
|
||||
[|withPrec SApp (prettyM s) <%> prettyTSubst th|]
|
||||
else
|
||||
prettyM $ pushSubstsWith' id th s
|
||||
|
||||
prettyM (DCloT s th) =
|
||||
if showSubsts then
|
||||
parensIfM SApp . hang 2 =<<
|
||||
|
@ -274,61 +301,97 @@ parameters (showSubsts : Bool)
|
|||
where
|
||||
prettyM (F x) =
|
||||
hl' Free <$> prettyM x
|
||||
|
||||
prettyM (B i) =
|
||||
prettyVar TVar TVarErr (!ask).tnames i
|
||||
|
||||
prettyM (e :@ s) =
|
||||
let GotArgs {fun, args, _} = getArgs' e [s] in
|
||||
prettyApps Nothing fun args
|
||||
|
||||
prettyM (CasePair pi p (S [< r] ret) (S [< x, y] body)) = do
|
||||
pat <- parens . separate commaD <$> traverse (hlF TVar . prettyM) [x, y]
|
||||
prettyCase pi p r ret.term [([< x, y], pat, body.term)]
|
||||
|
||||
prettyM (CaseEnum pi t (S [< r] ret) arms) =
|
||||
prettyCase pi t r ret.term
|
||||
[([<], prettyTag t, b) | (t, b) <- SortedMap.toList arms]
|
||||
|
||||
prettyM (CaseNat pi pi' nat (S [< r] ret) zer (S [< s, ih] suc)) =
|
||||
prettyCase pi nat r ret.term
|
||||
[([<], zeroD, eterm zer),
|
||||
([< s, ih], !succPat, eterm suc.term)]
|
||||
where
|
||||
succPat : m (Doc HL)
|
||||
succPat = case (ih, pi') of
|
||||
(Unused, Zero) => pure $ succD <++> !(pretty0M s)
|
||||
_ => pure $ asep [succD <++> !(pretty0M s) <+> comma,
|
||||
!(pretty0M $ MkWithQty pi' ih)]
|
||||
where
|
||||
succPat : m (Doc HL)
|
||||
succPat = case (ih, pi') of
|
||||
(Unused, Zero) => pure $ succD <++> !(pretty0M s)
|
||||
_ => pure $ asep [succD <++> !(pretty0M s) <+> comma,
|
||||
!(pretty0M $ MkWithQty pi' ih)]
|
||||
|
||||
prettyM (CaseBox pi box (S [< r] ret) (S [< u] body)) =
|
||||
prettyCase pi box r ret.term
|
||||
[([< u], !(prettyBoxVal $ TV u), body.term)]
|
||||
|
||||
prettyM (e :% d) =
|
||||
let GotDArgs {fun, args, _} = getDArgs' e [d] in
|
||||
prettyApps (Just "@") fun args
|
||||
|
||||
prettyM (s :# a) = do
|
||||
s <- withPrec AnnL $ prettyM s
|
||||
a <- withPrec AnnR $ prettyM a
|
||||
parensIfM AnnR $ hang 2 $ s <++> !annD <%%> a
|
||||
prettyM (TypeCase ty ret univ
|
||||
(S [< piA, piB] pi) (S [< sigA, sigB] sig) enum
|
||||
(S [< eqA0, eqA1, eqA, eqL, eqR] eq)
|
||||
nat (S [< boxA] box)) = do
|
||||
let v : BaseName -> Doc HL := pretty0 True . TV
|
||||
pipat <- pure $ parens $ hsep [v piA, !arrowD, v piB]
|
||||
sigpat <- pure $ parens $ hsep [v sigA, !timesD, v sigB]
|
||||
eqpat <- prettyApps Nothing (L eqD)
|
||||
[TV eqA0, TV eqA1, TV eqA, TV eqL, TV eqR]
|
||||
boxpat <- pure $ bracks $ v boxA
|
||||
prettyCase' typecaseD ty Unused ret
|
||||
[([<], !typeD, eterm univ),
|
||||
([< piA, piB], pipat, eterm pi.term),
|
||||
([< sigA, sigB], sigpat, eterm sig.term),
|
||||
([<], hl Syntax "{}", eterm enum),
|
||||
([< eqA0, eqA1, eqA, eqL, eqR], eqpat, eterm eq.term),
|
||||
([<], !natD, eterm nat),
|
||||
([< boxA], boxpat, eterm box.term)]
|
||||
|
||||
prettyM (Coe (S [< i] ty) p q val) =
|
||||
let ty = case ty of
|
||||
Y ty => epretty $ MkTypeLine i ty
|
||||
N ty => epretty ty
|
||||
in
|
||||
prettyApps' (L coeD)
|
||||
[(Nothing, ty),
|
||||
(Just "@", epretty p),
|
||||
(Just "@", epretty q),
|
||||
(Nothing, epretty val)]
|
||||
|
||||
prettyM (Comp ty p q val r (S [< z] zero) (S [< o] one)) = do
|
||||
apps <- prettyApps' (L compD)
|
||||
[(Nothing, epretty ty),
|
||||
(Just "@", epretty p),
|
||||
(Just "@", epretty q),
|
||||
(Nothing, epretty val)]
|
||||
arms <- prettyArms D
|
||||
[([< z], !(prettyCompPat r Zero z), zero.term),
|
||||
([< o], !(prettyCompPat r One o), one.term)]
|
||||
pure $ apps <++> arms
|
||||
|
||||
prettyM (TypeCase ty ret arms def) = do
|
||||
arms <- traverse fromArm (toList arms)
|
||||
prettyCase' typecaseD ty Unused ret $
|
||||
arms ++ [([<], hl Syntax "_", eterm def)]
|
||||
where
|
||||
v : BaseName -> Doc HL
|
||||
v = pretty0 True . TV
|
||||
|
||||
tyCasePat : (k : TyConKind) -> NContext (arity k) -> m (Doc HL)
|
||||
tyCasePat KTYPE [<] = typeD
|
||||
tyCasePat KPi [< a, b] = pure $ parens $ hsep [v a, !arrowD, v b]
|
||||
tyCasePat KSig [< a, b] = pure $ parens $ hsep [v a, !arrowD, v b]
|
||||
tyCasePat KEnum [<] = pure $ hl Syntax "{}"
|
||||
tyCasePat KEq vars = prettyApps Nothing (L eqD) $ map TV $ toList' vars
|
||||
tyCasePat KNat [<] = natD
|
||||
tyCasePat KBOX [< a] = pure $ bracks $ v a
|
||||
|
||||
fromArm : TypeCaseArm d n ->
|
||||
m (SnocList BaseName, Doc HL, Exists (Term d))
|
||||
fromArm (k ** S ns t) =
|
||||
pure (toSnocList' ns, !(tyCasePat k ns), eterm t.term)
|
||||
|
||||
prettyM (CloE e th) =
|
||||
if showSubsts then
|
||||
parensIfM SApp . hang 2 =<<
|
||||
[|withPrec SApp (prettyM e) <%> prettyTSubst th|]
|
||||
else
|
||||
prettyM $ pushSubstsWith' id th e
|
||||
|
||||
prettyM (DCloE e th) =
|
||||
if showSubsts then
|
||||
parensIfM SApp . hang 2 =<<
|
||||
|
|
|
@ -130,21 +130,21 @@ comp th ps ph = map (// th) ps . ph
|
|||
|
||||
|
||||
public export %inline
|
||||
dweakT : {by : Nat} -> Term d n -> Term (by + d) n
|
||||
dweakT t = t // shift by
|
||||
dweakT : (by : Nat) -> Term d n -> Term (by + d) n
|
||||
dweakT by t = t // shift by
|
||||
|
||||
public export %inline
|
||||
dweakE : {by : Nat} -> Elim d n -> Elim (by + d) n
|
||||
dweakE t = t // shift by
|
||||
dweakE : (by : Nat) -> Elim d n -> Elim (by + d) n
|
||||
dweakE by t = t // shift by
|
||||
|
||||
|
||||
public export %inline
|
||||
weakT : {default 1 by : Nat} -> Term d n -> Term d (by + n)
|
||||
weakT t = t // shift by
|
||||
weakT : (by : Nat) -> Term d n -> Term d (by + n)
|
||||
weakT by t = t // shift by
|
||||
|
||||
public export %inline
|
||||
weakE : {default 1 by : Nat} -> Elim d n -> Elim d (by + n)
|
||||
weakE t = t // shift by
|
||||
weakE : (by : Nat) -> Elim d n -> Elim d (by + n)
|
||||
weakE by t = t // shift by
|
||||
|
||||
|
||||
parameters {s : Nat}
|
||||
|
@ -152,7 +152,7 @@ parameters {s : Nat}
|
|||
export %inline
|
||||
(.term) : ScopedBody s (Term d) n -> Term d (s + n)
|
||||
(Y b).term = b
|
||||
(N b).term = weakT b {by = s}
|
||||
(N b).term = weakT s b
|
||||
|
||||
namespace ScopeTermN
|
||||
export %inline
|
||||
|
@ -163,7 +163,7 @@ parameters {s : Nat}
|
|||
export %inline
|
||||
(.term) : ScopedBody s (\d => Term d n) d -> Term (s + d) n
|
||||
(Y b).term = b
|
||||
(N b).term = dweakT b {by = s}
|
||||
(N b).term = dweakT s b
|
||||
|
||||
namespace DScopeTermN
|
||||
export %inline
|
||||
|
@ -305,13 +305,37 @@ mutual
|
|||
nclo $ (f // th // ph) :% (d // th)
|
||||
pushSubstsWith th ph (s :# a) =
|
||||
nclo $ (s // th // ph) :# (a // th // ph)
|
||||
pushSubstsWith th ph (TypeCase ty ret univ pi sig enum eq nat box) =
|
||||
nclo $ TypeCase (ty // th // ph) (ret // th // ph)
|
||||
(univ // th // ph) (pi // th // ph)
|
||||
(sig // th // ph) (enum // th // ph)
|
||||
(eq // th // ph) (nat // th // ph)
|
||||
(box // th // ph)
|
||||
pushSubstsWith th ph (Coe ty p q val) =
|
||||
nclo $ Coe (ty // th // ph) (p // th) (q // th) (val // th // ph)
|
||||
pushSubstsWith th ph (Comp ty p q val r zero one) =
|
||||
nclo $ Comp (ty // th // ph) (p // th) (q // th)
|
||||
(val // th // ph) (r // th)
|
||||
(zero // th // ph) (one // th // ph)
|
||||
pushSubstsWith th ph (TypeCase ty ret arms def) =
|
||||
nclo $ TypeCase (ty // th // ph) (ret // th // ph)
|
||||
(map (\t => t // th // ph) arms) (def // th // ph)
|
||||
pushSubstsWith th ph (CloE e ps) =
|
||||
pushSubstsWith th (comp th ps ph) e
|
||||
pushSubstsWith th ph (DCloE e ps) =
|
||||
pushSubstsWith (ps . th) ph e
|
||||
|
||||
|
||||
||| heterogeneous composition, using Comp and Coe (and subst)
|
||||
|||
|
||||
||| comp [i ⇒ A] @p @q s { (r=0) j ⇒ t₀; (r=1) j ⇒ t₁ }
|
||||
||| ≔
|
||||
||| comp [A‹q/i›] @p @q (coe [i ⇒ A] @p @q s) {
|
||||
||| (r=0) j ⇒ coe [i ⇒ A] @j @q t₀;
|
||||
||| (r=1) j ⇒ coe [i ⇒ A] @j @q t₁
|
||||
||| }
|
||||
public export %inline
|
||||
CompH : (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) ->
|
||||
(r : Dim d) -> (zero, one : DScopeTerm d n) -> Elim d n
|
||||
CompH {ty, p, q, val, r, zero, one} =
|
||||
let ty' = SY ty.names $ ty.term // (B VZ ::: shift 2) in
|
||||
Comp {
|
||||
ty = dsub1 ty q, p, q,
|
||||
val = E $ Coe ty p q val, r,
|
||||
zero = SY zero.names $ E $ Coe ty' (B VZ) (weakD 1 q) zero.term,
|
||||
one = SY one.names $ E $ Coe ty' (B VZ) (weakD 1 q) one.term
|
||||
}
|
||||
|
|
34
lib/Quox/Syntax/Term/TyConKind.idr
Normal file
34
lib/Quox/Syntax/Term/TyConKind.idr
Normal file
|
@ -0,0 +1,34 @@
|
|||
module Quox.Syntax.Term.TyConKind
|
||||
|
||||
import Decidable.Equality
|
||||
import Derive.Prelude
|
||||
import Generics.Derive
|
||||
|
||||
%language ElabReflection
|
||||
%default total
|
||||
|
||||
|
||||
public export
|
||||
data TyConKind = KTYPE | KPi | KSig | KEnum | KEq | KNat | KBOX
|
||||
%name TyConKind k
|
||||
%runElab derive "TyConKind" [Eq.Eq, Ord.Ord, Show.Show, Generic, Meta, DecEq]
|
||||
|
||||
||| a list of all `TyConKind`s
|
||||
public export %inline
|
||||
allKinds : List TyConKind
|
||||
allKinds = %runElab do
|
||||
-- as a script so it stays up to date if there are more tycon kinds
|
||||
cs <- getCons $ fst !(lookupName "TyConKind")
|
||||
traverse (check . var) cs
|
||||
|
||||
|
||||
||| in `type-case`, how many variables are bound in this branch
|
||||
public export %inline
|
||||
arity : TyConKind -> Nat
|
||||
arity KTYPE = 0
|
||||
arity KPi = 2
|
||||
arity KSig = 2
|
||||
arity KEnum = 0
|
||||
arity KEq = 5
|
||||
arity KNat = 0
|
||||
arity KBOX = 1
|
Loading…
Add table
Add a link
Reference in a new issue