Skip to content
Snippets Groups Projects
Commit ba544618 authored by Max New's avatar Max New
Browse files

Implement Nbe for solving substitution equations

parent bdbe1b23
No related branches found
No related tags found
No related merge requests found
......@@ -114,7 +114,7 @@ data Val where
-- Do these make sense?
injectN : Val [ nat ] dyn
injectArr : Val [ S ] (dyn ⇀ dyn) -> Val [ S ] dyn
injectArr : Val [ dyn ⇀ dyn ] dyn
mapDyn : Val [ nat ] nat → Val [ dyn ⇀ dyn ] (dyn ⇀ dyn) → Val [ dyn ] dyn
--mapDynβn : ?
--mapDynβf : ?
......@@ -134,7 +134,8 @@ data Val where
_[_]vP = _[_]v
varP = var
↑subst : Subst Δ Γ → Subst (R ∷ Δ) (R ∷ Γ)
↑subst γ = γ ∘s wk ,s var
data EvCtx where
∙E : EvCtx Γ S S
......@@ -195,9 +196,9 @@ data Comp where
-- match 0 Kz (x . Ks) ≡ Kz
matchNatβz : (Kz : Comp Γ S)(Ks : Comp (nat ∷ Γ) S)
→ matchNat Kz Ks [ ids ,s (zro [ !s ]v) ]c ≡ Kz
-- match (S V) Kz (x . Ks) ≡ Ks [ V / x ]
matchNatβs : (Kz : Comp Γ S)(Ks : Comp (nat ∷ Γ) S) (V : Val Γ nat)
→ matchNat Kz Ks [ ids ,s (suc [ !s ,s V ]v) ]c ≡ (Ks [ ids ,s V ]c)
-- match (S y) Kz (x . Ks) ≡ Ks [ y / x ]
matchNatβs : (Kz : Comp Γ S) (Ks : Comp (nat ∷ Γ) S)
→ matchNat Kz Ks [ wk ,s (suc [ !s ,s var ]v) ]c ≡ Ks
-- M[x] ≡ match x (M[0/x]) (x. M[S x/x])
matchNatη : M ≡ matchNat
(M [ ids ,s (zro [ !s ]v) ]c)
......@@ -213,9 +214,11 @@ data Comp where
matchDynβf : (Kn : Comp (nat ∷ Γ) S) (Kf : Comp ((dyn ⇀ dyn) ∷ Γ) S)
(V : Val Γ (dyn ⇀ dyn)) ->
matchDyn Kn Kf [ ids ,s ((injectArr var) [ !s ,s V ]v) ]c ≡
matchDyn Kn Kf [ ids ,s (injectArr [ !s ,s V ]v) ]c ≡
Kf [ ids ,s V ]c
matchDynSubst : ∀ (M : Comp (nat ∷ Γ) R) (N : Comp ((dyn ⇀ dyn) ∷ Γ) R) (γ : Subst Δ Γ)
→ matchDyn M N [ ↑subst γ ]c ≡ matchDyn (M [ ↑subst γ ]c) (N [ ↑subst γ ]c)
{-
matchDynβf' : (Kn : Comp (nat ∷ Γ) S) (Kf : Comp ((dyn ⇀ dyn) ∷ Γ) S)
(V : Val Γ T) (V-up : Val [ T ] (dyn ⇀ dyn)) ->
......@@ -227,6 +230,8 @@ data Comp where
-- tick commutes with ev ctxs
tick-strictness : ∀ {M} -> E [ tick M ]∙ ≡ tick (E [ M ]∙)
-- tick subs
tickSubst : (tick M) [ γ ]c ≡ tick (M [ γ ]c)
isSetComp : isSet (Comp Γ S)
-- isPosetComp : Comp⊑ (refl-⊑ctx Γ) (refl-⊑ S) M M'
......@@ -339,6 +344,9 @@ subst-naturality {S} {Γ} {Δ} {γ} {γ'} {V} {V'} =
!s-ext : {γ : Subst Γ []} → γ ≡ δ
!s-ext = []η ∙ sym []η
ids-sole : ∀ {R} → ids {Γ = R ∷ []} ≡ (!s ,s var)
ids-sole = ,sη ∙ cong₂ _,s_ []η substId
,s-nat : (γ ,s V) ∘s δ ≡ ((γ ∘s δ) ,s (V [ δ ]v))
,s-nat =
,sη ∙ cong₂ _,s_ (∘Assoc ∙ cong₂ (_∘s_) wkβ refl)
......@@ -381,6 +389,54 @@ fun-β' M V =
fun-η' : V ≡ lda (appVal (V [ wk ]v) var)
fun-η' = fun-η
matchNatβz' : matchNat M N [ γ ,s (zro [ !s ]v) ]c ≡ M [ γ ]c
matchNatβz' {M = M}{N = N}{γ = γ} =
cong (matchNat M N [_]c)
(sym (,s-nat ∙ cong₂ _,s_ ∘IdL (sym substAssoc ∙ cong (zro [_]v) []η)))
∙ substAssoc ∙ cong (_[ γ ]c) (matchNatβz _ N)
matchNatβs' : (matchNat M N [ γ ,s (suc [ !s ,s V ]v) ]c) ≡ (N [ γ ,s V ]c)
matchNatβs' {M = M}{N = N}{γ = γ}{V = V} =
cong (matchNat M N [_]c)
(sym (,s-nat ∙ cong₂ _,s_ wkβ (sym substAssoc ∙ cong (suc [_]v) (,s-nat ∙ cong₂ _,s_ []η varβ))))
∙ substAssoc ∙ cong (_[ γ ,s V ]c) (matchNatβs M N )
matchNat-nat :
matchNat M N [ ↑subst γ ]c
≡ matchNat (M [ γ ]c)
(N [ ↑subst γ ]c)
matchNat-nat {M = M}{N = N}{γ = γ} = matchNatη
∙ cong₂ matchNat
(sym substAssoc
∙ cong (matchNat M N [_]c)
(,s-nat ∙ cong₂ _,s_ (sym ∘Assoc ∙ cong (γ ∘s_) wkβ ∙ ∘IdR) varβ)
∙ matchNatβz')
(sym substAssoc
∙ cong (matchNat M N [_]c)
(,s-nat ∙ cong₂ _,s_ (sym ∘Assoc ∙ cong (γ ∘s_) wkβ) varβ)
∙ matchNatβs')
-- (γ , V) ≡ (γ o π1 , π2) ∘ (ids , V)
,s-separate : (γ ,s V) ≡ (↑subst γ) ∘s (ids ,s V)
,s-separate {γ = γ} =
sym (,s-nat ∙ cong₂ _,s_ (sym ∘Assoc ∙ cong (γ ∘s_) wkβ ∙ ∘IdR) varβ)
matchNat-nat' :
matchNat M N [ γ ,s V ]c
≡ matchNat (M [ γ ]c)
(N [ ↑subst γ ]c) [ ids ,s V ]c
matchNat-nat' {M = M}{N = N}{γ = γ}{V = V} =
cong (matchNat M N [_]c) ,s-separate
∙ substAssoc ∙ cong (_[ ids ,s V ]c) matchNat-nat
matchDyn-nat' :
matchDyn M N [ γ ,s V ]c
≡ matchDyn (M [ ↑subst γ ]c)
(N [ ↑subst γ ]c)
[ ids ,s V ]c
matchDyn-nat' {M = M}{N = N}{γ = γ}{V = V} =
cong (matchDyn M N [_]c) ,s-separate
∙ substAssoc ∙ cong (_[ ids ,s V ]c) (matchDynSubst _ _ _)
bind-nat : (bind M) [ γ ]e ≡ bind (M [ γ ∘s wk ,s var ]c)
bind-nat {M = M} {γ = γ} = ret-η ∙ cong bind (cong (_[ ret [ !s ,s var ]c ]∙) (sym substAssoc)
......@@ -441,7 +497,7 @@ Pertᴾ→EC (δi ⇀ δo) = bind (ret' (lda (
((Pertᴾ→EC δo [ !s ]e) [ ret' var ]∙)))))
Pertᴾ→EC (PertD δn δf) = bind (matchDyn
(((Pertᴾ→EC δn [ !s ]e) [ ret' var ]∙) >> ret' (injectN [ wk ]v))
(((Pertᴾ→EC δf [ !s ]e) [ ret' var ]∙) >> ret' (injectArr var [ wk ]v)))
(((Pertᴾ→EC δf [ !s ]e) [ ret' var ]∙) >> ret' (injectArr [ wk ]v)))
-- The four cast rule delay terms are admissible in the syntax
......@@ -483,7 +539,7 @@ emb (ci ⇀ co) =
((app [ drop2nd ]c) >>
((vToE (emb co) [ !s ]e) [ ret' var ]∙)))
emb inj-nat = injectN
emb (inj-arr c) = injectArr (emb c)
emb (inj-arr c) = injectArr [ !s ,s (emb c) ]v
-- We can show that emb (c ∘ d) ≡ emb d ∘ emb c
......
module Syntax.IntensionalTerms.Induction where
open import Cubical.Foundations.Prelude
open import Cubical.Foundations.HLevels
open import Cubical.Foundations.Structure
open import Cubical.Data.List
open import Cubical.Data.Unit
open import Cubical.Data.Sigma
open import Syntax.IntensionalTerms
open import Syntax.Types
private
variable
Δ Γ Θ Z Δ' Γ' Θ' Z' : Ctx
R S T U R' S' T' U' : Ty
γ γ' γ'' : Subst Δ Γ
δ δ' δ'' : Subst Θ Δ
θ θ' θ'' : Subst Z Θ
V V' V'' : Val Γ S
M M' M'' N N' : Comp Γ S
E E' E'' F F' : EvCtx Γ S T
ℓ ℓ' ℓ'' ℓ''' ℓ'''' : Level
module SynInd
(Ps : ∀ {Δ} {Γ} → Subst Δ Γ → hProp ℓ)
(Pv : ∀ {Γ} {R} → Val Γ R → hProp ℓ')
(Pc : ∀ {Γ} {R} → Comp Γ R → hProp ℓ'')
(Pe : ∀ {Γ R S} → EvCtx Γ R S → hProp ℓ''')
where
record indCases : Type (ℓ-max ℓ (ℓ-max ℓ' (ℓ-max ℓ'' ℓ'''))) where
field
indIds : ⟨ Ps {Γ} ids ⟩
ind∘s : ⟨ Ps γ ⟩ → ⟨ Ps θ ⟩ → ⟨ Ps (γ ∘s θ) ⟩
ind!s : ⟨ Ps {Δ} !s ⟩
ind,s : ⟨ Ps {Θ}{Γ} γ ⟩ → ⟨ Pv {R = R} V ⟩ → ⟨ Ps (γ ,s V) ⟩
indwk : ⟨ Ps (wk {S = S}{Γ = Γ}) ⟩
ind[]v : ⟨ Pv V ⟩ → ⟨ Ps γ ⟩ → ⟨ Pv (V [ γ ]v)⟩
indVar : ⟨ Pv {Γ = (R ∷ Γ)} var ⟩
indZero : ⟨ Pv zro ⟩
indSuc : ⟨ Pv suc ⟩
indLda : ⟨ Pc M ⟩ → ⟨ Pv (lda M) ⟩
indInjN : ⟨ Pv injectN ⟩
indInjArr : ⟨ Pv injectArr ⟩
indMapDyn : ⟨ Pv V ⟩ → ⟨ Pv V' ⟩
→ ⟨ Pv (mapDyn V V') ⟩
ind[]∙ : ⟨ Pe E ⟩ → ⟨ Pc M ⟩ → ⟨ Pc (E [ M ]∙)⟩
ind[]c : ⟨ Pc M ⟩ → ⟨ Ps γ ⟩ → ⟨ Pc (M [ γ ]c)⟩
indErr : ⟨ Pc {R = R} err ⟩
indTick : ⟨ Pc M ⟩ → ⟨ Pc (tick M) ⟩
indRet : ⟨ Pc {R = R} ret ⟩
indApp : ⟨ Pc (app {S = S}{T = T}) ⟩
indMatchNat : ⟨ Pc M ⟩ → ⟨ Pc M' ⟩ → ⟨ Pc (matchNat M M' )⟩
indMatchDyn : ⟨ Pc M ⟩ → ⟨ Pc M' ⟩ → ⟨ Pc (matchDyn M M' )⟩
ind∙ : ⟨ Pe {Γ}{R} ∙E ⟩
ind∘E : ⟨ Pe E ⟩ → ⟨ Pe F ⟩ → ⟨ Pe (E ∘E F) ⟩
ind[]e : ⟨ Pe E ⟩ → ⟨ Ps γ ⟩ → ⟨ Pe (E [ γ ]e)⟩
indbind : ⟨ Pc M ⟩ → ⟨ Pe (bind M) ⟩
module _ (ic : indCases) where
open indCases ic
indPs : ∀ (γ : Subst Δ Γ) → ⟨ Ps γ ⟩
indPv : ∀ (V : Val Γ R) → ⟨ Pv V ⟩
indPc : ∀ (M : Comp Γ R) → ⟨ Pc M ⟩
indPe : ∀ (E : EvCtx Γ R S) → ⟨ Pe E ⟩
indPs ids = indIds
indPs (γ ∘s γ₁) = ind∘s (indPs γ) (indPs γ₁)
indPs !s = ind!s
indPs (γ ,s V) = ind,s (indPs γ) (indPv V)
indPs wk = indwk
indPs (∘IdL {γ = γ} i) =
isProp→PathP ((λ i → Ps (∘IdL {γ = γ} i) .snd)) (ind∘s indIds (indPs γ)) (indPs γ) i
indPs (∘IdR {γ = γ} i) = isProp→PathP (((λ i → Ps (∘IdR {γ = γ} i) .snd)))
((ind∘s (indPs γ) indIds)) ((indPs γ)) i
indPs (∘Assoc {γ = γ}{δ = δ}{θ = θ} i) =
isProp→PathP (λ i → Ps (∘Assoc {γ = γ} {δ = δ} {θ = θ} i) .snd)
(ind∘s (indPs γ) (ind∘s (indPs δ) (indPs θ)))
(ind∘s (ind∘s (indPs γ) (indPs δ)) (indPs θ))
i
indPs ([]η {γ = γ} i) = isProp→PathP (λ i → Ps ([]η {γ = γ} i) .snd)
(indPs γ) ind!s i
indPs (wkβ {δ = δ}{V = V} i) = isProp→PathP (λ i → Ps (wkβ {δ = δ}{V = V} i) .snd)
(ind∘s indwk (ind,s (indPs δ) (indPv V))) (indPs δ) i
indPs (,sη {δ = δ} i) = isProp→PathP (λ i → Ps (,sη {δ = δ} i) .snd)
(indPs δ) (ind,s (ind∘s indwk (indPs δ)) (ind[]v indVar (indPs δ))) i
indPs (isSetSubst γ γ' p q i j) = isOfHLevel→isOfHLevelDep 2 (λ x → isProp→isSet (Ps x .snd))
(indPs γ)
(indPs γ')
(λ j → indPs (p j))
(λ j → indPs (q j))
(isSetSubst γ γ' p q) i j
indPv (V [ γ ]v) = ind[]v (indPv V) (indPs γ)
indPv var = indVar
indPv zro = indZero
indPv suc = indSuc
indPv (lda M) = indLda (indPc M)
indPv injectN = indInjN
indPv injectArr = indInjArr
indPv (mapDyn V V') = indMapDyn (indPv V) (indPv V')
-- avert your eyes
indPv (substId {V = V} i) =
isProp→PathP (λ i → Pv (substId i) .snd) (ind[]v (indPv V) indIds) (indPv V) i
indPv (substAssoc {V = V}{δ = δ}{γ = γ} i) = isProp→PathP (λ i → Pv (substAssoc i) .snd)
(ind[]v (indPv V) (ind∘s (indPs δ) (indPs γ)))
(ind[]v (ind[]v (indPv V) (indPs δ)) (indPs γ))
i
indPv (varβ {δ = δ}{V = V} i) = isProp→PathP (λ i → Pv (varβ i) .snd)
(ind[]v indVar (ind,s (indPs δ) (indPv V)))
(indPv V)
i
indPv (fun-η {V = V} i) = isProp→PathP (λ i → Pv (fun-η i) .snd)
(indPv V)
(indLda (ind[]c indApp (ind,s (ind,s ind!s (ind[]v (indPv V) indwk)) indVar)))
i
indPv (isSetVal V V' p q i j) = isOfHLevel→isOfHLevelDep 2 (λ x → isProp→isSet (Pv x .snd))
(indPv V)
(indPv V')
(λ j → indPv (p j))
(λ j → indPv (q j))
(isSetVal V V' p q)
i j
indPc (E [ M ]∙) = ind[]∙ (indPe E) (indPc M)
indPc (M [ γ ]c) = ind[]c (indPc M) (indPs γ)
indPc err = indErr
indPc (tick M) = indTick (indPc M)
indPc ret = indRet
indPc app = indApp
indPc (matchNat M M₁) = indMatchNat (indPc M) (indPc M₁)
indPc (matchDyn M M₁) = indMatchDyn (indPc M) (indPc M₁)
-- macros did this
indPc (plugId {M = M} i) = isProp→PathP (λ i → Pc (plugId i) .snd) (ind[]∙ ind∙ (indPc M)) (indPc M) i
indPc (plugAssoc {F = F}{E = E}{M = M} i) = isProp→PathP (λ i → Pc (plugAssoc i) .snd) (ind[]∙ (ind∘E (indPe F) (indPe E)) (indPc M)) (ind[]∙ (indPe F) (ind[]∙ (indPe E) (indPc M))) i
indPc (substId {M = M} i) = isProp→PathP (λ i → Pc (substId i) .snd) (ind[]c (indPc M) indIds) (indPc M) i
indPc (substAssoc {M = M}{δ = δ}{γ = γ} i) = isProp→PathP (λ i → Pc (substAssoc i) .snd) (ind[]c (indPc M) (ind∘s (indPs δ) (indPs γ))) (ind[]c (ind[]c (indPc M) (indPs δ)) (indPs γ)) i
indPc (substPlugDist {E = E}{M = M}{γ = γ} i) = isProp→PathP (λ i → Pc (substPlugDist i) .snd) (ind[]c (ind[]∙ (indPe E) (indPc M)) (indPs γ)) (ind[]∙ (ind[]e (indPe E) (indPs γ)) (ind[]c (indPc M) (indPs γ))) i
indPc (strictness {E = E} i) = isProp→PathP (λ i → Pc (strictness i) .snd) (ind[]∙ (indPe E) (ind[]c indErr ind!s)) (ind[]c indErr ind!s) i
indPc (ret-β {M = M} i) = isProp→PathP (λ i → Pc (ret-β i) .snd) (ind[]∙ (ind[]e (indbind (indPc M)) indwk) (ind[]c indRet (ind,s ind!s indVar))) (indPc M) i
indPc (fun-β {M = M} i) = isProp→PathP (λ i → Pc (fun-β i) .snd) (ind[]c indApp (ind,s (ind,s ind!s (ind[]v (indLda (indPc M)) indwk)) indVar)) (indPc M) i
indPc (matchNatβz M N i) = isProp→PathP (λ i → Pc (matchNatβz M N i) .snd) (ind[]c (indMatchNat (indPc M) (indPc N)) (ind,s indIds (ind[]v indZero ind!s))) (indPc M) i
indPc (matchNatβs M N i) = isProp→PathP (λ i → Pc (matchNatβs M N i) .snd) (ind[]c (indMatchNat (indPc M) (indPc N)) (ind,s indwk (ind[]v indSuc (ind,s ind!s indVar)))) (indPc N) i
indPc (matchNatη {M = M} i) = isProp→PathP (λ i → Pc (matchNatη i) .snd) (indPc M) (indMatchNat (ind[]c (indPc M) (ind,s indIds (ind[]v indZero ind!s))) (ind[]c (indPc M) (ind,s indwk (ind[]v indSuc (ind,s ind!s indVar))))) i
indPc (matchDynβn M N V i) = isProp→PathP (λ i → Pc (matchDynβn M N V i) .snd) (ind[]c (indMatchDyn (indPc M) (indPc N)) (ind,s indIds (ind[]v indInjN (ind,s ind!s (indPv V))))) (ind[]c (indPc M) (ind,s indIds (indPv V))) i
indPc (matchDynβf M N V i) = isProp→PathP (λ i → Pc (matchDynβf M N V i) .snd) (ind[]c (indMatchDyn (indPc M) (indPc N)) (ind,s indIds (ind[]v indInjArr (ind,s ind!s (indPv V))))) (ind[]c (indPc N) (ind,s indIds (indPv V))) i
indPc (matchDynSubst M N γ i) = isProp→PathP (λ i → Pc (matchDynSubst M N γ i) .snd) (ind[]c (indMatchDyn (indPc M) (indPc N)) (ind,s (ind∘s (indPs γ) indwk) indVar)) (indMatchDyn (ind[]c (indPc M) (ind,s (ind∘s (indPs γ) indwk) indVar)) (ind[]c (indPc N) (ind,s (ind∘s (indPs γ) indwk) indVar))) i
indPc (tick-strictness {E = E}{M = M} i) = isProp→PathP (λ i → Pc (tick-strictness i) .snd) (ind[]∙ (indPe E) (indTick (indPc M))) (indTick (ind[]∙ (indPe E) (indPc M))) i
indPc (tickSubst {M = M}{γ = γ} i) = isProp→PathP (λ i → Pc (tickSubst i) .snd) (ind[]c (indTick (indPc M)) (indPs γ)) (indTick (ind[]c (indPc M) (indPs γ))) i
indPc (isSetComp M N p q i j) = isOfHLevel→isOfHLevelDep 2 (λ x → isProp→isSet (Pc x .snd))
(indPc M)
(indPc N)
(λ j → indPc (p j))
(λ j → indPc (q j))
(isSetComp M N p q)
i j
indPe ∙E = ind∙
indPe (E ∘E E₁) = ind∘E (indPe E) (indPe E₁)
indPe (E [ γ ]e) = ind[]e (indPe E) (indPs γ)
indPe (bind M) = indbind (indPc M)
indPe (∘IdL {E = E} i) = isProp→PathP (λ i → Pe (∘IdL i) .snd) (ind∘E ind∙ (indPe E)) (indPe E) i
indPe (∘IdR {E = E} i) = isProp→PathP (λ i → Pe (∘IdR i) .snd) (ind∘E (indPe E) ind∙) (indPe E) i
indPe (∘Assoc {E = E}{F = F}{F' = F'} i) = isProp→PathP (λ i → Pe (∘Assoc i) .snd) (ind∘E (indPe E) (ind∘E (indPe F) (indPe F'))) (ind∘E (ind∘E (indPe E) (indPe F)) (indPe F')) i
indPe (substId {E = E} i) = isProp→PathP (λ i → Pe (substId i) .snd) (ind[]e (indPe E) indIds) (indPe E) i
indPe (substAssoc {E = E}{γ = γ}{δ = δ} i) = isProp→PathP (λ i → Pe (substAssoc i) .snd) (ind[]e (indPe E) (ind∘s (indPs γ) (indPs δ))) (ind[]e (ind[]e (indPe E) (indPs γ)) (indPs δ)) i
indPe (∙substDist {γ = γ} i) = isProp→PathP (λ i → Pe (∙substDist i) .snd) (ind[]e ind∙ (indPs γ)) (ind∙) i
indPe (∘substDist {E = E}{F = F}{γ = γ} i) = isProp→PathP (λ i → Pe (∘substDist i) .snd) (ind[]e (ind∘E (indPe E) (indPe F)) (indPs γ)) (ind∘E (ind[]e (indPe E) (indPs γ)) (ind[]e (indPe F) (indPs γ))) i
indPe (ret-η {E = E} i) = isProp→PathP (λ i → Pe (ret-η i) .snd) (indPe E) (indbind (ind[]∙ (ind[]e (indPe E) indwk) (ind[]c indRet (ind,s ind!s indVar)))) i
indPe (isSetEvCtx E F p q i j) = isOfHLevel→isOfHLevelDep 2 (λ x → isProp→isSet (Pe x .snd))
(indPe E)
(indPe F)
(λ j → indPe (p j))
(λ j → indPe (q j))
(isSetEvCtx E F p q)
i j
This diff is collapsed.
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment