{-# OPTIONS --safe #-}
module Cubical.HITs.RPn.Base where
open import Cubical.Foundations.Prelude
open import Cubical.Foundations.Function
open import Cubical.Foundations.Equiv
open import Cubical.Foundations.Equiv.Properties
open import Cubical.Foundations.Isomorphism
open import Cubical.Foundations.Univalence
open import Cubical.Foundations.HLevels
open import Cubical.Functions.Bundle
open import Cubical.Foundations.SIP
open import Cubical.Structures.Pointed
open import Cubical.Structures.TypeEqvTo
open import Cubical.Data.Bool hiding (elim ; Bool*)
open import Cubical.Data.Nat hiding (elim)
open import Cubical.Data.NatMinusOne
open import Cubical.Data.Sigma
open import Cubical.Data.Unit
open import Cubical.Data.Empty as ⊥ hiding (elim)
open import Cubical.Data.Sum as ⊎ hiding (elim)
open import Cubical.HITs.PropositionalTruncation as PropTrunc hiding (elim)
open import Cubical.HITs.Sn
open import Cubical.HITs.Susp
open import Cubical.HITs.Join
open import Cubical.HITs.Pushout
open import Cubical.HITs.Pushout.Flattening
private
variable
ℓ ℓ' ℓ'' : Level
A : Type ℓ
data RP² : Type₀ where
point : RP²
line : point ≡ point
square : line ≡ sym line
2-EltType₀ = TypeEqvTo ℓ-zero Bool
2-EltPointed₀ = PointedEqvTo ℓ-zero Bool
Bool* : 2-EltType₀
Bool* = Bool , ∣ idEquiv _ ∣₁
isContrBoolPointedEquiv : ∀ x → isContr ((Bool , false) ≃[ PointedEquivStr ] (Bool , x))
fst (isContrBoolPointedEquiv x) = ((λ y → x ⊕ y) , isEquiv-⊕ x) , ⊕-comm x false
snd (isContrBoolPointedEquiv x) (e , p)
= Σ≡Prop (λ e → isSetBool (equivFun e false) x)
(Σ≡Prop isPropIsEquiv (funExt λ { false → ⊕-comm x false ∙ sym p
; true → ⊕-comm x true ∙ sym q }))
where q : e .fst true ≡ not x
q with dichotomyBool (invEq e (not x))
... | inl q = invEq≡→equivFun≡ e q
... | inr q = ⊥.rec (not≢const x (sym (invEq≡→equivFun≡ e q) ∙ p))
isContr-2-EltPointedEquiv : (X∙ : 2-EltPointed₀)
→ isContr ((Bool , false , ∣ idEquiv Bool ∣₁) ≃[ PointedEqvToEquivStr Bool ] X∙)
isContr-2-EltPointedEquiv (X , x , ∣e∣)
= PropTrunc.rec isPropIsContr
(λ e → J (λ X∙ _ → isContr ((Bool , false) ≃[ PointedEquivStr ] X∙))
(isContrBoolPointedEquiv (e .fst x))
(sym (pointed-sip _ _ (e , refl))))
∣e∣
module ⊕* (X : 2-EltType₀) where
_⊕*_ : typ X → typ X → Bool
y ⊕* z = invEquiv (fst (fst (isContr-2-EltPointedEquiv (fst X , y , snd X)))) .fst z
isEquivʳ : (y : typ X) → isEquiv (y ⊕*_)
isEquivʳ y = invEquiv (fst (fst (isContr-2-EltPointedEquiv (fst X , y , snd X)))) .snd
Equivʳ : typ X → typ X ≃ Bool
Equivʳ y = (y ⊕*_) , isEquivʳ y
elim : ∀ {ℓ'} (P : (A : Type₀) (_⊕'_ : A → A → Bool) → Type ℓ') (propP : ∀ A _⊕'_ → isProp (P A _⊕'_))
→ P Bool _⊕_ → P (typ X) _⊕*_
elim {ℓ'} P propP r = PropTrunc.elim {P = λ ∣e∣ → P (typ X) (R₁ ∣e∣)} (λ _ → propP _ _)
(λ e → EquivJ (λ A e → P A (R₂ A e)) r e)
(snd X)
where R₁ : ∥ fst X ≃ Bool ∥₁ → typ X → typ X → Bool
R₁ ∣e∣ y = invEq (fst (fst (isContr-2-EltPointedEquiv (fst X , y , ∣e∣))))
R₂ : (B : Type₀) → B ≃ Bool → B → B → Bool
R₂ A e y = invEq (fst (fst (J (λ A∙ _ → isContr ((Bool , false) ≃[ PointedEquivStr ] A∙))
(isContrBoolPointedEquiv (e .fst y))
(sym (pointed-sip (A , y) (Bool , e .fst y) (e , refl))))))
comm : (y z : typ X) → y ⊕* z ≡ z ⊕* y
comm = elim (λ A _⊕'_ → (x y : A) → x ⊕' y ≡ y ⊕' x)
(λ _ _ → isPropΠ2 (λ _ _ → isSetBool _ _))
⊕-comm
isEquivˡ : (y : typ X) → isEquiv (_⊕* y)
isEquivˡ y = subst isEquiv (funExt (λ z → comm y z)) (isEquivʳ y)
Equivˡ : typ X → typ X ≃ Bool
Equivˡ y = (_⊕* y) , isEquivˡ y
isContr-2-EltPointed : isContr (2-EltPointed₀)
fst isContr-2-EltPointed = (Bool , false , ∣ idEquiv Bool ∣₁)
snd isContr-2-EltPointed A∙ = PointedEqvTo-sip Bool _ A∙ (fst (isContr-2-EltPointedEquiv A∙))
RP : ℕ₋₁ → Type₀
cov⁻¹ : (n : ℕ₋₁) → RP n → 2-EltType₀
RP neg1 = ⊥
RP (ℕ→ℕ₋₁ n) = Pushout (pr (cov⁻¹ (-1+ n))) (λ _ → tt)
cov⁻¹ neg1 x = Bool*
cov⁻¹ (ℕ→ℕ₋₁ n) (inl x) = cov⁻¹ (-1+ n) x
cov⁻¹ (ℕ→ℕ₋₁ n) (inr _) = Bool*
cov⁻¹ (ℕ→ℕ₋₁ n) (push (x , y) i) = ua ((λ z → y ⊕* z) , ⊕*.isEquivʳ (cov⁻¹ (-1+ n) x) y) i , ∣p∣ i
where open ⊕* (cov⁻¹ (-1+ n) x)
∣p∣ = isProp→PathP (λ i → squash₁ {A = ua (⊕*.Equivʳ (cov⁻¹ (-1+ n) x) y) i ≃ Bool})
(str (cov⁻¹ (-1+ n) x)) (∣ idEquiv _ ∣₁)
TotalCov≃Sn : ∀ n → Total (cov⁻¹ n) ≃ S n
TotalCov≃Sn neg1 = isoToEquiv (iso (λ { () }) (λ { () }) (λ { () }) (λ { () }))
TotalCov≃Sn (ℕ→ℕ₋₁ n) =
Total (cov⁻¹ (ℕ→ℕ₋₁ n)) ≃⟨ i ⟩
Pushout Σf Σg ≃⟨ ii ⟩
join (Total (cov⁻¹ (-1+ n))) Bool ≃⟨ iii ⟩
S (ℕ→ℕ₋₁ n) ■
where
open FlatteningLemma (λ x → pr (cov⁻¹ (-1+ n)) x) (λ _ → tt)
(λ x → typ (cov⁻¹ (-1+ n) x)) (λ _ → Bool)
(λ { (x , y) → ⊕*.Equivʳ (cov⁻¹ (-1+ n) x) y })
hiding (Σf ; Σg)
cov⁻¹≃E : ∀ x → typ (cov⁻¹ (ℕ→ℕ₋₁ n) x) ≃ E x
cov⁻¹≃E (inl x) = idEquiv _
cov⁻¹≃E (inr x) = idEquiv _
cov⁻¹≃E (push a i) = idEquiv _
Σf : Σ[ x ∈ Total (cov⁻¹ (-1+ n)) ] typ (cov⁻¹ (-1+ n) (fst x)) → Total (cov⁻¹ (-1+ n))
Σg : Σ[ x ∈ Total (cov⁻¹ (-1+ n)) ] typ (cov⁻¹ (-1+ n) (fst x)) → Unit × Bool
Σf ((x , y) , z) = (x , z)
Σg ((x , y) , z) = (tt , y ⊕* z)
where open ⊕* (cov⁻¹ (-1+ n) x)
i : Total (cov⁻¹ (ℕ→ℕ₋₁ n)) ≃ Pushout Σf Σg
i = (Σ[ x ∈ RP (ℕ→ℕ₋₁ n) ] typ (cov⁻¹ (ℕ→ℕ₋₁ n) x)) ≃⟨ Σ-cong-equiv-snd cov⁻¹≃E ⟩
(Σ[ x ∈ RP (ℕ→ℕ₋₁ n) ] E x) ≃⟨ flatten ⟩
Pushout Σf Σg ■
u : ∀ {n} → (Σ[ x ∈ Total (cov⁻¹ n) ] typ (cov⁻¹ n (fst x))) ≃ (Total (cov⁻¹ n) × Bool)
u {n} = Σ[ x ∈ Total (cov⁻¹ n) ] typ (cov⁻¹ n (fst x)) ≃⟨ Σ-assoc-≃ ⟩
Σ[ x ∈ RP n ] (typ (cov⁻¹ n x)) × (typ (cov⁻¹ n x)) ≃⟨ Σ-cong-equiv-snd (λ x → Σ-swap-≃) ⟩
Σ[ x ∈ RP n ] (typ (cov⁻¹ n x)) × (typ (cov⁻¹ n x)) ≃⟨ Σ-cong-equiv-snd
(λ x → Σ-cong-equiv-snd
(λ y → ⊕*.Equivˡ (cov⁻¹ n x) y)) ⟩
Σ[ x ∈ RP n ] (typ (cov⁻¹ n x)) × Bool ≃⟨ invEquiv Σ-assoc-≃ ⟩
Total (cov⁻¹ n) × Bool ■
H : ∀ x → u .fst x ≡ (Σf x , snd (Σg x))
H x = refl
nat : 3-span-equiv (3span Σf Σg) (3span {A2 = Total (cov⁻¹ (-1+ n)) × Bool} fst snd)
nat = record { e0 = idEquiv _ ; e2 = u ; e4 = ΣUnit _
; H1 = λ x → cong fst (H x)
; H3 = λ x → cong snd (H x) }
ii : Pushout Σf Σg ≃ join (Total (cov⁻¹ (-1+ n))) Bool
ii = compEquiv (pathToEquiv (spanEquivToPushoutPath nat)) (joinPushout≃join _ _)
iii : join (Total (cov⁻¹ (-1+ n))) Bool ≃ S (ℕ→ℕ₋₁ n)
iii = join (Total (cov⁻¹ (-1+ n))) Bool ≃⟨ invEquiv Susp≃joinBool ⟩
Susp (Total (cov⁻¹ (-1+ n))) ≃⟨ congSuspEquiv (TotalCov≃Sn (-1+ n)) ⟩
S (ℕ→ℕ₋₁ n) ■
cov : (n : ℕ₋₁) → S n → RP n
cov n = pr (cov⁻¹ n) ∘ invEq (TotalCov≃Sn n)
fibcov≡cov⁻¹ : ∀ n (x : RP n) → fiber (cov n) x ≡ cov⁻¹ n x .fst
fibcov≡cov⁻¹ n x =
fiber (cov n) x ≡[ i ]⟨ fiber {A = ua e i} (pr (cov⁻¹ n) ∘ ua-unglue e i) x ⟩
fiber (pr (cov⁻¹ n)) x ≡⟨ ua (fibPrEquiv (cov⁻¹ n) x) ⟩
cov⁻¹ n x .fst ∎
where e = invEquiv (TotalCov≃Sn n)
RP0≃Unit : RP 0 ≃ Unit
RP0≃Unit = isoToEquiv (iso (λ _ → tt) (λ _ → inr tt) (λ _ → refl) (λ { (inr tt) → refl }))
RP1≡S1 : RP 1 ≡ S 1
RP1≡S1 = Pushout {A = Total (cov⁻¹ 0)} {B = RP 0} (pr (cov⁻¹ 0)) (λ _ → tt) ≡⟨ i ⟩
Pushout {A = Total (cov⁻¹ 0)} {B = Unit} (λ _ → tt) (λ _ → tt) ≡⟨ ii ⟩
Pushout {A = S 0} {B = Unit} (λ _ → tt) (λ _ → tt) ≡⟨ PushoutSusp≡Susp ⟩
S 1 ∎
where i = λ i → Pushout {A = Total (cov⁻¹ 0)}
{B = ua RP0≃Unit i}
(λ x → ua-gluePt RP0≃Unit i (pr (cov⁻¹ 0) x))
(λ _ → tt)
ii = λ j → Pushout {A = ua (TotalCov≃Sn 0) j} (λ _ → tt) (λ _ → tt)
RP²Fun : (a : A) (p : a ≡ a) (p∼p⁻¹ : p ≡ sym p)
→ RP² → A
RP²Fun a p p∼p⁻¹ point = a
RP²Fun a p p∼p⁻¹ (line i) = p i
RP²Fun a p p∼p⁻¹ (square i i₁) = p∼p⁻¹ i i₁
elimSetRP² : {A : RP² → Type ℓ} → ((x : RP²) → isSet (A x))
→ (point* : A point)
→ PathP (λ i → A (line i)) point* point*
→ (x : RP²) → A x
elimSetRP² set point* p point = point*
elimSetRP² set point* p (line i) = p i
elimSetRP² {A = A} set point* p (square i j) =
isOfHLevel→isOfHLevelDep 2 {B = A} set point* point* p (symP p) square i j
elimPropRP² : {A : RP² → Type ℓ} → ((x : RP²) → isProp (A x))
→ (point* : A point)
→ (x : RP²) → A x
elimPropRP² pr point* =
elimSetRP² (λ x → isProp→isSet (pr _))
point* (isProp→PathP (λ _ → pr _) _ _)
characRP²Fun : ∀ {ℓ} {A : Type ℓ} (f : RP² → A)
→ RP²Fun (f point) (cong f line) (λ i j → f (square i j)) ≡ f
characRP²Fun f =
funExt λ { point → refl ; (line i) → refl ; (square i i₁) → refl}
RP²FunCharacIso : {A : RP² → Type ℓ}
→ Iso ((x : RP²) → A x)
(Σ[ x ∈ A point ]
Σ[ p ∈ PathP (λ i → A (line i)) x x ]
SquareP (λ i j → A (square i j))
p (λ i → p (~ i)) refl refl)
Iso.fun RP²FunCharacIso f = f point , cong f line , cong (cong f) square
Iso.inv RP²FunCharacIso (x , p , q) point = x
Iso.inv RP²FunCharacIso (x , p , q) (line i) = p i
Iso.inv RP²FunCharacIso (x , p , q) (square i j) = q i j
Iso.rightInv RP²FunCharacIso _ = refl
Iso.leftInv RP²FunCharacIso f =
funExt λ { point → refl ; (line i) → refl ; (square i i₁) → refl}