{-# OPTIONS --safe #-}
module Cubical.Core.Glue where
open import Cubical.Core.Primitives
open import Agda.Builtin.Cubical.Glue public
using ( isEquiv
; equiv-proof
; _≃_
; equivFun
; equivProof
; primGlue
; prim^unglue
)
renaming ( prim^glue to glue
)
private
variable
ℓ ℓ' : Level
Glue : (A : Type ℓ) {φ : I}
→ (Te : Partial φ (Σ[ T ∈ Type ℓ' ] T ≃ A))
→ Type ℓ'
Glue A Te = primGlue A (λ x → Te x .fst) (λ x → Te x .snd)
unglue : {A : Type ℓ} (φ : I) {T : Partial φ (Type ℓ')}
{e : PartialP φ (λ o → T o ≃ A)} → primGlue A T e → A
unglue φ = prim^unglue {φ = φ}
private
module GluePrims (A : Type ℓ) {φ : I} (Te : Partial φ (Σ[ T ∈ Type ℓ' ] T ≃ A)) where
T : Partial φ (Type ℓ')
T φ1 = Te φ1 .fst
e : PartialP φ (λ φ → T φ ≃ A)
e φ1 = Te φ1 .snd
Glue-S : Type ℓ' [ φ ↦ T ]
Glue-S = inS (Glue A Te)
coeT→G :
∀ (t : PartialP φ T)
→ Partial φ (Glue A Te)
coeT→G t (φ = i1) = t 1=1
coeG→T :
∀ (g : Glue A Te)
→ PartialP φ T
coeG→T g (φ = i1) = g
trans-e :
∀ (t : PartialP φ T)
→ Partial φ A
trans-e t ϕ1 = equivFun (e ϕ1) (t ϕ1)
glue-S :
∀ (t : PartialP φ T)
→ A [ φ ↦ trans-e t ]
→ Glue A Te [ φ ↦ coeT→G t ]
glue-S t s = inS (glue t (outS s))
unglue-S :
∀ (b : Glue A Te)
→ A [ φ ↦ trans-e (coeG→T b) ]
unglue-S b = inS (unglue φ b)
module GlueTransp (A : I → Type ℓ) (Te : (i : I) → Partial (i ∨ ~ i) (Σ[ T ∈ Type ℓ' ] T ≃ A i)) where
A0 A1 : Type ℓ
A0 = A i0
A1 = A i1
T : (i : I) → Partial (i ∨ ~ i) (Type ℓ')
T i φ = Te i φ .fst
e : (i : I) → PartialP (i ∨ ~ i) (λ φ → T i φ ≃ A i)
e i φ = Te i φ .snd
T0 T1 : Type ℓ'
T0 = T i0 1=1
T1 = T i1 1=1
e0 : T0 ≃ A0
e0 = e i0 1=1
e1 : T1 ≃ A1
e1 = e i1 1=1
open import Cubical.Foundations.Prelude using (transport)
transportA : A0 → A1
transportA = transport (λ i → A i)
invEq : ∀ {X : Type ℓ'} {ℓ''} {Y : Type ℓ''} (w : X ≃ Y) → Y → X
invEq w y = w .snd .equiv-proof y .fst .fst
transp-S : (t0 : T0) → T1 [ i1 ↦ (λ _ → invEq e1 (transportA (equivFun e0 t0))) ]
transp-S t0 = inS (transport (λ i → Glue (A i) (Te i)) t0)