{-# OPTIONS --safe #-}
module Cubical.Foundations.Pointed.Homotopy where
open import Cubical.Foundations.Prelude
open import Cubical.Foundations.GroupoidLaws
open import Cubical.Foundations.Isomorphism
open import Cubical.Foundations.Equiv
open import Cubical.Foundations.Equiv.Fiberwise
open import Cubical.Foundations.Equiv.Properties
open import Cubical.Foundations.Path
open import Cubical.Foundations.Transport
open import Cubical.Foundations.Univalence
open import Cubical.Foundations.Pointed.Base
open import Cubical.Foundations.Pointed.Properties
open import Cubical.Homotopy.Base
open import Cubical.Data.Sigma
private
variable
ℓ ℓ' : Level
module _ {A : Pointed ℓ} {B : typ A → Type ℓ'} {ptB : B (pt A)} where
⋆ = pt A
_∙∼_ : (f g : Π∙ A B ptB) → Type (ℓ-max ℓ ℓ')
(f₁ , f₂) ∙∼ (g₁ , g₂) = Π∙ A (λ x → f₁ x ≡ g₁ x) (f₂ ∙ g₂ ⁻¹)
_∙∼P_ : (f g : Π∙ A B ptB) → Type (ℓ-max ℓ ℓ')
(f₁ , f₂) ∙∼P (g₁ , g₂) = Σ[ h ∈ f₁ ∼ g₁ ] PathP (λ i → h ⋆ i ≡ ptB) f₂ g₂
private
module _ (f g : Π∙ A B ptB) (H : f .fst ∼ g .fst) where
f₁ = fst f
f₂ = snd f
g₁ = fst g
g₂ = snd g
P : Type ℓ'
P = H ⋆ ≡ f₂ ∙ g₂ ⁻¹
Q : Type ℓ'
Q = PathP (λ i → H ⋆ i ≡ ptB) f₂ g₂
p = H ⋆
r = f₂
s = g₂
P≡Q : P ≡ Q
P≡Q = p ≡ r ∙ s ⁻¹
≡⟨ isoToPath symIso ⟩
r ∙ s ⁻¹ ≡ p
≡⟨ cong (r ∙ s ⁻¹ ≡_) (rUnit p ∙∙ cong (p ∙_) (sym (rCancel s)) ∙∙ assoc p s (s ⁻¹)) ⟩
r ∙ s ⁻¹ ≡ (p ∙ s) ∙ s ⁻¹
≡⟨ sym (ua (compr≡Equiv r (p ∙ s) (s ⁻¹))) ⟩
r ≡ p ∙ s
≡⟨ ua (compl≡Equiv (p ⁻¹) r (p ∙ s)) ⟩
p ⁻¹ ∙ r ≡ p ⁻¹ ∙ (p ∙ s)
≡⟨ cong (p ⁻¹ ∙ r ≡_ ) (assoc (p ⁻¹) p s ∙∙ (cong (_∙ s) (lCancel p)) ∙∙ sym (lUnit s)) ⟩
p ⁻¹ ∙ r ≡ s
≡⟨ cong (λ z → p ⁻¹ ∙ z ≡ s) (rUnit r) ⟩
p ⁻¹ ∙ (r ∙ refl) ≡ s
≡⟨ cong (_≡ s) (sym (doubleCompPath-elim' (p ⁻¹) r refl)) ⟩
p ⁻¹ ∙∙ r ∙∙ refl ≡ s
≡⟨ sym (ua (Square≃doubleComp r s p refl)) ⟩
PathP (λ i → p i ≡ ptB) r s ∎
φ : P → Q
φ = transport P≡Q
totφ : (f g : Π∙ A B ptB) → f ∙∼ g → f ∙∼P g
totφ f g p .fst = p .fst
totφ f g p .snd = φ f g (p .fst) (p .snd)
∙∼→∙∼P : (f g : Π∙ A B ptB) → (f ∙∼ g) → (f ∙∼P g)
∙∼→∙∼P f g = totφ f g
∙∼≃∙∼P : (f g : Π∙ A B ptB) → (f ∙∼ g) ≃ (f ∙∼P g)
∙∼≃∙∼P f g = Σ-cong-equiv-snd (λ H → pathToEquiv (P≡Q f g H))
∙∼P→∙∼ : {f g : Π∙ A B ptB} → f ∙∼P g → f ∙∼ g
∙∼P→∙∼ {f = f} {g = g} = invEq (∙∼≃∙∼P f g)
∙∼≡∙∼P : (f g : Π∙ A B ptB) → (f ∙∼ g) ≡ (f ∙∼P g)
∙∼≡∙∼P f g = ua (∙∼≃∙∼P f g)
_∙∼Σ_ : (f g : Π∙ A B ptB) → Type (ℓ-max ℓ ℓ')
f ∙∼Σ g = Σ[ H ∈ f .fst ∼ g .fst ] (P f g H)
_∙∼PΣ_ : (f g : Π∙ A B ptB) → Type (ℓ-max ℓ ℓ')
f ∙∼PΣ g = Σ[ H ∈ f .fst ∼ g .fst ] (Q f g H)
∙∼≡∙∼Σ : (f g : Π∙ A B ptB) → f ∙∼ g ≡ f ∙∼Σ g
∙∼≡∙∼Σ f g = refl
∙∼P≡∙∼PΣ : (f g : Π∙ A B ptB) → f ∙∼P g ≡ f ∙∼PΣ g
∙∼P≡∙∼PΣ f g = refl