add source locations to inner syntax

This commit is contained in:
rhiannon morris 2023-05-02 03:06:25 +02:00
parent 30fa93ab4e
commit d5f4a012c5
35 changed files with 3210 additions and 2482 deletions

View file

@ -1,5 +1,6 @@
module Quox.Syntax.Dim
import Quox.Loc
import Quox.Name
import Quox.Syntax.Var
import Quox.Syntax.Subst
@ -17,7 +18,6 @@ import Derive.Prelude
public export
data DimConst = Zero | One
%name DimConst e
%runElab derive "DimConst" [Eq, Ord, Show]
||| `ends l r e` returns `l` if `e` is `Zero`, or `r` if it is `One`.
@ -26,23 +26,42 @@ ends : Lazy a -> Lazy a -> DimConst -> a
ends l r Zero = l
ends l r One = r
export Uninhabited (Zero = One) where uninhabited _ impossible
export Uninhabited (One = Zero) where uninhabited _ impossible
public export
DecEq DimConst where
decEq Zero Zero = Yes Refl
decEq Zero One = No absurd
decEq One Zero = No absurd
decEq One One = Yes Refl
public export
data Dim : Nat -> Type where
K : DimConst -> Dim d
B : Var d -> Dim d
K : DimConst -> Loc -> Dim d
B : Var d -> Loc -> Dim d
%name Dim.Dim p, q
%runElab deriveIndexed "Dim" [Eq, Ord, Show]
export
Located (Dim d) where
(K _ loc).loc = loc
(B _ loc).loc = loc
export
Relocatable (Dim d) where
setLoc loc (K e _) = K e loc
setLoc loc (B i _) = B i loc
export
PrettyHL DimConst where
prettyM = pure . hl Dim . ends "0" "1"
export
PrettyHL (Dim n) where
prettyM (K e) = prettyM e
prettyM (B i) = prettyVar DVar DVarErr (!ask).dnames i
prettyM (K e _) = prettyM e
prettyM (B i _) = prettyVar DVar DVarErr (!ask).dnames i
export
prettyDim : (dnames : NContext d) -> Dim d -> Doc HL
@ -60,13 +79,13 @@ prettyDim dnames p =
||| - `x` otherwise.
public export
endsOr : Lazy a -> Lazy a -> Lazy a -> Dim n -> a
endsOr l r x (K e) = ends l r e
endsOr l r x (B _) = x
endsOr l r x (K e _) = ends l r e
endsOr l r x (B _ _) = x
public export %inline
toConst : Dim 0 -> DimConst
toConst (K e) = e
toConst (K e _) = e
public export
@ -81,52 +100,55 @@ prettyDSubst th =
!(ifUnicode "" "<") !(ifUnicode "" ">") th
public export FromVar Dim where fromVar = B
public export FromVar Dim where fromVarLoc = B
export
CanShift Dim where
K e // _ = K e
B i // by = B (i // by)
K e loc // _ = K e loc
B i loc // by = B (i // by) loc
export
CanSubstSelf Dim where
K e // _ = K e
B i // th = th !! i
K e loc // _ = K e loc
B i loc // th = getLoc th i loc
export Uninhabited (Zero = One) where uninhabited _ impossible
export Uninhabited (One = Zero) where uninhabited _ impossible
export Uninhabited (B i = K e) where uninhabited _ impossible
export Uninhabited (K e = B i) where uninhabited _ impossible
public export %inline Injective Dim.B where injective Refl = Refl
public export %inline Injective Dim.K where injective Refl = Refl
export Uninhabited (B i loc1 = K e loc2) where uninhabited _ impossible
export Uninhabited (K e loc1 = B i loc2) where uninhabited _ impossible
public export
DecEq DimConst where
decEq Zero Zero = Yes Refl
decEq Zero One = No absurd
decEq One Zero = No absurd
decEq One One = Yes Refl
data Eqv : Dim d1 -> Dim d2 -> Type where
EK : K e _ `Eqv` K e _
EB : i `Eqv` j -> B i _ `Eqv` B j _
export Uninhabited (K e l1 `Eqv` B i l2) where uninhabited _ impossible
export Uninhabited (B i l1 `Eqv` K e l2) where uninhabited _ impossible
export
injectiveB : B i loc1 `Eqv` B j loc2 -> i `Eqv` j
injectiveB (EB e) = e
export
injectiveK : K e loc1 `Eqv` K f loc2 -> e = f
injectiveK EK = Refl
public export
DecEq (Dim d) where
decEq (K e) (K f) with (decEq e f)
_ | Yes prf = Yes $ cong K prf
_ | No contra = No $ contra . injective
decEq (K e) (B j) = No absurd
decEq (B i) (K f) = No absurd
decEq (B i) (B j) with (decEq i j)
_ | Yes prf = Yes $ cong B prf
_ | No contra = No $ contra . injective
decEqv : Dec2 Dim.Eqv
decEqv (K e _) (K f _) = case decEq e f of
Yes Refl => Yes EK
No n => No $ n . injectiveK
decEqv (B i _) (B j _) = case decEqv i j of
Yes y => Yes $ EB y
No n => No $ \(EB y) => n y
decEqv (B _ _) (K _ _) = No absurd
decEqv (K _ _) (B _ _) = No absurd
||| abbreviation for a bound variable like `BV 4` instead of
||| `B (VS (VS (VS (VS VZ))))`
public export %inline
BV : (i : Nat) -> (0 _ : LT i d) => Dim d
BV i = B $ V i
BV : (i : Nat) -> (0 _ : LT i d) => (loc : Loc) -> Dim d
BV i loc = B (V i) loc
export

View file

