add some needed ωs for w-types

i.o.u. linear trees. i'm still thinking
This commit is contained in:
rhiannon morris 2023-08-09 11:46:14 +02:00
parent 7cf3fa8bae
commit 5c4053e9d2
3 changed files with 21 additions and 20 deletions

View file

@ -232,13 +232,13 @@ parameters (defs : Definitions)
compare0' ctx ty@(W {shape, body, _}) s t = compare0' ctx ty@(W {shape, body, _}) s t =
case (s, t) of case (s, t) of
-- Γ ⊢ s₁ = t₁ : A -- Γ ⊢ s₁ = t₁ : A
-- Γ ⊢ s₂ = t₂ : 1.B[s₁∷A/x] → (x : A) ⊲ B -- Γ ⊢ s₂ = t₂ : ω.B[s₁∷A/x] → (x : A) ⊲ B
-- ----------------------------------------- -- -----------------------------------------
-- Γ ⊢ s₁⋄s₂ = t₁⋄t₂ : (x : A) ⊲ B -- Γ ⊢ s₁⋄s₂ = t₁⋄t₂ : (x : A) ⊲ B
(Sup sRoot sSub {}, Sup tRoot tSub {}) => do (Sup sRoot sSub {}, Sup tRoot tSub {}) => do
compare0 ctx shape sRoot tRoot compare0 ctx shape sRoot tRoot
let arg = sub1 body (Ann sRoot shape sRoot.loc) let arg = sub1 body (Ann sRoot shape sRoot.loc)
subTy = Arr One arg ty ty.loc subTy = Arr Any arg ty ty.loc
compare0 ctx subTy sSub tSub compare0 ctx subTy sSub tSub
(E e, E f) => Elim.compare0 ctx e f (E e, E f) => Elim.compare0 ctx e f
@ -481,7 +481,7 @@ parameters (defs : Definitions)
-- Ψ | Γ ⊢ e = f ⇒ (x : A) ⊲ B -- Ψ | Γ ⊢ e = f ⇒ (x : A) ⊲ B
-- Ψ | Γ, p : (x : A) ⊲ B ⊢ Q = R ⇐ Type -- Ψ | Γ, p : (x : A) ⊲ B ⊢ Q = R ⇐ Type
-- Ψ | Γ, x : A, y : 1.B → (x : A) ⊲ B, ih : 1.(z : B) → Q[y z/p] -- Ψ | Γ, x : A, y : ω.B → (x : A) ⊲ B, ih : 1.(z : B) → Q[y z/p]
-- ⊢ s = t ⇐ Q[(x⋄y ∷ (x : A) ⊲ B)/p] -- ⊢ s = t ⇐ Q[(x⋄y ∷ (x : A) ⊲ B)/p]
-- ---------------------------------------------------------------- -- ----------------------------------------------------------------
-- Ψ | Γ ⊢ caseπ e return Q of { x ⋄ y, ς.ih ⇒ s } -- Ψ | Γ ⊢ caseπ e return Q of { x ⋄ y, ς.ih ⇒ s }
@ -496,7 +496,7 @@ parameters (defs : Definitions)
let [< x, y, ih] = ebody.names let [< x, y, ih] = ebody.names
z <- mnb "z" ih.loc z <- mnb "z" ih.loc
let xbind = (epi, x, shape) let xbind = (epi, x, shape)
ybind = (epi, y, Arr One tbody.term (weakT 1 ety) y.loc) ybind = (epi, y, Arr Any tbody.term (weakT 1 ety) y.loc)
ihbind = (epi', ih, ihbind = (epi', ih,
PiY One z PiY One z
(sub1 (weakS 2 tbody) (BV 1 x.loc)) (sub1 (weakS 2 tbody) (BV 1 x.loc))

View file

@ -441,8 +441,8 @@ parameters {d, n : Nat} (defs : Definitions) (ctx : WhnfContext d n)
-- 𝒮𝑘 ≔ ((x : A) ⊲ B)𝑘/𝑖 -- 𝒮𝑘 ≔ ((x : A) ⊲ B)𝑘/𝑖
-- 𝒶𝑘 ≔ coe [𝑖 ⇒ A] @p @𝑘 a -- 𝒶𝑘 ≔ coe [𝑖 ⇒ A] @p @𝑘 a
-- : A𝑘/𝑖 -- : A𝑘/𝑖
-- 𝒷𝑘 ≔ coe [𝑖1.B[𝒶𝑖/x] → 𝒮𝑖] @p @𝑘 b -- 𝒷𝑘 ≔ coe [𝑖ω.B[𝒶𝑖/x] → 𝒮𝑖] @p @𝑘 b
-- : 1.B𝑘/𝑖[𝒶𝑘/x] → 𝒮𝑘 -- : ω.B𝑘/𝑖[𝒶𝑘/x] → 𝒮𝑘
-- 𝒾𝒽 ≔ coe [𝑖 ⇒ π.(z : B[𝒶𝑖/x]) → C[𝒷𝑖 z/p]] @p @q ih -- 𝒾𝒽 ≔ coe [𝑖 ⇒ π.(z : B[𝒶𝑖/x]) → C[𝒷𝑖 z/p]] @p @q ih
-- : π.(z : Bq/𝑖[𝒶q/x]) → C[𝒷q z/p] -- : π.(z : Bq/𝑖[𝒶q/x]) → C[𝒷q z/p]
-- -------------------------------------------------------------------- -- --------------------------------------------------------------------
@ -464,7 +464,7 @@ parameters {d, n : Nat} (defs : Definitions) (ctx : WhnfContext d n)
// (a' (BV 0 bi.loc) ::: shift 3) // (a' (BV 0 bi.loc) ::: shift 3)
ty' = ty // Shift (weak 1 by) // shift 3 ty' = ty // Shift (weak 1 by) // shift 3
in in
CoeT bi (Arr One tbody' ty' b.loc) (p // by) k (BVT 1 b.loc) b.loc CoeT bi (Arr Any tbody' ty' b.loc) (p // by) k (BVT 1 b.loc) b.loc
ih' : Elim d (3 + n) := ih' : Elim d (3 + n) :=
let tbody' = tbody.term // (a' (BV 0 ihi.loc) ::: shift 3) let tbody' = tbody.term // (a' (BV 0 ihi.loc) ::: shift 3)
ret' = sub1 (weakS 4 $ dweakS 1 ret) $ ret' = sub1 (weakS 4 $ dweakS 1 ret) $
@ -636,17 +636,18 @@ pushCoe defs ctx i ty p q s loc =
-- (coe [i ⇒ (x : A) ⊲ π.B] @p @q (s ⋄ t) ⇝ -- (coe [i ⇒ (x : A) ⊲ π.B] @p @q (s ⋄ t) ⇝
-- (coe [i ⇒ A] @p @q s ⋄ -- (coe [i ⇒ A] @p @q s ⋄
-- coe [i ⇒ 1.B[coe [j ⇒ Aj/i] @p @i s/x] → (x : A) ⊲ B] t) -- coe [i ⇒ ω.B[coe [j ⇒ Aj/i] @p @i s/x] → (x : A) ⊲ B] t)
-- ∷ ((x : A) ⊲ B)q/i -- ∷ ((x : A) ⊲ B)q/i
Sup {root, sub, loc = supLoc} => do Sup {root, sub, loc = supLoc} => do
let W {shape, body, loc = wLoc} = ty let W {shape, body, loc = wLoc} = ty
| _ => throw $ ExpectedW ty.loc (extendDim i ctx.names) ty | _ => throw $ ExpectedW ty.loc (extendDim i ctx.names) ty
let root' = E $ CoeT i shape p q root root.loc let root' = E $ CoeT i shape p q root root.loc
tsub' = sub1 body $ tsub0 = sub1 body $
CoeT !(fresh i) (shape // (B VZ root.loc ::: shift 2)) CoeT !(fresh i) (shape // (B VZ root.loc ::: shift 2))
(weakD 1 p) (BV 0 sub.loc) (weakD 1 p) (BV 0 sub.loc)
(dweakT 1 sub) sub.loc (dweakT 1 sub) sub.loc
sub' = E $ CoeT i tsub' p q sub sub.loc tsub' = Arr Any tsub0 ty sub.loc
sub' = E $ CoeT i tsub' p q sub sub.loc
pure $ pure $
Element (Ann (Sup root' sub' supLoc) Element (Ann (Sup root' sub' supLoc)
(W (shape // one q) (body // one q) wLoc) loc) Ah (W (shape // one q) (body // one q) wLoc) loc) Ah
@ -731,7 +732,7 @@ CanWhnf Elim Reduce.isRedexE where
pure $ Element (CasePair pi pair ret body caseLoc) $ pairnf `orNo` np pure $ Element (CasePair pi pair ret body caseLoc) $ pairnf `orNo` np
-- s' ≔ s ∷ A -- s' ≔ s ∷ A
-- t' ≔ t ∷ 1.B[s'/x] → (x : A) ⊲ B -- t' ≔ t ∷ ω.B[s'/x] → (x : A) ⊲ B
-- ih' ≔ (λ x ⇒ caseπ t x return p ⇒ C of { (a ⋄ b), ς.ih ⇒ u }) ∷ -- ih' ≔ (λ x ⇒ caseπ t x return p ⇒ C of { (a ⋄ b), ς.ih ⇒ u }) ∷
-- π.(y : B[s'/x]) → C[t' y/p] -- π.(y : B[s'/x]) → C[t' y/p]
-- st' ≔ s ⋄ t ∷ (x : A) ⊲ B -- st' ≔ s ⋄ t ∷ (x : A) ⊲ B
@ -748,7 +749,7 @@ CanWhnf Elim Reduce.isRedexE where
w@(W {shape, body = wbody, _}) treeLoc => w@(W {shape, body = wbody, _}) treeLoc =>
let root = Ann root shape root.loc let root = Ann root shape root.loc
wbody' = sub1 wbody root wbody' = sub1 wbody root
tsub = Arr One wbody' w sub.loc tsub = Arr Any wbody' w sub.loc
sub = Ann sub tsub sub.loc sub = Ann sub tsub sub.loc
ih' = LamY !(mnb "y" caseLoc) -- [todo] better name ih' = LamY !(mnb "y" caseLoc) -- [todo] better name
(E $ CaseW qty qtyIH (E $ CaseW qty qtyIH

View file

@ -85,12 +85,12 @@ where
(Zero, r, BVT 2 r.loc)] (Zero, r, BVT 2 r.loc)]
||| if a ⋄ b : (x : A) ⊲ B, then b : `supSubTy a A B _` ||| if a ⋄ b : (x : A) ⊲ B, then b : `supSubTy a A B _`
||| i.e. 1.B[a∷A/x] → ((x : A) ⊲ B) ||| i.e. ω.B[a∷A/x] → ((x : A) ⊲ B)
export export
supSubTy : (root, rootTy : Term d n) -> supSubTy : (root, rootTy : Term d n) ->
(body : ScopeTerm d n) -> Loc -> Term d n (body : ScopeTerm d n) -> Loc -> Term d n
supSubTy root rootTy body loc = supSubTy root rootTy body loc =
Arr One (sub1 body (Ann root rootTy root.loc)) (W rootTy body loc) loc Arr Any (sub1 body (Ann root rootTy root.loc)) (W rootTy body loc) loc
mutual mutual
@ -207,10 +207,10 @@ mutual
(shape, body) <- expectW !(askAt DEFS) ctx ty.loc ty (shape, body) <- expectW !(askAt DEFS) ctx ty.loc ty
-- if Ψ | Γ ⊢ σ · a ⇐ A ⊳ Σ₁ -- if Ψ | Γ ⊢ σ · a ⇐ A ⊳ Σ₁
qroot <- checkC ctx sg root shape qroot <- checkC ctx sg root shape
-- if Ψ | Γ ⊢ σ · b ⇐ 1.B[a∷A/x] → ((x : A) ⊲ B) ⊳ Σ₂ -- if Ψ | Γ ⊢ σ · b ⇐ ω.B[a∷A/x] → ((x : A) ⊲ B) ⊳ Σ₂
qsub <- checkC ctx sg sub (supSubTy root shape body sub.loc) qsub <- checkC ctx sg sub (supSubTy root shape body sub.loc)
-- then Ψ | Γ ⊢ σ · (a ⋄ b) ⇐ ((x : A) ⊲ B) ⊳ Σ₁+Σ₂ -- then Ψ | Γ ⊢ σ · (a ⋄ b) ⇐ ((x : A) ⊲ B) ⊳ Σ₁+ωΣ₂
pure $ qroot + qsub pure $ qroot + Any * qsub
check' ctx sg (Pair fst snd loc) ty = do check' ctx sg (Pair fst snd loc) ty = do
(tfst, tsnd) <- expectSig !(askAt DEFS) ctx ty.loc ty (tfst, tsnd) <- expectSig !(askAt DEFS) ctx ty.loc ty
@ -430,14 +430,14 @@ mutual
-- if Ψ | Γ, p : (x : A) ⊲ B ⊢₀ C ⇐ Type -- if Ψ | Γ, p : (x : A) ⊲ B ⊢₀ C ⇐ Type
checkTypeC (extendTy Zero ret.name w ctx) ret.term Nothing checkTypeC (extendTy Zero ret.name w ctx) ret.term Nothing
(shape, tbody) <- expectW !(askAt DEFS) ctx tree.loc w (shape, tbody) <- expectW !(askAt DEFS) ctx tree.loc w
-- if Ψ | Γ, x : A, y : 1.B → (x : A) ⊲ B, -- if Ψ | Γ, x : A, y : ω.B → (x : A) ⊲ B,
-- ih : π.(z : B) → C[y z/p] -- ih : π.(z : B) → C[y z/p]
-- ⊢ σ · u ⇐ C[((x ⋄ y) ∷ (x : A) ⊲ B)/p] -- ⊢ σ · u ⇐ C[((x ⋄ y) ∷ (x : A) ⊲ B)/p]
-- ⊳ Σ₂, π'.x, ς₁.y, ς₂.ih -- ⊳ Σ₂, π'.x, ς₁.y, ς₂.ih
-- with π' ≤ σπ, ς₂ ≤ σς, ς₁+ς₂ ≤ σπ -- with π' ≤ σπ, ς₂ ≤ σς, ς₁+ς₂ ≤ σπ
let [< x, y, ih] = body.names let [< x, y, ih] = body.names
-- 1.B → (x : A) ⊲ B -- ω.B → (x : A) ⊲ B
tsub = Arr One tbody.term (weakT 1 w) y.loc tsub = Arr Any tbody.term (weakT 1 w) y.loc
-- y z -- y z
ihret = App (BV 1 y.loc) (BVT 0 ih.loc) y.loc ihret = App (BV 1 y.loc) (BVT 0 ih.loc) y.loc
-- π.(z : B) → C[y z/p] -- π.(z : B) → C[y z/p]