intensionality hell
This commit is contained in:
parent
9c5ae69cbd
commit
b5b90ddd37
|
@ -152,7 +152,7 @@ For every context $\Gamma$ and $A ∈ Ty(\Gamma)$ there is a projection map $p_{
|
|||
Dubbing $p$ a projection is no coincidence, it projects a map out of a weakened map.
|
||||
```
|
||||
field
|
||||
p-∘ : {Δ Γ : Ob} {γ : Hom Δ Γ} {A : ∣ Ty Γ ∣} (a : ∣ Tr Δ (A ⟦ γ ⟧) ∣)
|
||||
p-∘ : {Δ Γ : Ob} (γ : Hom Δ Γ) {A : ∣ Ty Γ ∣} (a : ∣ Tr Δ (A ⟦ γ ⟧) ∣)
|
||||
→ p ∘ ⟨ γ , a ⟩ ≡ γ
|
||||
```
|
||||
|
||||
|
@ -162,7 +162,7 @@ Likewise, $q$ "projects" an object out of a weakened substitution.
|
|||
→ ∣ Tr Δ (A ⟦ p ⟧ ⟦ ⟨ γ , a ⟩ ⟧) ∣ ≡ ∣ Tr Δ (A ⟦ γ ⟧) ∣
|
||||
q-∘-pathp {Δ} {Γ} {γ} {A} a =
|
||||
∣ Tr Δ (A ⟦ p ⟧ ⟦ ⟨ γ , a ⟩ ⟧) ∣ ≡˘⟨ ⟦⟧-tr-comp p ⟨ γ , a ⟩ A ⟩
|
||||
∣ Tr Δ (A ⟦ p ∘ ⟨ γ , a ⟩ ⟧) ∣ ≡⟨ (λ i → ∣ Tr Δ (A ⟦ p-∘ a i ⟧) ∣) ⟩
|
||||
∣ Tr Δ (A ⟦ p ∘ ⟨ γ , a ⟩ ⟧) ∣ ≡⟨ (λ i → ∣ Tr Δ (A ⟦ p-∘ γ a i ⟧) ∣) ⟩
|
||||
∣ Tr Δ (A ⟦ γ ⟧) ∣ ∎
|
||||
|
||||
field
|
||||
|
|
|
@ -57,8 +57,7 @@ Substitution must also be closed under all of these.
|
|||
B ⟦ ⟨ id, x ⟩ ⟧ ⟦ γ ⟧
|
||||
typeof app f x [ γ ]
|
||||
|
||||
it would seem CwFs really don't want to be formalized :/
|
||||
not enough definitional equalities...
|
||||
Welcome to intensionality hell.
|
||||
|
||||
```
|
||||
app-subst-pathp : {Δ Γ : Ob} (γ : Hom Δ Γ) {A : ∣ Ty Γ ∣} {B : ∣ Ty (Γ ; A) ∣}
|
||||
|
@ -72,10 +71,23 @@ Substitution must also be closed under all of these.
|
|||
≡⟨⟩
|
||||
B ⟦ ⟨ γ ∘ p , transport (λ i → ⟦⟧-tr-comp γ p A (~ i)) q ⟩ ∘ ⟨id, x [ γ ] ⟩ ⟧
|
||||
≡⟨ (λ i → B ⟦ ⟨⟩-∘ ⟨id, x [ γ ] ⟩ (γ ∘ p) A (transport (λ i → ⟦⟧-tr-comp γ (p {Δ} {A ⟦ γ ⟧}) A (~ i)) q) i ⟧) ⟩
|
||||
B ⟦ ⟨ γ ∘ p ∘∘ ⟨id, x [ γ ] ⟩ , transport (λ i → ⟦⟧-tr-comp γ (p {Δ} {A ⟦ γ ⟧}) A (~ i)) q [ ⟨id, x [ γ ] ⟩ ] ⟩ ⟧
|
||||
≡⟨⟩
|
||||
B ⟦ ⟨ (γ ∘ p) ∘ ⟨id, x [ γ ] ⟩ , transport (λ i → ⟦⟧-tr-comp (γ ∘ p) ⟨id, x [ γ ] ⟩ A (~ i))
|
||||
(transport (λ j → ⟦⟧-tr-comp γ (p {Δ} {A ⟦ γ ⟧}) A (~ j)) q [ ⟨id, x [ γ ] ⟩ ]) ⟩ ⟧
|
||||
≡⟨ ((λ i → B ⟦ ⟨ assoc γ p ⟨id, x [ γ ] ⟩ (~ i) , {! !} ⟩ ⟧)) ⟩
|
||||
B ⟦ ⟨ γ ∘ (p ∘ ⟨id, x [ γ ] ⟩) , transport (λ i → ⟦⟧-tr-comp γ (p ∘ ⟨id, x [ γ ] ⟩) A (~ i))
|
||||
(transport (λ j → ⟦⟧-tr-comp p ⟨id, x [ γ ] ⟩ (A ⟦ γ ⟧) (~ j)) (q [ ⟨id, x [ γ ] ⟩ ])) ⟩ ⟧
|
||||
≡⟨⟩
|
||||
B ⟦ ⟨ γ ∘∘ p ∘ ⟨id, x [ γ ] ⟩ , transport (λ i → ⟦⟧-tr-comp p ⟨id, x [ γ ] ⟩ (A ⟦ γ ⟧) (~ i)) (q [ ⟨id, x [ γ ] ⟩ ]) ⟩ ⟧
|
||||
≡⟨ (λ i → B ⟦ ⟨ γ ∘∘ p-∘ id (transport (λ i → ⟦⟧-tr-id (A ⟦ γ ⟧) (~ i)) (x [ γ ])) i , {! !} ⟩ ⟧) ⟩
|
||||
B ⟦ ⟨ γ ∘∘ id , transport (λ i → ⟦⟧-tr-id (A ⟦ γ ⟧) (~ i)) (x [ γ ]) ⟩ ⟧
|
||||
≡⟨⟩
|
||||
B ⟦ ⟨ γ ∘ id , transport (λ i → ⟦⟧-tr-comp γ id A (~ i)) (transport (λ j → ⟦⟧-tr-id (A ⟦ γ ⟧) (~ j)) (x [ γ ])) ⟩ ⟧
|
||||
≡⟨ {! !} ⟩
|
||||
B ⟦ ⟨ γ ∘ id , transport ((λ j → ⟦⟧-tr-id (A ⟦ γ ⟧) (~ j)) ∙ (λ i → ⟦⟧-tr-comp γ id A (~ i))) (x [ γ ]) ⟩ ⟧
|
||||
≡⟨ {! !} ⟩
|
||||
B ⟦ ⟨ γ , x [ γ ] ⟩ ⟧
|
||||
≡⟨ {! !} ⟩
|
||||
B ⟦ ⟨ id ∘∘ γ , transport (λ i → ∣ Tr Δ (⟦⟧-id A (~ i) ⟦ γ ⟧) ∣) (x [ γ ]) ⟩ ⟧
|
||||
≡⟨ {! !} ⟩
|
||||
B ⟦ ⟨id, x ⟩ ⟧ ⟦ γ ⟧
|
||||
∎
|
||||
|
@ -95,3 +107,4 @@ Lastly comes the star of the show, $\beta$ reduction.
|
|||
(a : ∣ Tr Γ A ∣)
|
||||
→ app (Cλ A B b) a ≡ b [ ⟨id, a ⟩ ]
|
||||
```
|
||||
|
Loading…
Reference in New Issue
Block a user