@ -72,7 +72,7 @@ toMaybe (Just x) = Just x
export
fromGround' : Context' DimConst d -> DimEq' d
fromGround' [<] = [<]
fromGround' (ctx :< e) = fromGround' ctx :< Just (K e)
fromGround' (ctx :< e) = fromGround' ctx :< Just (K e noLoc)
export
fromGround : Context' DimConst d -> DimEq d
@ -98,8 +98,8 @@ get' : DimEq' d -> Var d -> Maybe (Dim d)
get' = getWith $ \p, by => map (// by) p
public export %inline
getVar : DimEq' d -> Var d -> Dim d
getVar eqs i = fromMaybe (B i) $ get' eqs i
getVar : DimEq' d -> Var d -> Loc -> Dim d
getVar eqs i loc = fromMaybe (B i loc) $ get' eqs i
public export %inline
getShift' : Shift len out -> DimEq' len -> Var len -> Maybe (Dim out)
@ -107,8 +107,8 @@ getShift' = getShiftWith $ \p, by => map (// by) p
public export %inline
get : DimEq' d -> Dim d -> Dim d
get _ (K e) = K e
get eqs (B i) = getVar eqs i
get _ (K e loc) = K e loc
get eqs (B i loc) = getVar eqs i loc
public export %inline
@ -126,7 +126,7 @@ C eqs :<? d = C $ eqs :< map (get eqs) d
private %inline
ifVar : Var d -> Dim d -> Maybe (Dim d) -> Maybe (Dim d)
ifVar i p = map $ \q => if q == B i then p else q
ifVar i p = map $ \q => if q == B i noLoc then p else q
-- (using decEq instead of (==) because of the proofs below)
private %inline
@ -135,39 +135,43 @@ checkConst e f eqs = if isYes $ e `decEq` f then C eqs else ZeroIsOne
export
setConst : Var d -> DimConst -> DimEq' d -> DimEq d
setConst VZ e (eqs :< Nothing) = C $ eqs :< Just (K e)
setConst VZ e (eqs :< Just (K f)) = checkConst e f $ eqs :< Just (K f)
setConst VZ e (eqs :< Just (B i)) = setConst i e eqs :<? Just (K e)
setConst (VS i) e (eqs :< p) = setConst i e eqs :<? ifVar i (K e) p
setConst : Var d -> DimConst -> Loc -> DimEq' d -> DimEq d
setConst VZ e loc (eqs :< Nothing) =
C $ eqs :< Just (K e loc)
setConst VZ e _ (eqs :< Just (K f loc)) =
checkConst e f $ eqs :< Just (K f loc)
setConst VZ e loc (eqs :< Just (B i _)) =
setConst i e loc eqs :<? Just (K e loc)
setConst (VS i) e loc (eqs :< p) =
setConst i e loc eqs :<? ifVar i (K e loc) p
mutual
private
setVar' : (i, j : Var d) -> (0 _ : i `LT` j) -> DimEq' d -> DimEq d
setVar' VZ (VS i) LTZ (eqs :< Nothing) =
C eqs :<? Just (B i)
setVar' VZ (VS i) LTZ (eqs :< Just (K e)) =
setConst i e eqs :<? Just (K e)
setVar' VZ (VS i) LTZ (eqs :< Just (B j)) =
setVar i j eqs :<? Just (B (max i j))
setVar' (VS i) (VS j) (LTS lt) (eqs :< p) =
setVar' i j lt eqs :<? ifVar i (B j) p
setVar' : (i, j : Var d) -> (0 _ : i `LT` j) -> Loc -> DimEq' d -> DimEq d
setVar' VZ (VS i) LTZ loc (eqs :< Nothing) =
C eqs :<? Just (B i loc)
setVar' VZ (VS i) LTZ loc (eqs :< Just (K e eloc)) =
setConst i e loc eqs :<? Just (K e eloc)
setVar' VZ (VS i) LTZ loc (eqs :< Just (B j jloc)) =
setVar i j loc jloc eqs :<? Just (if j > i then B j jloc else B i loc)
setVar' (VS i) (VS j) (LTS lt) loc (eqs :< p) =
setVar' i j lt loc eqs :<? ifVar i (B j loc) p
export %inline
setVar : (i, j : Var d) -> DimEq' d -> DimEq d
setVar i j eqs with (compareP i j) | (compare i.nat j.nat)
setVar i j eqs | IsLT lt | LT = setVar' i j lt eqs
setVar i i eqs | IsEQ | EQ = C eqs
setVar i j eqs | IsGT gt | GT = setVar' j i gt eqs
setVar : (i, j : Var d) -> Loc -> Loc -> DimEq' d -> DimEq d
setVar i j li lj eqs with (compareP i j) | (compare i.nat j.nat)
setVar i j li lj eqs | IsLT lt | LT = setVar' i j lt lj eqs
setVar i i li lj eqs | IsEQ | EQ = C eqs
setVar i j li lj eqs | IsGT gt | GT = setVar' j i gt li eqs
export %inline
set : (p, q : Dim d) -> DimEq d -> DimEq d
set _ _ ZeroIsOne = ZeroIsOne
set (K e) (K f) (C eqs) = checkConst e f eqs
set (K e) (B i) (C eqs) = setConst i e eqs
set (B i) (K e) (C eqs) = setConst i e eqs
set (B i) (B j) (C eqs) = setVar i j eqs
set (K e eloc) (K f floc) (C eqs) = checkConst e f eqs
set (K e eloc) (B i iloc) (C eqs) = setConst i e eloc eqs
set (B i iloc) (K e eloc) (C eqs) = setConst i e eloc eqs
set (B i iloc) (B j jloc) (C eqs) = setVar i j iloc jloc eqs
public export %inline
@ -175,25 +179,26 @@ Split : Nat -> Type
Split d = (DimEq' d, DSubst (S d) d)
export %inline
split1 : DimConst -> DimEq' (S d) -> Maybe (Split d)
split1 e eqs = case setConst VZ e eqs of
split1 : DimConst -> Loc -> DimEq' (S d) -> Maybe (Split d)
split1 e loc eqs = case setConst VZ e loc eqs of
ZeroIsOne => Nothing
C (eqs :< _) => Just (eqs, K e ::: id)
C (eqs :< _) => Just (eqs, K e loc ::: id)
export %inline
split : DimEq' (S d) -> List (Split d)
split eqs = toList (split1 Zero eqs) <+> toList (split1 One eqs)
split : Loc -> DimEq' (S d) -> List (Split d)
split loc eqs = toList (split1 Zero loc eqs) <+> toList (split1 One loc eqs)
export
splits' : DimEq' d -> List (DSubst d 0)
splits' [<] = [id]
splits' eqs@(_ :< _) = [th . ph | (eqs', th) <- split eqs, ph <- splits' eqs']
splits' : Loc -> DimEq' d -> List (DSubst d 0)
splits' _ [<] = [id]
splits' loc eqs@(_ :< _) =
[th . ph | (eqs', th) <- split loc eqs, ph <- splits' loc eqs']
||| the Loc is put into each of the DimConsts
export %inline
splits : DimEq d -> List (DSubst d 0)
splits ZeroIsOne = []
splits (C eqs) = splits' eqs
splits : Loc -> DimEq d -> List (DSubst d 0)
splits _ ZeroIsOne = []
splits loc (C eqs) = splits' loc eqs
private
@ -208,16 +213,16 @@ newGet' d i = newGetShift d i SZ
export
0 newGet : (d : Nat) -> (p : Dim d) -> get (new' {d}) p = p
newGet d (K e) = Refl
newGet d (B i) = rewrite newGet' d i in Refl
newGet d (K e _) = Refl
newGet d (B i _) = rewrite newGet' d i in Refl
export
0 setSelf : (p : Dim d) -> (eqs : DimEq d) -> set p p eqs = eqs
setSelf p ZeroIsOne = Refl
setSelf (K Zero) (C eqs) = Refl
setSelf (K One) (C eqs) = Refl
setSelf (B i) (C eqs) with (compareP i i) | (compare i.nat i.nat)
setSelf (K Zero _) (C eqs) = Refl
setSelf (K One _) (C eqs) = Refl
setSelf (B i _) (C eqs) with (compareP i i) | (compare i.nat i.nat)
_ | IsLT lt | LT = absurd lt
_ | IsEQ | EQ = Refl
_ | IsGT gt | GT = absurd gt
@ -250,7 +255,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 1 p)
eq <- prettyCst (BV {d = 1} 0 noLoc) (weakD 1 p)
eqs <- local {dnames $= tail} $ go eqs
pure $ eqs :< eq
@ -262,16 +267,16 @@ PrettyHL (DimEq d) where
prettyM (C eqs) = prettyM eqs
export
prettyDimEq : NContext d -> DimEq d -> Doc HL
prettyDimEq ds = pretty0With False (toSnocList' ds) [<]
prettyDimEq : BContext d -> DimEq d -> Doc HL
prettyDimEq ds = pretty0With False (toNames ds) [<]
public export
wf' : DimEq' d -> Bool
wf' [<] = True
wf' (eqs :< Nothing) = wf' eqs
wf' (eqs :< Just (K e)) = wf' eqs
wf' (eqs :< Just (B i)) = isNothing (get' eqs i) && wf' eqs
wf' [<] = True
wf' (eqs :< Nothing) = wf' eqs
wf' (eqs :< Just (K e _)) = wf' eqs
wf' (eqs :< Just (B i _)) = isNothing (get' eqs i) && wf' eqs
public export
wf : DimEq d -> Bool

View file

@ -35,27 +35,28 @@ public export
data Eqv : Shift from1 to1 -> Shift from2 to2 -> Type where
EqSZ : SZ `Eqv` SZ
EqSS : by `Eqv` bz -> SS by `Eqv` SS bz
%name Eqv e
%name Shift.Eqv e
||| two equivalent shifts are equal if they have the same indices.
export
0 fromEqv : by `Eqv` bz -> by = bz
fromEqv EqSZ = Refl
fromEqv (EqSS e) = cong SS $ fromEqv e
using (by : Shift from to, bz : Shift from to)
||| two equivalent shifts are equal if they have the same indices.
export
0 fromEqv : by `Eqv` bz -> by = bz
fromEqv EqSZ = Refl
fromEqv (EqSS e) = cong SS $ fromEqv e
||| two equal shifts are equivalent.
export
0 toEqv : by = bz -> by `Eqv` bz
toEqv Refl {by = SZ} = EqSZ
toEqv Refl {by = (SS by)} = EqSS $ toEqv Refl
||| two equal shifts are equivalent.
export
0 toEqv : by = bz -> by `Eqv` bz
toEqv Refl {by = SZ} = EqSZ
toEqv Refl {by = (SS by)} = EqSS $ toEqv Refl
export
eqLen : Shift from1 to -> Shift from2 to -> Maybe (from1 = from2)
eqLen SZ SZ = Just Refl
eqLen SZ (SS by) = Nothing
eqLen (SS by) SZ = Nothing
eqLen (SS by) (SS bz) = eqLen by bz
cmpLen : Shift from1 to -> Shift from2 to -> Either Ordering (from1 = from2)
cmpLen SZ SZ = Right Refl
cmpLen SZ (SS by) = Left LT
cmpLen (SS by) SZ = Left GT
cmpLen (SS by) (SS bz) = cmpLen by bz
export
0 shiftDiff : (by : Shift from to) -> to = by.nat + from

View file

@ -48,12 +48,16 @@ interface FromVar term => CanSubstSelf term where
(//) : term from -> Lazy (Subst term from to) -> term to
infixl 8 !!
public export
(!!) : FromVar term => Subst term from to -> Var from -> term to
(Shift by) !! i = fromVar $ shift by i
(t ::: th) !! VZ = t
(t ::: th) !! (VS i) = th !! i
getLoc : FromVar term => Subst term from to -> Var from -> Loc -> term to
getLoc (Shift by) i loc = fromVarLoc (shift by i) loc
getLoc (t ::: th) VZ _ = t
getLoc (t ::: th) (VS i) loc = getLoc th i loc
-- infixl 8 !!
-- public export
-- (!!) : FromVar term => Subst term from to -> Var from -> term to
-- th !! i = getLoc th i noLoc
public export
@ -160,12 +164,16 @@ PrettyHL (f to) => PrettyHL (Subst f from to) where
prettyM th = prettySubstM prettyM (!ask).tnames TVar "[" "]" th
||| whether two substitutions with the same codomain have the same shape
||| (the same number of terms and the same shift at the end). if so, they
||| also have the same domain
export
eqShape : Subst env from1 to -> Subst env from2 to -> Maybe (from1 = from2)
eqShape (Shift by) (Shift bz) = eqLen by bz
eqShape (Shift by) (t ::: th) = Nothing
eqShape (t ::: th) (Shift by) = Nothing
eqShape (t ::: th) (x ::: ph) = cong S <$> eqShape th ph
cmpShape : Subst env from1 to -> Subst env from2 to ->
Either Ordering (from1 = from2)
cmpShape (Shift by) (Shift bz) = cmpLen by bz
cmpShape (Shift _) (_ ::: _) = Left LT
cmpShape (_ ::: _) (Shift _) = Left GT
cmpShape (_ ::: th) (_ ::: ph) = cong S <$> cmpShape th ph
public export
@ -175,13 +183,20 @@ record WithSubst tm env n where
subst : Lazy (Subst env from n)
export
(forall n. Eq (tm n), Eq (env n)) => Eq (WithSubst tm env n) where
(Eq (env n), forall n. Eq (tm n)) => Eq (WithSubst tm env n) where
Sub t1 s1 == Sub t2 s2 =
case eqShape s1 s2 of
Just Refl => t1 == t2 && s1 == s2
Nothing => False
case cmpShape s1 s2 of
Left _ => False
Right Refl => t1 == t2 && s1 == s2
export
(Ord (env n), forall n. Ord (tm n)) => Ord (WithSubst tm env n) where
Sub t1 s1 `compare` Sub t2 s2 =
case cmpShape s1 s2 of
Left o => o
Right Refl => compare (t1, s1) (t2, s2)
export %hint
ShowWithSubst : (forall n. Show (tm n), Show (env n)) =>
ShowWithSubst : (Show (env n), forall n. Show (tm n)) =>
Show (WithSubst tm env n)
ShowWithSubst = deriveShow

View file

@ -7,6 +7,7 @@ import public Quox.Syntax.Qty
import public Quox.Syntax.Dim
import public Quox.Syntax.Term.TyConKind
import public Quox.Name
import public Quox.Loc
import public Quox.Context
import Quox.Pretty
@ -63,7 +64,7 @@ ShowScopedBody = deriveShow
public export
record Scoped (s : Nat) (f : Nat -> Type) (n : Nat) where
constructor S
names : NContext s
names : BContext s
body : ScopedBody s f n
%name Scoped body
@ -88,38 +89,38 @@ mutual
public export
data Term : (d, n : Nat) -> Type where
||| type of types
TYPE : (l : Universe) -> Term d n
TYPE : (l : Universe) -> (loc : Loc) -> Term d n
||| function type
Pi : (qty : Qty) -> (arg : Term d n) ->
(res : ScopeTerm d n) -> Term d n
(res : ScopeTerm d n) -> (loc : Loc) -> Term d n
||| function term
Lam : (body : ScopeTerm d n) -> Term d n
Lam : (body : ScopeTerm d n) -> (loc : Loc) -> Term d n
||| pair type
Sig : (fst : Term d n) -> (snd : ScopeTerm d n) -> Term d n
Sig : (fst : Term d n) -> (snd : ScopeTerm d n) -> (loc : Loc) -> Term d n
||| pair value
Pair : (fst, snd : Term d n) -> Term d n
Pair : (fst, snd : Term d n) -> (loc : Loc) -> Term d n
||| enumeration type
Enum : (cases : SortedSet TagVal) -> Term d n
Enum : (cases : SortedSet TagVal) -> (loc : Loc) -> Term d n
||| enumeration value
Tag : (tag : TagVal) -> Term d n
Tag : (tag : TagVal) -> (loc : Loc) -> Term d n
||| equality type
Eq : (ty : DScopeTerm d n) -> (l, r : Term d n) -> Term d n
Eq : (ty : DScopeTerm d n) -> (l, r : Term d n) -> (loc : Loc) -> Term d n
||| equality term
DLam : (body : DScopeTerm d n) -> Term d n
DLam : (body : DScopeTerm d n) -> (loc : Loc) -> Term d n
||| natural numbers (temporary until 𝐖 gets added)
Nat : Term d n
Nat : (loc : Loc) -> Term d n
-- [todo] can these be elims?
Zero : Term d n
Succ : (p : Term d n) -> Term d n
Zero : (loc : Loc) -> Term d n
Succ : (p : Term d n) -> (loc : Loc) -> Term d n
||| "box" (package a value up with a certain quantity)
BOX : (qty : Qty) -> (ty : Term d n) -> Term d n
Box : (val : Term d n) -> Term d n
BOX : (qty : Qty) -> (ty : Term d n) -> (loc : Loc) -> Term d n
Box : (val : Term d n) -> (loc : Loc) -> Term d n
||| elimination
E : (e : Elim d n) -> Term d n
@ -134,12 +135,12 @@ mutual
public export
data Elim : (d, n : Nat) -> Type where
||| free variable
F : (x : Name) -> Elim d n
F : (x : Name) -> (loc : Loc) -> Elim d n
||| bound variable
B : (i : Var n) -> Elim d n
B : (i : Var n) -> (loc : Loc) -> Elim d n
||| term application
(:@) : (fun : Elim d n) -> (arg : Term d n) -> Elim d n
App : (fun : Elim d n) -> (arg : Term d n) -> (loc : Loc) -> Elim d n
||| pair destruction
|||
@ -148,12 +149,14 @@ mutual
CasePair : (qty : Qty) -> (pair : Elim d n) ->
(ret : ScopeTerm d n) ->
(body : ScopeTermN 2 d n) ->
(loc : Loc) ->
Elim d n
||| enum matching
CaseEnum : (qty : Qty) -> (tag : Elim d n) ->
(ret : ScopeTerm d n) ->
(arms : CaseEnumArms d n) ->
(loc : Loc) ->
Elim d n
||| nat matching
@ -161,33 +164,36 @@ mutual
(ret : ScopeTerm d n) ->
(zero : Term d n) ->
(succ : ScopeTermN 2 d n) ->
(loc : Loc) ->
Elim d n
||| unboxing
CaseBox : (qty : Qty) -> (box : Elim d n) ->
(ret : ScopeTerm d n) ->
(body : ScopeTerm d n) ->
(loc : Loc) ->
Elim d n
||| dim application
(:%) : (fun : Elim d n) -> (arg : Dim d) -> Elim d n
DApp : (fun : Elim d n) -> (arg : Dim d) -> (loc : Loc) -> Elim d n
||| type-annotated term
(:#) : (tm, ty : Term d n) -> Elim d n
Ann : (tm, ty : Term d n) -> (loc : Loc) -> Elim 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
(val : Term d n) -> (loc : Loc) -> 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
(zero, one : DScopeTerm d n) -> (loc : Loc) -> 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) ->
(loc : Loc) ->
Elim d n
||| term closure/suspended substitution
@ -244,88 +250,211 @@ mutual
||| scope which ignores all its binders
public export %inline
SN : {s : Nat} -> f n -> Scoped s f n
SN = S (replicate s Unused) . N
SN = S (replicate s $ BN Unused noLoc) . N
||| scope which uses its binders
public export %inline
SY : NContext s -> f (s + n) -> Scoped s f n
SY : BContext s -> f (s + n) -> Scoped s f n
SY ns = S ns . Y
public export %inline
name : Scoped 1 f n -> BaseName
name : Scoped 1 f n -> BindName
name (S [< x] _) = x
public export %inline
(.name) : Scoped 1 f n -> BaseName
(.name) : Scoped 1 f n -> BindName
s.name = name s
||| more convenient Pi
public export %inline
PiY : (qty : Qty) -> (x : BaseName) ->
(arg : Term d n) -> (res : Term d (S n)) -> Term d n
PiY {qty, x, arg, res} = Pi {qty, arg, res = SY [< x] res}
PiY : (qty : Qty) -> (x : BindName) ->
(arg : Term d n) -> (res : Term d (S n)) -> (loc : Loc) -> Term d n
PiY {qty, x, arg, res, loc} = Pi {qty, arg, res = SY [< x] res, loc}
||| more convenient Lam
public export %inline
LamY : (x : BindName) -> (body : Term d (S n)) -> (loc : Loc) -> Term d n
LamY {x, body, loc} = Lam {body = SY [< x] body, loc}
public export %inline
LamN : (body : Term d n) -> (loc : Loc) -> Term d n
LamN {body, loc} = Lam {body = SN body, loc}
||| non dependent function type
public export %inline
Arr : (qty : Qty) -> (arg, res : Term d n) -> Term d n
Arr {qty, arg, res} = Pi {qty, arg, res = SN res}
Arr : (qty : Qty) -> (arg, res : Term d n) -> (loc : Loc) -> Term d n
Arr {qty, arg, res, loc} = Pi {qty, arg, res = SN res, loc}
||| more convenient Sig
public export %inline
SigY : (x : BaseName) -> (fst : Term d n) ->
(snd : Term d (S n)) -> Term d n
SigY {x, fst, snd} = Sig {fst, snd = SY [< x] snd}
SigY : (x : BindName) -> (fst : Term d n) ->
(snd : Term d (S n)) -> (loc : Loc) -> Term d n
SigY {x, fst, snd, loc} = Sig {fst, snd = SY [< x] snd, loc}
||| non dependent pair type
public export %inline
And : (fst, snd : Term d n) -> Term d n
And {fst, snd} = Sig {fst, snd = SN snd}
And : (fst, snd : Term d n) -> (loc : Loc) -> Term d n
And {fst, snd, loc} = Sig {fst, snd = SN snd, loc}
||| more convenient Eq
public export %inline
EqY : (i : BaseName) -> (ty : Term (S d) n) ->
(l, r : Term d n) -> Term d n
EqY {i, ty, l, r} = Eq {ty = SY [< i] ty, l, r}
EqY : (i : BindName) -> (ty : Term (S d) n) ->
(l, r : Term d n) -> (loc : Loc) -> Term d n
EqY {i, ty, l, r, loc} = Eq {ty = SY [< i] ty, l, r, loc}
||| more convenient DLam
public export %inline
DLamY : (i : BindName) -> (body : Term (S d) n) -> (loc : Loc) -> Term d n
DLamY {i, body, loc} = DLam {body = SY [< i] body, loc}
public export %inline
DLamN : (body : Term d n) -> (loc : Loc) -> Term d n
DLamN {body, loc} = DLam {body = SN body, loc}
||| non dependent equality type
public export %inline
Eq0 : (ty, l, r : Term d n) -> Term d n
Eq0 {ty, l, r} = Eq {ty = SN ty, l, r}
Eq0 : (ty, l, r : Term d n) -> (loc : Loc) -> Term d n
Eq0 {ty, l, r, loc} = Eq {ty = SN ty, l, r, loc}
||| same as `F` but as a term
public export %inline
FT : Name -> Term d n
FT = E . F
FT : Name -> (loc : Loc) -> Term d n
FT x loc = E $ F x loc
||| abbreviation for a bound variable like `BV 4` instead of
||| `B (VS (VS (VS (VS VZ))))`
public export %inline
BV : (i : Nat) -> (0 _ : LT i n) => Elim d n
BV i = B $ V i
BV : (i : Nat) -> (0 _ : LT i n) => (loc : Loc) -> Elim d n
BV i loc = B (V i) loc
||| same as `BV` but as a term
public export %inline
BVT : (i : Nat) -> (0 _ : LT i n) => Term d n
BVT i = E $ BV i
BVT : (i : Nat) -> (0 _ : LT i n) => (loc : Loc) -> Term d n
BVT i loc = E $ BV i loc
public export
makeNat : Nat -> Term d n
makeNat 0 = Zero
makeNat (S k) = Succ $ makeNat k
makeNat : Nat -> Loc -> Term d n
makeNat 0 loc = Zero loc
makeNat (S k) loc = Succ (makeNat k loc) loc
public export %inline
enum : List TagVal -> Term d n
enum = Enum . SortedSet.fromList
enum : List TagVal -> Loc -> Term d n
enum ts loc = Enum (SortedSet.fromList ts) loc
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
List (TypeCaseArm d n) -> Term d n -> Loc -> Elim d n
typeCase ty ret arms def loc = TypeCase ty ret (fromList arms) def loc
public export %inline
typeCase1Y : Elim d n -> Term d n ->
(k : TyConKind) -> NContext (arity k) -> Term d (arity k + n) ->
{default Nat def : Term d n} ->
(k : TyConKind) -> BContext (arity k) -> Term d (arity k + n) ->
(loc : Loc) ->
{default (Nat loc) def : Term d n} ->
Elim d n
typeCase1Y ty ret k ns body {def} =
typeCase ty ret [(k ** SY ns body)] def
typeCase1Y ty ret k ns body loc = typeCase ty ret [(k ** SY ns body)] def loc
export
Located (Elim d n) where
(F _ loc).loc = loc
(B _ loc).loc = loc
(App _ _ loc).loc = loc
(CasePair _ _ _ _ loc).loc = loc
(CaseEnum _ _ _ _ loc).loc = loc
(CaseNat _ _ _ _ _ _ loc).loc = loc
(CaseBox _ _ _ _ loc).loc = loc
(DApp _ _ loc).loc = loc
(Ann _ _ loc).loc = loc
(Coe _ _ _ _ loc).loc = loc
(Comp _ _ _ _ _ _ _ loc).loc = loc
(TypeCase _ _ _ _ loc).loc = loc
(CloE (Sub e _)).loc = e.loc
(DCloE (Sub e _)).loc = e.loc
export
Located (Term d n) where
(TYPE _ loc).loc = loc
(Pi _ _ _ loc).loc = loc
(Lam _ loc).loc = loc
(Sig _ _ loc).loc = loc
(Pair _ _ loc).loc = loc
(Enum _ loc).loc = loc
(Tag _ loc).loc = loc
(Eq _ _ _ loc).loc = loc
(DLam _ loc).loc = loc
(Nat loc).loc = loc
(Zero loc).loc = loc
(Succ _ loc).loc = loc
(BOX _ _ loc).loc = loc
(Box _ loc).loc = loc
(E e).loc = e.loc
(CloT (Sub t _)).loc = t.loc
(DCloT (Sub t _)).loc = t.loc
export
Located1 f => Located (ScopedBody s f n) where
(Y t).loc = t.loc
(N t).loc = t.loc
export
Located1 f => Located (Scoped s f n) where
t.loc = t.body.loc
export
Relocatable (Elim d n) where
setLoc loc (F x _) = F x loc
setLoc loc (B i _) = B i loc
setLoc loc (App fun arg _) = App fun arg loc
setLoc loc (CasePair qty pair ret body _) =
CasePair qty pair ret body loc
setLoc loc (CaseEnum qty tag ret arms _) =
CaseEnum qty tag ret arms loc
setLoc loc (CaseNat qty qtyIH nat ret zero succ _) =
CaseNat qty qtyIH nat ret zero succ loc
setLoc loc (CaseBox qty box ret body _) =
CaseBox qty box ret body loc
setLoc loc (DApp fun arg _) =
DApp fun arg loc
setLoc loc (Ann tm ty _) =
Ann tm ty loc
setLoc loc (Coe ty p q val _) =
Coe ty p q val loc
setLoc loc (Comp ty p q val r zero one _) =
Comp ty p q val r zero one loc
setLoc loc (TypeCase ty ret arms def _) =
TypeCase ty ret arms def loc
setLoc loc (CloE (Sub term subst)) =
CloE $ Sub (setLoc loc term) subst
setLoc loc (DCloE (Sub term subst)) =
DCloE $ Sub (setLoc loc term) subst
export
Relocatable (Term d n) where
setLoc loc (TYPE l _) = TYPE l loc
setLoc loc (Pi qty arg res _) = Pi qty arg res loc
setLoc loc (Lam body _) = Lam body loc
setLoc loc (Sig fst snd _) = Sig fst snd loc
setLoc loc (Pair fst snd _) = Pair fst snd loc
setLoc loc (Enum cases _) = Enum cases loc
setLoc loc (Tag tag _) = Tag tag loc
setLoc loc (Eq ty l r _) = Eq ty l r loc
setLoc loc (DLam body _) = DLam body loc
setLoc loc (Nat _) = Nat loc
setLoc loc (Zero _) = Zero loc
setLoc loc (Succ p _) = Succ p loc
setLoc loc (BOX qty ty _) = BOX qty ty loc
setLoc loc (Box val _) = Box val loc
setLoc loc (E e) = E $ setLoc loc e
setLoc loc (CloT (Sub term subst)) = CloT $ Sub (setLoc loc term) subst
setLoc loc (DCloT (Sub term subst)) = DCloT $ Sub (setLoc loc term) subst
export
Relocatable1 f => Relocatable (ScopedBody s f n) where
setLoc loc (Y body) = Y $ setLoc loc body
setLoc loc (N body) = N $ setLoc loc body
export
Relocatable1 f => Relocatable (Scoped s f n) where
setLoc loc (S names body) = S (setLoc loc <$> names) (setLoc loc body)

View file

@ -82,8 +82,8 @@ PrettyHL a => PrettyHL (Binder a) where
export
prettyBindType : PrettyHL a => PrettyHL b =>
Pretty.HasEnv m =>
Maybe Qty -> BaseName -> a -> Doc HL -> b -> m (Doc HL)
prettyBindType q x s arr t = do
Maybe Qty -> BindName -> a -> Doc HL -> b -> m (Doc HL)
prettyBindType q (BN x _) s arr t = do
bind <- case q of
Nothing => pretty0M $ MkBinder x s
Just q => pretty0M $ MkWithQty q $ MkBinder x s
@ -92,14 +92,15 @@ prettyBindType q x s arr t = do
export
prettyArm : PrettyHL a => Pretty.HasEnv m =>
BinderSort -> SnocList BaseName -> Doc HL -> a -> m (Doc HL)
BinderSort -> SnocList BindName -> Doc HL -> a -> m (Doc HL)
prettyArm sort xs pat body = do
let xs = map name xs
body <- withPrec Outer $ unders sort xs $ prettyM body
pure $ hang 2 $ sep [pat <++> !darrowD, body]
export
prettyLams : PrettyHL a => Pretty.HasEnv m =>
Maybe (Doc HL) -> BinderSort -> SnocList BaseName -> a ->
Maybe (Doc HL) -> BinderSort -> SnocList BindName -> a ->
m (Doc HL)
prettyLams lam sort names body = do
let var = case sort of T => TVar; D => DVar
@ -109,14 +110,15 @@ prettyLams lam sort names body = do
public export
data TypeLine a = MkTypeLine BaseName a
data TypeLine a = MkTypeLine BindName a
export
PrettyHL a => PrettyHL (TypeLine a) where
prettyM (MkTypeLine Unused ty) =
bracks <$> pretty0M ty
prettyM (MkTypeLine i ty) =
map bracks $ withPrec Outer $ prettyLams Nothing D [< i] ty
if i.name == Unused then
bracks <$> pretty0M ty
else
map bracks $ withPrec Outer $ prettyLams Nothing D [< i] ty
export
@ -142,28 +144,28 @@ prettyTuple = map (parens . align . separate commaD) . traverse prettyM
export
prettyArms : PrettyHL a => Pretty.HasEnv m =>
BinderSort -> List (SnocList BaseName, Doc HL, a) -> m (Doc HL)
BinderSort -> List (SnocList BindName, Doc HL, a) -> m (Doc HL)
prettyArms s =
map (braces . aseparate semiD) .
traverse (\(xs, l, r) => prettyArm s xs l r)
export
prettyCase' : (PrettyHL a, PrettyHL b, PrettyHL c, Pretty.HasEnv m) =>
Doc HL -> a -> BaseName -> b ->
List (SnocList BaseName, Doc HL, c) ->
Doc HL -> a -> BindName -> b ->
List (SnocList BindName, Doc HL, c) ->
m (Doc HL)
prettyCase' intro elim r ret arms = do
elim <- pretty0M elim
ret <- case r of
Unused => under T r $ pretty0M ret
ret <- case r.name of
Unused => under T r.name $ pretty0M ret
_ => prettyLams Nothing T [< r] ret
arms <- prettyArms T arms
pure $ asep [intro <++> elim, returnD <++> ret, ofD <++> arms]
export
prettyCase : (PrettyHL a, PrettyHL b, PrettyHL c, Pretty.HasEnv m) =>
Qty -> a -> BaseName -> b ->
List (SnocList BaseName, Doc HL, c) ->
Qty -> a -> BindName -> b ->
List (SnocList BindName, Doc HL, c) ->
m (Doc HL)
prettyCase pi elim r ret arms = do
caseq <- (caseD <+>) <$> prettySuffix pi
@ -197,14 +199,14 @@ prettyBoxVal : PrettyHL a => Pretty.HasEnv m => a -> m (Doc HL)
prettyBoxVal val = bracks <$> pretty0M val
export
prettyCompPat : Pretty.HasEnv m => DimConst -> BaseName -> m (Doc HL)
prettyCompPat e j = hsep <$> sequence [pretty0M e, pretty0M $ DV j]
prettyCompPat : Pretty.HasEnv m => DimConst -> BindName -> m (Doc HL)
prettyCompPat e j = hsep <$> sequence [pretty0M e, pretty0M $ DV j.name]
export
toNatLit : Term d n -> Maybe Nat
toNatLit Zero = Just 0
toNatLit (Succ n) = [|S $ toNatLit n|]
toNatLit _ = Nothing
toNatLit (Zero _) = Just 0
toNatLit (Succ n _) = [|S $ toNatLit n|]
toNatLit _ = Nothing
private
eterm : Term d n -> Exists (Term d)
@ -216,69 +218,69 @@ parameters (showSubsts : Bool)
export covering
[TermSubst] PrettyHL (Term d n) using ElimSubst
where
prettyM (TYPE l) =
prettyM (TYPE l _) =
pure $ !typeD <+> hl Syntax !(prettyUnivSuffix l)
prettyM (Pi qty s (S _ (N t))) = do
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))) =
prettyM (Pi qty s (S [< x] (Y t)) _) =
prettyBindType (Just qty) x s !arrowD t
prettyM (Lam (S x 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
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))) =
prettyM (Sig s (S [< x] (Y t)) _) =
prettyBindType Nothing x s !timesD t
prettyM (Pair s t) =
prettyM (Pair s t _) =
let GotPairs {init, last, _} = getPairs' [< s] t in
prettyTuple $ toList $ init :< last
prettyM (Enum tags) =
prettyM (Enum tags _) =
pure $ delims "{" "}" . aseparate comma $ map prettyTagBare $
Prelude.toList tags
prettyM (Tag t) =
prettyM (Tag t _) =
pure $ prettyTag t
prettyM (Eq (S _ (N ty)) l r) = do
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
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)) =
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 (Nat _) = natD
prettyM Zero = pure $ hl Syntax "0"
prettyM (Zero _) = pure $ hl Syntax "0"
prettyM (Succ n) =
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
prettyM (BOX pi ty _) = do
pi <- pretty0M pi
ty <- pretty0M ty
pure $ bracks $ hcat [pi, dotD, align ty]
prettyM (Box val) = prettyBoxVal val
prettyM (Box val _) = prettyBoxVal val
prettyM (E e) = prettyM e
@ -299,49 +301,49 @@ parameters (showSubsts : Bool)
export covering
[ElimSubst] PrettyHL (Elim d n) using TermSubst
where
prettyM (F x) =
prettyM (F x _) =
hl' Free <$> prettyM x
prettyM (B i) =
prettyM (B i _) =
prettyVar TVar TVarErr (!ask).tnames i
prettyM (e :@ s) =
prettyM (App 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
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) =
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)) =
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)
(BN 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)) =
prettyM (CaseBox pi box (S [< r] ret) (S [< u] body) _) =
prettyCase pi box r ret.term
[([< u], !(prettyBoxVal $ TV u), body.term)]
[([< u], !(prettyBoxVal $ TV u.name), body.term)]
prettyM (e :% d) =
prettyM (DApp e d _) =
let GotDArgs {fun, args, _} = getDArgs' e [d] in
prettyApps (Just "@") fun args
prettyM (s :# a) = do
prettyM (Ann s a _) = do
s <- withPrec AnnL $ prettyM s
a <- withPrec AnnR $ prettyM a
parensIfM AnnR $ hang 2 $ s <++> !annD <%%> a
prettyM (Coe (S [< i] ty) p q val) =
prettyM (Coe (S [< i] ty) p q val _) =
let ty = case ty of
Y ty => epretty $ MkTypeLine i ty
N ty => epretty ty
@ -352,9 +354,9 @@ parameters (showSubsts : Bool)
(Just "@", epretty q),
(Nothing, epretty val)]
prettyM (Comp ty p q val r (S [< z] zero) (S [< o] one)) = do
prettyM (Comp ty p q val r (S [< z] zero) (S [< o] one) _) = do
apps <- prettyApps' (L compD)
[(Nothing, epretty $ MkTypeLine Unused ty),
[(Nothing, epretty $ MkTypeLine (BN Unused noLoc) ty),
(Just "@", epretty p),
(Just "@", epretty q),
(Nothing, epretty val),
@ -364,29 +366,30 @@ parameters (showSubsts : Bool)
([< o], !(prettyCompPat One o), one.term)]
pure $ apps <++> arms
prettyM (TypeCase ty ret arms def) = do
prettyM (TypeCase ty ret arms def _) = do
arms <- traverse fromArm (toList arms)
prettyCase' typecaseD ty Unused ret $
prettyCase' typecaseD ty (BN Unused noLoc) ret $
arms ++ [([<], hl Syntax "_", eterm def)]
where
v : BaseName -> Doc HL
v = pretty0 True . TV
v : BindName -> Doc HL
v = pretty0 True . TV . name
tyCasePat : (k : TyConKind) -> NContext (arity k) -> m (Doc HL)
tyCasePat : (k : TyConKind) -> BContext (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
tyCasePat KEq vars =
prettyApps Nothing (L eqD) $ map (TV . name) $ toList' vars
fromArm : TypeCaseArm d n ->
m (SnocList BaseName, Doc HL, Exists (Term d))
m (SnocList BindName, Doc HL, Exists (Term d))
fromArm (k ** S ns t) =
pure (toSnocList' ns, !(tyCasePat k ns), eterm t.term)
prettyM (CloE (Sub e th)) =
prettyM (CloE (Sub e th)) =
if showSubsts then
parensIfM SApp . hang 2 =<<
[|withPrec SApp (prettyM e) <%> prettyTSubst th|]
@ -414,7 +417,7 @@ PrettyHL (Elim d n) where prettyM = prettyM @{ElimSubst False}
export covering
prettyTerm : (unicode : Bool) ->
(dnames : NContext d) -> (tnames : NContext n) ->
(dnames : BContext d) -> (tnames : BContext n) ->
Term d n -> Doc HL
prettyTerm unicode dnames tnames term =
pretty0With unicode (toSnocList' dnames) (toSnocList' tnames) term
pretty0With unicode (toNames dnames) (toNames tnames) term

View file

@ -13,8 +13,8 @@ import public Data.Vect
public export %inline
isLam : Term {} -> Bool
isLam (Lam _) = True
isLam _ = False
isLam (Lam {}) = True
isLam _ = False
public export
0 NotLam : Pred $ Term {}
@ -23,8 +23,8 @@ NotLam = No . isLam
public export %inline
isDLam : Term {} -> Bool
isDLam (DLam _) = True
isDLam _ = False
isDLam (DLam {}) = True
isDLam _ = False
public export
0 NotDLam : Pred $ Term {}
@ -43,7 +43,7 @@ NotPair = No . isPair
public export %inline
isApp : Elim {} -> Bool
isApp (_ :@ _) = True
isApp (App {}) = True
isApp _ = False
public export
@ -53,19 +53,21 @@ NotApp = No . isApp
public export %inline
isDApp : Elim {} -> Bool
isDApp (_ :% _) = True
isDApp _ = False
isDApp (DApp {}) = True
isDApp _ = False
public export
0 NotDApp : Pred $ Elim {}
NotDApp = No . isDApp
infixl 9 :@@
||| apply multiple arguments at once
public export %inline
(:@@) : Elim d n -> List (Term d n) -> Elim d n
f :@@ ss = foldl (:@) f ss
-- infixl 9 :@@
-- ||| apply multiple arguments at once
-- public export %inline
-- (:@@) : Elim d n -> List (Term d n) -> Elim d n
-- f :@@ ss = foldl app f ss where
-- app : Elim d n -> Term d n -> Elim d n
-- app f s = App f s (f.loc `extend'` s.loc.bounds)
public export
record GetArgs d n where
@ -85,7 +87,7 @@ mutual
getArgsNc : (e : Elim d n) -> (0 nc : NotClo e) =>
List (Term d n) -> GetArgs d n
getArgsNc fun args = case nchoose $ isApp fun of
Left y => let f :@ a = fun in getArgs' f (a :: args)
Left y => let App f a _ = fun in getArgs' f (a :: args)
Right n => GotArgs {fun, args, notApp = n}
||| splits an application into its head and arguments. if it's not an
@ -96,11 +98,13 @@ getArgs : Elim d n -> GetArgs d n
getArgs e = getArgs' e []
infixl 9 :%%
||| apply multiple dimension arguments at once
public export %inline
(:%%) : Elim d n -> List (Dim d) -> Elim d n
f :%% ss = foldl (:%) f ss
-- infixl 9 :%%
-- ||| apply multiple dimension arguments at once
-- public export %inline
-- (:%%) : Elim d n -> List (Dim d) -> Elim d n
-- f :%% ss = foldl dapp f ss where
-- dapp : Elim d n -> Dim d -> Elim d n
-- dapp f p = DApp f p (f.loc `extend'` p.loc.bounds)
public export
record GetDArgs d n where
@ -120,7 +124,7 @@ mutual
getDArgsNc : (e : Elim d n) -> (0 nc : NotClo e) =>
List (Dim d) -> GetDArgs d n
getDArgsNc fun args = case nchoose $ isDApp fun of
Left y => let f :% d = fun in getDArgs' f (d :: args)
Left y => let DApp f d _ = fun in getDArgs' f (d :: args)
Right n => GotDArgs {fun, args, notDApp = n}
||| splits a dimension application into its head and arguments. if it's not an
@ -130,33 +134,33 @@ getDArgs : Elim d n -> GetDArgs d n
getDArgs e = getDArgs' e []
infixr 1 :\\
public export
absN : NContext m -> Term d (m + n) -> Term d n
absN [<] s = s
absN (xs :< x) s = absN xs $ Lam $ ST [< x] s
-- infixr 1 :\\
-- public export
-- absN : BContext m -> Term d (m + n) -> Term d n
-- absN [<] s = s
-- absN (xs :< x) s = absN xs $ Lam (ST [< x] s) ?absloc
public export %inline
(:\\) : NContext m -> Term d (m + n) -> Term d n
(:\\) = absN
-- public export %inline
-- (:\\) : BContext m -> Term d (m + n) -> Term d n
-- (:\\) = absN
infixr 1 :\\%
public export
dabsN : NContext m -> Term (m + d) n -> Term d n
dabsN [<] s = s
dabsN (xs :< x) s = dabsN xs $ DLam $ DST [< x] s
-- infixr 1 :\\%
-- public export
-- dabsN : BContext m -> Term (m + d) n -> Term d n
-- dabsN [<] s = s
-- dabsN (xs :< x) s = dabsN xs $ DLam (DST [< x] s) ?dabsLoc
public export %inline
(:\\%) : NContext m -> Term (m + d) n -> Term d n
(:\\%) = dabsN
-- public export %inline
-- (:\\%) : BContext m -> Term (m + d) n -> Term d n
-- (:\\%) = dabsN
public export
record GetLams d n where
constructor GotLams
{0 lams, rest : Nat}
names : NContext lams
names : BContext lams
body : Term d rest
0 eq : lams + n = rest
0 notLam : NotLam body
@ -164,7 +168,7 @@ record GetLams d n where
mutual
export %inline
getLams' : forall lams, rest.
NContext lams -> Term d rest -> (0 eq : lams + n = rest) ->
BContext lams -> Term d rest -> (0 eq : lams + n = rest) ->
GetLams d n
getLams' xs s0 eq =
let Element s nc = pushSubsts s0 in
@ -172,12 +176,12 @@ mutual
private
getLamsNc : forall lams, rest.
NContext lams ->
BContext lams ->
(t : Term d rest) -> (0 nc : NotClo t) =>
(0 eq : lams + n = rest) ->
GetLams d n
getLamsNc xs s Refl = case nchoose $ isLam s of
Left y => let Lam (S [< x] body) = s in
Left y => let Lam (S [< x] body) _ = s in
getLams' (xs :< x) (assert_smaller s body.term) Refl
Right n => GotLams xs s Refl n
@ -190,7 +194,7 @@ public export
record GetDLams d n where
constructor GotDLams
{0 lams, rest : Nat}
names : NContext lams
names : BContext lams
body : Term rest n
0 eq : lams + d = rest
0 notDLam : NotDLam body
@ -198,7 +202,7 @@ record GetDLams d n where
mutual
export %inline
getDLams' : forall lams, rest.
NContext lams -> Term rest n -> (0 eq : lams + d = rest) ->
BContext lams -> Term rest n -> (0 eq : lams + d = rest) ->
GetDLams d n
getDLams' xs s0 eq =
let Element s nc = pushSubsts s0 in
@ -206,12 +210,12 @@ mutual
private
getDLamsNc : forall lams, rest.
NContext lams ->
BContext lams ->
(t : Term rest n) -> (0 nc : NotClo t) =>
(0 eq : lams + d = rest) ->
GetDLams d n
getDLamsNc is s Refl = case nchoose $ isDLam s of
Left y => let DLam (S [< i] body) = s in
Left y => let DLam (S [< i] body) _ = s in
getDLams' (is :< i) (assert_smaller s body.term) Refl
Right n => GotDLams is s Refl n
@ -238,7 +242,7 @@ mutual
(t : Term d n) -> (0 nc : NotClo t) =>
GetPairs d n
getPairsNc ss t = case nchoose $ isPair t of
Left y => let Pair s t = t in
Left y => let Pair s t _ = t in
getPairs' (ss :< s) t
Right n => GotPairs ss t n

View file

@ -20,14 +20,14 @@ namespace CanDSubst
export
CanDSubst Term where
s // Shift SZ = s
TYPE l // _ = TYPE l
TYPE l loc // _ = TYPE l loc
DCloT (Sub s ph) // th = DCloT $ Sub s $ ph . th
s // th = DCloT $ Sub s th
private
subDArgs : Elim dfrom n -> DSubst dfrom dto -> Elim dto n
subDArgs (f :% d) th = subDArgs f th :% (d // th)
subDArgs e th = DCloE $ Sub e th
subDArgs (DApp f d loc) th = DApp (subDArgs f th) (d // th) loc
subDArgs e th = DCloE $ Sub e th
||| does the minimal reasonable work:
||| - deletes the closure around a term variable
@ -39,9 +39,9 @@ subDArgs e th = DCloE $ Sub e th
export
CanDSubst Elim where
e // Shift SZ = e
F x // _ = F x
B i // _ = B i
f :% d // th = subDArgs (f :% d) th
F x loc // _ = F x loc
B i loc // _ = B i loc
e@(DApp {}) // th = subDArgs e th
DCloE (Sub e ph) // th = DCloE $ Sub e $ ph . th
e // th = DCloE $ Sub e th
@ -61,8 +61,8 @@ namespace DSubst.DScopeTermN
S ns (N body) // th = S ns $ N $ body // th
export %inline FromVar (Elim d) where fromVar = B
export %inline FromVar (Term d) where fromVar = E . fromVar
export %inline FromVar (Elim d) where fromVarLoc = B
export %inline FromVar (Term d) where fromVarLoc = E .: fromVar
||| does the minimal reasonable work:
@ -73,8 +73,8 @@ export %inline FromVar (Term d) where fromVar = E . fromVar
||| - otherwise, wraps in a new closure
export
CanSubstSelf (Elim d) where
F x // _ = F x
B i // th = th !! i
F x loc // _ = F x loc
B i loc // th = getLoc th i loc
CloE (Sub e ph) // th = assert_total CloE $ Sub e $ ph . th
e // th = case force th of
Shift SZ => e
@ -93,7 +93,7 @@ namespace CanTSubst
||| - otherwise, wraps in a new closure
export
CanTSubst Term where
TYPE l // _ = TYPE l
TYPE l loc // _ = TYPE l loc
E e // th = E $ e // th
CloT (Sub s ph) // th = CloT $ Sub s $ ph . th
s // th = case force th of
@ -192,12 +192,12 @@ dsub1 t p = dsubN t [< p]
public export %inline
(.zero) : DScopeTerm d n -> Term d n
body.zero = dsub1 body $ K Zero
(.zero) : DScopeTerm d n -> {default noLoc loc : Loc} -> Term d n
body.zero = dsub1 body $ K Zero loc
public export %inline
(.one) : DScopeTerm d n -> Term d n
body.one = dsub1 body $ K One
(.one) : DScopeTerm d n -> {default noLoc loc : Loc} -> Term d n
body.one = dsub1 body $ K One loc
public export
@ -251,29 +251,34 @@ mutual
mutual
export
PushSubsts Term Subst.isCloT where
pushSubstsWith th ph (TYPE l) =
nclo $ TYPE l
pushSubstsWith th ph (Pi qty a body) =
nclo $ Pi qty (a // th // ph) (body // th // ph)
pushSubstsWith th ph (Lam body) =
nclo $ Lam (body // th // ph)
pushSubstsWith th ph (Sig a b) =
nclo $ Sig (a // th // ph) (b // th // ph)
pushSubstsWith th ph (Pair s t) =
nclo $ Pair (s // th // ph) (t // th // ph)
pushSubstsWith th ph (Enum tags) =
nclo $ Enum tags
pushSubstsWith th ph (Tag tag) =
nclo $ Tag tag
pushSubstsWith th ph (Eq ty l r) =
nclo $ Eq (ty // th // ph) (l // th // ph) (r // th // ph)
pushSubstsWith th ph (DLam body) =
nclo $ DLam (body // th // ph)
pushSubstsWith _ _ Nat = nclo Nat
pushSubstsWith _ _ Zero = nclo Zero
pushSubstsWith th ph (Succ n) = nclo $ Succ $ n // th // ph
pushSubstsWith th ph (BOX pi ty) = nclo $ BOX pi $ ty // th // ph
pushSubstsWith th ph (Box val) = nclo $ Box $ val // th // ph
pushSubstsWith th ph (TYPE l loc) =
nclo $ TYPE l loc
pushSubstsWith th ph (Pi qty a body loc) =
nclo $ Pi qty (a // th // ph) (body // th // ph) loc
pushSubstsWith th ph (Lam body loc) =
nclo $ Lam (body // th // ph) loc
pushSubstsWith th ph (Sig a b loc) =
nclo $ Sig (a // th // ph) (b // th // ph) loc
pushSubstsWith th ph (Pair s t loc) =
nclo $ Pair (s // th // ph) (t // th // ph) loc
pushSubstsWith th ph (Enum tags loc) =
nclo $ Enum tags loc
pushSubstsWith th ph (Tag tag loc) =
nclo $ Tag tag loc
pushSubstsWith th ph (Eq ty l r loc) =
nclo $ Eq (ty // th // ph) (l // th // ph) (r // th // ph) loc
pushSubstsWith th ph (DLam body loc) =
nclo $ DLam (body // th // ph) loc
pushSubstsWith _ _ (Nat loc) =
nclo $ Nat loc
pushSubstsWith _ _ (Zero loc) =
nclo $ Zero loc
pushSubstsWith th ph (Succ n loc) =
nclo $ Succ (n // th // ph) loc
pushSubstsWith th ph (BOX pi ty loc) =
nclo $ BOX pi (ty // th // ph) loc
pushSubstsWith th ph (Box val loc) =
nclo $ Box (val // th // ph) loc
pushSubstsWith th ph (E e) =
let Element e nc = pushSubstsWith th ph e in nclo $ E e
pushSubstsWith th ph (CloT (Sub s ps)) =
@ -283,38 +288,38 @@ mutual
export
PushSubsts Elim Subst.isCloE where
pushSubstsWith th ph (F x) =
nclo $ F x
pushSubstsWith th ph (B i) =
let res = ph !! i in
pushSubstsWith th ph (F x loc) =
nclo $ F x loc
pushSubstsWith th ph (B i loc) =
let res = getLoc ph i loc in
case nchoose $ isCloE res of
Left yes => assert_total pushSubsts res
Right no => Element res no
pushSubstsWith th ph (f :@ s) =
nclo $ (f // th // ph) :@ (s // th // ph)
pushSubstsWith th ph (CasePair pi p r b) =
nclo $ CasePair pi (p // th // ph) (r // th // ph) (b // th // ph)
pushSubstsWith th ph (CaseEnum pi t r arms) =
pushSubstsWith th ph (App f s loc) =
nclo $ App (f // th // ph) (s // th // ph) loc
pushSubstsWith th ph (CasePair pi p r b loc) =
nclo $ CasePair pi (p // th // ph) (r // th // ph) (b // th // ph) loc
pushSubstsWith th ph (CaseEnum pi t r arms loc) =
nclo $ CaseEnum pi (t // th // ph) (r // th // ph)
(map (\b => b // th // ph) arms)
pushSubstsWith th ph (CaseNat pi pi' n r z s) =
(map (\b => b // th // ph) arms) loc
pushSubstsWith th ph (CaseNat pi pi' n r z s loc) =
nclo $ CaseNat pi pi' (n // th // ph) (r // th // ph)
(z // th // ph) (s // th // ph)
pushSubstsWith th ph (CaseBox pi x r b) =
nclo $ CaseBox pi (x // th // ph) (r // th // ph) (b // th // ph)
pushSubstsWith th ph (f :% d) =
nclo $ (f // th // ph) :% (d // th)
pushSubstsWith th ph (s :# a) =
nclo $ (s // th // ph) :# (a // 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) =
(z // th // ph) (s // th // ph) loc
pushSubstsWith th ph (CaseBox pi x r b loc) =
nclo $ CaseBox pi (x // th // ph) (r // th // ph) (b // th // ph) loc
pushSubstsWith th ph (DApp f d loc) =
nclo $ DApp (f // th // ph) (d // th) loc
pushSubstsWith th ph (Ann s a loc) =
nclo $ Ann (s // th // ph) (a // th // ph) loc
pushSubstsWith th ph (Coe ty p q val loc) =
nclo $ Coe (ty // th // ph) (p // th) (q // th) (val // th // ph) loc
pushSubstsWith th ph (Comp ty p q val r zero one loc) =
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) =
(zero // th // ph) (one // th // ph) loc
pushSubstsWith th ph (TypeCase ty ret arms def loc) =
nclo $ TypeCase (ty // th // ph) (ret // th // ph)
(map (\t => t // th // ph) arms) (def // th // ph)
(map (\t => t // th // ph) arms) (def // th // ph) loc
pushSubstsWith th ph (CloE (Sub e ps)) =
pushSubstsWith th (comp th ps ph) e
pushSubstsWith th ph (DCloE (Sub e ps)) =
@ -323,14 +328,19 @@ mutual
private %inline
CompHY : (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) ->
(r : Dim d) -> (zero, one : DScopeTerm d n) -> Elim d n
CompHY {ty, p, q, val, r, zero, one} =
let ty' = SY ty.names $ ty.term // (B VZ ::: shift 2) in
(r : Dim d) -> (zero, one : DScopeTerm d n) -> (loc : Loc) -> Elim d n
CompHY {ty, p, q, val, r, zero, one, loc} =
-- [fixme] maintain location of existing B VZ
let ty' = SY ty.names $ ty.term // (B VZ noLoc ::: 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
val = E $ Coe ty p q val val.loc, r,
-- [fixme] better locations for these vars?
zero = SY zero.names $ E $
Coe ty' (B VZ zero.loc) (weakD 1 q) zero.term zero.loc,
one = SY one.names $ E $
Coe ty' (B VZ one.loc) (weakD 1 q) one.term one.loc,
loc
}
public export %inline
@ -338,26 +348,28 @@ CompH' : (ty : DScopeTerm d n) ->
(p, q : Dim d) -> (val : Term d n) -> (r : Dim d) ->
(zero : DScopeTerm d n) ->
(one : DScopeTerm d n) ->
(loc : Loc) ->
Elim d n
CompH' {ty, p, q, val, r, zero, one} =
CompH' {ty, p, q, val, r, zero, one, loc} =
case dsqueeze ty of
S _ (N ty) => Comp {ty, p, q, val, r, zero, one}
S _ (Y _) => CompHY {ty, p, q, val, r, zero, one}
S _ (N ty) => Comp {ty, p, q, val, r, zero, one, loc}
S _ (Y _) => CompHY {ty, p, q, val, r, zero, one, loc}
||| heterogeneous composition, using Comp and Coe (and subst)
|||
||| comp [i ⇒ A] @p @q s { (r=0) j ⇒ t₀; (r=1) j ⇒ t₁ }
||| comp [i ⇒ A] @p @q s @r { 0 j ⇒ t₀; 1 j ⇒ t₁ }
||| ≔
||| comp [Aq/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₁
||| comp [Aq/i] @p @q (coe [i ⇒ A] @p @q s) @r {
||| 0 j ⇒ coe [i ⇒ A] @j @q t₀;
||| 1 j ⇒ coe [i ⇒ A] @j @q t₁
||| }
public export %inline
CompH : (i : BaseName) -> (ty : Term (S d) n) ->
CompH : (i : BindName) -> (ty : Term (S d) n) ->
(p, q : Dim d) -> (val : Term d n) -> (r : Dim d) ->
(j0 : BaseName) -> (zero : Term (S d) n) ->
(j1 : BaseName) -> (one : Term (S d) n) ->
(j0 : BindName) -> (zero : Term (S d) n) ->
(j1 : BindName) -> (one : Term (S d) n) ->
(loc : Loc) ->
Elim d n
CompH {i, ty, p, q, val, r, j0, zero, j1, one} =
CompH {i, ty, p, q, val, r, j0, zero, j1, one, loc} =
CompH' {ty = SY [< i] ty, p, q, val, r,
zero = SY [< j0] zero, one = SY [< j0] one}
zero = SY [< j0] zero, one = SY [< j0] one, loc}

View file

@ -18,6 +18,12 @@ Tighten (Shift from) where
tighten (Keep p) (SS by) = [|SS $ tighten p by|]
export
Tighten Dim where
tighten p (K e loc) = pure $ K e loc
tighten p (B i loc) = B <$> tighten p i <*> pure loc
export
tightenSub : (forall m, n. OPE m n -> env n -> Maybe (env m)) ->
OPE to1 to2 -> Subst env from to2 -> Maybe (Subst env from to1)
@ -46,23 +52,35 @@ tightenDScope f p (S names (N body)) = S names . N <$> f p body
mutual
private
tightenT : OPE n1 n2 -> Term d n2 -> Maybe (Term d n1)
tightenT p (TYPE l) = pure $ TYPE l
tightenT p (Pi qty arg res) =
Pi qty <$> tightenT p arg <*> tightenS p res
tightenT p (Lam body) = Lam <$> tightenS p body
tightenT p (Sig fst snd) = Sig <$> tightenT p fst <*> tightenS p snd
tightenT p (Pair fst snd) = Pair <$> tightenT p fst <*> tightenT p snd
tightenT p (Enum cases) = pure $ Enum cases
tightenT p (Tag tag) = pure $ Tag tag
tightenT p (Eq ty l r) =
Eq <$> tightenDS p ty <*> tightenT p l <*> tightenT p r
tightenT p (DLam body) = DLam <$> tightenDS p body
tightenT p Nat = pure Nat
tightenT p Zero = pure Zero
tightenT p (Succ s) = Succ <$> tightenT p s
tightenT p (BOX qty ty) = BOX qty <$> tightenT p ty
tightenT p (Box val) = Box <$> tightenT p val
tightenT p (E e) = assert_total $ E <$> tightenE p e
tightenT p (TYPE l loc) = pure $ TYPE l loc
tightenT p (Pi qty arg res loc) =
Pi qty <$> tightenT p arg <*> tightenS p res <*> pure loc
tightenT p (Lam body loc) =
Lam <$> tightenS p body <*> pure loc
tightenT p (Sig fst snd loc) =
Sig <$> tightenT p fst <*> tightenS p snd <*> pure loc
tightenT p (Pair fst snd loc) =
Pair <$> tightenT p fst <*> tightenT p snd <*> pure loc
tightenT p (Enum cases loc) =
pure $ Enum cases loc
tightenT p (Tag tag loc) =
pure $ Tag tag loc
tightenT p (Eq ty l r loc) =
Eq <$> tightenDS p ty <*> tightenT p l <*> tightenT p r <*> pure loc
tightenT p (DLam body loc) =
DLam <$> tightenDS p body <*> pure loc
tightenT p (Nat loc) =
pure $ Nat loc
tightenT p (Zero loc) =
pure $ Zero loc
tightenT p (Succ s loc) =
Succ <$> tightenT p s <*> pure loc
tightenT p (BOX qty ty loc) =
BOX qty <$> tightenT p ty <*> pure loc
tightenT p (Box val loc) =
Box <$> tightenT p val <*> pure loc
tightenT p (E e) =
assert_total $ E <$> tightenE p e
tightenT p (CloT (Sub tm th)) = do
th <- assert_total $ tightenSub tightenE p th
pure $ CloT $ Sub tm th
@ -72,45 +90,57 @@ mutual
private
tightenE : OPE n1 n2 -> Elim d n2 -> Maybe (Elim d n1)
tightenE p (F x) = pure $ F x
tightenE p (B i) = [|B $ tighten p i|]
tightenE p (fun :@ arg) = [|tightenE p fun :@ tightenT p arg|]
tightenE p (CasePair qty pair ret body) =
tightenE p (F x loc) =
pure $ F x loc
tightenE p (B i loc) =
B <$> tighten p i <*> pure loc
tightenE p (App fun arg loc) =
App <$> tightenE p fun <*> tightenT p arg <*> pure loc
tightenE p (CasePair qty pair ret body loc) =
CasePair qty <$> tightenE p pair
<*> tightenS p ret
<*> tightenS p body
tightenE p (CaseEnum qty tag ret arms) =
<*> pure loc
tightenE p (CaseEnum qty tag ret arms loc) =
CaseEnum qty <$> tightenE p tag
<*> tightenS p ret
<*> traverse (tightenT p) arms
tightenE p (CaseNat qty qtyIH nat ret zero succ) =
<*> pure loc
tightenE p (CaseNat qty qtyIH nat ret zero succ loc) =
CaseNat qty qtyIH
<$> tightenE p nat
<*> tightenS p ret
<*> tightenT p zero
<*> tightenS p succ
tightenE p (CaseBox qty box ret body) =
<*> pure loc
tightenE p (CaseBox qty box ret body loc) =
CaseBox qty <$> tightenE p box
<*> tightenS p ret
<*> tightenS p body
tightenE p (fun :% arg) = (:% arg) <$> tightenE p fun
tightenE p (tm :# ty) = [|tightenT p tm :# tightenT p ty|]
tightenE p (Coe ty q0 q1 val) =
<*> pure loc
tightenE p (DApp fun arg loc) =
DApp <$> tightenE p fun <*> pure arg <*> pure loc
tightenE p (Ann tm ty loc) =
Ann <$> tightenT p tm <*> tightenT p ty <*> pure loc
tightenE p (Coe ty q0 q1 val loc) =
Coe <$> tightenDS p ty
<*> pure q0 <*> pure q1
<*> tightenT p val
tightenE p (Comp ty q0 q1 val r zero one) =
<*> pure loc
tightenE p (Comp ty q0 q1 val r zero one loc) =
Comp <$> tightenT p ty
<*> pure q0 <*> pure q1
<*> tightenT p val
<*> pure r
<*> tightenDS p zero
<*> tightenDS p one
tightenE p (TypeCase ty ret arms def) =
<*> pure loc
tightenE p (TypeCase ty ret arms def loc) =
TypeCase <$> tightenE p ty
<*> tightenT p ret
<*> traverse (tightenS p) arms
<*> tightenT p def
<*> pure loc
tightenE p (CloE (Sub el th)) = do
th <- assert_total $ tightenSub tightenE p th
pure $ CloE $ Sub el th
@ -130,35 +160,40 @@ mutual
export Tighten (Elim d) where tighten p e = tightenE p e
export Tighten (Term d) where tighten p t = tightenT p t
export
Tighten Dim where
tighten p (K e) = pure $ K e
tighten p (B i) = B <$> tighten p i
mutual
export
dtightenT : OPE d1 d2 -> Term d2 n -> Maybe (Term d1 n)
dtightenT p (TYPE l) = pure $ TYPE l
dtightenT p (Pi qty arg res) =
Pi qty <$> dtightenT p arg <*> dtightenS p res
dtightenT p (Lam body) =
Lam <$> dtightenS p body
dtightenT p (Sig fst snd) =
Sig <$> dtightenT p fst <*> dtightenS p snd
dtightenT p (Pair fst snd) =
Pair <$> dtightenT p fst <*> dtightenT p snd
dtightenT p (Enum cases) = pure $ Enum cases
dtightenT p (Tag tag) = pure $ Tag tag
dtightenT p (Eq ty l r) =
Eq <$> dtightenDS p ty <*> dtightenT p l <*> dtightenT p r
dtightenT p (DLam body) = DLam <$> dtightenDS p body
dtightenT p Nat = pure Nat
dtightenT p Zero = pure Zero
dtightenT p (Succ s) = Succ <$> dtightenT p s
dtightenT p (BOX qty ty) = BOX qty <$> dtightenT p ty
dtightenT p (Box val) = Box <$> dtightenT p val
dtightenT p (E e) = assert_total $ E <$> dtightenE p e
dtightenT p (TYPE l loc) =
pure $ TYPE l loc
dtightenT p (Pi qty arg res loc) =
Pi qty <$> dtightenT p arg <*> dtightenS p res <*> pure loc
dtightenT p (Lam body loc) =
Lam <$> dtightenS p body <*> pure loc
dtightenT p (Sig fst snd loc) =
Sig <$> dtightenT p fst <*> dtightenS p snd <*> pure loc
dtightenT p (Pair fst snd loc) =
Pair <$> dtightenT p fst <*> dtightenT p snd <*> pure loc
dtightenT p (Enum cases loc) =
pure $ Enum cases loc
dtightenT p (Tag tag loc) =
pure $ Tag tag loc
dtightenT p (Eq ty l r loc) =
Eq <$> dtightenDS p ty <*> dtightenT p l <*> dtightenT p r <*> pure loc
dtightenT p (DLam body loc) =
DLam <$> dtightenDS p body <*> pure loc
dtightenT p (Nat loc) =
pure $ Nat loc
dtightenT p (Zero loc) =
pure $ Zero loc
dtightenT p (Succ s loc) =
Succ <$> dtightenT p s <*> pure loc
dtightenT p (BOX qty ty loc) =
BOX qty <$> dtightenT p ty <*> pure loc
dtightenT p (Box val loc) =
Box <$> dtightenT p val <*> pure loc
dtightenT p (E e) =
assert_total $ E <$> dtightenE p e
dtightenT p (CloT (Sub tm th)) = do
tm <- dtightenT p tm
th <- assert_total $ traverse (dtightenE p) th
@ -169,38 +204,48 @@ mutual
export
dtightenE : OPE d1 d2 -> Elim d2 n -> Maybe (Elim d1 n)
dtightenE p (F x) = pure $ F x
dtightenE p (B i) = pure $ B i
dtightenE p (fun :@ arg) = [|dtightenE p fun :@ dtightenT p arg|]
dtightenE p (CasePair qty pair ret body) =
dtightenE p (F x loc) =
pure $ F x loc
dtightenE p (B i loc) =
pure $ B i loc
dtightenE p (App fun arg loc) =
App <$> dtightenE p fun <*> dtightenT p arg <*> pure loc
dtightenE p (CasePair qty pair ret body loc) =
CasePair qty <$> dtightenE p pair
<*> dtightenS p ret
<*> dtightenS p body
dtightenE p (CaseEnum qty tag ret arms) =
<*> pure loc
dtightenE p (CaseEnum qty tag ret arms loc) =
CaseEnum qty <$> dtightenE p tag
<*> dtightenS p ret
<*> traverse (dtightenT p) arms
dtightenE p (CaseNat qty qtyIH nat ret zero succ) =
<*> pure loc
dtightenE p (CaseNat qty qtyIH nat ret zero succ loc) =
CaseNat qty qtyIH
<$> dtightenE p nat
<*> dtightenS p ret
<*> dtightenT p zero
<*> dtightenS p succ
dtightenE p (CaseBox qty box ret body) =
<*> pure loc
dtightenE p (CaseBox qty box ret body loc) =
CaseBox qty <$> dtightenE p box
<*> dtightenS p ret
<*> dtightenS p body
dtightenE p (fun :% arg) = [|dtightenE p fun :% tighten p arg|]
dtightenE p (tm :# ty) = [|dtightenT p tm :# dtightenT p ty|]
dtightenE p (Coe ty q0 q1 val) =
[|Coe (dtightenDS p ty) (tighten p q0) (tighten p q1) (dtightenT p val)|]
dtightenE p (Comp ty q0 q1 val r zero one) =
<*> pure loc
dtightenE p (DApp fun arg loc) =
DApp <$> dtightenE p fun <*> tighten p arg <*> pure loc
dtightenE p (Ann tm ty loc) =
Ann <$> dtightenT p tm <*> dtightenT p ty <*> pure loc
dtightenE p (Coe ty q0 q1 val loc) =
[|Coe (dtightenDS p ty) (tighten p q0) (tighten p q1) (dtightenT p val)
(pure loc)|]
dtightenE p (Comp ty q0 q1 val r zero one loc) =
[|Comp (dtightenT p ty) (tighten p q0) (tighten p q1)
(dtightenT p val) (tighten p r)
(dtightenDS p zero) (dtightenDS p one)|]
dtightenE p (TypeCase ty ret arms def) =
(dtightenDS p zero) (dtightenDS p one) (pure loc)|]
dtightenE p (TypeCase ty ret arms def loc) =
[|TypeCase (dtightenE p ty) (dtightenT p ret)
(traverse (dtightenS p) arms) (dtightenT p def)|]
(traverse (dtightenS p) arms) (dtightenT p def) (pure loc)|]
dtightenE p (CloE (Sub el th)) = do
el <- dtightenE p el
th <- assert_total $ traverse (dtightenE p) th
@ -226,46 +271,55 @@ export [ElimD] Tighten (\d => Elim d n) where tighten p e = dtightenE p e
-- versions of SY, etc, that try to tighten and use SN automatically
public export
ST : Tighten f => {s : Nat} -> NContext s -> f (s + n) -> Scoped s f n
ST : Tighten f => {s : Nat} -> BContext s -> f (s + n) -> Scoped s f n
ST names body =
case tightenN s body of
Just body => S names $ N body
Nothing => S names $ Y body
public export
DST : {s : Nat} -> NContext s -> Term (s + d) n -> DScopeTermN s d n
DST : {s : Nat} -> BContext s -> Term (s + d) n -> DScopeTermN s d n
DST names body =
case tightenN @{TermD} s body of
Just body => S names $ N body
Nothing => S names $ Y body
public export %inline
PiT : (qty : Qty) -> (x : BaseName) ->
(arg : Term d n) -> (res : Term d (S n)) -> Term d n
PiT {qty, x, arg, res} = Pi {qty, arg, res = ST [< x] res}
PiT : (qty : Qty) -> (x : BindName) ->
(arg : Term d n) -> (res : Term d (S n)) -> (loc : Loc) -> Term d n
PiT {qty, x, arg, res, loc} = Pi {qty, arg, res = ST [< x] res, loc}
public export %inline
SigT : (x : BaseName) -> (fst : Term d n) ->
(snd : Term d (S n)) -> Term d n
SigT {x, fst, snd} = Sig {fst, snd = ST [< x] snd}
LamT : (x : BindName) -> (body : Term d (S n)) -> (loc : Loc) -> Term d n
LamT {x, body, loc} = Lam {body = ST [< x] body, loc}
public export %inline
EqT : (i : BaseName) -> (ty : Term (S d) n) ->
(l, r : Term d n) -> Term d n
EqT {i, ty, l, r} = Eq {ty = DST [< i] ty, l, r}
SigT : (x : BindName) -> (fst : Term d n) ->
(snd : Term d (S n)) -> (loc : Loc) -> Term d n
SigT {x, fst, snd, loc} = Sig {fst, snd = ST [< x] snd, loc}
public export %inline
CoeT : (i : BaseName) -> (ty : Term (S d) n) ->
(p, q : Dim d) -> (val : Term d n) -> Elim d n
CoeT {i, ty, p, q, val} = Coe {ty = DST [< i] ty, p, q, val}
EqT : (i : BindName) -> (ty : Term (S d) n) ->
(l, r : Term d n) -> (loc : Loc) -> Term d n
EqT {i, ty, l, r, loc} = Eq {ty = DST [< i] ty, l, r, loc}
public export %inline
DLamT : (i : BindName) -> (body : Term (S d) n) -> (loc : Loc) -> Term d n
DLamT {i, body, loc} = DLam {body = DST [< i] body, loc}
public export %inline
CoeT : (i : BindName) -> (ty : Term (S d) n) ->
(p, q : Dim d) -> (val : Term d n) -> (loc : Loc) -> Elim d n
CoeT {i, ty, p, q, val, loc} = Coe {ty = DST [< i] ty, p, q, val, loc}
public export %inline
typeCase1T : Elim d n -> Term d n ->
(k : TyConKind) -> NContext (arity k) -> Term d (arity k + n) ->
{default Nat def : Term d n} ->
(k : TyConKind) -> BContext (arity k) -> Term d (arity k + n) ->
(loc : Loc) ->
{default (Nat loc) def : Term d n} ->
Elim d n
typeCase1T ty ret k ns body {def} =
typeCase ty ret [(k ** ST ns body)] def
typeCase1T ty ret k ns body loc {def} =
typeCase ty ret [(k ** ST ns body)] def loc
export

View file

@ -1,6 +1,7 @@
module Quox.Syntax.Var
import Quox.Name
import public Quox.Loc
import public Quox.Name
import Quox.Pretty
import Quox.OPE
@ -42,6 +43,23 @@ export Uninhabited (VZ = VS i) where uninhabited _ impossible
export Uninhabited (VS i = VZ) where uninhabited _ impossible
public export
data Eqv : Var m -> Var n -> Type where
EZ : VZ `Eqv` VZ
ES : i `Eqv` j -> VS i `Eqv` VS j
%name Var.Eqv e
export
decEqv : Dec2 Eqv
decEqv VZ VZ = Yes EZ
decEqv VZ (VS i) = No $ \case _ impossible
decEqv (VS i) VZ = No $ \case _ impossible
decEqv (VS i) (VS j) =
case decEqv i j of
Yes y => Yes $ ES y
No n => No $ \(ES y) => n y
private
lookupS : Nat -> SnocList a -> Maybe a
lookupS _ [<] = Nothing
@ -148,9 +166,13 @@ weakIsSpec p i = toNatInj $ trans (weakCorrect p i) (sym $ weakSpecCorrect p i)
public export
interface FromVar f where %inline fromVar : Var n -> f n
interface FromVar f where %inline fromVarLoc : Var n -> Loc -> f n
public export FromVar Var where fromVar = id
public export %inline
fromVar : FromVar f => Var n -> {default noLoc loc : Loc} -> f n
fromVar x = fromVarLoc x loc
public export FromVar Var where fromVarLoc x _ = x
export
tabulateV : {0 tm : Nat -> Type} -> (forall n. Var n -> tm n) ->