From 4da7d053210483b56ef07e2e7447345577012592 Mon Sep 17 00:00:00 2001 From: Axel Date: Wed, 12 Feb 2020 23:43:03 +0100 Subject: [PATCH 01/89] First update --- Cubical/Foundations/HLevels.agda | 10 + Cubical/Foundations/Pointed/Base.agda | 5 + Cubical/Foundations/Prelude.agda | 1 + .../PropositionalTruncation/Properties.agda | 9 + Cubical/HITs/SetTruncation/Properties.agda | 20 ++ Cubical/HITs/Truncation/Base.agda | 1 + Cubical/HITs/Truncation/Properties.agda | 188 +++++++++++++++++- Cubical/ZCohomology/Base.agda | 51 +++++ Cubical/ZCohomology/Everything.agda | 7 + Cubical/ZCohomology/Properties.agda | 88 ++++++++ Cubical/ZCohomology/S1/S1.agda | 57 ++++++ 11 files changed, 435 insertions(+), 2 deletions(-) create mode 100644 Cubical/ZCohomology/Base.agda create mode 100644 Cubical/ZCohomology/Everything.agda create mode 100644 Cubical/ZCohomology/Properties.agda create mode 100644 Cubical/ZCohomology/S1/S1.agda diff --git a/Cubical/Foundations/HLevels.agda b/Cubical/Foundations/HLevels.agda index 0be65a054..b7e2c3324 100644 --- a/Cubical/Foundations/HLevels.agda +++ b/Cubical/Foundations/HLevels.agda @@ -246,6 +246,16 @@ hLevelSuc 0 A = isContr→isProp hLevelSuc 1 A = isProp→isSet hLevelSuc (suc (suc n)) A h a b = hLevelSuc (suc n) (a ≡ b) (h a b) +hLevelAdd : (n : ℕ) (m : ℕ) (A : Type ℓ) → isOfHLevel n A → isOfHLevel (n + m) A +hLevelAdd n zero A h = transport (λ i → isOfHLevel (helper n i) A) h where + helper : (n : ℕ) → n ≡ n + zero + helper zero = refl + helper (suc n) = cong (λ x → (suc x)) (helper n) +hLevelAdd n (suc m) A h = transport (λ i → (isOfHLevel (helper n m i) A )) (hLevelSuc (n + m) A (hLevelAdd n m A h)) where + helper : (n m : ℕ) → (suc (n + m)) ≡ (n + (suc m)) + helper zero m = refl + helper (suc n) m = cong (λ x → (suc x)) (helper n m) + hLevelPath : (n : ℕ) → isOfHLevel n A → (x y : A) → isOfHLevel n (x ≡ y) hLevelPath 0 h x y = isContrPath h x y hLevelPath 1 h x y = hLevelSuc zero _ (isProp→isContr≡ h x y) diff --git a/Cubical/Foundations/Pointed/Base.agda b/Cubical/Foundations/Pointed/Base.agda index 83ac03040..2c2d7e6d4 100644 --- a/Cubical/Foundations/Pointed/Base.agda +++ b/Cubical/Foundations/Pointed/Base.agda @@ -11,3 +11,8 @@ typ = fst pt : ∀ {ℓ} (A∙ : Pointed ℓ) → typ A∙ pt = snd + +{- Pointed functions -} +_→*_ : ∀{ℓ ℓ'} → (A : Pointed ℓ) → (B : Pointed ℓ') → Type (ℓ-max ℓ ℓ') +_→*_ A B = Σ (typ A → typ B) λ f → f (pt A) ≡ pt B + diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index b35b3af20..31086087a 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -131,6 +131,7 @@ transport p a = transp (λ i → p i) i0 a transportRefl : (x : A) → transport refl x ≡ x transportRefl {A = A} x i = transp (λ _ → A) i x + -- We want B to be explicit in subst subst : (B : A → Type ℓ') (p : x ≡ y) → B x → B y subst B p pa = transport (λ i → B (p i)) pa diff --git a/Cubical/HITs/PropositionalTruncation/Properties.agda b/Cubical/HITs/PropositionalTruncation/Properties.agda index 26c2b0dfe..a22739907 100644 --- a/Cubical/HITs/PropositionalTruncation/Properties.agda +++ b/Cubical/HITs/PropositionalTruncation/Properties.agda @@ -40,6 +40,15 @@ elimPropTrunc {A = A} {P = P} Pprop f (squash x y i) = PpropOver : {a b : ∥ A ∥} → (sq : a ≡ b) → ∀ x y → PathP (λ i → P (sq i)) x y PpropOver {a} = J (λ b (sq : a ≡ b) → ∀ x y → PathP (λ i → P (sq i)) x y) (Pprop a) + +propId : isProp A → ∥ A ∥ ≡ A +propId {A = A} hA = isoToPath (iso (recPropTrunc hA (idfun A)) (λ x → ∣ x ∣) (λ _ → refl) rinv) + where + rinv : ∀ (x : ∥ A ∥) → ∣ recPropTrunc hA (idfun A) x ∣ ≡ x + rinv x = elimPropTrunc {P = λ x → ∣ recPropTrunc hA (idfun A) x ∣ ≡ x} + (λ _ → isProp→isSet propTruncIsProp _ _) + (λ _ → refl) x + -- We could also define the eliminator using the recursor elimPropTrunc' : ∀ {P : ∥ A ∥ → Type ℓ} → ((a : ∥ A ∥) → isProp (P a)) → ((x : A) → P ∣ x ∣) → (a : ∥ A ∥) → P a diff --git a/Cubical/HITs/SetTruncation/Properties.agda b/Cubical/HITs/SetTruncation/Properties.agda index 6a9f364c7..ebe21aa20 100644 --- a/Cubical/HITs/SetTruncation/Properties.agda +++ b/Cubical/HITs/SetTruncation/Properties.agda @@ -56,3 +56,23 @@ elimSetTrunc3 : {B : (x y z : ∥ A ∥₀) → Type ℓ} (x y z : ∥ A ∥₀) → B x y z elimSetTrunc3 Bset g = elimSetTrunc2 (λ _ _ → hLevelPi 2 λ _ → Bset _ _ _) (λ a b → elimSetTrunc (λ _ → Bset _ _ _) (λ c → g a b c)) + + +setTruncIsSet : isSet ∥ A ∥₀ +setTruncIsSet a b p q = squash₀ a b p q + +setId : isSet A → A ≡ ∥ A ∥₀ +setId {A = A} isset = isoToPath (iso (λ x → ∣ x ∣₀) + (elimSetTrunc {A = A} + (λ _ → isset) + (λ x → x)) + (λ a → idLemma a) + (λ _ → refl)) + where + idLemma : ∀ (b : ∥ A ∥₀) → ∣ elimSetTrunc (λ x → isset) (λ x → x) b ∣₀ ≡ b + idLemma b = elimSetTrunc {B = (λ x → ∣ elimSetTrunc (λ _ → isset) (λ x → x) x ∣₀ ≡ x)} + (λ x → (hLevelSuc 2 ∥ A ∥₀ (setTruncIsSet {A = A})) + ∣ elimSetTrunc (λ _ → isset) (λ x₁ → x₁) x ∣₀ + x) + (λ _ → refl) + b diff --git a/Cubical/HITs/Truncation/Base.agda b/Cubical/HITs/Truncation/Base.agda index ef1993225..f69ddacb9 100644 --- a/Cubical/HITs/Truncation/Base.agda +++ b/Cubical/HITs/Truncation/Base.agda @@ -22,3 +22,4 @@ open import Cubical.HITs.Sn ∥_∥_ : ∀ {ℓ} → Type ℓ → ℕ₋₂ → Type ℓ ∥ A ∥ n = Null (S (1+ n)) A + diff --git a/Cubical/HITs/Truncation/Properties.agda b/Cubical/HITs/Truncation/Properties.agda index 231f19698..d720fe948 100644 --- a/Cubical/HITs/Truncation/Properties.agda +++ b/Cubical/HITs/Truncation/Properties.agda @@ -13,9 +13,9 @@ open Modality open import Cubical.Data.Empty open import Cubical.Data.Nat -open import Cubical.Data.NatMinusOne using (ℕ₋₁; neg1; suc; ℕ→ℕ₋₁) +open import Cubical.Data.NatMinusOne using (ℕ₋₁; neg1; suc; ℕ→ℕ₋₁) renaming (-1+_ to -1+₋₁_ ; 1+_ to 1+₋₁_) import Cubical.Data.NatMinusOne as ℕ₋₁ -open import Cubical.Data.NatMinusTwo +open import Cubical.Data.NatMinusTwo open import Cubical.HITs.Sn open import Cubical.HITs.Susp open import Cubical.HITs.Nullification @@ -124,10 +124,20 @@ isSnNull→isOfHLevel {n = suc n} nA = isSphereFilled→isOfHLevelSuc (λ f → isOfHLevel∥∥ : (n : ℕ₋₂) → isOfHLevel (2+ n) (∥ A ∥ n) isOfHLevel∥∥ neg2 = hub ⊥-elim , λ _ → ≡hub ⊥-elim isOfHLevel∥∥ (suc n) = isSphereFilled→isOfHLevelSuc isSphereFilled∥∥ + + + -- isOfHLevel∥∥ n = isSnNull→isOfHLevel isNull-Null -- ∥_∥ n is a modality +rec : {n : ℕ₋₂} + {B : Type ℓ'} → + (isOfHLevel (2+ n) B) → + (g : (a : A) → B) → + (∥ A ∥ n → B) +rec {B = B} h = Null-ind {B = λ _ → B} λ x → isOfHLevel→isSnNull h + ind : {n : ℕ₋₂} {B : ∥ A ∥ n → Type ℓ'} (hB : (x : ∥ A ∥ n) → isOfHLevel (2+ n) (B x)) @@ -145,6 +155,9 @@ ind2 : {n : ℕ₋₂} ind2 {n = n} hB g = ind (λ _ → hLevelPi (2+ n) (λ _ → hB _ _)) λ a → ind (λ _ → hB _ _) (λ b → g a b) + + + ind3 : {n : ℕ₋₂} {B : (x y z : ∥ A ∥ n) → Type ℓ'} (hB : ((x y z : ∥ A ∥ n) → isOfHLevel (2+ n) (B x y z))) @@ -204,3 +217,174 @@ groupoidTrunc≃Trunc1 = (ind (λ _ → squash₂) ∣_∣₂) (ind (λ _ → hLevelPath 4 (isOfHLevel∥∥ 2) _ _) (λ _ → refl)) (g2TruncElim _ _ (λ _ → hLevelPath 4 squash₂ _ _) (λ _ → refl))) + + + +---- ∥ Ω A ∥ ₙ ≡ Ω ∥ A ∥ₙ₊₁ ---- + + + {- Proofs of Theorem 7.3.12. and Corollary 7.3.13. in the HoTT book -} + +private + {- We define the fibration P to show a more general result -} + P : ∀ {ℓ} {B : Type ℓ}{n : ℕ₋₂} → ∥ B ∥ (suc n) → ∥ B ∥ (suc n) → Type ℓ + P x y = fst (P₁ x y) + + where + P₁ : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} → ∥ B ∥ (suc n) → ∥ B ∥ (suc n) → (HLevel ℓ (2+ n)) + P₁ {ℓ} {n = n} x y = ind2 (λ _ _ → hLevelHLevel (2+ n)) + (λ a b → (∥ a ≡ b ∥ n , isOfHLevel∥∥ n )) + x + y + + {- We will need P to be of hLevel n + 3 -} + hLevelP : ∀{ℓ} {n : ℕ₋₂} {B : Type ℓ} + (a b : ∥ B ∥ (suc n)) → + isOfHLevel (2+ (suc n)) (P a b ) + hLevelP {n = n} {B = B} = ind2 {A = B} + {n = (suc n)} + {B = λ x y → isOfHLevel (2+ (suc n)) (P x y)} + (λ x y → hLevelAdd 1 (2+ n) (isOfHLevel (2+ suc n) (P x y)) (isPropIsOfHLevel (2+ suc n) (P x y)) ) + λ a b → ( hLevelSuc (2+ n) (P ∣ a ∣ ∣ b ∣) ) + (isOfHLevel∥∥ {A = a ≡ b} n) + + {- decode function from P x y to x ≡ y -} + decode-fun : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (x y : ∥ B ∥ (suc n)) → + P x y → + _≡_ {A = ∥ B ∥ (suc n)} x y + + decode-fun {B = B} {n = n} x y = ind2 {B = λ u v → P u v → _≡_ {A = ∥ B ∥ (suc n)} u v } + (λ u v → hLevelPi {A = P u v} {B = λ _ → u ≡ v} + (2+ suc n) + λ _ → (((hLevelSuc (2+ suc n) (∥ B ∥ suc n) (isOfHLevel∥∥ {A = B} (suc n))) u v)) ) + decode* x y + where + decode* : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂}(u v : B) → + (P {n = n} + (∣_∣ {S = (S (1+ (suc n)))} u) + (∣_∣ {S = (S (1+ (suc n)))} v)) → + _≡_ {A = ∥ B ∥ (suc n) } + ∣ u ∣ + ∣ v ∣ + decode* {B = B} {n = neg2} u v = rec {A = u ≡ v} {n = neg2} + {B = _≡_ {A = ∥ B ∥ (suc (neg2)) } (∣ u ∣) (∣ v ∣)} + ((isOfHLevel∥∥ {A = B} (suc neg2) ∣ u ∣ ∣ v ∣) , + λ y → (hLevelSuc (2+ suc neg2) (∥ B ∥ suc neg2) + (isOfHLevel∥∥ {A = B} (suc neg2)) ∣ u ∣ ∣ v ∣) + (isOfHLevel∥∥ {A = B} (suc neg2) ∣ u ∣ ∣ v ∣) y ) + (λ p → cong (λ z → ∣ z ∣) p) + decode* {B = B} {n = suc n} u v = rec {A = u ≡ v} {n = suc n} + {B = _≡_ {A = ∥ B ∥ (suc (suc n)) } (∣ u ∣) (∣ v ∣)} + (isOfHLevel∥∥ {A = B} (suc (suc n)) ∣ u ∣ ∣ v ∣) + (λ p → cong (λ z → ∣ z ∣) p) + + + {- auxilliary function r used to define encode -} + r : ∀ {ℓ} {B : Type ℓ} {m : ℕ₋₂} (u : ∥ B ∥ (suc m)) → P u u + r {m = m} = ind {B = (λ u → P u u)} + (λ x → hLevelP x x) + (λ a → ∣ refl {x = a} ∣) + + {- encode function from x ≡ y to P x y -} + encode-fun : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (x y : ∥ B ∥ (suc n)) → + _≡_ {A = ∥ B ∥ (suc n)} x y → + P x y + encode-fun x y p = transport (λ i → P x (p i )) (r x) + + {- We need the following two lemmas on the functions behaviour for refl -} + dec-refl : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (x : ∥ B ∥ (suc n)) → + decode-fun x x (r x) ≡ refl {x = x} + dec-refl {B = B} {n = neg2} x = ind {A = B} {n = suc neg2} + {B = λ x → decode-fun x x (r x) ≡ refl {x = x} } + (λ x → (hLevelSuc (2+ (suc neg2)) (x ≡ x) + (hLevelSuc (2+ (suc neg2)) + (∥ B ∥ suc neg2) + (isOfHLevel∥∥ {A = B} (suc neg2)) x x)) + (decode-fun x x (r x)) refl) + (λ a → refl) x + dec-refl {B = B} {n = suc n} = ind {A = B} {n = suc (suc n)} + {B = λ x → decode-fun x x (r x) ≡ refl {x = x} } + (λ x → hLevelSuc (2+ suc n) (decode-fun x x (r x) ≡ refl) + (hLevelSuc (2+ suc n) (x ≡ x) + (isOfHLevel∥∥ {A = B} (suc (suc n)) x x ) + (decode-fun x x (r x)) refl)) + λ c → refl + + + enc-refl : ∀ {ℓ} {B : Type ℓ} + {n : ℕ₋₂} + (x : ∥ B ∥ (suc n)) → + encode-fun x x (refl {x = x}) ≡ r x + enc-refl x j = transp (λ i → P x (refl {x = x} i)) j (r x) + + + + {- decode-fun is a right-inverse -} + P-rinv : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (u v : ∥ B ∥ (suc n)) → + (x : _≡_ {A = ∥ B ∥ (suc n)} u v) → + decode-fun u v (encode-fun u v x) ≡ x + P-rinv {ℓ = ℓ} {B = B} {n = n} u v = J {ℓ} { ∥ B ∥ (suc n)} {u} {ℓ} + (λ y p → decode-fun u y (encode-fun u y p) ≡ p) + ((λ i → (decode-fun u u (enc-refl u i))) ∙ dec-refl u) + {v} + + {- decode-fun is a left-inverse -} + P-linv : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (u v : ∥ B ∥ (suc n )) → + (x : P u v) → + encode-fun u v (decode-fun u v x) ≡ x + P-linv {ℓ = ℓ} {B = B} {n = n} u v = ind2 {A = B} + {n = (suc n)} + {B = λ u v → (x : P u v) → encode-fun u v (decode-fun u v x) ≡ x} + (λ x y → hLevelPi {A = P x y} + (2+ suc n) + λ z → hLevelSuc (2+ suc n) (P x y) + (hLevelP {n = n} x y) (encode-fun x y (decode-fun x y z)) z) + helper u v + + + where + helper : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} + (a b : B) + (x : P {n = n} ∣ a ∣ ∣ b ∣) → + encode-fun ∣ a ∣ ∣ b ∣ (decode-fun ∣ a ∣ ∣ b ∣ x) ≡ x + helper {ℓ = ℓ} {B = B} {n = neg2} a b = ind {A = (a ≡ b)} + {n = neg2} + {B = λ x → encode-fun ∣ a ∣ ∣ b ∣ (decode-fun ∣ a ∣ ∣ b ∣ x) ≡ x} + (λ x → (sym ((snd (isOfHLevel∥∥ {A = a ≡ b} neg2)) + (encode-fun ∣ a ∣ ∣ b ∣ (decode-fun ∣ a ∣ ∣ b ∣ x))) + ∙ ((snd (isOfHLevel∥∥ {A = a ≡ b} neg2)) x)) + , + λ y → hLevelSuc (2+ (suc neg2)) (∥ a ≡ b ∥ neg2) + (hLevelSuc (2+ neg2) (∥ a ≡ b ∥ neg2) (isOfHLevel∥∥ {A = a ≡ b} (neg2))) + (encode-fun ∣ a ∣ ∣ b ∣ (decode-fun ∣ a ∣ ∣ b ∣ x)) x + ((sym ((snd (isOfHLevel∥∥ {A = a ≡ b} neg2)) + (encode-fun ∣ a ∣ ∣ b ∣ (decode-fun ∣ a ∣ ∣ b ∣ x))) + ∙ ((snd (isOfHLevel∥∥ {A = a ≡ b} neg2)) x))) + y) + (J {ℓ}{B}{a}{ℓ} + ((λ y p → (encode-fun {n = neg2}) ∣ a ∣ ∣ y ∣ + ((decode-fun ∣ a ∣ ∣ y ∣) ∣ p ∣) ≡ ∣ p ∣)) + (enc-refl {n = neg2} ∣ a ∣ ) + {b}) + helper {ℓ = ℓ} {B = B} {n = suc n} a b = ind {A = (a ≡ b)} + {n = suc n} + {B = λ x → encode-fun ∣ a ∣ ∣ b ∣ (decode-fun ∣ a ∣ ∣ b ∣ x) ≡ x} + (λ x → (hLevelP {n = suc n} ∣ a ∣ ∣ b ∣ (encode-fun ∣ a ∣ ∣ b ∣ (decode-fun ∣ a ∣ ∣ b ∣ x)) x) ) + (J {ℓ}{B}{a}{ℓ} ((λ y p → (encode-fun {n = suc n}) ∣ a ∣ ∣ y ∣ ((decode-fun ∣ a ∣ ∣ y ∣) ∣ p ∣) ≡ ∣ p ∣)) + (enc-refl {n = suc n} ∣ a ∣ ) + {b}) + + {- The final Iso established -} + IsoFinal : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (x y : ∥ B ∥ (suc n)) → Iso (x ≡ y) (P x y) + IsoFinal x y = iso (encode-fun x y ) (decode-fun x y) (P-linv x y) (P-rinv x y) + +IsoIdTrunc : {a b : A} (n : ℕ₋₂) → Iso (_≡_ {A = ∥ A ∥ (suc n)} ∣ a ∣ ∣ b ∣) (∥ a ≡ b ∥ n) +IsoIdTrunc {a = a} {b = b} n = IsoFinal {n = n} ∣ a ∣ ∣ b ∣ + +IsoΩ : {a : A} (n : ℕ₋₂) → Iso (_≡_ {A = ∥ A ∥ (suc n)} ∣ a ∣ ∣ a ∣) (∥ a ≡ a ∥ n) +IsoΩ {a = a} n = IsoIdTrunc {a = a} {b = a} n + + + + + diff --git a/Cubical/ZCohomology/Base.agda b/Cubical/ZCohomology/Base.agda new file mode 100644 index 000000000..8732a2aae --- /dev/null +++ b/Cubical/ZCohomology/Base.agda @@ -0,0 +1,51 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.Base where + + +open import Cubical.Data.Int.Base +open import Cubical.Data.Nat.Base +open import Cubical.Data.NatMinusTwo.Base +open import Cubical.Data.Sigma.Base + +open import Cubical.Foundations.Pointed.Base + +open import Cubical.HITs.Nullification.Base +open import Cubical.HITs.SetTruncation.Base +open import Cubical.HITs.Sn.Base +open import Cubical.HITs.Susp.Base +open import Cubical.HITs.Truncation.Base + +private + variable + ℓ : Level + A : Type ℓ + +{- +Definition of cohomology with cocefficients in the integers. For now, this definition uses ℕ₋₂, for instance defining H⁰ as +coHom neg2, H¹ as coHom (suc neg2), and so on. This has so far been easier to work with than a definition relying on ℕ→ℕ₋₂. +-} + + +--- Cohomology --- + +{- Types Kₙ from Brunerie 2016 -} +coHomK : (n : ℕ₋₂) → Type₀ +coHomK neg2 = Int +coHomK (suc n) = ∥ S₊ (2+ suc n) ∥ (suc (suc (suc n))) + +{- Cohomology -} +coHom : (n : ℕ₋₂) → Type ℓ → Type ℓ +coHom n A = ∥ (A → coHomK n) ∥₀ + + +--- Reduced cohomology --- + +{- Pointed version of Kₙ -} +coHomK-ptd : (n : ℕ₋₂) → Pointed (ℓ-zero) +coHomK-ptd neg2 = coHomK neg2 , (pos 0) +coHomK-ptd (suc n) = (coHomK (suc n) , ∣ north ∣) + +{- Reduced cohomology -} +coHomRed : ∀ {ℓ} → (n : ℕ₋₂) → (A : Pointed ℓ) → Type ℓ +coHomRed neg2 A = ∥ (A →* (coHomK-ptd neg2)) ∥₀ +coHomRed (suc n) A = ∥ (A →* (coHomK-ptd (suc n))) ∥₀ diff --git a/Cubical/ZCohomology/Everything.agda b/Cubical/ZCohomology/Everything.agda new file mode 100644 index 000000000..160104dce --- /dev/null +++ b/Cubical/ZCohomology/Everything.agda @@ -0,0 +1,7 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.Everything where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties + +open import Cubical.ZCohomology.S1.S1 diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda new file mode 100644 index 000000000..b4a2583d1 --- /dev/null +++ b/Cubical/ZCohomology/Properties.agda @@ -0,0 +1,88 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.Properties where + + +open import Cubical.ZCohomology.Base +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Data.NatMinusTwo.Base renaming (-1+_ to -1+₋₂_ ; 1+_ to 1+₋₂_) +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation +open import Cubical.HITs.Nullification +open import Cubical.Data.Int hiding (_+_) +open import Cubical.Data.Nat +open import Cubical.HITs.Truncation + +open import Cubical.HITs.Pushout +open import Cubical.Data.Sum.Base +open import Cubical.Data.HomotopyGroup +private + variable + ℓ : Level + A : Type ℓ + + + + + +{- Equivalence between cohomology and the reduced version -} +coHom↔coHomRed : ∀ {ℓ} → (n : ℕ₋₂) → (A : Set ℓ) → (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) +coHom↔coHomRed neg2 A i = ∥ helplemma {C = (Int , pos 0)} i ∥₀ + module ClaireCaron where + helplemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →* C))) + helplemma {C = C} = isoToPath (iso map1 map2 (λ b → linvPf b) (λ a → refl)) where + map1 : (A → typ C) → ((((A ⊎ Unit) , inr (tt)) →* C)) + map1 f = map1' , refl module helpmap where + map1' : A ⊎ Unit → fst C + map1' (inl x) = f x + map1' (inr x) = pt C + + map2 : ((((A ⊎ Unit) , inr (tt)) →* C)) → (A → typ C) + map2 (g , pf) x = g (inl x) + linvPf : (b :((((A ⊎ Unit) , inr (tt)) →* C))) → map1 (map2 b) ≡ b + linvPf (f , snd) i = (λ x → helper x i) , λ j → snd ((~ i) ∨ j) where + helper : (x : A ⊎ Unit) → ((helpmap.map1') (map2 (f , snd)) x) ≡ f x + helper (inl x) = refl + helper (inr tt) = sym snd + +coHom↔coHomRed (suc n) A i = ∥ ClaireCaron.helplemma A i {C = ((K (suc n)) , ∣ north ∣)} i ∥₀ + + + + +{- Unfinished proof that Kₙ ≡ Ω Kₙ₊₁ -} + +invPathCancel : {a b : A} → (p : a ≡ b) → p ∙ (sym p) ≡ refl +invPathCancel {a = a} {b = b} p = J {x = a} (λ y p → p ∙ (sym p) ≡ refl ) (sym (lUnit refl)) p + +φ : (a : A) → A → (north {A = A} ≡ north {A = A}) +φ x = (λ a → ((merid a) ∙ sym ((merid x)))) + +φ* : (A : Pointed ℓ) → A →* Ω ((Susp (typ A)) , north) +φ* A = (φ (pt A)) , invPathCancel (merid (pt A)) + +σ : (n : ℕ₋₂) → (typ (K-ptd n)) → (typ (Ω (K-ptd (suc n)))) +σ neg2 k = looper k ( cong {B = λ _ → (∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1)} + (λ x → ∣ x ∣) + ((merid north) ∙ (sym (merid south))) ) + where + looper : ∀ {ℓ} → {A : Type ℓ} → {a : A} → (n : Int) → (a ≡ a) → (a ≡ a) -- defining loopⁿ + looper {A = A} {a = a} (pos zero) p = refl + looper {A = A} {a = a} (pos (suc n)) p = (looper (pos n) p) ∙ p + looper {A = A} {a = a} (negsuc zero) p = sym p + looper {A = A} {a = a} (negsuc (suc n)) p = (sym p) ∙ (looper (negsuc n) p) +σ (suc n) x = rec {n = suc (suc (suc n))} + {B = (typ (Ω (K-ptd (suc (suc n)))))} + (isOfHLevel∥∥ {A = S₊ (2+ suc (suc n))} (suc (suc (suc (suc n)))) ∣ north ∣ ∣ north ∣) + (λ y → cong {B = λ _ → Null (S (1+₋₂ (suc (suc (suc (suc n)))))) (S₊ (2+ (suc (suc n))))} + (λ z → ∣ z ∣) + (φ north y)) + x + + + + diff --git a/Cubical/ZCohomology/S1/S1.agda b/Cubical/ZCohomology/S1/S1.agda new file mode 100644 index 000000000..605253014 --- /dev/null +++ b/Cubical/ZCohomology/S1/S1.agda @@ -0,0 +1,57 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.S1.S1 where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.Data.Sigma +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.Data.Empty.Base +open import Cubical.Data.Unit.Base +open import Cubical.Foundations.Everything +open import Cubical.Data.NatMinusTwo.Base renaming (-1+_ to -1+₋₂_ ; 1+_ to 1+₋₂_) +open import Cubical.Foundations.Equiv +open import Cubical.Data.Prod +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation +open import Cubical.HITs.Nullification +open import Cubical.Data.Int hiding (_+_) +open import Cubical.Data.Nat +open import Cubical.HITs.Truncation +open import Cubical.HITs.Pushout +open import Cubical.Data.Sum.Base +open import Cubical.Data.Unit.Base +open import Cubical.Data.HomotopyGroup +private + variable + ℓ : Level + A : Type ℓ + + + +---- H⁰(S¹) ---- + +coHom0-S1 : (coHom neg2 S¹) ≡ Int +coHom0-S1 = (λ i → ∥ helpLemma i ∥₀ ) ∙ sym (setId isSetInt) + where + helpLemma : (S¹ → Int) ≡ Int + helpLemma = isoToPath (iso fun funinv (λ b → refl) λ f → funExt (λ x → rinvLemma f x)) where + fun : (S¹ → Int) → Int + fun f = f base + + funinv : Int → (S¹ → Int) + funinv a base = a + funinv a (loop i) = refl {x = a} i + + rinvLemma : ( f : (S¹ → Int)) → (x : S¹) → funinv (fun f) x ≡ f x + rinvLemma f base = refl + rinvLemma f (loop i) = sym (helper f i) where + helper : (f : (S¹ → Int)) → (i : I) → (f (loop i) ≡ f base) + helper f i = (λ l → ((isSetInt (f base) (f base) (λ k → (f (loop k))) refl) l) i) + +------------------------- + + + + + From 56a5e0f5b5e0c271a482a03673ea354acb693807 Mon Sep 17 00:00:00 2001 From: Axel Date: Thu, 13 Feb 2020 00:06:24 +0100 Subject: [PATCH 02/89] Minor fixes --- Cubical/ZCohomology/Properties.agda | 19 ++++++++++--------- Cubical/ZCohomology/S1/S1.agda | 4 ++-- 2 files changed, 12 insertions(+), 11 deletions(-) diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index b4a2583d1..de90ec943 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -13,7 +13,7 @@ open import Cubical.Data.NatMinusTwo.Base renaming (-1+_ to -1+₋₂_ ; 1+_ to open import Cubical.HITs.Susp open import Cubical.HITs.SetTruncation open import Cubical.HITs.Nullification -open import Cubical.Data.Int hiding (_+_) +open import Cubical.Data.Int open import Cubical.Data.Nat open import Cubical.HITs.Truncation @@ -27,12 +27,10 @@ private - - {- Equivalence between cohomology and the reduced version -} coHom↔coHomRed : ∀ {ℓ} → (n : ℕ₋₂) → (A : Set ℓ) → (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) coHom↔coHomRed neg2 A i = ∥ helplemma {C = (Int , pos 0)} i ∥₀ - module ClaireCaron where + module coHom↔coHomRed where helplemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →* C))) helplemma {C = C} = isoToPath (iso map1 map2 (λ b → linvPf b) (λ a → refl)) where map1 : (A → typ C) → ((((A ⊎ Unit) , inr (tt)) →* C)) @@ -49,12 +47,12 @@ coHom↔coHomRed neg2 A i = ∥ helplemma {C = (Int , pos 0)} i ∥₀ helper (inl x) = refl helper (inr tt) = sym snd -coHom↔coHomRed (suc n) A i = ∥ ClaireCaron.helplemma A i {C = ((K (suc n)) , ∣ north ∣)} i ∥₀ +coHom↔coHomRed (suc n) A i = ∥ coHom↔coHomRed.helplemma A i {C = ((coHomK (suc n)) , ∣ north ∣)} i ∥₀ -{- Unfinished proof that Kₙ ≡ Ω Kₙ₊₁ -} +{- TODO: Prove Kₙ ≡ Ω Kₙ₊₁ -} invPathCancel : {a b : A} → (p : a ≡ b) → p ∙ (sym p) ≡ refl invPathCancel {a = a} {b = b} p = J {x = a} (λ y p → p ∙ (sym p) ≡ refl ) (sym (lUnit refl)) p @@ -65,18 +63,21 @@ invPathCancel {a = a} {b = b} p = J {x = a} (λ y p → p ∙ (sym p) ≡ refl ) φ* : (A : Pointed ℓ) → A →* Ω ((Susp (typ A)) , north) φ* A = (φ (pt A)) , invPathCancel (merid (pt A)) -σ : (n : ℕ₋₂) → (typ (K-ptd n)) → (typ (Ω (K-ptd (suc n)))) +σ : (n : ℕ₋₂) → (typ (coHomK-ptd n)) → (typ (Ω (coHomK-ptd (suc n)))) σ neg2 k = looper k ( cong {B = λ _ → (∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1)} (λ x → ∣ x ∣) ((merid north) ∙ (sym (merid south))) ) where - looper : ∀ {ℓ} → {A : Type ℓ} → {a : A} → (n : Int) → (a ≡ a) → (a ≡ a) -- defining loopⁿ + looper : ∀ {ℓ}{A : Type ℓ}{a : A} -- defining loopⁿ + (n : Int) → + (a ≡ a) → + (a ≡ a) looper {A = A} {a = a} (pos zero) p = refl looper {A = A} {a = a} (pos (suc n)) p = (looper (pos n) p) ∙ p looper {A = A} {a = a} (negsuc zero) p = sym p looper {A = A} {a = a} (negsuc (suc n)) p = (sym p) ∙ (looper (negsuc n) p) σ (suc n) x = rec {n = suc (suc (suc n))} - {B = (typ (Ω (K-ptd (suc (suc n)))))} + {B = (typ (Ω (coHomK-ptd (suc (suc n)))))} (isOfHLevel∥∥ {A = S₊ (2+ suc (suc n))} (suc (suc (suc (suc n)))) ∣ north ∣ ∣ north ∣) (λ y → cong {B = λ _ → Null (S (1+₋₂ (suc (suc (suc (suc n)))))) (S₊ (2+ (suc (suc n))))} (λ z → ∣ z ∣) diff --git a/Cubical/ZCohomology/S1/S1.agda b/Cubical/ZCohomology/S1/S1.agda index 605253014..c349d04ed 100644 --- a/Cubical/ZCohomology/S1/S1.agda +++ b/Cubical/ZCohomology/S1/S1.agda @@ -15,7 +15,7 @@ open import Cubical.Data.Prod open import Cubical.HITs.Susp open import Cubical.HITs.SetTruncation open import Cubical.HITs.Nullification -open import Cubical.Data.Int hiding (_+_) +open import Cubical.Data.Int open import Cubical.Data.Nat open import Cubical.HITs.Truncation open import Cubical.HITs.Pushout @@ -51,7 +51,7 @@ coHom0-S1 = (λ i → ∥ helpLemma i ∥₀ ) ∙ sym (setId isSetInt) ------------------------- - +{- TODO : give Hᵏ(S¹) for all k -} From 098153c949c8731c6948f4c03756d32d934e37b5 Mon Sep 17 00:00:00 2001 From: Axel Date: Thu, 13 Feb 2020 00:14:40 +0100 Subject: [PATCH 03/89] Cleanup --- Cubical/ZCohomology/Base.agda | 2 +- Cubical/ZCohomology/Properties.agda | 22 +++++++++++++++------- 2 files changed, 16 insertions(+), 8 deletions(-) diff --git a/Cubical/ZCohomology/Base.agda b/Cubical/ZCohomology/Base.agda index 8732a2aae..1c02e2057 100644 --- a/Cubical/ZCohomology/Base.agda +++ b/Cubical/ZCohomology/Base.agda @@ -46,6 +46,6 @@ coHomK-ptd neg2 = coHomK neg2 , (pos 0) coHomK-ptd (suc n) = (coHomK (suc n) , ∣ north ∣) {- Reduced cohomology -} -coHomRed : ∀ {ℓ} → (n : ℕ₋₂) → (A : Pointed ℓ) → Type ℓ +coHomRed : (n : ℕ₋₂) → (A : Pointed ℓ) → Type ℓ coHomRed neg2 A = ∥ (A →* (coHomK-ptd neg2)) ∥₀ coHomRed (suc n) A = ∥ (A →* (coHomK-ptd (suc n))) ∥₀ diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index de90ec943..139514488 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -28,11 +28,17 @@ private {- Equivalence between cohomology and the reduced version -} -coHom↔coHomRed : ∀ {ℓ} → (n : ℕ₋₂) → (A : Set ℓ) → (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) +coHom↔coHomRed : (n : ℕ₋₂) → + (A : Set ℓ) → + (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) coHom↔coHomRed neg2 A i = ∥ helplemma {C = (Int , pos 0)} i ∥₀ module coHom↔coHomRed where helplemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →* C))) - helplemma {C = C} = isoToPath (iso map1 map2 (λ b → linvPf b) (λ a → refl)) where + helplemma {C = C} = isoToPath (iso map1 + map2 + (λ b → linvPf b) + (λ _ → refl)) + where map1 : (A → typ C) → ((((A ⊎ Unit) , inr (tt)) →* C)) map1 f = map1' , refl module helpmap where map1' : A ⊎ Unit → fst C @@ -41,8 +47,10 @@ coHom↔coHomRed neg2 A i = ∥ helplemma {C = (Int , pos 0)} i ∥₀ map2 : ((((A ⊎ Unit) , inr (tt)) →* C)) → (A → typ C) map2 (g , pf) x = g (inl x) + linvPf : (b :((((A ⊎ Unit) , inr (tt)) →* C))) → map1 (map2 b) ≡ b - linvPf (f , snd) i = (λ x → helper x i) , λ j → snd ((~ i) ∨ j) where + linvPf (f , snd) i = (λ x → helper x i) , λ j → snd ((~ i) ∨ j) + where helper : (x : A ⊎ Unit) → ((helpmap.map1') (map2 (f , snd)) x) ≡ f x helper (inl x) = refl helper (inr tt) = sym snd @@ -68,10 +76,10 @@ invPathCancel {a = a} {b = b} p = J {x = a} (λ y p → p ∙ (sym p) ≡ refl ) (λ x → ∣ x ∣) ((merid north) ∙ (sym (merid south))) ) where - looper : ∀ {ℓ}{A : Type ℓ}{a : A} -- defining loopⁿ - (n : Int) → - (a ≡ a) → - (a ≡ a) + looper : {a : A} -- defining loopⁿ + (n : Int) → + (a ≡ a) → + (a ≡ a) looper {A = A} {a = a} (pos zero) p = refl looper {A = A} {a = a} (pos (suc n)) p = (looper (pos n) p) ∙ p looper {A = A} {a = a} (negsuc zero) p = sym p From 67a450d928db1674957a1596a50074ac425dccdf Mon Sep 17 00:00:00 2001 From: Axel Date: Thu, 13 Feb 2020 00:19:05 +0100 Subject: [PATCH 04/89] Cleanup --- Cubical/ZCohomology/S1/S1.agda | 14 +++++++++++--- 1 file changed, 11 insertions(+), 3 deletions(-) diff --git a/Cubical/ZCohomology/S1/S1.agda b/Cubical/ZCohomology/S1/S1.agda index c349d04ed..311fc0590 100644 --- a/Cubical/ZCohomology/S1/S1.agda +++ b/Cubical/ZCohomology/S1/S1.agda @@ -35,7 +35,11 @@ coHom0-S1 : (coHom neg2 S¹) ≡ Int coHom0-S1 = (λ i → ∥ helpLemma i ∥₀ ) ∙ sym (setId isSetInt) where helpLemma : (S¹ → Int) ≡ Int - helpLemma = isoToPath (iso fun funinv (λ b → refl) λ f → funExt (λ x → rinvLemma f x)) where + helpLemma = isoToPath (iso fun + funinv + (λ b → refl) + λ f → funExt (λ x → rinvLemma f x)) + where fun : (S¹ → Int) → Int fun f = f base @@ -43,10 +47,14 @@ coHom0-S1 = (λ i → ∥ helpLemma i ∥₀ ) ∙ sym (setId isSetInt) funinv a base = a funinv a (loop i) = refl {x = a} i - rinvLemma : ( f : (S¹ → Int)) → (x : S¹) → funinv (fun f) x ≡ f x + rinvLemma : ( f : (S¹ → Int)) → + (x : S¹) → + funinv (fun f) x ≡ f x rinvLemma f base = refl rinvLemma f (loop i) = sym (helper f i) where - helper : (f : (S¹ → Int)) → (i : I) → (f (loop i) ≡ f base) + helper : (f : (S¹ → Int)) → + (i : I) → + (f (loop i) ≡ f base) helper f i = (λ l → ((isSetInt (f base) (f base) (λ k → (f (loop k))) refl) l) i) ------------------------- From 1df167b1d7c5c02474d213ec9129b2a8ff811897 Mon Sep 17 00:00:00 2001 From: Axel Date: Mon, 24 Feb 2020 16:19:36 +0100 Subject: [PATCH 05/89] fixes --- Cubical/Foundations/HLevels.agda | 10 ---- Cubical/Foundations/Pointed/Base.agda | 4 +- Cubical/Foundations/Prelude.agda | 1 - Cubical/HITs/SetTruncation/Properties.agda | 10 ++-- Cubical/HITs/Truncation/Base.agda | 1 - Cubical/HITs/Truncation/Properties.agda | 17 ++---- Cubical/ZCohomology/Base.agda | 24 +++----- Cubical/ZCohomology/Properties.agda | 65 ++++++---------------- Cubical/ZCohomology/S1/S1.agda | 24 ++------ 9 files changed, 44 insertions(+), 112 deletions(-) diff --git a/Cubical/Foundations/HLevels.agda b/Cubical/Foundations/HLevels.agda index b7e2c3324..0be65a054 100644 --- a/Cubical/Foundations/HLevels.agda +++ b/Cubical/Foundations/HLevels.agda @@ -246,16 +246,6 @@ hLevelSuc 0 A = isContr→isProp hLevelSuc 1 A = isProp→isSet hLevelSuc (suc (suc n)) A h a b = hLevelSuc (suc n) (a ≡ b) (h a b) -hLevelAdd : (n : ℕ) (m : ℕ) (A : Type ℓ) → isOfHLevel n A → isOfHLevel (n + m) A -hLevelAdd n zero A h = transport (λ i → isOfHLevel (helper n i) A) h where - helper : (n : ℕ) → n ≡ n + zero - helper zero = refl - helper (suc n) = cong (λ x → (suc x)) (helper n) -hLevelAdd n (suc m) A h = transport (λ i → (isOfHLevel (helper n m i) A )) (hLevelSuc (n + m) A (hLevelAdd n m A h)) where - helper : (n m : ℕ) → (suc (n + m)) ≡ (n + (suc m)) - helper zero m = refl - helper (suc n) m = cong (λ x → (suc x)) (helper n m) - hLevelPath : (n : ℕ) → isOfHLevel n A → (x y : A) → isOfHLevel n (x ≡ y) hLevelPath 0 h x y = isContrPath h x y hLevelPath 1 h x y = hLevelSuc zero _ (isProp→isContr≡ h x y) diff --git a/Cubical/Foundations/Pointed/Base.agda b/Cubical/Foundations/Pointed/Base.agda index 2c2d7e6d4..d796dee28 100644 --- a/Cubical/Foundations/Pointed/Base.agda +++ b/Cubical/Foundations/Pointed/Base.agda @@ -13,6 +13,8 @@ pt : ∀ {ℓ} (A∙ : Pointed ℓ) → typ A∙ pt = snd {- Pointed functions -} -_→*_ : ∀{ℓ ℓ'} → (A : Pointed ℓ) → (B : Pointed ℓ') → Type (ℓ-max ℓ ℓ') +_→*_ : ∀{ℓ ℓ'} → (A : Pointed ℓ) (B : Pointed ℓ') → Type (ℓ-max ℓ ℓ') _→*_ A B = Σ (typ A → typ B) λ f → f (pt A) ≡ pt B +_→*_/_ : ∀{ℓ ℓ'} → (A : Pointed ℓ) (B : Pointed ℓ') → (A →* B) → Pointed (ℓ-max ℓ ℓ') +A →* B / f = (A →* B) , f diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index 31086087a..b35b3af20 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -131,7 +131,6 @@ transport p a = transp (λ i → p i) i0 a transportRefl : (x : A) → transport refl x ≡ x transportRefl {A = A} x i = transp (λ _ → A) i x - -- We want B to be explicit in subst subst : (B : A → Type ℓ') (p : x ≡ y) → B x → B y subst B p pa = transport (λ i → B (p i)) pa diff --git a/Cubical/HITs/SetTruncation/Properties.agda b/Cubical/HITs/SetTruncation/Properties.agda index ebe21aa20..9c7393674 100644 --- a/Cubical/HITs/SetTruncation/Properties.agda +++ b/Cubical/HITs/SetTruncation/Properties.agda @@ -61,13 +61,13 @@ elimSetTrunc3 Bset g = elimSetTrunc2 (λ _ _ → hLevelPi 2 λ _ → Bset _ _ _) setTruncIsSet : isSet ∥ A ∥₀ setTruncIsSet a b p q = squash₀ a b p q -setId : isSet A → A ≡ ∥ A ∥₀ -setId {A = A} isset = isoToPath (iso (λ x → ∣ x ∣₀) - (elimSetTrunc {A = A} +setId : isSet A → ∥ A ∥₀ ≡ A +setId {A = A} isset = isoToPath (iso (elimSetTrunc {A = A} (λ _ → isset) (λ x → x)) - (λ a → idLemma a) - (λ _ → refl)) + (λ x → ∣ x ∣₀) + (λ b → refl) + λ b → idLemma b) where idLemma : ∀ (b : ∥ A ∥₀) → ∣ elimSetTrunc (λ x → isset) (λ x → x) b ∣₀ ≡ b idLemma b = elimSetTrunc {B = (λ x → ∣ elimSetTrunc (λ _ → isset) (λ x → x) x ∣₀ ≡ x)} diff --git a/Cubical/HITs/Truncation/Base.agda b/Cubical/HITs/Truncation/Base.agda index f69ddacb9..ef1993225 100644 --- a/Cubical/HITs/Truncation/Base.agda +++ b/Cubical/HITs/Truncation/Base.agda @@ -22,4 +22,3 @@ open import Cubical.HITs.Sn ∥_∥_ : ∀ {ℓ} → Type ℓ → ℕ₋₂ → Type ℓ ∥ A ∥ n = Null (S (1+ n)) A - diff --git a/Cubical/HITs/Truncation/Properties.agda b/Cubical/HITs/Truncation/Properties.agda index d720fe948..aa9a1f7f5 100644 --- a/Cubical/HITs/Truncation/Properties.agda +++ b/Cubical/HITs/Truncation/Properties.agda @@ -125,7 +125,6 @@ isOfHLevel∥∥ : (n : ℕ₋₂) → isOfHLevel (2+ n) (∥ A ∥ n) isOfHLevel∥∥ neg2 = hub ⊥-elim , λ _ → ≡hub ⊥-elim isOfHLevel∥∥ (suc n) = isSphereFilled→isOfHLevelSuc isSphereFilled∥∥ - -- isOfHLevel∥∥ n = isSnNull→isOfHLevel isNull-Null @@ -244,7 +243,8 @@ private hLevelP {n = n} {B = B} = ind2 {A = B} {n = (suc n)} {B = λ x y → isOfHLevel (2+ (suc n)) (P x y)} - (λ x y → hLevelAdd 1 (2+ n) (isOfHLevel (2+ suc n) (P x y)) (isPropIsOfHLevel (2+ suc n) (P x y)) ) + (λ x y → transport (λ i → isOfHLevel (+-comm (2+ n) 1 i) + (isOfHLevel (2+ suc n) (P x y))) (hLevelLift (2+ n) (isPropIsOfHLevel (2+ suc n) (P x y))) ) λ a b → ( hLevelSuc (2+ n) (P ∣ a ∣ ∣ b ∣) ) (isOfHLevel∥∥ {A = a ≡ b} n) @@ -378,13 +378,8 @@ private IsoFinal : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (x y : ∥ B ∥ (suc n)) → Iso (x ≡ y) (P x y) IsoFinal x y = iso (encode-fun x y ) (decode-fun x y) (P-linv x y) (P-rinv x y) -IsoIdTrunc : {a b : A} (n : ℕ₋₂) → Iso (_≡_ {A = ∥ A ∥ (suc n)} ∣ a ∣ ∣ b ∣) (∥ a ≡ b ∥ n) -IsoIdTrunc {a = a} {b = b} n = IsoFinal {n = n} ∣ a ∣ ∣ b ∣ - -IsoΩ : {a : A} (n : ℕ₋₂) → Iso (_≡_ {A = ∥ A ∥ (suc n)} ∣ a ∣ ∣ a ∣) (∥ a ≡ a ∥ n) -IsoΩ {a = a} n = IsoIdTrunc {a = a} {b = a} n - - - - +PathIdTrunc : {a b : A} (n : ℕ₋₂) → (_≡_ {A = ∥ A ∥ (suc n)} ∣ a ∣ ∣ b ∣) ≡ (∥ a ≡ b ∥ n) +PathIdTrunc {a = a} {b = b} n = isoToPath (IsoFinal {n = n} ∣ a ∣ ∣ b ∣) +PathΩ : {a : A} (n : ℕ₋₂) → (_≡_ {A = ∥ A ∥ (suc n)} ∣ a ∣ ∣ a ∣) ≡ (∥ a ≡ a ∥ n) +PathΩ {a = a} n = PathIdTrunc {a = a} {b = a} n diff --git a/Cubical/ZCohomology/Base.agda b/Cubical/ZCohomology/Base.agda index 1c02e2057..1b9fc98ea 100644 --- a/Cubical/ZCohomology/Base.agda +++ b/Cubical/ZCohomology/Base.agda @@ -5,7 +5,7 @@ module Cubical.ZCohomology.Base where open import Cubical.Data.Int.Base open import Cubical.Data.Nat.Base open import Cubical.Data.NatMinusTwo.Base -open import Cubical.Data.Sigma.Base +open import Cubical.Data.Sigma open import Cubical.Foundations.Pointed.Base @@ -20,32 +20,26 @@ private ℓ : Level A : Type ℓ -{- -Definition of cohomology with cocefficients in the integers. For now, this definition uses ℕ₋₂, for instance defining H⁰ as -coHom neg2, H¹ as coHom (suc neg2), and so on. This has so far been easier to work with than a definition relying on ℕ→ℕ₋₂. --} - --- Cohomology --- {- Types Kₙ from Brunerie 2016 -} -coHomK : (n : ℕ₋₂) → Type₀ -coHomK neg2 = Int -coHomK (suc n) = ∥ S₊ (2+ suc n) ∥ (suc (suc (suc n))) +coHomK : (n : ℕ) → Type₀ +coHomK zero = Int +coHomK (suc n) = ∥ S₊ (suc n) ∥ (suc (ℕ→ℕ₋₂ n)) {- Cohomology -} -coHom : (n : ℕ₋₂) → Type ℓ → Type ℓ +coHom : (n : ℕ) → Type ℓ → Type ℓ coHom n A = ∥ (A → coHomK n) ∥₀ --- Reduced cohomology --- {- Pointed version of Kₙ -} -coHomK-ptd : (n : ℕ₋₂) → Pointed (ℓ-zero) -coHomK-ptd neg2 = coHomK neg2 , (pos 0) +coHomK-ptd : (n : ℕ) → Pointed (ℓ-zero) +coHomK-ptd zero = coHomK zero , (pos 0) coHomK-ptd (suc n) = (coHomK (suc n) , ∣ north ∣) {- Reduced cohomology -} -coHomRed : (n : ℕ₋₂) → (A : Pointed ℓ) → Type ℓ -coHomRed neg2 A = ∥ (A →* (coHomK-ptd neg2)) ∥₀ -coHomRed (suc n) A = ∥ (A →* (coHomK-ptd (suc n))) ∥₀ +coHomRed : (n : ℕ) → (A : Pointed ℓ) → Type ℓ +coHomRed n A = ∥ (A →* (coHomK-ptd n)) ∥₀ diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index 139514488..04ded167c 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -5,11 +5,17 @@ module Cubical.ZCohomology.Properties where open import Cubical.ZCohomology.Base open import Cubical.HITs.Sn open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Everything open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.GroupoidLaws open import Cubical.Data.NatMinusTwo.Base renaming (-1+_ to -1+₋₂_ ; 1+_ to 1+₋₂_) +open import Cubical.Data.NatMinusOne.Base +open import Cubical.Data.Empty +open import Cubical.Data.Sigma +open import Cubical.Data.Prod.Base open import Cubical.HITs.Susp open import Cubical.HITs.SetTruncation open import Cubical.HITs.Nullification @@ -20,21 +26,21 @@ open import Cubical.HITs.Truncation open import Cubical.HITs.Pushout open import Cubical.Data.Sum.Base open import Cubical.Data.HomotopyGroup +open import Cubical.Data.Bool private variable - ℓ : Level + ℓ ℓ' : Level A : Type ℓ + B : Type ℓ' - - -{- Equivalence between cohomology and the reduced version -} -coHom↔coHomRed : (n : ℕ₋₂) → +{- Equivalence between cohomology of A and reduced cohomology of (A + 1) -} +coHomRed+1Equiv : (n : ℕ) → (A : Set ℓ) → (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) -coHom↔coHomRed neg2 A i = ∥ helplemma {C = (Int , pos 0)} i ∥₀ - module coHom↔coHomRed where - helplemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →* C))) - helplemma {C = C} = isoToPath (iso map1 +coHomRed+1Equiv zero A i = ∥ helpLemma {C = (Int , pos 0)} i ∥₀ + module coHomRed+1 where + helpLemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →* C))) + helpLemma {C = C} = isoToPath (iso map1 map2 (λ b → linvPf b) (λ _ → refl)) @@ -55,43 +61,4 @@ coHom↔coHomRed neg2 A i = ∥ helplemma {C = (Int , pos 0)} i ∥₀ helper (inl x) = refl helper (inr tt) = sym snd -coHom↔coHomRed (suc n) A i = ∥ coHom↔coHomRed.helplemma A i {C = ((coHomK (suc n)) , ∣ north ∣)} i ∥₀ - - - - -{- TODO: Prove Kₙ ≡ Ω Kₙ₊₁ -} - -invPathCancel : {a b : A} → (p : a ≡ b) → p ∙ (sym p) ≡ refl -invPathCancel {a = a} {b = b} p = J {x = a} (λ y p → p ∙ (sym p) ≡ refl ) (sym (lUnit refl)) p - -φ : (a : A) → A → (north {A = A} ≡ north {A = A}) -φ x = (λ a → ((merid a) ∙ sym ((merid x)))) - -φ* : (A : Pointed ℓ) → A →* Ω ((Susp (typ A)) , north) -φ* A = (φ (pt A)) , invPathCancel (merid (pt A)) - -σ : (n : ℕ₋₂) → (typ (coHomK-ptd n)) → (typ (Ω (coHomK-ptd (suc n)))) -σ neg2 k = looper k ( cong {B = λ _ → (∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1)} - (λ x → ∣ x ∣) - ((merid north) ∙ (sym (merid south))) ) - where - looper : {a : A} -- defining loopⁿ - (n : Int) → - (a ≡ a) → - (a ≡ a) - looper {A = A} {a = a} (pos zero) p = refl - looper {A = A} {a = a} (pos (suc n)) p = (looper (pos n) p) ∙ p - looper {A = A} {a = a} (negsuc zero) p = sym p - looper {A = A} {a = a} (negsuc (suc n)) p = (sym p) ∙ (looper (negsuc n) p) -σ (suc n) x = rec {n = suc (suc (suc n))} - {B = (typ (Ω (coHomK-ptd (suc (suc n)))))} - (isOfHLevel∥∥ {A = S₊ (2+ suc (suc n))} (suc (suc (suc (suc n)))) ∣ north ∣ ∣ north ∣) - (λ y → cong {B = λ _ → Null (S (1+₋₂ (suc (suc (suc (suc n)))))) (S₊ (2+ (suc (suc n))))} - (λ z → ∣ z ∣) - (φ north y)) - x - - - - +coHomRed+1Equiv (suc n) A i = ∥ coHomRed+1.helpLemma A i {C = ((coHomK (suc n)) , ∣ north ∣)} i ∥₀ diff --git a/Cubical/ZCohomology/S1/S1.agda b/Cubical/ZCohomology/S1/S1.agda index 311fc0590..0516dbdce 100644 --- a/Cubical/ZCohomology/S1/S1.agda +++ b/Cubical/ZCohomology/S1/S1.agda @@ -3,36 +3,25 @@ module Cubical.ZCohomology.S1.S1 where open import Cubical.ZCohomology.Base open import Cubical.ZCohomology.Properties -open import Cubical.Data.Sigma open import Cubical.HITs.S1 -open import Cubical.HITs.Sn -open import Cubical.Data.Empty.Base -open import Cubical.Data.Unit.Base -open import Cubical.Foundations.Everything -open import Cubical.Data.NatMinusTwo.Base renaming (-1+_ to -1+₋₂_ ; 1+_ to 1+₋₂_) -open import Cubical.Foundations.Equiv -open import Cubical.Data.Prod -open import Cubical.HITs.Susp +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism open import Cubical.HITs.SetTruncation open import Cubical.HITs.Nullification open import Cubical.Data.Int open import Cubical.Data.Nat open import Cubical.HITs.Truncation -open import Cubical.HITs.Pushout -open import Cubical.Data.Sum.Base -open import Cubical.Data.Unit.Base -open import Cubical.Data.HomotopyGroup + private variable ℓ : Level A : Type ℓ - ---- H⁰(S¹) ---- -coHom0-S1 : (coHom neg2 S¹) ≡ Int -coHom0-S1 = (λ i → ∥ helpLemma i ∥₀ ) ∙ sym (setId isSetInt) +coHom0-S1 : (coHom zero S¹) ≡ Int +coHom0-S1 = (λ i → ∥ helpLemma i ∥₀ ) ∙ (setId isSetInt) where helpLemma : (S¹ → Int) ≡ Int helpLemma = isoToPath (iso fun @@ -60,6 +49,3 @@ coHom0-S1 = (λ i → ∥ helpLemma i ∥₀ ) ∙ sym (setId isSetInt) ------------------------- {- TODO : give Hᵏ(S¹) for all k -} - - - From 95d1c5bca2babd2b9c235a33e03604811f78fd66 Mon Sep 17 00:00:00 2001 From: Axel Date: Mon, 24 Feb 2020 16:24:55 +0100 Subject: [PATCH 06/89] fixes1 --- Cubical/AndersFile.agda | 41 ++ Cubical/Data/NatMinusTwo/Base.agda | 17 + Cubical/ZCohomology/FunLoopSpace.agda | 205 +++++++ Cubical/ZCohomology/FunLoopSpace2.agda | 34 ++ Cubical/ZCohomology/PropertiesTrash.agda | 660 +++++++++++++++++++++++ 5 files changed, 957 insertions(+) create mode 100644 Cubical/AndersFile.agda create mode 100644 Cubical/ZCohomology/FunLoopSpace.agda create mode 100644 Cubical/ZCohomology/FunLoopSpace2.agda create mode 100644 Cubical/ZCohomology/PropertiesTrash.agda diff --git a/Cubical/AndersFile.agda b/Cubical/AndersFile.agda new file mode 100644 index 000000000..59285462e --- /dev/null +++ b/Cubical/AndersFile.agda @@ -0,0 +1,41 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.AndersFile where + + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Data.Nat +open import Cubical.Data.HomotopyGroup + +private + variable + ℓ ℓ' : Level + A : Type ℓ + B : Type ℓ' + + +FunLR : (n : ℕ) {f : A → B} → + ((typ ((Ω^ (suc n)) ((A → B) , f))) → + (((x : A) → (typ ((Ω^ (suc n)) (B , f x)))))) +LR-id : (n : ℕ) {f : A → B} → FunLR n {f} refl ≡ (λ x → refl) +FunLR n {f} = {!!} +LR-id n {f} = {!!} + + +FunRL : (n : ℕ) {f : A → B} → + ((x : A) → (typ ((Ω^ (suc n)) (B , f x)))) → + ((typ ((Ω^ (suc n)) ((A → B) , f)))) +RL-id : (n : ℕ) {f : A → B} → + FunRL n (λ x → pt (((Ω^ (suc n)) (B , f x)))) ≡ pt ((Ω^ (suc n)) ((A → B) , f)) +FunRL n {f} = {!!} +RL-id n {f} = {!!} + + +wantToShow : (n : ℕ) (f : A → B) → + ((x : A) → (typ ((Ω^ n) (B , f x)))) ≡ (typ ((Ω^ n) ((A → B) , f))) +wantToShow n f = {!!} + + diff --git a/Cubical/Data/NatMinusTwo/Base.agda b/Cubical/Data/NatMinusTwo/Base.agda index 35e5b3ac8..4ca66c900 100644 --- a/Cubical/Data/NatMinusTwo/Base.agda +++ b/Cubical/Data/NatMinusTwo/Base.agda @@ -3,6 +3,7 @@ module Cubical.Data.NatMinusTwo.Base where open import Cubical.Core.Primitives open import Cubical.Data.Nat +open import Cubical.Data.Bool open import Cubical.Data.Empty import Cubical.Data.NatMinusOne as ℕ₋₁ @@ -32,6 +33,22 @@ data ℕ₋₂ : Set where ℕ→ℕ₋₂ : ℕ → ℕ₋₂ ℕ→ℕ₋₂ n = ℕ₋₁→ℕ₋₂ (ℕ→ℕ₋₁ n) +_<₋₂_ : ℕ₋₂ → ℕ₋₂ → Bool +neg2 <₋₂ m = true +suc n <₋₂ neg2 = false +suc n <₋₂ suc m = n <₋₂ m + +_+₋₂_ : ℕ₋₂ → ℕ₋₂ → ℕ₋₂ +neg2 +₋₂ m = m +suc n +₋₂ neg2 = suc n +suc n +₋₂ suc m = suc (suc ( n +₋₂ m )) + +_-₋₂_ : ℕ₋₂ → ℕ₋₂ → ℕ₋₂ +neg2 -₋₂ m = neg2 +suc n -₋₂ neg2 = suc n +suc n -₋₂ suc m = n -₋₂ m + + -- Natural number and negative integer literals for ℕ₋₂ instance diff --git a/Cubical/ZCohomology/FunLoopSpace.agda b/Cubical/ZCohomology/FunLoopSpace.agda new file mode 100644 index 000000000..215e12a0b --- /dev/null +++ b/Cubical/ZCohomology/FunLoopSpace.agda @@ -0,0 +1,205 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.FunLoopSpace where + + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.FunExtEquiv +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Data.Nat +open import Cubical.Data.HomotopyGroup + +private + variable + ℓ ℓ' : Level + A : Type ℓ + B : Type ℓ' + + + + + +{- Beware: Use of J -} +cancelReflMid : ∀ {ℓ} {A : Type ℓ} {a b : A} (p : a ≡ b) (q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q +cancelReflMid {ℓ = ℓ}{A = A} {a = a} {b = b} p q = J {ℓ} {A} {a} {ℓ} (λ b p → ((q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q)) (λ r → sym (lUnit (refl ∙ r ))) {b} p q + +--- + + +FunLR : (n : ℕ) {f : A → B} → + ((typ ((Ω^ (suc n)) ((A → B) , f))) → + (((x : A) → (typ ((Ω^ (suc n)) (B , f x)))))) +LR-id : (n : ℕ) {f : A → B} → FunLR n {f} refl ≡ (λ x → refl) +FunLR zero {f} = funExt⁻ +FunLR (suc n) {f} p x = cong (λ y → y x) (sym (LR-id n)) ∙ (funExt⁻ (cong (FunLR n) p ) x) ∙ cong (λ y → y x) (LR-id n) +LR-id zero {f} = refl +LR-id (suc n) {f} = funExt λ y → (cancelReflMid (λ i → (LR-id n) (~ i) y) + (λ i → (LR-id n {f}) i y) ∙ + rCancel (λ i → (LR-id n {f}) (~ i) y)) + +{- +lemma2 : (n : ℕ) {f : A → B} {x : A} → PathP (λ i → LR-id n {f} i x ≡ λ _ → pt ((Ω^ n) (B , f x))) (cong (λ y → y x) ((LR-id n {f}))) refl +lemma2 zero = refl +lemma2 (suc n) {f} {x} j = hcomp {!!} {!(cong (λ y → y x) (funExt (λ y → cancelReflMid (λ i → LR-id n {f} (~ (i)) y) (λ i → LR-id n {f} (i ∨ j) y) ∙ rCancel (λ i → LR-id n ((~ i)) y)))) !} +-} + + +FunRL : (n : ℕ) {f : A → B} → + ((x : A) → (typ ((Ω^ (suc n)) (B , f x)))) → + ((typ ((Ω^ (suc n)) ((A → B) , f)))) +RL-id : (n : ℕ) {f : A → B} → + FunRL n (λ x → pt (((Ω^ (suc n)) (B , f x)))) ≡ pt ((Ω^ (suc n)) ((A → B) , f)) +FunRL zero = funExt +FunRL (suc n) {f} = λ g → (sym (RL-id n {f}) ∙ cong (FunRL n {f}) (funExt g) ∙ (RL-id n {f})) +RL-id zero = refl +RL-id (suc n) {f} = cancelReflMid (λ i → (RL-id n {f}) (~ i)) (RL-id n) ∙ + rCancel (λ i → (RL-id n) (~ i)) + + + +---- + + + +wantToShow : (n : ℕ) (f : A → B) → + ((x : A) → (typ ((Ω^ n) (B , f x)))) ≡ (typ ((Ω^ n) ((A → B) , f))) +wantToShow zero f = refl +wantToShow (suc zero) f = ua funExtEquiv +wantToShow (suc (suc n)) f = isoToPath (iso (FunRL (suc n)) (FunLR (suc n)) (λ b → {!!}) {!!}) + where + hauptLemma : (n : ℕ) (f : A → B) (b : typ (Ω ((Ω^ n) ((A → B) , f)))) → FunRL n (FunLR n b ) ≡ b + hauptLemma zero f b = refl + hauptLemma (suc n) f b = {!!} + where + + oi : {!λ j → ((λ i → RL-id n {f} (~ i)) ∙ + (λ i → (hauptLemma n f (b i) (i))) ∙ + (λ i → RL-id n {f} i))!} + oi = {!!} + + + firstTransp : PathP (λ j → RL-id n (~ j) ≡ RL-id n (~ j)) + (FunRL (suc n) (FunLR (suc n) b )) + ((λ i → FunRL n {f} (funExt (λ x → (λ i₁ → LR-id n {f} (~ i₁) x) ∙ (λ i₁ → FunLR n {f} (b i₁) x) ∙ (λ i₁ → LR-id n {f} i₁ x)) i))) + firstTransp j = hcomp (λ k → λ{(j = i0) → FunRL (suc n) (FunLR (suc n) b ) ; + (j = i1) → (sym (lUnit ((λ i → FunRL n (funExt (λ x → (λ i₁ → LR-id n (~ i₁) x) ∙ + (λ i₁ → FunLR n {f} (b i₁) x) ∙ (λ i₁ → LR-id n {f} i₁ x)) i)) ∙ + (λ i → RL-id n {f} (i0)))) ∙ + sym (rUnit (λ i → FunRL n (funExt (λ x → (λ i₁ → LR-id n (~ i₁) x) ∙ + (λ i₁ → FunLR n {f} (b i₁) x) ∙ + (λ i₁ → LR-id n {f} i₁ x)) i)))) k }) + ((λ i → RL-id n {f} (~ i ∧ ( ~ j))) ∙ + (λ i → FunRL n (funExt (λ x → (λ i₁ → LR-id n (~ i₁) x) ∙ + (λ i₁ → FunLR n {f} (b i₁) x) ∙ + (λ i₁ → LR-id n {f} i₁ x)) i)) ∙ + (λ i → RL-id n {f} (i ∧ (~ j)))) + + SNDtransp : PathP (λ j → (FunRL n {f} (λ x → LR-id n {f} (~ j) x)) ≡ (FunRL n {f} (λ x → LR-id n {f} (~ j) x))) + ((λ i → FunRL n {f} (funExt (λ x → (λ i₁ → LR-id n (~ i₁) x) ∙ + (λ i₁ → FunLR n {f} (b i₁) x) ∙ + (λ i₁ → LR-id n {f} i₁ x)) i))) λ i → FunRL n {f} (FunLR n {f} (b i)) + SNDtransp j = hcomp (λ k → λ{ (j = i0) → + ((λ i → FunRL n (funExt (λ x → (λ i₁ → LR-id n (~ i₁) x) ∙ + (λ i₁ → FunLR n {f} (b i₁) x) ∙ + (λ i₁ → LR-id n {f} i₁ x)) i))) + ; (j = i1) → {!(FunRL n {f} (λ x → LR-id n {f} (~ i1) x))!}}) + (λ i → FunRL n {f} (funExt (λ x → (λ i₁ → LR-id n {f} (~ i₁ ∧ (~ j)) x) ∙ + (λ i₁ → FunLR n {f} (b i₁) x) ∙ + (λ i₁ → LR-id n {f} (i₁ ∧ (~ j)) x)) i)) + + 3transp : PathP (λ j → (hauptLemma n f refl j) ≡ (hauptLemma n f refl j)) (λ i → FunRL n {f} (FunLR n {f} (b i))) b + 3transp j = (λ i → (hauptLemma n f (b i) j)) + + test : (λ j₂ → RL-id n {f} (~ j₂) ≡ RL-id n {f} (~ j₂)) ∙ (λ i → FunRL n {f} (LR-id n {f} (~ i)) ≡ FunRL n {f} (LR-id n {f} (~ i))) ∙ (λ i → hauptLemma n f refl i ≡ hauptLemma n f refl i) ≡ refl + test j = {!((λ j₂ → RL-id n {f} ((~ j₂) ∨ j) ≡ RL-id n {f} ((~ j₂) ∨ j)) ∙ (λ i → FunRL n {f} (LR-id n {f} (?)) ≡ FunRL n {f} (LR-id n {f} (?))))!} + where + + + + postulate + wowiecomp : (i : I) → FunRL (n) (funExt (λ x → (λ i₁ → LR-id (n) (~ i₁) x) ∙ + (λ i₁ → FunLR (n) (b i₁) x) ∙ + (λ i₁ → LR-id (n) {f} i₁ x)) i) ≡ b i + + + + needToShow : ((λ j₁ → ((λ j₂ → RL-id n {f} (~ j₂) ≡ RL-id n {f} (~ j₂)) ∙ (λ i → FunRL n {f} (LR-id n (~ i)) ≡ FunRL n {f} (LR-id n {f} (~ i)))) + j₁) ∙ (λ i → hauptLemma n f refl i ≡ hauptLemma n f refl i)) ≡ refl + needToShow j k = {!((λ j₁ → ((λ j₂ → RL-id n {f} (~ j₂) ≡ RL-id n {f} (~ j₂)) ∙ (λ i → FunRL n {f} (LR-id n (~ i)) ≡ FunRL n {f} (LR-id n {f} (~ i))))!} + + +{- +where + + linvLemma : (n : ℕ) {f : A → B} {b : typ ((Ω^ (suc (suc (suc n)))) ((A → B) , f))} → + (λ i → RL-id (suc n) (~ i)) ∙ + (λ i → FunRL (suc n) (funExt (λ x → (λ i₁ → LR-id (suc n) (~ i₁) x) ∙ + (λ i₁ → FunLR (suc n) (b i₁) x) ∙ + (λ i₁ → LR-id (suc n) {f} i₁ x)) i)) ∙ + RL-id (suc n) {f} + ≡ b + linvLemma n {f} {b} j = {!hcomp ? ( (λ i → RL-id (suc n) (~ i ∧ j)) ∙ + (λ i → FunRL (suc n) (funExt (λ x → (λ i₁ → LR-id (suc n) ((~ i₁)) x) ∙ + (λ i₁ → FunLR (suc n) (b i₁) x) ∙ + (λ i₁ → LR-id (suc n) {f} (i₁ ) x)) (i))) ∙ + (λ i → (RL-id (suc n) {f} (i ∧ j)))) !} + where + cunt : (i : I) → (funExt (λ x → (λ i₁ → LR-id (suc n) (~ i₁) x) ∙ (λ i₁ → FunLR (suc n) {f} (b i₁) x) ∙ (λ i₁ → LR-id (suc n) {f} i₁ x )) i ) + ≡ (funExt (λ x → (λ i₁ → LR-id (suc n) (~ i₁) x)) ∙ funExt (λ x → (λ i₁ → FunLR (suc n) {f} (b i₁) x)) ∙ funExt (λ x → (λ i₁ → LR-id (suc n) {f} i₁ x ))) i + cunt i = refl + + cunt2 : (i : I) → (funExt (λ x → (λ i₁ → LR-id (suc n) (~ i₁) x) ∙ (λ i₁ → FunLR (suc n) {f} (b i₁) x) ∙ (λ i₁ → LR-id (suc n) {f} i₁ x )) i ) + ≡ (funExt (λ x → (λ i₁ → LR-id (suc n) (~ i₁) x)) ∙ (λ j → FunLR (suc n) {f} (b j)) ∙ funExt (λ x → (λ i₁ → LR-id (suc n) {f} i₁ x ))) i + cunt2 i = refl +{- + cunt3 : (i : I) → FunRL n (funExt (λ x → (λ i₁ → LR-id n (~ i₁) x) ∙ (λ i₁ → FunLR n (b i₁) x) ∙ (λ i₁ → LR-id n {f} i₁ x)) i) ≡ FunRL n (((λ i₁ → LR-id n {f} (~ i₁)) ∙ (λ j → FunLR n {f} (b j)) ∙ (λ i₁ → LR-id n {f} i₁)) i) + cunt3 i = refl + + cunt4 : (i : I) → FunRL (suc n) (((λ i₁ → LR-id (suc n) {f} (~ i₁)) ∙ (λ j → FunLR (suc n) {f} (b j)) ∙ (λ i₁ → LR-id (suc n) {f} i₁)) i) ≡ FunRL n {f} ( LR-id n {f} (~ i)) ∙ FunRL n (FunLR n {f} (b i)) ∙ FunRL n {f} ( LR-id n {f} i) + cunt4 i = {!!} + + +{- + main2 : (λ i → RL-id n (~ i)) ∙ (λ i → FunRL n (((λ i₁ → LR-id n {f} (~ i₁)) ∙ (λ j → FunLR n {f} (b j)) ∙ (λ i₁ → LR-id n {f} i₁)) i)) ∙ (λ i → RL-id n i) ≡ {!!} + main2 j i = {!FunRL n (((λ i₁ → LR-id n {f} (~ i₁ ∧ i0)) ∙ (λ k → FunLR n {f} (b (k))) ∙ (λ i₁ → LR-id n {f} (i₁ ∧ i0))) i0)!} + + main3 : PathP (λ j → (RL-id n {f} j ≡ RL-id n j)) (λ i → FunRL n (((λ i₁ → LR-id n {f} (~ i₁)) ∙ (λ j → FunLR n {f} (b j)) ∙ (λ i₁ → LR-id n {f} i₁)) i)) ((λ i → RL-id n (~ i)) ∙ (λ i → FunRL n (((λ i₁ → LR-id n {f} (~ i₁)) ∙ (λ j → FunLR n {f} (b j)) ∙ (λ i₁ → LR-id n {f} i₁)) i)) ∙ (λ i → RL-id n {f} i)) + main3 j = hcomp (λ k → λ{(j = i0) → {!((λ i → RL-id n ((~ i) ∧ j))!} ; (j = i1) → {!!}}) ((λ i → RL-id n ((~ i) ∧ j)) ∙ (λ i → FunRL n (((λ i₁ → LR-id n {f} (~ i₁)) ∙ (λ j → FunLR n {f} (b j)) ∙ (λ i₁ → LR-id n {f} i₁)) i)) ∙ (λ i → RL-id n (i ∧ j))) + + main4 : PathP (λ j → FunRL n (LR-id n j) ≡ FunRL n (LR-id n j)) {!LR-id n {f} i0!} (λ i → FunRL n (((λ i₁ → LR-id n {f} (~ i₁)) ∙ (λ j → FunLR n {f} (b j)) ∙ (λ i₁ → LR-id n {f} i₁)) i)) + main4 j = hcomp (λ k → {! LR-id n {f} (i0)!}) (λ i → FunRL n {f} (((λ i₁ → LR-id n {f} (~ i₁ ∧ j)) ∙ (λ k → FunLR n {f} (b (k))) ∙ (λ i₁ → LR-id n {f} (i₁ ∧ j))) i)) +-} + + wowie : (i : I) → ( FunRL (suc n) (((λ i₁ → LR-id (suc n) {f} (~ i₁)) ∙ (λ j → FunLR (suc n) {f} (b j)) ∙ (λ i₁ → LR-id (suc n) {f} i₁)) i)) ≡ (FunRL (suc n) (((λ i₁ → (FunLR (suc n) refl)) ∙ (λ j → FunLR (suc n) {f} (b j)) ∙ (λ i₁ → (FunLR (suc n) refl))) i)) + wowie i = sym (λ j → FunRL (suc n) (((λ i₁ → LR-id (suc n) {f} (~ i₁ ∧ j)) ∙ (λ k → FunLR (suc n) {f} (b (k))) ∙ (λ i₁ → LR-id (suc n) {f} (i₁ ∧ j))) i)) + + wowie2 : (i : I) → (FunRL (suc n) (((λ i₁ → (FunLR (suc n) refl)) ∙ (λ j → FunLR (suc n) {f} (b j)) ∙ (λ i₁ → (FunLR (suc n) refl))) i)) ≡ (FunRL (suc n) ((( FunLR (suc n) {f} (b i))) )) + wowie2 i = (λ j → FunRL (suc n) {f} ((sym (lUnit ((λ j → FunLR (suc n) {f} (b j)) ∙ refl)) j) i) ) ∙ ((λ j → FunRL (suc n) {f} ((sym (rUnit ((λ j → FunLR (suc n) {f} (b j)))) j) i))) + + main4 : PathP (λ j → FunRL n (LR-id n j) ≡ FunRL n (LR-id n j)) {!LR-id n {f} i0!} (λ i → FunRL n (((λ i₁ → LR-id n {f} (~ i₁)) ∙ (λ j → FunLR n {f} (b j)) ∙ (λ i₁ → LR-id n {f} i₁)) i)) + main4 j = hcomp (λ k → {! LR-id n {f} (i0)!}) (λ i → FunRL n {f} (((λ i₁ → LR-id n {f} (~ i₁ ∧ j)) ∙ (λ k → FunLR n {f} (b (k))) ∙ (λ i₁ → LR-id n {f} (i₁ ∧ j))) i)) + + postulate + fixLater : (i : I) → (FunRL (suc n) ((( FunLR (suc n) {f} (b i))) )) ≡ b i + + wowie3 : + (λ i → FunRL (suc n) (funExt (λ x → (λ i₁ → LR-id (suc n) (~ i₁) x) ∙ + (λ i₁ → FunLR (suc n) (b i₁) x) ∙ + (λ i₁ → LR-id (suc n) {f} i₁ x)) i)) + ≡ + {!b!} + wowie3 j = {!FunRL!} + + wowiecomp : (i : I) → FunRL (suc n) (funExt (λ x → (λ i₁ → LR-id (suc n) (~ i₁) x) ∙ + (λ i₁ → FunLR (suc n) (b i₁) x) ∙ + (λ i₁ → LR-id (suc n) {f} i₁ x)) i) ≡ b i + wowiecomp i = wowie i ∙ wowie2 i ∙ fixLater {!wowiecomp i0!} + + +-} + +-} diff --git a/Cubical/ZCohomology/FunLoopSpace2.agda b/Cubical/ZCohomology/FunLoopSpace2.agda new file mode 100644 index 000000000..ae60e32b6 --- /dev/null +++ b/Cubical/ZCohomology/FunLoopSpace2.agda @@ -0,0 +1,34 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.FunLoopSpace2 where + + + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.FunExtEquiv +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Data.Nat +open import Cubical.Data.HomotopyGroup +open import Cubical.HITs.Nullification +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.Truncation +open import Cubical.Data.NatMinusTwo + +private + variable + ℓ ℓ' : Level + A : Type ℓ + B : Type ℓ' + + + + + +{- Beware: Use of J -} +cancelReflMid : ∀ {ℓ} {A : Type ℓ} {a b : A} (p : a ≡ b) (q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q +cancelReflMid {ℓ = ℓ}{A = A} {a = a} {b = b} p q = J {ℓ} {A} {a} {ℓ} (λ b p → ((q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q)) (λ r → sym (lUnit (refl ∙ r ))) {b} p q + diff --git a/Cubical/ZCohomology/PropertiesTrash.agda b/Cubical/ZCohomology/PropertiesTrash.agda new file mode 100644 index 000000000..5a9005989 --- /dev/null +++ b/Cubical/ZCohomology/PropertiesTrash.agda @@ -0,0 +1,660 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.Properties where + + +open import Cubical.ZCohomology.Base +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Everything +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Data.NatMinusTwo.Base renaming (-1+_ to -1+₋₂_ ; 1+_ to 1+₋₂_) +open import Cubical.Data.NatMinusOne.Base +open import Cubical.Data.Empty +open import Cubical.Data.Sigma +open import Cubical.Data.Prod.Base +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation +open import Cubical.HITs.Nullification +open import Cubical.Data.Int +open import Cubical.Data.Nat +open import Cubical.HITs.Truncation + +open import Cubical.HITs.Pushout +open import Cubical.Data.Sum.Base +open import Cubical.Data.HomotopyGroup +open import Cubical.Data.Bool +private + variable + ℓ ℓ' : Level + A : Type ℓ + B : Type ℓ' + + +{- Equivalence between cohomology and the reduced version -} +coHom↔coHomRed : (n : ℕ₋₂) → + (A : Set ℓ) → + (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) +coHom↔coHomRed neg2 A i = ∥ helplemma {C = (Int , pos 0)} i ∥₀ + module coHom↔coHomRed where + helplemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →* C))) + helplemma {C = C} = isoToPath (iso map1 + map2 + (λ b → linvPf b) + (λ _ → refl)) + where + map1 : (A → typ C) → ((((A ⊎ Unit) , inr (tt)) →* C)) + map1 f = map1' , refl module helpmap where + map1' : A ⊎ Unit → fst C + map1' (inl x) = f x + map1' (inr x) = pt C + + map2 : ((((A ⊎ Unit) , inr (tt)) →* C)) → (A → typ C) + map2 (g , pf) x = g (inl x) + + linvPf : (b :((((A ⊎ Unit) , inr (tt)) →* C))) → map1 (map2 b) ≡ b + linvPf (f , snd) i = (λ x → helper x i) , λ j → snd ((~ i) ∨ j) + where + helper : (x : A ⊎ Unit) → ((helpmap.map1') (map2 (f , snd)) x) ≡ f x + helper (inl x) = refl + helper (inr tt) = sym snd + +coHom↔coHomRed (suc n) A i = ∥ coHom↔coHomRed.helplemma A i {C = ((coHomK (suc n)) , ∣ north ∣)} i ∥₀ + + + +{- Test -} +SigmaEqHelper : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) (b : B) (p1 p2 : Σ A (λ a → (f a) ≡ b)) → (fst p1 ≡ fst p2) → p1 ≡ p2 +SigmaEqHelper {ℓ = ℓ} {ℓ' = ℓ'} {A = A} {B = B} f b p1 p2 path = J {ℓ'} {B} {f (fst p1)} (λ b pf → ( (pair2 : Σ A (λ a → (f a) ≡ b)) → (fst p1) ≡ (fst pair2) → ((fst p1 , pf) ≡ pair2) ) ) (λ pr id → ΣPathP (id , {!transport (λ i → f (id i) ≡ ((snd pr) i))!})) {y = b} (snd p1) (p2) path {- J {ℓ} {A} {fst p1} {ℓ-max ℓ ℓ'} (λ y pf → ( (s : (f y ≡ b)) → (p1 ≡ (y , s)) )) + (λ s → ΣPathP (refl , {!J {ℓ'} {B} {x = (f (fst p1))} {ℓ'} (λ c → )!})) + {fst p2} path (snd p2) -} + + +{- TODO: Prove Kₙ ≡ Ω Kₙ₊₁ -} + +invPathCancel : {a b : A} → (p : a ≡ b) → p ∙ (sym p) ≡ refl +invPathCancel {a = a} {b = b} p = J {x = a} (λ y p → p ∙ (sym p) ≡ refl ) (sym (lUnit refl)) p +invPathCancel2 : {a b : A} → (p : a ≡ b) → (sym p) ∙ p ≡ refl +invPathCancel2 {a = a} {b = b} p = J {x = a} (λ y p → (sym p) ∙ p ≡ refl ) (sym (lUnit refl)) p + +cancelReflMid : ∀ {ℓ} {A : Type ℓ} {a b : A} (p : a ≡ b) (q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q +cancelReflMid {ℓ = ℓ}{A = A} {a = a} {b = b} p q = J {ℓ} {A} {a} {ℓ} (λ b p → ((q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q)) (λ r → sym (lUnit (refl ∙ r ))) {b} p q + +φ' : (a : A) → A → typ (Ω (Susp A , north )) +φ' x = (λ a → ((merid a) ∙ sym ((merid x)))) + +φ'' : (b : Susp A) → b ≡ south → A → north ≡ b +φ'' b p a = (merid a) ∙ (sym p) + +--φ''' : (a : A) (b : Susp A) → A → north ≡ b +--φ''' (a : A) b p = (merid a) ∙ (sym p) + + + + +looper : {a : A} -- defining loopⁿ + (n : Int) → + (a ≡ a) → + (a ≡ a) +looper {A = A} {a = a} (pos zero) p = refl +looper {A = A} {a = a} (pos (suc n)) p = (looper (pos n) p) ∙ p +looper {A = A} {a = a} (negsuc zero) p = sym p +looper {A = A} {a = a} (negsuc (suc n)) p = (sym p) ∙ (looper (negsuc n) p) + + + +φ : (a : A) → A → (north {A = A} ≡ north {A = A}) +φ x = (λ a → ((merid a) ∙ sym ((merid x)))) + +φ* : (A : Pointed ℓ) → A →* Ω ((Susp (typ A)) , north) +φ* A = (φ (pt A)) , invPathCancel (merid (pt A)) + +σ : (n : ℕ₋₂) → (typ (coHomK-ptd n)) → (typ (Ω (coHomK-ptd (suc n)))) +σ neg2 k = looper k ( cong {B = λ _ → (∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1)} + (λ x → ∣ x ∣) + ((merid north) ∙ (sym (merid south))) ) +σ (suc n) x = rec {n = suc (suc (suc n))} + {B = (typ (Ω (coHomK-ptd (suc (suc n)))))} + (isOfHLevel∥∥ {A = S₊ (2+ suc (suc n))} (suc (suc (suc (suc n)))) ∣ north ∣ ∣ north ∣) + (λ y → cong {B = λ _ → Null (S (1+₋₂ (suc (suc (suc (suc n)))))) (S₊ (2+ (suc (suc n))))} + (λ z → ∣ z ∣) + (φ north y)) + x + +testing : (n : ℕ) → (a : typ (Ω (S₊ (suc n) , north))) → isContr (∥ (Σ (S₊ n) λ k → (φ' {A = (S₊ n)} north k ) ≡ a) ∥ (ℕ→ℕ₋₂ n) ) +testing zero a = {!!} +testing (suc zero) a = {!!} +testing (suc (suc n)) a = ∣ {!!} , {!a!} ∣ , λ y → {!!} + + +-- φ : (a : A) → A → (north {A = A} ≡ north {A = A}) +-- φ x = (λ a → ((merid a) ∙ sym ((merid x)))) + +isConnectedT : (n : ℕ₋₂) (A : Type ℓ) → Type ℓ +isConnectedT n A = isContr (∥ A ∥ n) + +isConnectedF : ∀{ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n : ℕ₋₂) → (A → B) → Type ((ℓ-max ℓ ℓ')) +isConnectedF {A = A} {B = B} n f = ((b : B) → isContr (∥ fiber f b ∥ n)) + +inducedFun : ∀{ℓ''} {k : ℕ₋₂} + (f : A → B) → + (P : (B → HLevel ℓ'' (2+ k))) → + (((b : B) → fst (P b)) → (a : A) → fst (P (f a)) ) +inducedFun f P = λ g a → g (f a) + +Lemma861 : ∀{ℓ''} (n k : ℕ₋₂) → + Σ ℕ₋₂ (λ num → (suc (suc n)) +₋₂ num ≡ (suc (suc k)) ) → + (f : A → B) → + (isConnectedF (suc (suc n)) f) → + (P : (B → HLevel ℓ'' (2+ (suc (suc k))))) → + isConnectedF (((suc (suc n)) -₋₂ (suc (suc k))) -₋₂ (suc (suc neg2))) (inducedFun {k = (suc (suc k))} f P) +Lemma861 n k (neg2 , pf) f conf P b = {!fiber!} +Lemma861 {A = A} {B = B} n k (suc j , pf) f conf P g = {!!} where + helper : (l : (x : A) → (fst (P (f x)) )) → + (fiber (inducedFun {k = (suc (suc k))} f P) l) ≡ Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a) + helper l = isoToPath (iso (fun) finv (λ b → refl) λ b → refl) where + fun : (fiber (inducedFun {k = (suc (suc k))} f P) l) → Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a) + fun (fs , sn) = fs , λ a → cong (λ x → x a) sn + + finv : (Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a)) → (fiber (inducedFun {k = (suc (suc k))} f P) l) + finv (fs , sn) = fs , (funExt sn) + + helper'3 : (l : (x : A) → (fst (P (f x)) )) + (gp hq : Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a)) → + (gp ≡ hq) + ≡ Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) ( λ r → ((λ x → (r (f x))) ≡ λ (x : A) → (snd gp x) ∙ sym (snd hq x) )) + helper'3 l gp hq = isoToPath (iso fun finv (λ b i → ({!!})) {!!}) where + fun : (gp ≡ hq) → Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) ( λ r → ((λ x → (r (f x))) ≡ λ (x : A) → (snd gp x) ∙ sym (snd hq x) )) + fun path = (λ x i → fst (path i) x ) , funExt λ x → pathFinal {x} where + pathHelper : {x : A} → PathP (λ i → fst gp (f x) ≡ fst (path i) (f x)) refl (snd gp x ∙ sym (snd hq x)) + pathHelper {x = x} i = hcomp (λ j → λ{ (i = i0) → rCancel (snd gp x) j ; (i = i1) → (snd gp x ∙ sym (snd hq x)) } ) (snd gp x ∙ sym (snd (path i) x)) + + pathHelper2 : {x : A} → PathP (λ i → fst gp (f x) ≡ fst (path (~ i)) (f x)) (λ i → fst (path i) (f x)) refl + pathHelper2 {x = x} i = λ i₁ → fst (path (i₁ ∧ (~ i))) (f x) + + pathFinal : {x : A} → (λ i → fst (path i) (f x)) ≡ snd gp x ∙ sym (snd hq x) + pathFinal {x = x} = transport (λ i → PathP (λ j₁ → (lCancel (λ i → fst gp (f x) ≡ fst (path i) (f x))) i j₁) + (λ i → fst (path i) (f x)) + (snd gp x ∙ sym (snd hq x))) + (compPathP pathHelper2 pathHelper) + + finv : Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) ( λ r → ((λ x → (r (f x))) ≡ λ (x : A) → (snd gp x) ∙ sym (snd hq x) )) → (gp ≡ hq) + finv (fs , sn) = ΣPathP {x = gp} {y = hq} (funExt (λ b → fs b) , {! !}) -- (λ b → (fs b) i) , λ a → hcomp (λ k → λ{(i = i0) → (snd gp a) ; (i = i1) → {!!}}) (transp (λ j → fs (f a) (i ∧ j) ≡ l a) (~ i) ((snd gp) a)) + where + potTerm : {a : A} → PathP (λ i → (a : A) → funExt (λ b → fs b) i (f a) ≡ l a) (snd gp) (snd hq) + potTerm {a = a} j = λ a → hcomp {!(transp (λ i → funExt (λ b → fs b) (i) (f a) ≡ l a) i0 ((snd gp a))) !} (transp (λ i → funExt (λ b → fs b) (i ∧ j) (f a) ≡ l a) i0 ((snd gp a))) + + helper2 : {a : A} → (fs (f a)) ∙ (snd hq a) ≡ snd gp a + helper2 {a = a} = (cong (λ x → x ∙ (snd hq a)) (cong (λ x → x a) sn)) ∙ sym (assoc (snd gp a) (sym (snd hq a)) (snd hq a)) ∙ (λ j → (snd gp a) ∙ (lCancel (snd hq a)) j) ∙ sym (rUnit (snd gp a)) + + + -- fst gp (f x) ≡ fst hq (f x) + -- fst gp (f x) ≡ fst gp (f x) + + + + + helper2 : (l : (x : A) → (fst (P (f x)) )) + (gp hq : Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a)) → + Σ (fst gp ≡ fst hq) (λ a≡ → PathP (λ i → (a : A) → a≡ i (f a) ≡ l a) (snd gp) (snd hq)) + ≡ Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) ( λ r → ((λ x → (r (f x))) ≡ λ (x : A) → (snd gp x) ∙ sym (snd hq x) )) + helper2 l gp hq = isoToPath (iso fun finv (λ b i → {!fst gp x!} , {!!}) {!!}) where + + + finv : Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) (λ r → (inducedFun {k = (suc (suc k))} f (λ b → (fst gp b ≡ fst hq b) , hLevelSuc (suc (suc (1+ (1+₋₂ k)))) (fst (P b)) (snd (P b)) (fst gp b) (fst hq b) ) r ≡ (λ (x : A) → (snd gp x) ∙ sym (snd hq x)))) → Σ (fst gp ≡ fst hq) (λ a≡ → PathP (λ i → (a : A) → a≡ i (f a) ≡ l a) (snd gp) (snd hq)) + finv (fs , sn) = {!!} {- (funExt fs) , {!(snd hq) !}-} where + propLemma : (a : A) → isProp (fst hq (f a) ≡ l a) + propLemma a f g = {!!} + + + fun : Σ (fst gp ≡ fst hq) (λ a≡ → PathP (λ i → (a : A) → a≡ i (f a) ≡ l a) (snd gp) (snd hq)) → Σ ((x : B) → (((fst gp) x) ≡ ((fst hq) x))) (λ r → (inducedFun {k = (suc (suc k))} f (λ b → (fst gp b ≡ fst hq b) , hLevelSuc (suc (suc (1+ (1+₋₂ k)))) (fst (P b)) (snd (P b)) (fst gp b) (fst hq b) ) r ≡ (λ (x : A) → (snd gp x) ∙ sym (snd hq x)))) + fun (fs , sn) = (funExt⁻ fs) , (funExt (λ x → test5)) where + + test5 : {x : A} → PathP (λ _ → fst gp (f x) ≡ fst hq (f x) ) (λ i → (fs i) (f x)) (snd gp x ∙ sym (snd hq x)) + test5 {x = x} i j = hcomp ((λ k → λ{ (i = i0) → fs j (f x) ; (i = i1) → test3 {x = x} k j ; (j = i0) → fst gp (f x) ; (j = i1) → fs (~ i ∨ k) (f x) })) (test4 {x = x} i j) where + test3 : {x : A} → PathP (λ i → fst gp (f x) ≡ fs i (f x)) (snd gp x ∙ sym (snd gp x)) (snd gp x ∙ sym (snd hq x)) + test3 {x = x} i = (snd gp x) ∙ (sym (sn i x)) + + test4 : {x : A} → PathP (λ i → fst gp (f x) ≡ fs (~ i) (f x)) (λ i → (fs i) (f x)) (snd gp x ∙ (sym (snd gp x))) + test4 {x = x} i = hcomp (λ j → λ { (i = i0) → (λ i → (fs i) (f x)) ; (i = i1) → (sym (invPathCancel (snd gp x)) j) } ) (test i) where + test : {x : A} → PathP (λ i → fst gp (f x) ≡ fs (~ i) (f x)) (λ i → (fs i) (f x)) (refl {x = (fst gp (f x))}) + test {x = x} j i = fs (i ∧ ( ~ j)) (f x) + + helper3 : (l : (x : A) → (fst (P (f x)) )) (gp hq : Σ ((b : B) → (fst (P b))) λ g → (a : A) → (g (f a)) ≡ (l a)) → (gp ≡ hq) ≡ Σ (fst gp ≡ fst hq) + (λ a≡ → PathP (λ i → (a : A) → a≡ i (f a) ≡ l a) (snd gp) (snd hq)) + helper3 l gp hq = sym (ua (Σ≡ {A = ((a : B) → fst (P a))} {B = (λ x → (a : A) → x (f a) ≡ l a)} {x = gp} {y = hq})) + + + + +------------------ +{- Sn ≡ Ω(Sn+1) -} + + + + + + + + +-------------------- Guillaume Test -------------------- + + + + + + + + + + +sphereSuc : (n : ℕ) → (S₊ n) → (S₊ (suc n)) +sphereSuc n north = north +sphereSuc n south = south +sphereSuc (suc n) (merid a i) = merid (sphereSuc n a) i + +helplemma : {A : Type ℓ} {a : A} {p q : a ≡ a} {P : p ≡ q} → (C : (a ≡ a) → Type ℓ) → (a1 a2 : C (P i0)) → (b1 b2 : C (P (i1))) → PathP (λ i → C (P i)) a1 b2 → (a1 ≡ a2) → (b1 ≡ b2) → PathP (λ i → C (P i)) a2 b2 +helplemma {ℓ = ℓ} {a = a}{p = p} {q = q} {P = P} C a1 a2 b1 b2 depP a1a2 b1b2 = {!J {ℓ} {?} !} + +pathExtender : (B : A → Type ℓ) → {!!} +pathExtender = {!!} + +promote : (n : ℕ) → (S₊ n) → typ (Ω ((S₊ (suc n)) , north) ) +promote zero north = refl {x = north} +promote zero south = refl {x = north} +promote (suc n) north = refl {x = north} +promote (suc n) south = refl {x = north} +promote (suc n) (merid a i) = (( sym (rCancel (merid (merid a i0))) ∙ (λ i → (merid (merid a i) ∙ (sym (merid (merid north i)))))) ∙ rCancel (merid (merid a i1)) ) i + +decode' : (n : ℕ) → ∥ (S₊ n) ∥ (ℕ→ℕ₋₂ n) → ∥ typ (Ω (S₊ (suc n) , north)) ∥ (ℕ→ℕ₋₂ n) +decode' n = rec {A = (S₊ n)} {B = ∥ typ (Ω (S₊ (suc n) , north)) ∥ (ℕ→ℕ₋₂ n) } (isOfHLevel∥∥ {A = typ (Ω (S₊ (suc n) , north))} (ℕ→ℕ₋₂ n)) λ x → ∣ promote n x ∣ + +CODES : (n : ℕ) → S₊ (suc n) → Type₀ +CODES n north = ∥ S₊ n ∥ (ℕ→ℕ₋₂ n) +CODES n south = ∥ S₊ n ∥ (ℕ→ℕ₋₂ n) +CODES n (merid a i) = {!(typ ((Ω^ (suc n)) (Type₀ , ∥ S₊ n ∥ (ℕ→ℕ₋₂ n))))!} + +encode' : (n : ℕ) → ∥ typ (Ω ((S₊ (suc n)) , north)) ∥ (ℕ→ℕ₋₂ n) → ∥ S₊ n ∥ (ℕ→ℕ₋₂ n) +encode' n = rec {A = typ (Ω ((S₊ (suc n)) , north))} + {B = ∥ S₊ n ∥ (ℕ→ℕ₋₂ n)} + (isOfHLevel∥∥ {A = (S₊ n )} (ℕ→ℕ₋₂ n)) + λ x → {!(cong (CODES n) x) ?!} + +{- +functions : (n : ℕ) (f : A → B) → + (typ ((Ω^ n) ((A → B) , f))) ≡ + ((x : A) → (typ ((Ω^ n) (B , f x)))) +functions zero = {!!} +functions {A = A} {B = B}(suc zero) f = isoToPath (iso (λ g x i → (g i) x) (λ g → funExt {f = f} {g = f} g ) (λ b → refl) λ b → refl) +functions (suc (suc zero)) f = isoToPath (iso {!!} {!!} {!!} {!!}) + where + helper' : {A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + typ (Ω (Ω ((A → B) , f))) ≡ ( _≡_ {A = ( _≡_ f f)} (refl {x = f}) (refl {x = f})) + helper' f = refl + helper : {A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + (typ (Ω (Ω ((A → B) , f)))) ≡ ( _≡_ {A = ((x : A) → (f x ≡ f x) )} + (λ x _ → f x) + (λ x _ → f x)) + helper f = isoToPath (iso (λ g i x j → (g i j) x) (λ x k j y → ((x k) y) j) (λ b → refl) λ b → refl) + + helper2 :{A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + ( _≡_ {A = ((x : A) → (f x ≡ f x) )} (λ x _ → f x) (λ x _ → f x)) + ≡ + ((x : A) → (_≡_ {A = f x ≡ f x} (λ _ → f x) (λ _ → f x ))) + helper2 = {!!} + + +functions (suc (suc (suc n))) f = {!!} -- isoToPath (iso {!!} (λ x → {!!}) {!!} {!!}) + where + helper : {A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + typ (Ω (Ω (Ω ((Ω^ n) ((A → B) , f))))) ≡ ( _≡_ {A = typ (Ω (Ω ((Ω^ n) ((A → B) , f))))} + (λ _ → pt (Ω ((Ω^ n) ((A → B) , f)))) + (λ _ → pt (Ω ((Ω^ n) ((A → B) , f))))) + helper f = refl + + where + Ω^n : {A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + typ (Ω (Ω ((Ω^ n) ((A → B) , f)))) ≡ ((x : A) → typ (Ω (Ω ((Ω^ n) (B , f x))))) + Ω^n f = functions (suc (suc n)) f + Ω^n' : {A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + typ (Ω (Ω (Ω ((Ω^ n) ((A → B) , f))))) ≡ typ (Ω (((x : A) → typ (Ω (Ω ((Ω^ n) (B , f x))))) , (λ _ → refl))) + Ω^n' f i = typ (Ω (( Ω^n f i) , {!!})) + + +-} + +Fun1000 : (n : ℕ) {f : A → B} → + ((typ ((Ω^ (suc n)) ((A → B) , f))) → + (((x : A) → (typ ((Ω^ (suc n)) (B , f x)))))) +-- Fun1001 : (n : ℕ) {f : A → B} → Fun1000 n {f} refl ≡ (λ x → refl) +Fun1000 zero p x i = p i x +Fun1000 (suc n) p x = {!Fun1000 n (p j) x i!} +-- Fun1001 = {!!} + + +lrFun2* : (n : ℕ) {f : A → B} → + Σ ((typ ((Ω^ (suc n)) ((A → B) , f))) → + (((x : A) → (typ ((Ω^ (suc n)) (B , f x)))))) + (λ g → (g refl) ≡ (λ x → refl)) +lrFun2* zero {f = f} = funExt⁻ , refl +lrFun2* (suc n) {f = f} = (λ p x → {!(funExt⁻ (cong (fst (lrFun2* n)) p) x)!}) + , + funExt (λ y → cancelReflMid (λ i → snd (lrFun2* n) (~ i) y) + (λ i → snd (lrFun2* n {f = f}) i y) ∙ + invPathCancel (λ i → snd (lrFun2* n) (~ i) y)) + +lrFun : (n : ℕ) {f : A → B} → + ((typ ((Ω^ (suc n)) ((A → B) , f))) → + (((x : A) → (typ ((Ω^ (suc n)) (B , f x)))))) +lrFun n = fst (lrFun2* n) + + + where + + +rlFun* : (n : ℕ) {f : A → B} → + Σ ((((x : A) → (typ ((Ω^ (suc n)) (B , f x))))) → + ((typ ((Ω^ (suc n)) ((A → B) , f)))) ) + (λ g → (g (λ x → snd (((Ω^ (suc n)) (B , f x)))) ≡ pt ((Ω^ (suc n)) ((A → B) , f)))) + +rlFun* zero {f = f} = (λ g → funExt g) , refl +rlFun* {A = A} {B = B}(suc n) {f = f} = (λ g → (sym (snd (rlFun* n)) ∙ cong (fst (rlFun* n)) (funExt g) ∙ snd (rlFun* n))) + , + (cancelReflMid (λ i → snd (rlFun* n) (~ i)) (snd (rlFun* n))) ∙ + invPathCancel (λ i → snd (rlFun* n) (~ i)) + +rlFun : (n : ℕ) {f : A → B} → + (((x : A) → (typ ((Ω^ (suc n)) (B , f x))))) → + ((typ ((Ω^ (suc n)) ((A → B) , f)))) +rlFun n {f = f} = fst (rlFun* n {f = f}) + where + +functions2 : (n : ℕ) (f : A → B) → + ((x : A) → (typ ((Ω^ n) (B , f x)))) ≃ (typ ((Ω^ n) ((A → B) , f))) + +functions2 {A = A} {B = B} zero f = idEquiv (A → B) +functions2 {A = A} {B = B} (suc n) f = rlFun n , record { equiv-proof = λ y → ((fst (lrFun2* n)) y , {!!}) , {!!} } where + + helper : (n : ℕ) (y : (typ ((Ω^ (suc n)) ((A → B) , f)))) → (rlFun n ((fst (lrFun2* n)) y) ≡ y) + helper zero y = refl + helper (suc n) y = (λ i → (sym (snd (rlFun* n))) ∙ ((helper5 (funExt (λ x → sym (cong (λ g → g x ) (snd (lrFun2* n))))) (cong (fst (lrFun2* n)) y) (funExt (λ x → (cong (λ g → g x ) (snd (lrFun2* n))))) (fst (rlFun* n))) i ) ∙ (snd (rlFun* n))) ∙ {!helper !} + where + + + idTest : (n : ℕ) → (rlFun n {f} (λ x → snd (lrFun2* n {f}) i1 x) ≡ rlFun n (λ x → snd (lrFun2* n {f}) i0 x)) ≡ ((fst (rlFun* n) (λ x → snd (Ω ((Ω^ n) (B , f x))))) ≡ (pt (Ω ((Ω^ n) ((A → B) , f))))) + idTest n i = refl {x = rlFun n {f} (λ x → snd (lrFun2* n {f}) i1 x)} i ≡ idTest1 i where + idTest1 : rlFun n (λ x → snd (lrFun2* n {f}) i0 x) ≡ (pt (Ω ((Ω^ n) ((A → B) , f)))) + idTest1 = helper n (pt (Ω ((Ω^ n) ((A → B) , f)))) + + + test : (n : ℕ) → PathP (λ i → {!!}) ( (cong (rlFun n {f}) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x)))))) (snd (rlFun* n {f})) + test n i = {!!} where + + + maybe2 : {n : ℕ} (i : I) → snd (rlFun* n) (~ i) ≡ (pt (Ω ((Ω^ n) ((A → B) , f)))) + maybe2 {n} i j = snd (rlFun* n) ((~ i) ∨ j) + + maybe3 : (n : ℕ) → (pt (Ω ((Ω^ n) ((A → B) , f))) ≡ rlFun n (λ x → snd (lrFun2* n {f}) i0 x)) + maybe3 n = sym (snd (rlFun* n {f})) ∙ λ i → rlFun n (hihih ( ~ i)) where + hihih : (λ (x : A) → snd (lrFun2* n {f}) i0 x) ≡ (λ x → snd (Ω ((Ω^ n) (B , f x)))) + hihih = funExt λ x → cong (λ g → g x) (snd (lrFun2* n)) + + maybe : (n : ℕ) (y : (typ (Ω ((Ω^ (suc n)) ((A → B) , f))))) → ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n) (~ i) x)))))) ∙ (cong (rlFun n) (funExt (λ x → (λ i → fst (lrFun2* n) (y i) x)))) ∙ ((sym (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n) (~ i) x))))))) ≡ (snd (rlFun* n)) ∙ y ∙ (sym (snd (rlFun* n))) + maybe n y i = test n i ∙ (cong (λ x → (helper n x) i) y ) ∙ sym (test n i) + -- (λ j → (snd (rlFun* n)) (~ j ∨ i)) ∙ {!!} + + -- ∙ {!λ i → (cong (λ x → helper n x i ) y)!} ∙ {!!} -- (cong (λ x → helper n x i ) ?) + + compTest : (n : ℕ) (y : (typ (Ω ((Ω^ (suc n)) ((A → B) , f))))) → + (sym (snd (rlFun* n))) ∙ + ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n) (~ i) x)))))) ∙ + (cong (rlFun n) (funExt (λ x → (λ i → fst (lrFun2* n) (y i) x)))) ∙ + ((sym (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n) (~ i) x))))))) ∙ + (snd (rlFun* n)) ≡ + refl ∙ {!!} ∙ y ∙ {!!} ∙ refl + compTest n y i = (λ i₁ → snd (rlFun* n {f}) (~ i₁ ∨ i)) ∙ + {!snd (rlFun* n {f}) (~ i1 ∨ i)!} {- hcomp (λ k → λ{(i = i0) → {!maxZeuner i0!} ; (i = i1) → (maxZeuner i1)}) (maxZeuner i) -} ∙ + (cong (λ x → (helper n x) i) y ) ∙ + {!snd (rlFun* n {f}) (~ i1 ∨ i)!} + ∙ λ i₁ → snd (rlFun* n {f}) (i₁ ∨ i) + + where + + maxZeuner : (i : I) → snd (rlFun* n) i ≡ helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) i + maxZeuner i j = hcomp (λ k → λ{(j = i0) → snd (rlFun* n) i ; (j = i1) → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i ∨ (~ k)) }) (snd (rlFun* n {f}) (i ∨ j)) + maxZeunerHelper : (i : I) → maxZeuner i0 ≡ {!snd (rlFun* n) i0 ≡ helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) i0!} + maxZeunerHelper = {!!} + + gimmeMore : (i : I) → rlFun n (λ x → snd (lrFun2* n {f}) i0 x) ≡ helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) i + gimmeMore i = (λ j → rlFun n (λ x → snd (lrFun2* n {f}) j x)) ∙ (snd (rlFun* n {f})) ∙ λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i ∨ (~ j)) + + gimmeLeft : (i : I) → rlFun n (λ x → snd (lrFun2* n {f}) i1 x) ≡ snd (rlFun* n) i + gimmeLeft i = λ j → snd (rlFun* n {f}) (j ∧ i) + + gimmei0 : ((sym (gimmeLeft i0)) ∙ ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x)))))) ∙ (gimmeMore i0)) ≡ ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x)))))) + gimmei0 = {!((sym (gimmeLeft i0)))!} where + + gimmei0' : ((sym (gimmeLeft i0))) ≡ refl + gimmei0' j = refl + + + + + gimmei1 : ((sym (gimmeLeft i1)) ∙ ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x)))))) ∙ (gimmeMore i1)) ≡ refl + gimmei1 = gimmei1' ∙ more!! where + + gimmei1' : ((sym (gimmeLeft i1)) ∙ ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x)))))) ∙ (gimmeMore i1)) ≡ (sym (gimmeLeft i1)) ∙ (snd (rlFun* n {f})) ∙ λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i1 ∨ (~ j)) + gimmei1' = ∙-assoc (sym (gimmeLeft i1)) (( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x))))))) ((snd (rlFun* n {f})) ∙ λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i1 ∨ (~ j))) where + + ∙-assoc : {A : Type ℓ}{a b c d : A} (p : a ≡ b) (q : b ≡ c) (t : b ≡ d) → p ∙ q ∙ (sym q ∙ t) ≡ p ∙ t + ∙-assoc {ℓ = ℓ} {A = A} {a = a} {b = b} {c = c} {d = d} p q t = cong (λ x → p ∙ x) {!q ∙ (sym q ∙ t) ≡ t!} + + + more!! : (sym (gimmeLeft i1)) ∙ (snd (rlFun* n {f})) ∙ (λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i1 ∨ (~ j))) ≡ refl + more!! = (assoc (sym (gimmeLeft i1)) (snd (rlFun* n {f})) (λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i1 ∨ (~ j)))) ∙ (λ j → (invPathCancel (sym (gimmeLeft i1)) j) ∙ (λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i1 ∨ (~ j)))) ∙ lCancel (λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i1 ∨ (~ j))) + + more!!! : (λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i1 ∨ (~ j))) ≡ refl + more!!! = refl + + + + hello : (i : I) → snd (rlFun* n) i ≡ helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) i + hello i = (λ j → snd (rlFun* n {f}) (i ∨ j)) ∙ λ j → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (i ∨ (~ j)) + + kittyHelper : PathP (λ i → (snd (rlFun* n {f})) i ≡ helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) i) (hello i0) refl + kittyHelper j = hcomp (λ k → λ{ (j = i0) → hello i0 ; + (j = i1) → invPathCancel refl k }) + ((λ i → snd (rlFun* n {f}) (j ∨ i)) ∙ λ i → helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) (j ∨ (~ i))) +{- + kittyHelperB : PathP (λ i → (snd (rlFun* n {f})) i ≡ helper n (snd (Ω ((Ω^ n) ((A → B) , f)))) i) ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n) (~ i) x)))))) refl + kittyHelperB i j = hcomp (λ k → λ{ (j = i0) → {!!} + ; (j = i1) → {!(cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n) (~ i) x)))))!}} ) + {!( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ j) x))))))!} +-} + + -- helloKitty : hello i0 ≡ ( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n) (~ i) x)))))) + -- helloKitty k = {!!} + + -- helloKitty2 : hello i1 ≡ refl + -- helloKitty2 = {!!} + + mini : (n : ℕ) → rlFun n (λ x → snd (lrFun2* n {f}) i0 x) ≡ pt (Ω ((Ω^ n) ((A → B) , f))) + mini n = helper n refl + + mini2 : (n : ℕ) → rlFun n (λ x → snd (lrFun2* n {f}) i1 x) ≡ pt (Ω ((Ω^ n) ((A → B) , f))) + mini2 n = (λ i → rlFun n (λ x → snd (lrFun2* n {f}) (~ i) x)) ∙ helper n refl + + leftMost : PathP (λ i → (pt (Ω ((Ω^ n) ((A → B) , f))) ≡ (snd (rlFun* n)) i)) (sym (snd (rlFun* n))) refl + leftMost i = λ i₁ → snd (rlFun* n) (~ i₁ ∨ i) + + + + notLeftMost : (n : ℕ) → PathP (λ i → mini2 n i ≡ mini n i ) (( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x))))))) refl + notLeftMost zero i j = {!fst (rlFun* (suc n)) (λ x → snd (lrFun2* (suc n) {f}) (i1) x)!} + notLeftMost (suc n) i = {!(( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x)))))))!} + where + ffs2 : (n : ℕ) (i : I) → (fst (rlFun* n {f}) (λ x → snd (lrFun2* n {f}) (~ i1 ∨ i) x)) ≡ mini n i + ffs2 n i = ((λ j → fst (rlFun* n) (λ x → snd (lrFun2* n {f}) (i ∨ j) x)) ∙ snd (rlFun* n {f}) ) ∙ λ j → mini n (i ∨ (~ j)) + + ffs : (n : ℕ) (i : I) → (fst (rlFun* n {f}) (λ x → snd (lrFun2* n {f}) (~ i0 ∨ i) x)) ≡ mini2 n i + ffs n i = snd (rlFun* n {f}) ∙ {!mini2 i1!} + + {- + + tester : ((sym (ffs i0)) ∙ (λ j → fst (rlFun* n) (snd (lrFun2* n {f}) (i0 ∨ (~ j)))) ∙ (ffs2 i0)) ≡ (( (cong (rlFun n) (funExt ((λ x → (λ i → snd (lrFun2* n {f}) (~ i) x))))))) + tester = (({!((sym (ffs i1)) ∙ (λ j → fst (rlFun* n) (snd (lrFun2* n {f}) (i1 ∨ (~ j)))) ∙ (ffs2 i1))!}) ∙ sym (rUnit (refl ∙ (λ j → fst (rlFun* n) (snd (lrFun2* n {f}) (~ j)))))) ∙ sym (lUnit (λ j → fst (rlFun* n) (snd (lrFun2* n {f}) (i0 ∨ (~ j))))) + where + helpi : (ffs2 i0) ≡ refl + helpi = {!!} + + snLeftMost : {!(λ j → fst (rlFun* n {f}) (snd (lrFun2* n) (~ j)))!} + snLeftMost = {!!} + + -} + + + helper2 : (n : ℕ) (y : typ (Ω (Ω ((Ω^ n) ((A → B) , f))))) → funExt ((λ (x : A) → (λ i → snd (lrFun2* n) (~ i) x) ∙ (λ i → fst (lrFun2* n) (y i) x) ∙ (λ i → snd (lrFun2* n {f = f}) i x))) ≡ (funExt (λ (x : A) → (λ i → snd (lrFun2* n) (~ i) x))) ∙ funExt (λ (x : A) → (λ i → fst (lrFun2* n) (y i) x)) ∙ funExt λ (x : A) → (λ i → snd (lrFun2* n {f = f}) i x) + helper2 zero y = refl + helper2 (suc n) y = refl + + helper4 : {A : Type ℓ} {B : Type ℓ} {a b c : A} (p : a ≡ b) (q : b ≡ c) (g : A → B) → cong g (p ∙ q) ≡ ((cong g p) ∙ (cong g q)) + helper4 {ℓ = ℓ} {A} {B} {a} {b} {c} p q g = J {ℓ} {A} {a} {ℓ} (λ b1 p → ((q1 : b1 ≡ c) → cong g (p ∙ q1) ≡ ((cong g p) ∙ (cong g q1)))) + (λ q1 → J {ℓ} {A} {a} {ℓ} (λ c2 q2 → (λ i → g ((refl ∙ q2) i)) ≡ (λ i → g (refl i)) ∙ (λ i → g (q2 i))) ( (λ j → ((λ i → g (((rUnit (λ _ → a)) (~ j) ) i)))) ∙ lUnit ((λ i → g (refl i)))) {c} q1 ) + p q + + helper5 : {A : Type ℓ} {B : Type ℓ} {a b c d : A} (p : a ≡ b) (q : b ≡ c) (r : c ≡ d) (g : A → B) → cong g (p ∙ q ∙ r) ≡ ((cong g p) ∙ (cong g q) ∙ (cong g r)) + helper5 {ℓ = ℓ} {A} {B} {a} {b} {c} {d} p q r g = helper4 p (q ∙ r) g ∙ ((cong (λ x → (cong g p) ∙ x )) (helper4 q r g)) + + helper3 : (n : ℕ) (y : typ (Ω (Ω ((Ω^ n) ((A → B) , f))))) → (cong (rlFun n) (funExt (λ x → (λ i₁ → fst (lrFun2* n) (y i₁) x)))) ≡ {!!} + helper3 zero y = {!!} + helper3 (suc n) y = {!(cong (rlFun (suc n)) (funExt (λ x → (λ i₁ → fst (lrFun2* (suc n)) (y i₁) x))))!} + +{- +functions' {A = A} {B = B}(suc zero) f = isoToPath (iso (λ g x i → (g i) x) (λ g → funExt {f = f} {g = f} g ) (λ b → refl) λ b → refl) + where + helper : transport (λ i → (sym (functions' (suc zero) f)) i) (λ x → refl) ≡ (pt ((Ω ) ((A → B) , f))) + helper = {!!} +functions' (suc (suc zero)) f = isoToPath (iso {!!} {!!} {!!} {!!}) + where + helper' : {A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + typ (Ω (Ω ((A → B) , f))) ≡ ( _≡_ {A = ( _≡_ f f)} (refl {x = f}) (refl {x = f})) + helper' f = refl + helper : {A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + (typ (Ω (Ω ((A → B) , f)))) ≡ ( _≡_ {A = ((x : A) → (f x ≡ f x) )} + (λ x _ → f x) + (λ x _ → f x)) + helper f = isoToPath (iso (λ g i x j → (g i j) x) (λ x k j y → ((x k) y) j) (λ b → refl) λ b → refl) + + helper2 :{A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + ( _≡_ {A = ((x : A) → (f x ≡ f x) )} (λ x _ → f x) (λ x _ → f x)) + ≡ + ((x : A) → (_≡_ {A = f x ≡ f x} (λ _ → f x) (λ _ → f x ))) + helper2 = {!!} +functions' (suc (suc (suc n))) f = {!!} + where + helper : {A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + typ (Ω (Ω (Ω ((Ω^ n) ((A → B) , f))))) ≡ ( _≡_ {A = typ (Ω (Ω ((Ω^ n) ((A → B) , f))))} + (λ _ → pt (Ω ((Ω^ n) ((A → B) , f)))) + (λ _ → pt (Ω ((Ω^ n) ((A → B) , f))))) + helper f = refl + + where + Ω^n : {A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + typ (Ω (Ω ((Ω^ n) ((A → B) , f)))) ≡ ((x : A) → typ (Ω (Ω ((Ω^ n) (B , f x))))) + Ω^n f = functions (suc (suc n)) f + Ω^n' : {A : Type ℓ} {B : Type ℓ'} → + (f : A → B) → + (Ω (Ω (Ω ((Ω^ n) ((A → B) , f))))) ≡ (Ω (((x : A) → typ (Ω (Ω ((Ω^ n) (B , f x))))) , (λ _ → refl))) + Ω^n' f i = (Ω (( Ω^n f i) , {!(transport (λ j → ( Ω^n f j)) (refl i)) !})) where + lemma : transport (λ i → Ω^n f i) refl ≡ λ _ → refl + lemma = funExt {f = transport (λ i → Ω^n f i) refl} {g = λ _ → refl} λ x → {!!} + + + + + ---------- + +lrFun3* : (n : ℕ) {f : A → B} → + Σ ((typ ((Ω^ (suc n)) ((A → B) , f))) → + (((x : A) → (typ ((Ω^ (suc n)) (B , f x)))))) (λ g → (g refl) ≡ (λ x → refl)) +lrFun3* zero {f = f} = funExt⁻ , refl +lrFun3* {A = A} {B = B} (suc n) {f = f} = (λ p x → transport (λ i → (cong (λ g → g x) (snd (lrFun3* n {f = f}))) i ≡ (cong (λ g → g x) (snd (lrFun3* n {f = f}))) i) (funExt⁻ (cong (λ x → (fst (lrFun3* n) x )) p) x)) , funExt (λ x → transpLemma x) + module lr where + helpLemma : (x : A) → ((fst (lrFun3* n) (snd (Ω ((Ω^ n) ((A → B) , f)))) x ≡ (fst (lrFun3* n) (snd (Ω ((Ω^ n) ((A → B) , f)))) x))) ≡ (typ (Ω (Ω ((Ω^ n) (B , f x))))) + helpLemma x i = helper x i ≡ helper x i where + + helper : (x : A) → fst (lrFun3* n) (snd (Ω ((Ω^ n) ((A → B) , f)))) x ≡ pt ((Ω ((Ω^ n) (B , f x)))) + helper x = cong (λ g → g x) (snd (lrFun3* n {f = f})) + + transpLemma : (x : A) → transp (λ i → helpLemma x i) i0 refl ≡ refl + transpLemma x j = transp (λ i → helpLemma x (i ∨ j)) j refl + +lrFun3 : (n : ℕ) {f : A → B} → ((typ ((Ω^ (suc n)) ((A → B) , f))) → (((x : A) → (typ ((Ω^ (suc n)) (B , f x)))))) +lrFun3 n {f = f} = (fst (lrFun3* n {f = f})) + + +rlFun3* : (n : ℕ) {f : A → B} → + Σ ((((x : A) → (typ ((Ω^ (suc n)) (B , f x))))) → + ((typ ((Ω^ (suc n)) ((A → B) , f)))) ) + (λ g → (g (λ x → snd (((Ω^ (suc n)) (B , f x)))) ≡ pt ((Ω^ (suc n)) ((A → B) , f)))) + + + + +rlFun3* zero {f = f} = (λ g → funExt g) , refl +rlFun3* {A = A} {B = B}(suc n) {f = f} = (λ g → transport (λ i → (snd (rlFun3* n {f})) i ≡ (snd (rlFun3* n {f})) i) (cong (fst (rlFun3* n)) (funExt g))) , transpLemma -- (cancelReflMid (λ i → snd (rlFun* n) (~ i)) (snd (rlFun* n))) ∙ invPathCancel (λ i → snd (rlFun* n) (~ i)) + module rl where + helpLemma : (fst (rlFun3* n) (λ x → snd (Ω ((Ω^ n) (B , f x)))) ≡ fst (rlFun3* n) (λ x → snd (Ω ((Ω^ n) (B , f x))))) ≡ typ (Ω (Ω ((Ω^ n) ((A → B) , f)))) + helpLemma i = (snd (rlFun3* n {f})) i ≡ (snd (rlFun3* n {f})) i + + transpLemma : transport (λ i → helpLemma i) refl ≡ refl + transpLemma j = transp (λ i → helpLemma (i ∨ j)) j refl -- transp (λ i → helpLemma (i ∨ j)) j refl + +rlFun3 : (n : ℕ) {f : A → B} → + (((x : A) → (typ ((Ω^ (suc n)) (B , f x))))) → + ((typ ((Ω^ (suc n)) ((A → B) , f)))) +rlFun3 n {f = f} = fst (rlFun3* n {f = f}) + where + +functions3 : (n : ℕ) (f : A → B) → + ((x : A) → (typ ((Ω^ n) (B , f x)))) ≃ (typ ((Ω^ n) ((A → B) , f))) + +functions3 {A = A} {B = B} zero f = idEquiv (A → B) +functions3 {A = A} {B = B} (suc n) f = rlFun3 n , record { equiv-proof = λ y → (( (fst (lrFun3* n)) y) , λ i → {!!}) , {!!} } where + linv : (n : ℕ) (y : typ (Ω ((Ω^ n) ((A → B) , f)))) → rlFun3 n (fst (lrFun3* n {f = f}) y) ≡ y + linv zero y = refl + linv (suc n) y = {!!} where + maybe : rlFun3 (suc n) (fst (lrFun3* (suc n) {f = f}) y) ≡ λ j x → {!((cong (λ x → (fst (lrFun3* n) x )) y) (~ j)) x!} + maybe = {!!} + + +inv1 : (n : ℕ) {f : A → B} (g : ((typ ((Ω^ (suc n)) ((A → B) , f))) )) → (rlFun3 n) ((lrFun3 n) g ) ≡ g +inv1 zero {f = f} g = refl +inv1 {A = A} {B = B}(suc n) {f = f} g j = {! -- transp (λ i → (snd (rlFun3* n {f})) (i ∨ j) ≡ (snd (rlFun3* n {f})) (i ∨ j)) i0 + (cong (fst (rlFun3* n)) (funExt ( + (λ x → transp (λ i → (cong (λ g → g x) (snd (lrFun3* n {f = f}))) (i ∨ j) + ≡ (cong (λ g → g x) (snd (lrFun3* n {f = f}))) (i ∨ j)) (~ j) + (funExt⁻ (cong (λ x → (fst (lrFun3* n) x )) g) x)) ) ) )!} + + where + maybe : PathP (λ i → {! !} ≡ {!!} ) (cong (fst (rlFun3* n)) (funExt ( + (λ x → transport (λ i → (cong (λ g → g x) (snd (lrFun3* n {f = f}))) i + ≡ (cong (λ g → g x) (snd (lrFun3* n {f = f}))) i) + (funExt⁻ (cong (λ x → (fst (lrFun3* n {f}) x )) g) x)) ) ) ) + {!(cong (fst (rlFun3* n)) (funExt ( (λ x → (funExt⁻ (cong (λ x → (fst (lrFun3* n {f}) x )) g) x))))) !} + maybe j = {!fst (rlFun3* n) (λ x → snd (lrFun3* n {f}) i1 x)!} +-} From abb8a5246bd3cf47e3d67419f363c6619a34f2d5 Mon Sep 17 00:00:00 2001 From: Axel Date: Tue, 25 Feb 2020 10:57:25 +0100 Subject: [PATCH 07/89] Pointed funs change --- Cubical/Foundations/Pointed/Base.agda | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Cubical/Foundations/Pointed/Base.agda b/Cubical/Foundations/Pointed/Base.agda index d796dee28..79ce3184b 100644 --- a/Cubical/Foundations/Pointed/Base.agda +++ b/Cubical/Foundations/Pointed/Base.agda @@ -16,5 +16,5 @@ pt = snd _→*_ : ∀{ℓ ℓ'} → (A : Pointed ℓ) (B : Pointed ℓ') → Type (ℓ-max ℓ ℓ') _→*_ A B = Σ (typ A → typ B) λ f → f (pt A) ≡ pt B -_→*_/_ : ∀{ℓ ℓ'} → (A : Pointed ℓ) (B : Pointed ℓ') → (A →* B) → Pointed (ℓ-max ℓ ℓ') -A →* B / f = (A →* B) , f +_→*_* : ∀{ℓ ℓ'} → (A : Pointed ℓ) (B : Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') +A →* B * = (A →* B) , (λ x → pt B) , refl From 6a49b56db5ca703760866f9f9196d0704b270fa7 Mon Sep 17 00:00:00 2001 From: aljungstrom <49276137+aljungstrom@users.noreply.github.com> Date: Wed, 25 Mar 2020 23:56:55 +0100 Subject: [PATCH 08/89] remove bool import --- Cubical/Data/NatMinusTwo/Base.agda | 1 - 1 file changed, 1 deletion(-) diff --git a/Cubical/Data/NatMinusTwo/Base.agda b/Cubical/Data/NatMinusTwo/Base.agda index 7e156d0a2..c28455ecb 100644 --- a/Cubical/Data/NatMinusTwo/Base.agda +++ b/Cubical/Data/NatMinusTwo/Base.agda @@ -3,7 +3,6 @@ module Cubical.Data.NatMinusTwo.Base where open import Cubical.Core.Primitives open import Cubical.Data.Nat -open import Cubical.Data.Bool open import Cubical.Data.Empty record ℕ₋₂ : Type₀ where From 0bf84a72af8462b034f65fb025d3b5b268abba18 Mon Sep 17 00:00:00 2001 From: aljungstrom <49276137+aljungstrom@users.noreply.github.com> Date: Wed, 25 Mar 2020 23:58:43 +0100 Subject: [PATCH 09/89] newline --- Cubical/HITs/SetTruncation/Properties.agda | 1 - 1 file changed, 1 deletion(-) diff --git a/Cubical/HITs/SetTruncation/Properties.agda b/Cubical/HITs/SetTruncation/Properties.agda index de11a0f93..a9f795e55 100644 --- a/Cubical/HITs/SetTruncation/Properties.agda +++ b/Cubical/HITs/SetTruncation/Properties.agda @@ -39,7 +39,6 @@ setTruncUniversal Bset = isoToEquiv (iso intro out leftInv rightInv) leftInv : ∀ g → intro (out g) ≡ g leftInv g = refl - rightInv : ∀ h → out (intro h) ≡ h rightInv h i x = elim (λ x → isProp→isSet (Bset (out (intro h) x) (h x))) (λ a → refl) x i From fe50520027f67fba02b8168a0d2c88c49e313188 Mon Sep 17 00:00:00 2001 From: aljungstrom <49276137+aljungstrom@users.noreply.github.com> Date: Thu, 26 Mar 2020 00:00:12 +0100 Subject: [PATCH 10/89] delete everything @:-) --- .../HITs/Truncation/Connected/FreudenthalProof/Everything.agda | 0 1 file changed, 0 insertions(+), 0 deletions(-) delete mode 100644 Cubical/HITs/Truncation/Connected/FreudenthalProof/Everything.agda diff --git a/Cubical/HITs/Truncation/Connected/FreudenthalProof/Everything.agda b/Cubical/HITs/Truncation/Connected/FreudenthalProof/Everything.agda deleted file mode 100644 index e69de29bb..000000000 From 6dddf6ca05daefd54aa7ac903d73f9a44649efa4 Mon Sep 17 00:00:00 2001 From: aljungstrom <49276137+aljungstrom@users.noreply.github.com> Date: Thu, 26 Mar 2020 01:44:21 +0100 Subject: [PATCH 11/89] whitespace... --- .../Connected/FreudenthalProof/Prelim.agda | 22 +++++++++---------- 1 file changed, 11 insertions(+), 11 deletions(-) diff --git a/Cubical/HITs/Truncation/Connected/FreudenthalProof/Prelim.agda b/Cubical/HITs/Truncation/Connected/FreudenthalProof/Prelim.agda index 92b37481f..983b345db 100644 --- a/Cubical/HITs/Truncation/Connected/FreudenthalProof/Prelim.agda +++ b/Cubical/HITs/Truncation/Connected/FreudenthalProof/Prelim.agda @@ -1,7 +1,7 @@ {-# OPTIONS --cubical --safe #-} module Cubical.HITs.Truncation.Connected.FreudenthalProof.Prelim where -open import Cubical.HITs.Truncation.Connected.Base +open import Cubical.HITs.Truncation.Connected.Base open import Cubical.Foundations.Path open import Cubical.Foundations.Transport open import Cubical.Foundations.HLevels @@ -33,7 +33,7 @@ congComp2 {A = A}{a = a} {b = b} {c = c} f p q = J (λ b p → (c : A) (q : b (λ c → J (λ c q → (λ i → f a) ∙ (λ i → f (q i)) ≡ (λ i → f ((refl ∙ q) i))) ((λ j → rUnit (refl {x = f a}) (~ j)) ∙ λ j i → f ((lUnit (refl {x = a}) j) i))) - p c q + p c q congComp3 : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {a b c d : A} (f : A → B) (p : a ≡ b) (q : b ≡ c) (r : c ≡ d) → (cong f p ∙ cong f q) ∙ cong f r ≡ cong f (p ∙ q ∙ r) @@ -76,15 +76,15 @@ canceller : ∀ {ℓ} {A : Type ℓ} {a b c : A} → (r : b ≡ c) (p q : a ≡ canceller {ℓ} {A} {a} {b} {c} = J {ℓ} {A} {b} (λ c r → ((p q : a ≡ b) → p ∙ r ≡ q ∙ r → p ≡ q)) λ p q id → (rUnit p) ∙ id ∙ sym (rUnit q) cancellerReflCase : ∀ {ℓ} {A : Type ℓ} {a b : A} → (p q : a ≡ b) → canceller refl p q ≡ λ id → (rUnit p) ∙ id ∙ sym (rUnit q) -cancellerReflCase {a = a} {b = b} p q = cong (λ x → x p q) (JRefl (λ c r → ((p q : a ≡ b) → p ∙ r ≡ q ∙ r → p ≡ q)) λ p q id → (rUnit p) ∙ id ∙ sym (rUnit q)) +cancellerReflCase {a = a} {b = b} p q = cong (λ x → x p q) (JRefl (λ c r → ((p q : a ≡ b) → p ∙ r ≡ q ∙ r → p ≡ q)) λ p q id → (rUnit p) ∙ id ∙ sym (rUnit q)) cancellerInv : ∀ {ℓ} {A : Type ℓ} {a b c : A} → (r : b ≡ c) (p q : a ≡ b) → (p ≡ q) → p ∙ r ≡ q ∙ r cancellerInv {a = a} {b = b} = J (λ c r → (p q : a ≡ b) → (p ≡ q) → p ∙ r ≡ q ∙ r) λ p q id → sym (rUnit p) ∙ id ∙ rUnit q cancellerInvReflCase : ∀ {ℓ} {A : Type ℓ} {a b : A} → (p q : a ≡ b) → cancellerInv refl p q ≡ λ id → sym (rUnit p) ∙ id ∙ rUnit q -cancellerInvReflCase {a = a} {b = b} p q = cong (λ x → x p q) (JRefl (λ c r → (p q : a ≡ b) → (p ≡ q) → p ∙ r ≡ q ∙ r) λ p q id → sym (rUnit p) ∙ id ∙ rUnit q) +cancellerInvReflCase {a = a} {b = b} p q = cong (λ x → x p q) (JRefl (λ c r → (p q : a ≡ b) → (p ≡ q) → p ∙ r ≡ q ∙ r) λ p q id → sym (rUnit p) ∙ id ∙ rUnit q) -cancellerSection : ∀ {ℓ} {A : Type ℓ} {a b c : A} → (r : b ≡ c) (p q : a ≡ b) → section (canceller r p q) (cancellerInv r p q) +cancellerSection : ∀ {ℓ} {A : Type ℓ} {a b c : A} → (r : b ≡ c) (p q : a ≡ b) → section (canceller r p q) (cancellerInv r p q) cancellerSection {a = a} {b = b} {c = c} = J (λ c r → (p q : a ≡ b) → section (canceller r p q) (cancellerInv r p q) ) λ p q → transport (λ i → section (cancellerReflCase p q (~ i)) (cancellerInvReflCase p q (~ i))) (λ b → assoc (rUnit p) ((λ i → rUnit p (~ i)) ∙ b ∙ rUnit q) (λ i → rUnit q (~ i)) ∙ (λ i → ((assoc (rUnit p) (sym (rUnit p)) (b ∙ rUnit q)) i) ∙ (λ i → rUnit q (~ i))) ∙ (λ i → ((rCancel (rUnit p) i) ∙ (b ∙ rUnit q)) ∙ (sym (rUnit q))) ∙ @@ -93,7 +93,7 @@ cancellerSection {a = a} {b = b} {c = c} = J (λ c r → (p q : a ≡ b) → sec (λ i → b ∙ (rCancel (rUnit q) i)) ∙ sym (rUnit b)) -cancellerRetract : ∀ {ℓ} {A : Type ℓ} {a b c : A} → (r : b ≡ c) (p q : a ≡ b) → retract (canceller r p q) (cancellerInv r p q) +cancellerRetract : ∀ {ℓ} {A : Type ℓ} {a b c : A} → (r : b ≡ c) (p q : a ≡ b) → retract (canceller r p q) (cancellerInv r p q) cancellerRetract {a = a} {b = b} {c = c} = J (λ c r → (p q : a ≡ b) → retract (canceller r p q) (cancellerInv r p q) ) λ p q → transport (λ i → retract (cancellerReflCase p q (~ i)) (cancellerInvReflCase p q (~ i))) λ b → assoc (sym (rUnit p)) ((λ i → rUnit p (i)) ∙ b ∙ (sym (rUnit q))) (rUnit q) ∙ (λ i → ((assoc (sym (rUnit p)) (rUnit p) (b ∙ sym (rUnit q))) i) ∙ (rUnit q)) ∙ (λ i → ((lCancel (rUnit p) i) ∙ (b ∙ sym (rUnit q))) ∙ ((rUnit q))) ∙ @@ -101,7 +101,7 @@ cancellerRetract {a = a} {b = b} {c = c} = J (λ c r → (p q : a ≡ b) → ret (sym (assoc b (sym (rUnit q)) (rUnit q))) ∙ (λ i → b ∙ (rCancel (sym (rUnit q)) i)) ∙ sym (rUnit b) - + cancellerIsEquiv : ∀ {ℓ} {A : Type ℓ} {a b c : A} → (r : b ≡ c) (p q : a ≡ b) → isEquiv (canceller r p q) cancellerIsEquiv {ℓ} {A} {a} {b} {c} = J {ℓ} {A} {b} (λ c r → ((p q : a ≡ b) → isEquiv (canceller r p q))) λ p q → transport (λ i → isEquiv (cancellerHelp p q (~ i))) (helper p q) @@ -180,7 +180,7 @@ assocJRefl {x = x} = (cong (λ x → x refl refl) (JRefl (λ y p → (q : y ≡ transpFunct : ∀ {ℓ ℓ'} {A : Type ℓ} {x y z : A} {B : A → Type ℓ'} (p : x ≡ y) (q : y ≡ z) (u : B x) → - transport (λ i → B (q i)) (transport (λ i → B (p i)) u) ≡ transport (λ i → B ((p ∙ q) i)) u + transport (λ i → B (q i)) (transport (λ i → B (p i)) u) ≡ transport (λ i → B ((p ∙ q) i)) u transpFunct {A = A} {x = x} {y = y} {z = z} {B = B} p = J (λ y p → (z : A) (q : y ≡ z) (u : B x) → transport (λ i → B (q i)) (transport (λ i → B (p i)) u) ≡ transport (λ i → B ((p ∙ q) i)) u) @@ -204,7 +204,7 @@ transpFunctRefl {A = A} {x = x} {B = B} u = ∙ cong (λ x → x u) (JRefl (λ z q → (u : B x) → transport (λ i → B (q i)) (transport (λ i → B (refl {x = x} i)) u) ≡ transport (λ i → B ((refl ∙ q) i)) u) (λ u → transportRefl ((transport (λ i → B x) u)) ∙ - λ j → transport (λ i → B ((lUnit ((λ _ → x)) j) i)) u)) + λ j → transport (λ i → B ((lUnit ((λ _ → x)) j) i)) u)) cancelReflMid : ∀ {ℓ} {A : Type ℓ} {a b : A} (p : a ≡ b) (q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q cancelReflMid {ℓ = ℓ}{A = A} {a = a} {b = b} p q = J {ℓ} {A} {a} {ℓ} (λ b p → ((q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q)) (λ r → sym (lUnit (refl ∙ r ))) {b} p q @@ -234,7 +234,7 @@ abstract p f g Lemma296-funRefl : ∀ {ℓ ℓ' ℓ''} {X : Type ℓ} {x : X} {A : X → Type ℓ'} {B : X → Type ℓ''} (f : (A x) → (B x)) - (g : (A x) → (B x)) → + (g : (A x) → (B x)) → Lemma296-fun {A = A} {B = B} (refl {x = x}) f g ≡ λ h → transportRefl f ∙ funExt λ z → sym (transportRefl (f z)) ∙ (h z) ∙ @@ -283,7 +283,7 @@ toPathCancel {A = A} {x = x} {y = y} z = cong (λ x → fst x) (equiv-proof (toP transportLemma : {a b : A} {B : (z : A) → Type ℓ'} (p : a ≡ b) (x : B a) (y : B b) → transport (λ i → B (p i)) x ≡ y → transport (λ i → B (p (~ i))) y ≡ x transportLemma {a = a} {B = B} x y = J (λ b p → (x : B a) (y : B b) → transport (λ i → B (p i)) x ≡ y → transport (λ i → B (p (~ i))) y ≡ x) (λ x y id → transportRefl y ∙ (sym id) ∙ transportRefl x) - x y + x y transportLemmaRefl : {a : A} {B : (z : A) → Type ℓ'} → (x y : B a) → transportLemma {B = B} (λ _ → a) ≡ λ x y id → transportRefl y ∙ (sym id) ∙ transportRefl x transportLemmaRefl {ℓ} {A = A} {a = a} {B = B} x y = JRefl {ℓ} {A} {a} (λ b p → (x y : B a) → transport (λ i → B a) x ≡ y → transport (λ i → B a) y ≡ x) (λ x y id → transportRefl y ∙ (sym id) ∙ transportRefl x) From 182c5d6ff3872a859221db719228fe18b82bc829 Mon Sep 17 00:00:00 2001 From: aljungstrom <49276137+aljungstrom@users.noreply.github.com> Date: Thu, 26 Mar 2020 02:35:38 +0100 Subject: [PATCH 12/89] Update 8-6-1.agda --- Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-1.agda | 4 ---- 1 file changed, 4 deletions(-) diff --git a/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-1.agda b/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-1.agda index 08a5bf18f..9699a89c5 100644 --- a/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-1.agda +++ b/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-1.agda @@ -29,10 +29,6 @@ private A : Type ℓ B : Type ℓ' -private - 0* : ℕ₋₂ - 0* = ℕ→ℕ₋₂ 0 - private Lemma861-fibId : ∀{ℓ} (n : ℕ₋₂) (k : ℕ) (f : A → B) → (is- n -Connected f) → From 1816ffcdf927a87c56f37f6f721aa229ffcd1366 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Wed, 1 Apr 2020 00:28:51 +0200 Subject: [PATCH 13/89] testing --- Cubical/HITs/Truncation/Properties.agda | 8 +- Cubical/ZCohomology/S1Loop.agda | 352 ++++++++++++++++++++++++ 2 files changed, 359 insertions(+), 1 deletion(-) create mode 100644 Cubical/ZCohomology/S1Loop.agda diff --git a/Cubical/HITs/Truncation/Properties.agda b/Cubical/HITs/Truncation/Properties.agda index 97604aecf..0bb4784da 100644 --- a/Cubical/HITs/Truncation/Properties.agda +++ b/Cubical/HITs/Truncation/Properties.agda @@ -327,8 +327,14 @@ private IsoFinal : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (x y : ∥ B ∥ (suc₋₂ n)) → Iso (x ≡ y) (P x y) IsoFinal x y = iso (encode-fun x y ) (decode-fun x y) (P-linv x y) (P-rinv x y) + IsoFinal2 : ∀ {ℓ} {B : Type ℓ} {n : ℕ₋₂} (x y : ∥ B ∥ (suc₋₂ n)) → Iso (P x y) (x ≡ y) + IsoFinal2 x y = iso (decode-fun x y) (encode-fun x y ) (P-rinv x y) (P-linv x y) + PathIdTrunc : {a b : A} (n : ℕ₋₂) → (_≡_ {A = ∥ A ∥ (suc₋₂ n)} ∣ a ∣ ∣ b ∣) ≡ (∥ a ≡ b ∥ n) -PathIdTrunc {a = a} {b = b} n = isoToPath (IsoFinal {n = n} ∣ a ∣ ∣ b ∣) +PathIdTrunc {a = a} {b = b} n = isoToPath (IsoFinal ∣ a ∣ ∣ b ∣) PathΩ : {a : A} (n : ℕ₋₂) → (_≡_ {A = ∥ A ∥ (suc₋₂ n)} ∣ a ∣ ∣ a ∣) ≡ (∥ a ≡ a ∥ n) PathΩ {a = a} n = PathIdTrunc {a = a} {b = a} n + +truncEquivΩ : {a : A} (n : ℕ₋₂) → (∥ a ≡ a ∥ n) ≃ (_≡_ {A = ∥ A ∥ (suc₋₂ n)} ∣ a ∣ ∣ a ∣) +truncEquivΩ {a = a} n = isoToEquiv (IsoFinal2 ∣ a ∣ ∣ a ∣) diff --git a/Cubical/ZCohomology/S1Loop.agda b/Cubical/ZCohomology/S1Loop.agda new file mode 100644 index 000000000..1159465a3 --- /dev/null +++ b/Cubical/ZCohomology/S1Loop.agda @@ -0,0 +1,352 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.S1Loop where + +open import Cubical.ZCohomology.Base +open import Cubical.HITs.S1 +open import Cubical.HITs.S3 +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Everything +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Data.NatMinusTwo.Base +open import Cubical.Data.Empty +open import Cubical.Data.Sigma +open import Cubical.Data.Prod.Base +open import Cubical.HITs.Susp +open import Cubical.HITs.Nullification +open import Cubical.Data.Int +open import Cubical.Data.Nat +open import Cubical.HITs.Truncation renaming (elim to trElim) +open import Cubical.HITs.Hopf + +open import Cubical.HITs.Pushout +open import Cubical.Data.Sum.Base +open import Cubical.Data.HomotopyGroup +open import Cubical.Data.Bool +private + variable + ℓ ℓ' : Level + A : Type ℓ + B : Type ℓ' + + +pred₋₂ : ℕ₋₂ → ℕ₋₂ +pred₋₂ neg2 = neg2 +pred₋₂ neg1 = neg2 +pred₋₂ (ℕ→ℕ₋₂ zero) = neg1 +pred₋₂ (ℕ→ℕ₋₂ (suc n)) = ℕ→ℕ₋₂ n + +is-_-Connected : (n : ℕ₋₂) (f : A → B) → Type _ +is-_-Connected {B = B} n f = (x : B) → isContr (∥ fiber f x ∥ n) + + +Pr2310 : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} (n : ℕ₋₂) + (f : C → A) (g : C → B) → + is- n -Connected f → + is-_-Connected {A = B} {B = Pushout f g} n inr +Pr2310 {A = A} {B = B} {C = C} n f g iscon = {!!} + where + module hej {ℓ : Level} (P : (Pushout f g) → HLevel ℓ (2+ n)) + (h : (b : B) → typ (P (inr b))) + where + Q : A → Type _ + Q a = typ (P (inl a)) + + fun : (c : C) → Q (f c) + fun = transport (λ i → (c : C) → typ (P (push c (~ i)))) λ c → h (g c) + + typeFam : (d : Pushout f g) → typ (P d) + typeFam (inl x) = {!!} + typeFam (inr x) = {!!} + typeFam (push a i) = {!!} + + + +Pr242 : (n : ℕ₋₂) → isContr (∥ {!S₊ n!} ∥ pred₋₂ n) +Pr242 = {!!} + + + + + +-- congFunct : {a b c : A} (f : A → B) (p : a ≡ b) (q : b ≡ c) → cong f (p ∙ q) ≡ cong f p ∙ cong f q +-- congFunct f p q i = hcomp (λ j → λ{(i = i0) → rUnit (cong f (p ∙ q)) (~ j) ; +-- (i = i1) → cong f (rUnit p (~ j)) ∙ cong f q}) +-- (cong f (p ∙ (λ k → q (k ∧ (~ i)))) ∙ cong f λ k → q ((~ i) ∨ k) ) +-- -- +-- symDistr : {a b c : A} (p : a ≡ b) (q : b ≡ c) → sym (p ∙ q) ≡ sym q ∙ sym p +-- symDistr p q i = hcomp (λ j → λ{(i = i0) → rUnit (sym (p ∙ q)) (~ j) ; +-- (i = i1) → sym (lUnit q (~ j)) ∙ sym p}) +-- (sym ((λ k → p (k ∨ i)) ∙ q) ∙ sym λ k → p (i ∧ k)) + +-- {- We want to prove that Kn≃ΩKn+1. For this we need the map ϕ-} + +-- private +-- ϕ : (pt a : A) → typ (Ω (Susp A , north)) +-- ϕ pt a = (merid a) ∙ sym (merid pt) + +-- {- To define the map for n=0 we use the λ k → loopᵏ map for S₊ 1. The loop is given by ϕ south north -} + + +-- looper : Int → _≡_ {A = S₊ 1} north north +-- looper (pos zero) = refl +-- looper (pos (suc n)) = looper (pos n) ∙ (ϕ south north) +-- looper (negsuc zero) = sym (ϕ south north) +-- looper (negsuc (suc n)) = looper (negsuc n) ∙ sym (ϕ south north) + +-- {- The map of Kn≃ΩKn+1 is given as follows. -} +-- Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) +-- Kn→ΩKn+1 zero x = cong ∣_∣ (looper x) +-- Kn→ΩKn+1 (suc n) = trElim (λ x → (isOfHLevel∥∥ (ℕ→ℕ₋₂ (suc (suc n))) ∣ north ∣ ∣ north ∣)) +-- λ a → cong ∣_∣ ((merid a) ∙ (sym (merid north))) + +-- {- +-- We want to show that the function (looper : Int → S₊ 1) defined by λ k → loopᵏ is an equivalece. We already know that the corresponding function (intLoop : Int → S¹ is) an equivalence, +-- so the idea is to show that when intLoop is transported along a suitable path S₊ 1 ≡ S¹ we get looper. Instead of using S₊ 1 straight away, we begin by showing this for the equivalent Susp Bool. +-- -} + +-- -- loop for Susp Bool -- +-- loop* : _≡_ {A = Susp Bool} north north +-- loop* = merid false ∙ sym (merid true) + +-- -- the loop function -- +-- intLoop* : Int → _≡_ {A = Susp Bool} north north +-- intLoop* (pos zero) = refl +-- intLoop* (pos (suc n)) = intLoop* (pos n) ∙ loop* +-- intLoop* (negsuc zero) = sym loop* +-- intLoop* (negsuc (suc n)) = intLoop* (negsuc n) ∙ sym loop* + +-- -- we show that the loop spaces agree -- +-- loopSpId : ΩS¹ ≡ _≡_ {A = Susp Bool} north north +-- loopSpId i = typ (Ω (S¹≡SuspBool i , transp ((λ j → S¹≡SuspBool (j ∧ i))) (~ i) base)) + +-- -- the transport map -- +-- altMap2 : Int → _≡_ {A = Susp Bool} north north +-- altMap2 n i = transport S¹≡SuspBool (intLoop n i) + +-- -- We show that the transporting intLoop over S¹≡SuspBool gives intLoop* (modulo function extensionality) -- +-- altMap≡intLoop*2 : (x : Int) → intLoop* x ≡ altMap2 x +-- altMap≡intLoop*2 (pos zero) = refl +-- altMap≡intLoop*2 (pos (suc n)) = (λ i → (altMap≡intLoop*2 (pos n) i) ∙ loop*) ∙ +-- sym (helper n) + +-- where +-- helper : (n : ℕ) → altMap2 (pos (suc n)) ≡ altMap2 (pos n) ∙ loop* +-- helper zero = (λ i j → transport S¹≡SuspBool (lUnit loop (~ i) j)) ∙ +-- (λ i j → transport S¹≡SuspBool (loop j)) ∙ +-- (λ i j → transportRefl ((merid false ∙ (sym (merid true))) j) i ) ∙ +-- lUnit loop* +-- helper (suc n) = anotherHelper n ∙ +-- (λ i → altMap2 (pos (suc n)) ∙ helper zero i) ∙ +-- (λ i → altMap2 (pos (suc n)) ∙ lUnit loop* (~ i)) +-- where +-- anotherHelper : (n : ℕ) → altMap2 (pos (suc (suc n))) ≡ altMap2 (pos (suc n)) ∙ altMap2 (pos 1) +-- anotherHelper n = ((λ i j → transport S¹≡SuspBool ((intLoop (pos (suc n)) ∙ loop) j))) ∙ +-- rUnit (λ j → transport S¹≡SuspBool ((intLoop (pos (suc n)) ∙ loop) j) ) ∙ +-- (λ i → (λ j → transport S¹≡SuspBool ((intLoop (pos (suc n)) ∙ λ k → loop (k ∧ (~ i))) j)) ∙ λ j → transport S¹≡SuspBool (loop (j ∨ (~ i))) ) ∙ +-- (λ i → (λ j → transport S¹≡SuspBool (rUnit (intLoop (pos (suc n))) (~ i) j)) ∙ λ j → transport S¹≡SuspBool ((lUnit loop i) j)) + +-- altMap≡intLoop*2 (negsuc zero) = sym ((λ i j → transport S¹≡SuspBool (loop (~ j))) ∙ +-- λ i j → transportRefl (loop* (~ j)) i ) +-- altMap≡intLoop*2 (negsuc (suc n)) = helper n +-- where +-- altMapneg1 : altMap2 (negsuc zero) ≡ sym (loop*) +-- altMapneg1 i j = transportRefl (loop* (~ j)) i + +-- anotherHelper : (n : ℕ) → altMap2 (negsuc (suc n)) ≡ altMap2 (negsuc n) ∙ altMap2 (negsuc zero) +-- anotherHelper n = ((λ i → rUnit (λ j → (transport S¹≡SuspBool ((intLoop (negsuc n) ∙ sym loop) j))) i)) ∙ +-- ((λ i → (λ j → transport S¹≡SuspBool ((intLoop (negsuc n) ∙ (λ k → loop ((~ k) ∨ i))) j)) ∙ λ j → transport S¹≡SuspBool (loop ((~ j) ∧ i)))) ∙ +-- (λ i → ((λ j → transport S¹≡SuspBool (rUnit (intLoop (negsuc n)) (~ i) j))) ∙ altMap2 (negsuc zero)) + +-- helper : (n : ℕ) → intLoop* (negsuc n) ∙ (sym loop*) ≡ altMap2 (negsuc (suc n)) +-- helper zero = (λ i → altMapneg1 (~ i) ∙ altMapneg1 (~ i)) ∙ sym (anotherHelper zero) +-- helper (suc n) = (λ i → (helper n i) ∙ altMapneg1 (~ i) ) ∙ +-- sym (anotherHelper (suc n)) + + +-- {- We have that the transported map is an equivalence -} +-- mapIsEquiv : isEquiv altMap2 +-- mapIsEquiv = compEquiv (intLoop , (isoToIsEquiv (iso intLoop winding (decodeEncode base) windingIntLoop))) ((λ x i → transport S¹≡SuspBool (x i)) , helper) .snd +-- where +-- helper : isEquiv {A = ΩS¹} {B = _≡_ {A = Susp Bool} north north} (λ x i → transport S¹≡SuspBool (x i)) +-- helper = congEquiv (transport S¹≡SuspBool , helper3) .snd +-- where +-- helper2 : transport S¹≡SuspBool ≡ S¹→SuspBool +-- helper2 = funExt λ z → transportRefl (S¹→SuspBool z) +-- helper3 : isEquiv (transport S¹≡SuspBool ) +-- helper3 = transport (λ i → isEquiv (helper2 (~ i))) (S¹≃SuspBool .snd) + +-- {- From this we get that intLoop* is an equivalence-} +-- intLoop*Equiv : isEquiv intLoop* +-- intLoop*Equiv = transport (λ i → isEquiv (funExt (altMap≡intLoop*2) (~ i))) mapIsEquiv + +-- {- We now only need to work with Susp Bool and S₊ 1. We transport intLoop* along a path S1≡SuspBool and show that this gives us looper. -} +-- SuspBool→S1 : Susp Bool → S₊ 1 +-- SuspBool→S1 north = north +-- SuspBool→S1 south = south +-- SuspBool→S1 (merid false i) = merid north i +-- SuspBool→S1 (merid true i) = merid south i + +-- {- We need to spell out the trivial equivalence S1≃SuspBool. Of course it's important to make sure that we choose the right version of it. -} +-- S1≃SuspBool : Susp Bool ≃ S₊ 1 +-- S1≃SuspBool = isoToEquiv (iso SuspBool→S1 S1→SuspBool retrHelper sectHelper) +-- where +-- S1→SuspBool : S₊ 1 → Susp Bool +-- S1→SuspBool north = north +-- S1→SuspBool south = south +-- S1→SuspBool (merid north i) = merid false i +-- S1→SuspBool (merid south i) = merid true i + +-- sectHelper : section S1→SuspBool SuspBool→S1 +-- sectHelper north = refl +-- sectHelper south = refl +-- sectHelper (merid false i) = refl +-- sectHelper (merid true i) = refl + +-- retrHelper : retract S1→SuspBool SuspBool→S1 +-- retrHelper north = refl +-- retrHelper south = refl +-- retrHelper (merid north i) = refl +-- retrHelper (merid south i) = refl + +-- {- We show that transporting intLoop* along (ua S1≃SuspBool) gives us looper (again, modulo function extensionality) -} +-- looperIntoBool : (x : Int) → looper x ≡ λ i → transport ((ua (S1≃SuspBool))) (intLoop* x i) +-- looperIntoBool (pos zero) = refl +-- looperIntoBool (pos (suc n)) = (λ j → looperIntoBool (pos n) j ∙ merid north ∙ sym (merid south)) ∙ +-- (λ j → (λ i → transportRefl (SuspBool→S1 (intLoop* (pos n) i)) j ) ∙ merid north ∙ sym (merid south)) ∙ +-- (λ j → cong SuspBool→S1 (intLoop* (pos n)) ∙ congFunct SuspBool→S1 (merid false) (sym (merid true)) (~ j)) ∙ +-- sym (congFunct SuspBool→S1 (intLoop* (pos n)) loop*) ∙ +-- λ j i → transportRefl (SuspBool→S1 ((intLoop* (pos n) ∙ loop*) i)) (~ j) +-- looperIntoBool (negsuc zero) = symDistr (merid north) (sym (merid south)) ∙ +-- sym (congFunct SuspBool→S1 (merid true) (sym (merid false))) ∙ +-- (λ j → cong SuspBool→S1 (merid true ∙ sym (merid false))) ∙ +-- (λ j → cong SuspBool→S1 (symDistr (merid false) (sym (merid true)) (~ j))) ∙ +-- λ j i → transportRefl (SuspBool→S1 (loop* (~ i))) (~ j) +-- looperIntoBool (negsuc (suc n)) = (λ i → looperIntoBool (negsuc n) i ∙ sym ((merid north) ∙ (sym (merid south))) ) ∙ +-- (λ i → looperIntoBool (negsuc n) i1 ∙ symDistr (merid north) (sym (merid south)) i) ∙ +-- (λ j → (λ i → transportRefl (SuspBool→S1 (intLoop* (negsuc n) i)) j) ∙ merid south ∙ sym (merid north)) ∙ +-- (λ j → cong SuspBool→S1 (intLoop* (negsuc n)) ∙ congFunct SuspBool→S1 (merid true) (sym (merid false)) (~ j)) ∙ +-- ((λ j → cong SuspBool→S1 (intLoop* (negsuc n)) ∙ cong SuspBool→S1 (symDistr (merid false) (sym (merid true)) (~ j)))) ∙ +-- sym (congFunct SuspBool→S1 (intLoop* (negsuc n)) (sym loop*)) ∙ +-- λ j i → transportRefl (SuspBool→S1 ((intLoop* (negsuc n) ∙ sym loop*) i)) (~ j) + +-- {- From this, we can finally deduce that looper is an equivalence-} +-- isEquivLooper : isEquiv looper +-- isEquivLooper = transport (λ i → isEquiv (funExt (looperIntoBool) (~ i))) isEquivTranspIntLoop +-- where +-- isEquivTranspIntLoop : isEquiv λ x → cong (transport ((ua (S1≃SuspBool)))) (intLoop* x) +-- isEquivTranspIntLoop = compEquiv (intLoop* , intLoop*Equiv) +-- (cong (transport (ua S1≃SuspBool)) , +-- congEquiv (transport (ua S1≃SuspBool) , +-- transportEquiv (ua S1≃SuspBool) .snd) .snd) .snd + +-- isSetS1 : isSet (_≡_ {A = S₊ 1} north north) +-- isSetS1 = transport (λ i → isSet (helper i)) isSetInt +-- where +-- helper : Int ≡ (_≡_ {A = S₊ 1} north north) +-- helper = sym ΩS¹≡Int ∙ +-- (λ i → typ (Ω (S¹≡SuspBool i , transport (λ j → S¹≡SuspBool (j ∧ i)) base))) ∙ +-- (λ i → typ (Ω (ua S1≃SuspBool i , transport (λ j → ua S1≃SuspBool (i ∧ j)) north))) + +-- isEquivHelper2 : isOfHLevel 3 A → isEquiv {B = ∥ A ∥ 1} ∣_∣ +-- isEquivHelper2 ofHlevl = +-- isoToIsEquiv (iso ∣_∣ +-- (trElim (λ _ → ofHlevl) (λ a → a)) +-- (trElim {B = λ b → ∣ trElim (λ _ → ofHlevl) (λ a₁ → a₁) b ∣ ≡ b} (λ b → isOfHLevelSuc 3 (isOfHLevel∥∥ 1) ∣ trElim (λ _ → ofHlevl) (λ a₁ → a₁) b ∣ b) λ a → refl) +-- λ b → refl) + +-- isEquivHelper : {a b : A} → isOfHLevel 3 A → isEquiv {B = _≡_ {A = ∥ A ∥ 1} ∣ a ∣ ∣ b ∣ } (cong ∣_∣) +-- isEquivHelper {A = A} {a = a} {b = b} ofHlevl = congEquiv (∣_∣ , isEquivHelper2 ofHlevl) .snd + + +-- S1≡S¹ : S₊ 1 ≡ S¹ +-- S1≡S¹ = {!!} + + +-- d-map : typ (Ω ((Susp S¹) , north)) → S¹ +-- d-map p = subst HopfSuspS¹ p base + +-- d-mapId : (r : S¹) → d-map (ϕ base r) ≡ r +-- d-mapId r = (λ i → subst HopfSuspS¹ (□≡∙ (merid r) (sym (merid base)) (~ i)) base) ∙ +-- (substComposite-□ HopfSuspS¹ (merid r) (sym (merid base)) base) ∙ +-- test r +-- where +-- test : (r : S¹) → rot r base ≡ r +-- test base = refl +-- test (loop i) = refl + + +-- testProp : (B : A → Type ℓ') → {!x !} +-- testProp = {!!} + + +-- S³≡SuspSuspS¹ : S³ ≡ Susp (Susp S¹) +-- S³≡SuspSuspS¹ = S³≡SuspS² ∙ λ i → Susp (S²≡SuspS¹ i) + +-- d-mapComp : fiber d-map base ≡ (_≡_ {A = Susp (Susp S¹)} north north) +-- d-mapComp = sym (pathSigma≡sigmaPath {B = HopfSuspS¹} _ _) ∙ helper +-- where +-- helper : (_≡_ {A = Σ (Susp S¹) λ x → HopfSuspS¹ x} (north , base) (north , base)) ≡ (_≡_ {A = Susp (Susp S¹)} north north) +-- helper = (λ i → (_≡_ {A = S³≡TotalHopf (~ i)} +-- (transp (λ j → S³≡TotalHopf (~ i ∨ ~ j)) (~ i) (north , base)) +-- ((transp (λ j → S³≡TotalHopf (~ i ∨ ~ j)) (~ i) (north , base))))) ∙ +-- (λ i → _≡_ {A = S³} base base) ∙ +-- (λ i → _≡_ {A = S³≡SuspSuspS¹ i} (transp (λ j → S³≡SuspSuspS¹ (i ∧ j)) (~ i) base) ((transp (λ j → S³≡SuspSuspS¹ (i ∧ j)) (~ i) base))) + +-- n=1Equiv : isEquiv {A = S₊ 1} {B = typ (Ω ((S₊ 2) , north))} (ϕ north) +-- n=1Equiv = isoToIsEquiv (iso (ϕ north) +-- (λ x → {!(cong (transport (λ i → Susp (sym (S¹≡SuspBool) i))) (cong (transport (λ i → Susp ((sym (ua S1≃SuspBool)) i))) x))!}) +-- {!Hopf!} +-- {!HopfSuspS¹ !}) +-- where +-- helper : transport (λ i → Susp ((S¹≡SuspBool) (~ i))) (transport (λ i → Susp (ua S1≃SuspBool (~ i))) north) ≡ north +-- helper = refl -- fantastic + +-- is1Connected-dmap : is- 1 -Connected d-map +-- is1Connected-dmap base = transport (λ j → isContr (∥ d-mapComp (~ j) ∥ ℕ→ℕ₋₂ 1)) {!!} +-- is1Connected-dmap (loop i) = {!!} + +-- isOfHLevelLemma : (n : ℕ₋₂) → isOfHLevel (2+ n) A → isOfHLevel (suc (2+ n)) (Susp A) +-- isOfHLevelLemma neg2 hlvl = λ x y → {!!} +-- isOfHLevelLemma neg1 hlvl = {!!} +-- isOfHLevelLemma (ℕ→ℕ₋₂ n) hlvl = {!!} +-- where +-- helper : (n : ℕ) → isOfHLevel n A → (x y : Susp A) → isOfHLevel (suc n) (x ≡ y) +-- helper zero hlvl north north = {!s!} +-- helper zero hlvl north south = {!!} +-- helper zero hlvl north (merid a i) = {!!} +-- helper zero hlvl south north = {!!} +-- helper zero hlvl south south = {!!} +-- helper zero hlvl south (merid a i) = {!!} +-- helper zero hlvl (merid a i) y = {!!} +-- helper (suc n) hlvl x y = {!!} + +-- -- Kn→ΩKn+1Sucn : (n : ℕ) → Kn→ΩKn+1 (suc n) ≡ λ x → truncEquivΩ (ℕ→ℕ₋₂ (suc n)) .fst (trElim (λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (suc n))) (λ a → ∣ ϕ north a ∣) x) -- +-- -- Kn→ΩKn+1Sucn n = funExt (trElim (λ x → isOfHLevelSuc (suc (suc n)) +-- -- ((isOfHLevel∥∥ (ℕ→ℕ₋₂ (suc (suc n))) ∣ north ∣ ∣ north ∣) +-- -- (Kn→ΩKn+1 (suc n) x) +-- -- (truncEquivΩ (ℕ→ℕ₋₂ (suc n)) .fst (trElim (λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (suc n))) (λ a → ∣ ϕ north a ∣) x)))) +-- -- λ a → refl) + + +-- -- isEquivKn→ΩKn+1 : (n : ℕ) → isEquiv (Kn→ΩKn+1 n) +-- -- isEquivKn→ΩKn+1 zero = compEquiv (looper , isEquivLooper) (cong ∣_∣ , isEquivHelper hLevl3) .snd +-- -- where +-- -- hLevl3 : (x y : S₊ 1) (p q : x ≡ y) → isProp (p ≡ q) +-- -- hLevl3 x y = J (λ y p → (q : x ≡ y) → isProp (p ≡ q) ) +-- -- (transport (λ i → isSet (helper (~ i))) isSetInt refl) +-- -- where +-- -- helper : (x ≡ x) ≡ Int +-- -- helper = (λ i → transp (λ j → ua S1≃SuspBool (~ j ∨ ~ i)) (~ i) x ≡ transp (λ j → ua S1≃SuspBool (~ j ∨ ~ i)) (~ i) x) ∙ +-- -- (λ i → transp (λ j → S¹≡SuspBool (~ j ∨ ~ i)) (~ i) (transport (sym (ua S1≃SuspBool)) x) ≡ transp (λ j → S¹≡SuspBool (~ j ∨ ~ i)) (~ i) (transport (sym (ua S1≃SuspBool)) x)) ∙ +-- -- basedΩS¹≡Int (transport (sym S¹≡SuspBool) (transport (sym (ua S1≃SuspBool)) x)) +-- -- isEquivKn→ΩKn+1 (suc zero) = transport (λ i → isEquiv (Kn→ΩKn+1Sucn zero (~ i))) +-- -- (compEquiv (trElim (λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (suc zero))) (λ a → ∣ ϕ north a ∣) , +-- -- {!!}) +-- -- (truncEquivΩ (ℕ→ℕ₋₂ (suc zero))) .snd) +-- -- isEquivKn→ΩKn+1 (suc (suc n)) = {!!} From 81c36689c78e07c2ba7a094dfe4467b96ecfbc89 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 7 Apr 2020 02:05:44 +0200 Subject: [PATCH 14/89] Application --- Cubical/HITs/Truncation.agda | 3 - Cubical/HITs/Truncation/Connected/Base.agda | 14 +- .../Connected/FreudenthalProof/7-5-11.agda | 64 +- .../HITs/Truncation/Connected/Properties.agda | 41 +- Cubical/HITs/Truncation/Properties.agda | 25 + Cubical/ZCohomology/S1Loop.agda | 712 ++++++++++-------- 6 files changed, 526 insertions(+), 333 deletions(-) diff --git a/Cubical/HITs/Truncation.agda b/Cubical/HITs/Truncation.agda index 11296b12a..35e9c7d66 100644 --- a/Cubical/HITs/Truncation.agda +++ b/Cubical/HITs/Truncation.agda @@ -5,6 +5,3 @@ open import Cubical.HITs.Truncation.Base public open import Cubical.HITs.Truncation.Properties public import Cubical.HITs.Truncation.FromNegOne - -open import Cubical.HITs.Truncation.Connected.Base public -open import Cubical.HITs.Truncation.Connected.Properties public diff --git a/Cubical/HITs/Truncation/Connected/Base.agda b/Cubical/HITs/Truncation/Connected/Base.agda index 7eaf96d90..5535b29a8 100644 --- a/Cubical/HITs/Truncation/Connected/Base.agda +++ b/Cubical/HITs/Truncation/Connected/Base.agda @@ -5,6 +5,7 @@ open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Transport open import Cubical.Foundations.Isomorphism open import Cubical.Data.NatMinusTwo.Base open import Cubical.Data.Unit @@ -21,12 +22,17 @@ is- n -ConnectedType A = isContr (∥ A ∥ n) is-_-ConnectedType2 : ∀ {ℓ} (n : ℕ₋₂) (A : Type ℓ) → Type ℓ is- n -ConnectedType2 A = is- n -Connected (λ (x : A) → tt) +private + conTypEqHelper : ∀ {ℓ} (A : Type ℓ) (n : ℕ₋₂) (b : Unit) → A ≡ fiber (λ (x : A) → b) b + conTypEqHelper A n tt = isoToPath (iso (λ x → x , refl) (λ x → fst x) (λ b i → (refl {x = fst b} i) , ((isOfHLevelSuc 1 (isPropUnit) tt tt (snd b) refl) i)) λ a → refl) + {- The first def implies the second -} conTyp→conTyp2 : ∀ {ℓ} (n : ℕ₋₂) (A : Type ℓ) → is- n -ConnectedType A → is- n -ConnectedType2 A -conTyp→conTyp2 n A iscon tt = transport (λ i → isContr (∥ helper i ∥ n)) iscon - where - helper : A ≡ fiber (λ (x : A) → tt) tt - helper = isoToPath (iso (λ x → x , refl) (λ x → fst x) (λ b i → (refl {x = fst b} i) , ((isOfHLevelSuc 1 (isPropUnit) tt tt (snd b) refl) i)) λ a → refl) +conTyp→conTyp2 n A iscon tt = transport (λ i → isContr (∥ conTypEqHelper A n tt i ∥ n)) iscon + +{- The second def implies the first -} +conTyp2→conTyp : ∀ {ℓ} (n : ℕ₋₂) (A : Type ℓ) → is- n -ConnectedType2 A → is- n -ConnectedType A +conTyp2→conTyp n A iscon = transport (λ i → isContr (∥ conTypEqHelper A n tt (~ i) ∥ n)) (iscon tt) {- n-truncated types -} is-_-Truncated : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} diff --git a/Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-11.agda b/Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-11.agda index f1d6e7533..ec56315f4 100644 --- a/Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-11.agda +++ b/Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-11.agda @@ -6,7 +6,9 @@ open import Cubical.HITs.Truncation.Connected.FreudenthalProof.7-5-7 open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Univalence open import Cubical.Data.HomotopyGroup open import Cubical.Data.NatMinusTwo.Base @@ -15,6 +17,26 @@ open import Cubical.Data.Unit open import Cubical.HITs.Truncation.Base open import Cubical.HITs.Truncation.Properties +private + + Cor759 : ∀ {ℓ ℓ'} {A : Type ℓ} (n : ℕ₋₂) → is- n -ConnectedType A → (B : HLevel ℓ' (2+ n)) → isEquiv (λ (b : (fst B)) → λ (a : A) → b) + Cor759 {A = A} n isCon B = isEquivLemma (conInd-i→ii (λ (x : A) → tt) n (conTyp→conTyp2 n A isCon) (λ _ → B)) + where + + isEquivLemma : isEquiv (indConFun {A = A} {B = Unit} (λ x → tt) (λ _ → B)) → isEquiv (λ (b : (fst B)) → λ (a : A) → b) + isEquivLemma record { equiv-proof = eq } = record { equiv-proof = λ y → ((eq y .fst .fst tt) , (eq y .fst .snd)) , + λ z i → ((cong (λ y → (fst y) tt) ((eq y .snd) ((λ x → fst z) , snd z))) i) , + (cong (λ y → (snd y)) ((eq y .snd) ((λ x → fst z) , snd z))) i} + + Cor759← : ∀ {ℓ} {A : Type ℓ} (n : ℕ₋₂) → (∀ {ℓ'} (B : HLevel ℓ' (2+ n)) → isEquiv (λ (b : (fst B)) → λ (a : A) → b)) → is- n -ConnectedType A + Cor759← {A = A} n BEq = conTyp2→conTyp n A (conInd-iii→i (λ x → tt) n λ {ℓ'} P → conInd-ii→iii (λ x → tt) n P (compEquiv ((λ Q → Q tt) , helper P) (_ , BEq (P tt)) .snd)) + where + helper : ∀ {ℓ'} → (P : Unit → HLevel ℓ' (2+ n)) → isEquiv {A = ((b : Unit) → fst (P b))} λ Q → Q tt -- ∀ {ℓ'} P (ℓ'' : Unit) → fst (P₁ ℓ'') + helper P = isoToIsEquiv (iso (λ Q → Q tt) invFun (λ b → refl) λ b → refl) + where + invFun : fst (P tt) → (b : Unit) → fst (P b) + invFun Q tt = Q + Lemma7-5-11 : ∀{ℓ} {A : Type ℓ} {a : A} → (n : ℕ₋₂) → (is- (suc₋₂ n) -ConnectedType A) → is- n -Connected (λ (x : Unit) → a) Lemma7-5-11 {A = A} {a = a} n isCon = conInd-iii→i (λ _ → a) _ λ P → ((λ g → fst (helper P (g tt)))) , (λ g → funExt (λ x → snd (helper P (g x)))) @@ -27,15 +49,6 @@ Lemma7-5-11 {A = A} {a = a} n isCon = conInd-iii→i (λ _ → a) _ λ P → (( (transport (λ i → Bs3 a (~ i ∧ ~ j)) (transport (λ i → Bs3 a (i ∧ (~ j))) u)) where - Cor759 : ∀ {ℓ'} (n : ℕ₋₂) → is- n -ConnectedType A → (B : HLevel ℓ' (2+ n)) → isEquiv (λ (b : (fst B)) → λ (a : A) → b) - Cor759 n isCon B = isEquivLemma (conInd-i→ii (λ (x : A) → tt) n (conTyp→conTyp2 n A isCon) (λ _ → B)) - where - - isEquivLemma : isEquiv (indConFun {A = A} {B = Unit} (λ x → tt) (λ _ → B)) → isEquiv (λ (b : (fst B)) → λ (a : A) → b) - isEquivLemma record { equiv-proof = eq } = record { equiv-proof = λ y → ((eq y .fst .fst tt) , (eq y .fst .snd)) , - λ z i → ((cong (λ y → (fst y) tt) ((eq y .snd) ((λ x → fst z) , snd z))) i) , - (cong (λ y → (snd y)) ((eq y .snd) ((λ x → fst z) , snd z))) i} - Bs2 : Σ (HLevel ℓ (2+ n)) λ B → P ≡ λ (a : A) → B Bs2 = (equiv-proof (Cor759 (suc₋₂ n) isCon (HLevel ℓ (2+ n) , isOfHLevelHLevel _ ))) P .fst .fst , @@ -46,3 +59,36 @@ Lemma7-5-11 {A = A} {a = a} n isCon = conInd-iii→i (λ _ → a) _ λ P → (( Bs3 : (a : A) → fst (P a) ≡ B* Bs3 a = cong (λ x → fst (x a)) (snd Bs2) + + + +Lemma7-5-11← : ∀{ℓ} {A : Type ℓ} {a : A} → (n : ℕ₋₂) → is- n -Connected (λ (x : Unit) → a) → (is- (suc₋₂ n) -ConnectedType A) +Lemma7-5-11← {A = A} {a = a} n iscon = Cor759← (suc₋₂ n) lemmas.isEquivMap + where + module lemmas {ℓ' : Level} (B : HLevel ℓ' (2+ (suc₋₂ n))) where + + P : (f : A → fst B) → A → Type _ + P f a2 = (f a2 ≡ f a) + + hLevelP : (f : A → fst B) → (a : A) → isOfHLevel (2+ n) (P f a) + hLevelP f a2 = helper n (f a2) (f a) (snd B) + where + helper : (n : ℕ₋₂) → (b1 b2 : (fst B)) → isOfHLevel (2+ suc₋₂ n) (fst B) → isOfHLevel (2+ n) (b1 ≡ b2) + helper neg2 b1 b2 hlvl = (hlvl b1 b2) , λ y → isProp→isSet hlvl b1 b2 _ _ + helper (-1+ n) b1 b2 hlvl = hlvl b1 b2 + + forAllP : (f : A → fst B) → (a : A) → P f a + forAllP f = equiv-proof (conInd-i→ii (λ x → a) n iscon (λ a → P f a , hLevelP f a)) (λ a → refl) .fst .fst + where + helper : (n : ℕ₋₂) → (b1 b2 : (fst B)) → isOfHLevel (2+ suc₋₂ n) (fst B) → isOfHLevel (2+ n) (b1 ≡ b2) + helper neg2 b1 b2 hlvl = (hlvl b1 b2) , λ y → isProp→isSet hlvl b1 b2 _ _ + helper (-1+ n) b1 b2 hlvl = hlvl b1 b2 + + sect : section (λ (b : fst B) (a : A) → b) λ f → f a + sect f = funExt λ y → sym ((forAllP f) y) + + isEquivMap : isEquiv (λ (b : fst B) (a : A) → b) + isEquivMap = isoToIsEquiv (iso (λ b a → b) + (λ f → f a) + sect + λ h → refl) diff --git a/Cubical/HITs/Truncation/Connected/Properties.agda b/Cubical/HITs/Truncation/Connected/Properties.agda index 1a57c14e2..40d366c15 100644 --- a/Cubical/HITs/Truncation/Connected/Properties.agda +++ b/Cubical/HITs/Truncation/Connected/Properties.agda @@ -14,12 +14,14 @@ open import Cubical.Foundations.Equiv.Properties open import Cubical.Foundations.Function open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Univalence open import Cubical.Data.NatMinusTwo.Base open import Cubical.Data.Prod.Base open import Cubical.HITs.Susp open import Cubical.Data.Nat open import Cubical.HITs.Truncation.Base -open import Cubical.HITs.Truncation.Properties +open import Cubical.HITs.Truncation.Properties renaming (elim to trElim) +open import Cubical.HITs.Nullification open import Cubical.Data.HomotopyGroup @@ -29,6 +31,13 @@ private A : Type ℓ B : Type ℓ' +isConnectedSubtr : (n m : ℕ) (f : A → B) → is- (-2+ (n + m)) -Connected f → is- (-2+ n) -Connected f +isConnectedSubtr n m f iscon b = transport (λ i → isContr (ua (truncOfTruncEq {A = fiber f b} n m) (~ i) )) + (∣ iscon b .fst ∣ , + trElim (λ x → isOfHLevelPath n (isOfHLevel∥∥ (-2+ n)) _ _) + λ a → cong ∣_∣ (iscon b .snd a)) +-- + {- The "induction principle" for n-connected functions from the HoTT book (see Lemma 7.5.7): @@ -58,10 +67,19 @@ module elim (f : A → B) (n : ℕ₋₂) where trivFunCon : ∀{ℓ} {A : Type ℓ} {a : A} → (n : ℕ₋₂) → (is- (suc₋₂ n) -ConnectedType A) → is- n -Connected (λ (x : Unit) → a) trivFunCon = Lemma7-5-11 +trivFunCon← : ∀{ℓ} {A : Type ℓ} {a : A} → (n : ℕ₋₂) → is- n -Connected (λ (x : Unit) → a) → (is- (suc₋₂ n) -ConnectedType A) +trivFunCon← = Lemma7-5-11← + {- n-connected functions induce equivalences between n-truncations -} conEquiv : (n : ℕ₋₂) (f : A → B) → (is- n -Connected f) → ∥ A ∥ n ≃ ∥ B ∥ n conEquiv = Lemma7-5-14 +conEquiv2 : (n m : ℕ) (f : A → B) → (is- (-2+ (n + m)) -Connected f) → ∥ A ∥ (-2+ n) ≃ ∥ B ∥ (-2+ n) +conEquiv2 n m f iscon = conEquiv (-2+ n) f (isConnectedSubtr n m f iscon) + +conEquiv3 : (n m : ℕ) (f : A → B) → Σ[ x ∈ ℕ ] n + x ≡ m → (is- (-2+ m) -Connected f) → ∥ A ∥ (-2+ n) ≃ ∥ B ∥ (-2+ n) +conEquiv3 n m f (x , pf) iscon = conEquiv (-2+ n) f (isConnectedSubtr n x f (transport (λ i → is- (-2+ pf (~ i)) -Connected f) iscon)) -- + {- Wedge connectivity lemma (Lemma 8.6.2 in the HoTT book) -} WedgeConn : (A : Pointed ℓ) (B : Pointed ℓ') (n m : ℕ) → (is- (ℕ→ℕ₋₂ n) -ConnectedType (typ A)) → @@ -80,12 +98,17 @@ WedgeConn = Lemma8-6-2 FthalFun : A → A → typ (Ω ((Susp A) , north)) FthalFun a x = merid x ∙ sym (merid a) -FthalFun-2nConnected : (n : ℕ) (a : A) - (iscon : is- (ℕ→ℕ₋₂ n) -ConnectedType A) → - is- ℕ→ℕ₋₂ (n + n) -Connected (FthalFun a) -FthalFun-2nConnected n a iscon = Thm8-6-4 n a iscon +abstract + FthalFun-2nConnected : (n : ℕ) (a : A) + (iscon : is- (ℕ→ℕ₋₂ n) -ConnectedType A) → + is- ℕ→ℕ₋₂ (n + n) -Connected (FthalFun a) + FthalFun-2nConnected n a iscon = Thm8-6-4 n a iscon + + Freudenthal : (n : ℕ) (A : Pointed ℓ) → + is- (ℕ→ℕ₋₂ n) -ConnectedType (typ A) → + ∥ typ A ∥ (ℕ→ℕ₋₂ (n + n)) ≃ ∥ typ (Ω (Susp (typ A) , north)) ∥ ((ℕ→ℕ₋₂ (n + n))) + Freudenthal n A iscon = conEquiv _ (FthalFun (pt A)) (FthalFun-2nConnected n (pt A) iscon) + + +------------------------------ -Freudenthal : (n : ℕ) (A : Pointed ℓ) → - is- (ℕ→ℕ₋₂ n) -ConnectedType (typ A) → - ∥ typ A ∥ (ℕ→ℕ₋₂ (n + n)) ≃ ∥ typ (Ω (Susp (typ A) , north)) ∥ ((ℕ→ℕ₋₂ (n + n))) -Freudenthal n A iscon = conEquiv _ (FthalFun (pt A)) (FthalFun-2nConnected n (pt A) iscon) diff --git a/Cubical/HITs/Truncation/Properties.agda b/Cubical/HITs/Truncation/Properties.agda index d60c10f12..465d3e1a8 100644 --- a/Cubical/HITs/Truncation/Properties.agda +++ b/Cubical/HITs/Truncation/Properties.agda @@ -124,6 +124,7 @@ isSnNull→isOfHLevel {n = -1+ n} nA = isSphereFilled→isOfHLevelSuc (λ f → isOfHLevel∥∥ : (n : ℕ₋₂) → isOfHLevel (2+ n) (∥ A ∥ n) isOfHLevel∥∥ neg2 = hub ⊥.rec , λ _ → ≡hub ⊥.rec isOfHLevel∥∥ (-1+ n) = isSphereFilled→isOfHLevelSuc isSphereFilled∥∥ + -- isOfHLevel∥∥ n = isSnNull→isOfHLevel isNull-Null -- ∥_∥ n is a modality @@ -363,3 +364,27 @@ PathΩ {a = a} n = PathIdTrunc {a = a} {b = a} n truncEquivΩ : {a : A} (n : ℕ₋₂) → (∥ a ≡ a ∥ n) ≃ (_≡_ {A = ∥ A ∥ (suc₋₂ n)} ∣ a ∣ ∣ a ∣) truncEquivΩ {a = a} n = isoToEquiv (IsoFinal2 ∣ a ∣ ∣ a ∣) + +-------------------------- + + +truncOfTruncEq : (n m : ℕ) → ∥ A ∥ (-2+ n) ≃ ∥ ∥ A ∥ (-2+ (n + m)) ∥ (-2+ n) +truncOfTruncEq {A = A} n m = isoToEquiv (iso fun funInv sect retr) + where + fun : ∥ A ∥ (-2+ n) → ∥ ∥ A ∥ (-2+ (n + m)) ∥ (-2+ n) + fun = elim (λ _ → isOfHLevel∥∥ (-2+ n)) λ a → ∣ ∣ a ∣ ∣ + funInv : ∥ ∥ A ∥ (-2+ (n + m)) ∥ (-2+ n) → ∥ A ∥ (-2+ n) + funInv = elim (λ _ → isOfHLevel∥∥ (-2+ n)) + (elim (λ _ → transport (λ i → isOfHLevel (+-comm n m (~ i)) (∥ A ∥ (-2+ n))) + (isOfHLevelPlus m (isOfHLevel∥∥ (-2+ n)))) + λ a → ∣ a ∣) + + sect : section fun funInv + sect = elim (λ x → isOfHLevelPath n (isOfHLevel∥∥ (-2+ n)) _ _ ) + (elim (λ x → isOfHLevelPath (n + m) (transport (λ i → isOfHLevel (+-comm n m (~ i)) + (∥ ∥ A ∥ (-2+ (n + m)) ∥ (-2+ n))) + (isOfHLevelPlus m (isOfHLevel∥∥ (-2+ n)))) _ _ ) + λ a → refl) + + retr : retract fun funInv + retr = elim (λ x → isOfHLevelPath n (isOfHLevel∥∥ (-2+ n)) _ _) λ a → refl diff --git a/Cubical/ZCohomology/S1Loop.agda b/Cubical/ZCohomology/S1Loop.agda index 1159465a3..9c5e0f845 100644 --- a/Cubical/ZCohomology/S1Loop.agda +++ b/Cubical/ZCohomology/S1Loop.agda @@ -3,6 +3,7 @@ module Cubical.ZCohomology.S1Loop where open import Cubical.ZCohomology.Base open import Cubical.HITs.S1 +open import Cubical.HITs.S2 open import Cubical.HITs.S3 open import Cubical.HITs.Sn open import Cubical.Foundations.HLevels @@ -18,11 +19,14 @@ open import Cubical.Data.Empty open import Cubical.Data.Sigma open import Cubical.Data.Prod.Base open import Cubical.HITs.Susp -open import Cubical.HITs.Nullification -open import Cubical.Data.Int +open import Cubical.HITs.Nullification +open import Cubical.Data.Unit +open import Cubical.Data.Int renaming (_+_ to +Int) open import Cubical.Data.Nat open import Cubical.HITs.Truncation renaming (elim to trElim) open import Cubical.HITs.Hopf +open import Cubical.HITs.Truncation.Connected.Base public +open import Cubical.HITs.Truncation.Connected.Properties public open import Cubical.HITs.Pushout open import Cubical.Data.Sum.Base @@ -35,318 +39,410 @@ private B : Type ℓ' -pred₋₂ : ℕ₋₂ → ℕ₋₂ -pred₋₂ neg2 = neg2 -pred₋₂ neg1 = neg2 -pred₋₂ (ℕ→ℕ₋₂ zero) = neg1 -pred₋₂ (ℕ→ℕ₋₂ (suc n)) = ℕ→ℕ₋₂ n - -is-_-Connected : (n : ℕ₋₂) (f : A → B) → Type _ -is-_-Connected {B = B} n f = (x : B) → isContr (∥ fiber f x ∥ n) - +compToIdEquiv : (f : A → B) (g : B → A) → f ∘ g ≡ idfun B → isEquiv f → isEquiv g +compToIdEquiv f g id iseqf = + isoToIsEquiv (iso g + f + (λ b → (λ i → (equiv-proof iseqf (f b) .snd (g (f b) , cong (λ h → h (f b)) id) (~ i)) .fst ) ∙ + cong (λ x → (equiv-proof iseqf (f b) .fst .fst )) id ∙ + λ i → (equiv-proof iseqf (f b) .snd) (b , refl) i .fst ) + λ a → cong (λ f → f a) id) + +Susp≡Push : Susp A ≡ Pushout {A = A} (λ a → tt) λ a → tt +Susp≡Push {A = A} = isoToPath (iso fun inverse sect retr) + where + fun : Susp A → Pushout {A = A} (λ a → tt) (λ a → tt) + fun north = inl tt + fun south = inr tt + fun (merid a i) = push a i + inverse : Pushout {A = A} (λ a → tt) (λ a → tt) → Susp A + inverse (inl tt) = north + inverse (inr tt) = south + inverse (push a i) = merid a i + + sect : section fun inverse + sect (inl tt) = refl + sect (inr tt) = refl + sect (push a i) = refl + + retr : retract fun inverse + retr north = refl + retr south = refl + retr (merid a i) = refl Pr2310 : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} (n : ℕ₋₂) (f : C → A) (g : C → B) → is- n -Connected f → is-_-Connected {A = B} {B = Pushout f g} n inr -Pr2310 {A = A} {B = B} {C = C} n f g iscon = {!!} +Pr2310 {A = A} {B = B} {C = C} n f g iscon = elim.3→1 inr n λ P → (λ k → helpLemmas.k P k) , λ b → refl where - module hej {ℓ : Level} (P : (Pushout f g) → HLevel ℓ (2+ n)) + module helpLemmas {ℓ : Level} (P : (Pushout f g) → HLevel ℓ (2+ n)) (h : (b : B) → typ (P (inr b))) where - Q : A → Type _ - Q a = typ (P (inl a)) - - fun : (c : C) → Q (f c) - fun = transport (λ i → (c : C) → typ (P (push c (~ i)))) λ c → h (g c) - - typeFam : (d : Pushout f g) → typ (P d) - typeFam (inl x) = {!!} - typeFam (inr x) = {!!} - typeFam (push a i) = {!!} - - - -Pr242 : (n : ℕ₋₂) → isContr (∥ {!S₊ n!} ∥ pred₋₂ n) -Pr242 = {!!} - - - - - --- congFunct : {a b c : A} (f : A → B) (p : a ≡ b) (q : b ≡ c) → cong f (p ∙ q) ≡ cong f p ∙ cong f q --- congFunct f p q i = hcomp (λ j → λ{(i = i0) → rUnit (cong f (p ∙ q)) (~ j) ; --- (i = i1) → cong f (rUnit p (~ j)) ∙ cong f q}) --- (cong f (p ∙ (λ k → q (k ∧ (~ i)))) ∙ cong f λ k → q ((~ i) ∨ k) ) --- -- --- symDistr : {a b c : A} (p : a ≡ b) (q : b ≡ c) → sym (p ∙ q) ≡ sym q ∙ sym p --- symDistr p q i = hcomp (λ j → λ{(i = i0) → rUnit (sym (p ∙ q)) (~ j) ; --- (i = i1) → sym (lUnit q (~ j)) ∙ sym p}) --- (sym ((λ k → p (k ∨ i)) ∙ q) ∙ sym λ k → p (i ∧ k)) - --- {- We want to prove that Kn≃ΩKn+1. For this we need the map ϕ-} - --- private --- ϕ : (pt a : A) → typ (Ω (Susp A , north)) --- ϕ pt a = (merid a) ∙ sym (merid pt) - --- {- To define the map for n=0 we use the λ k → loopᵏ map for S₊ 1. The loop is given by ϕ south north -} - - --- looper : Int → _≡_ {A = S₊ 1} north north --- looper (pos zero) = refl --- looper (pos (suc n)) = looper (pos n) ∙ (ϕ south north) --- looper (negsuc zero) = sym (ϕ south north) --- looper (negsuc (suc n)) = looper (negsuc n) ∙ sym (ϕ south north) - --- {- The map of Kn≃ΩKn+1 is given as follows. -} --- Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) --- Kn→ΩKn+1 zero x = cong ∣_∣ (looper x) --- Kn→ΩKn+1 (suc n) = trElim (λ x → (isOfHLevel∥∥ (ℕ→ℕ₋₂ (suc (suc n))) ∣ north ∣ ∣ north ∣)) --- λ a → cong ∣_∣ ((merid a) ∙ (sym (merid north))) - --- {- --- We want to show that the function (looper : Int → S₊ 1) defined by λ k → loopᵏ is an equivalece. We already know that the corresponding function (intLoop : Int → S¹ is) an equivalence, --- so the idea is to show that when intLoop is transported along a suitable path S₊ 1 ≡ S¹ we get looper. Instead of using S₊ 1 straight away, we begin by showing this for the equivalent Susp Bool. --- -} - --- -- loop for Susp Bool -- --- loop* : _≡_ {A = Susp Bool} north north --- loop* = merid false ∙ sym (merid true) - --- -- the loop function -- --- intLoop* : Int → _≡_ {A = Susp Bool} north north --- intLoop* (pos zero) = refl --- intLoop* (pos (suc n)) = intLoop* (pos n) ∙ loop* --- intLoop* (negsuc zero) = sym loop* --- intLoop* (negsuc (suc n)) = intLoop* (negsuc n) ∙ sym loop* - --- -- we show that the loop spaces agree -- --- loopSpId : ΩS¹ ≡ _≡_ {A = Susp Bool} north north --- loopSpId i = typ (Ω (S¹≡SuspBool i , transp ((λ j → S¹≡SuspBool (j ∧ i))) (~ i) base)) - --- -- the transport map -- --- altMap2 : Int → _≡_ {A = Susp Bool} north north --- altMap2 n i = transport S¹≡SuspBool (intLoop n i) - --- -- We show that the transporting intLoop over S¹≡SuspBool gives intLoop* (modulo function extensionality) -- --- altMap≡intLoop*2 : (x : Int) → intLoop* x ≡ altMap2 x --- altMap≡intLoop*2 (pos zero) = refl --- altMap≡intLoop*2 (pos (suc n)) = (λ i → (altMap≡intLoop*2 (pos n) i) ∙ loop*) ∙ --- sym (helper n) - --- where --- helper : (n : ℕ) → altMap2 (pos (suc n)) ≡ altMap2 (pos n) ∙ loop* --- helper zero = (λ i j → transport S¹≡SuspBool (lUnit loop (~ i) j)) ∙ --- (λ i j → transport S¹≡SuspBool (loop j)) ∙ --- (λ i j → transportRefl ((merid false ∙ (sym (merid true))) j) i ) ∙ --- lUnit loop* --- helper (suc n) = anotherHelper n ∙ --- (λ i → altMap2 (pos (suc n)) ∙ helper zero i) ∙ --- (λ i → altMap2 (pos (suc n)) ∙ lUnit loop* (~ i)) --- where --- anotherHelper : (n : ℕ) → altMap2 (pos (suc (suc n))) ≡ altMap2 (pos (suc n)) ∙ altMap2 (pos 1) --- anotherHelper n = ((λ i j → transport S¹≡SuspBool ((intLoop (pos (suc n)) ∙ loop) j))) ∙ --- rUnit (λ j → transport S¹≡SuspBool ((intLoop (pos (suc n)) ∙ loop) j) ) ∙ --- (λ i → (λ j → transport S¹≡SuspBool ((intLoop (pos (suc n)) ∙ λ k → loop (k ∧ (~ i))) j)) ∙ λ j → transport S¹≡SuspBool (loop (j ∨ (~ i))) ) ∙ --- (λ i → (λ j → transport S¹≡SuspBool (rUnit (intLoop (pos (suc n))) (~ i) j)) ∙ λ j → transport S¹≡SuspBool ((lUnit loop i) j)) - --- altMap≡intLoop*2 (negsuc zero) = sym ((λ i j → transport S¹≡SuspBool (loop (~ j))) ∙ --- λ i j → transportRefl (loop* (~ j)) i ) --- altMap≡intLoop*2 (negsuc (suc n)) = helper n --- where --- altMapneg1 : altMap2 (negsuc zero) ≡ sym (loop*) --- altMapneg1 i j = transportRefl (loop* (~ j)) i - --- anotherHelper : (n : ℕ) → altMap2 (negsuc (suc n)) ≡ altMap2 (negsuc n) ∙ altMap2 (negsuc zero) --- anotherHelper n = ((λ i → rUnit (λ j → (transport S¹≡SuspBool ((intLoop (negsuc n) ∙ sym loop) j))) i)) ∙ --- ((λ i → (λ j → transport S¹≡SuspBool ((intLoop (negsuc n) ∙ (λ k → loop ((~ k) ∨ i))) j)) ∙ λ j → transport S¹≡SuspBool (loop ((~ j) ∧ i)))) ∙ --- (λ i → ((λ j → transport S¹≡SuspBool (rUnit (intLoop (negsuc n)) (~ i) j))) ∙ altMap2 (negsuc zero)) + Q : A → HLevel _ (2+ n) + Q a = (P (inl a)) + + fun : (c : C) → typ (Q (f c)) + fun c = transport (λ i → typ (P (push c (~ i)))) (h (g c)) + + k : (d : Pushout f g) → typ (P d) + k (inl x) = elim.2→3 f n Q (elim.1→2 f n iscon Q) .fst fun x + k (inr x) = h x + k (push a i) = hcomp (λ k → λ{(i = i0) → elim.2→3 f n Q + (elim.1→2 f n iscon Q) .fst + fun (f a) ; + (i = i1) → transportTransport⁻ (λ j → typ (P (push a j))) (h (g a)) k}) + (transp (λ j → typ (P (push a (i ∧ j)))) + (~ i) + (elim.2→3 f n Q + (elim.1→2 f n iscon Q) .snd fun i a)) + + +Pr242 : (n : ℕ) → is- (-1+ n) -ConnectedType (S₊ n) +Pr242 zero = ∣ north ∣ , (isOfHLevel∥∥ neg1 ∣ north ∣) +Pr242 (suc n) = transport (λ i → is- ℕ→ℕ₋₂ n -ConnectedType (Susp≡Push {A = S₊ n} (~ i))) + (trivFunCon← {A = Pushout {A = S₊ n} (λ x → tt) λ x → tt} {a = inr tt } _ + (transport (λ i → is- (-1+ n) -Connected (mapsAgree (~ i))) + (Pr2310 _ (λ x → tt) (λ x → tt) λ{tt → transport (λ i → isContr (∥ fibUnit (~ i) ∥ (-1+ n))) (Pr242 n)}))) + where + mapsAgree : Path ((x : Unit) → Pushout {A = S₊ n} (λ x → tt) (λ x → tt)) (λ (x : Unit) → inr tt) inr + mapsAgree = funExt λ {tt → refl} + fibUnit : fiber (λ (x : S₊ n) → tt) tt ≡ S₊ n + fibUnit = isoToPath (iso (λ b → fst b) (λ a → a , refl) (λ b → refl) λ b i → (fst b) , isProp→isSet isPropUnit tt tt refl (snd b) i) + + + -- need to show susp and pushout are same here + +Pr242SpecCase : is- 2 -ConnectedType (Susp (Susp S¹)) +Pr242SpecCase = transport (λ i → is- 2 -ConnectedType (helper i)) (Pr242 3) + where + helper : S₊ 3 ≡ Susp (Susp S¹) + helper = (λ i → Susp (Susp (Susp (ua Bool≃Susp⊥ (~ i))))) ∙ λ i → Susp (Susp (S¹≡SuspBool (~ i))) + + +congFunct : {a b c : A} (f : A → B) (p : a ≡ b) (q : b ≡ c) → cong f (p ∙ q) ≡ cong f p ∙ cong f q +congFunct f p q i = hcomp (λ j → λ{(i = i0) → rUnit (cong f (p ∙ q)) (~ j) ; + (i = i1) → cong f (rUnit p (~ j)) ∙ cong f q}) + (cong f (p ∙ (λ k → q (k ∧ (~ i)))) ∙ cong f λ k → q ((~ i) ∨ k) ) + -- +symDistr : {a b c : A} (p : a ≡ b) (q : b ≡ c) → sym (p ∙ q) ≡ sym q ∙ sym p +symDistr p q i = hcomp (λ j → λ{(i = i0) → rUnit (sym (p ∙ q)) (~ j) ; + (i = i1) → sym (lUnit q (~ j)) ∙ sym p}) + (sym ((λ k → p (k ∨ i)) ∙ q) ∙ sym λ k → p (i ∧ k)) + +{- We want to prove that Kn≃ΩKn+1. For this we need the map ϕ-} + +private + ϕ : (pt a : A) → typ (Ω (Susp A , north)) + ϕ pt a = (merid a) ∙ sym (merid pt) + +{- To define the map for n=0 we use the λ k → loopᵏ map for S₊ 1. The loop is given by ϕ south north -} + + +looper : Int → _≡_ {A = S₊ 1} north north +looper (pos zero) = refl +looper (pos (suc n)) = looper (pos n) ∙ (ϕ south north) +looper (negsuc zero) = sym (ϕ south north) +looper (negsuc (suc n)) = looper (negsuc n) ∙ sym (ϕ south north) + +{- The map of Kn≃ΩKn+1 is given as follows. -} +Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) +Kn→ΩKn+1 zero x = cong ∣_∣ (looper x) +Kn→ΩKn+1 (suc n) = trElim (λ x → (isOfHLevel∥∥ (ℕ→ℕ₋₂ (suc (suc n))) ∣ north ∣ ∣ north ∣)) + λ a → cong ∣_∣ ((merid a) ∙ (sym (merid north))) + +{- +We want to show that the function (looper : Int → S₊ 1) defined by λ k → loopᵏ is an equivalece. We already know that the corresponding function (intLoop : Int → S¹ is) an equivalence, +so the idea is to show that when intLoop is transported along a suitable path S₊ 1 ≡ S¹ we get looper. Instead of using S₊ 1 straight away, we begin by showing this for the equivalent Susp Bool. +-} + +-- loop for Susp Bool -- +loop* : _≡_ {A = Susp Bool} north north +loop* = merid false ∙ sym (merid true) + +-- the loop function -- +intLoop* : Int → _≡_ {A = Susp Bool} north north +intLoop* (pos zero) = refl +intLoop* (pos (suc n)) = intLoop* (pos n) ∙ loop* +intLoop* (negsuc zero) = sym loop* +intLoop* (negsuc (suc n)) = intLoop* (negsuc n) ∙ sym loop* + +-- we show that the loop spaces agree -- +loopSpId : ΩS¹ ≡ _≡_ {A = Susp Bool} north north +loopSpId i = typ (Ω (S¹≡SuspBool i , transp ((λ j → S¹≡SuspBool (j ∧ i))) (~ i) base)) + +-- the transport map -- +altMap2 : Int → _≡_ {A = Susp Bool} north north +altMap2 n i = transport S¹≡SuspBool (intLoop n i) + +-- We show that the transporting intLoop over S¹≡SuspBool gives intLoop* (modulo function extensionality) -- +altMap≡intLoop*2 : (x : Int) → intLoop* x ≡ altMap2 x +altMap≡intLoop*2 (pos zero) = refl +altMap≡intLoop*2 (pos (suc n)) = (λ i → (altMap≡intLoop*2 (pos n) i) ∙ loop*) ∙ + sym (helper n) + + where + helper : (n : ℕ) → altMap2 (pos (suc n)) ≡ altMap2 (pos n) ∙ loop* + helper zero = (λ i j → transport S¹≡SuspBool (lUnit loop (~ i) j)) ∙ + (λ i j → transport S¹≡SuspBool (loop j)) ∙ + (λ i j → transportRefl ((merid false ∙ (sym (merid true))) j) i ) ∙ + lUnit loop* + helper (suc n) = anotherHelper n ∙ + (λ i → altMap2 (pos (suc n)) ∙ helper zero i) ∙ + (λ i → altMap2 (pos (suc n)) ∙ lUnit loop* (~ i)) + where + anotherHelper : (n : ℕ) → altMap2 (pos (suc (suc n))) ≡ altMap2 (pos (suc n)) ∙ altMap2 (pos 1) + anotherHelper n = ((λ i j → transport S¹≡SuspBool ((intLoop (pos (suc n)) ∙ loop) j))) ∙ + rUnit (λ j → transport S¹≡SuspBool ((intLoop (pos (suc n)) ∙ loop) j) ) ∙ + (λ i → (λ j → transport S¹≡SuspBool ((intLoop (pos (suc n)) ∙ λ k → loop (k ∧ (~ i))) j)) ∙ λ j → transport S¹≡SuspBool (loop (j ∨ (~ i))) ) ∙ + (λ i → (λ j → transport S¹≡SuspBool (rUnit (intLoop (pos (suc n))) (~ i) j)) ∙ λ j → transport S¹≡SuspBool ((lUnit loop i) j)) + +altMap≡intLoop*2 (negsuc zero) = sym ((λ i j → transport S¹≡SuspBool (loop (~ j))) ∙ + λ i j → transportRefl (loop* (~ j)) i ) +altMap≡intLoop*2 (negsuc (suc n)) = helper n + where + altMapneg1 : altMap2 (negsuc zero) ≡ sym (loop*) + altMapneg1 i j = transportRefl (loop* (~ j)) i + + anotherHelper : (n : ℕ) → altMap2 (negsuc (suc n)) ≡ altMap2 (negsuc n) ∙ altMap2 (negsuc zero) + anotherHelper n = ((λ i → rUnit (λ j → (transport S¹≡SuspBool ((intLoop (negsuc n) ∙ sym loop) j))) i)) ∙ + ((λ i → (λ j → transport S¹≡SuspBool ((intLoop (negsuc n) ∙ (λ k → loop ((~ k) ∨ i))) j)) ∙ λ j → transport S¹≡SuspBool (loop ((~ j) ∧ i)))) ∙ + (λ i → ((λ j → transport S¹≡SuspBool (rUnit (intLoop (negsuc n)) (~ i) j))) ∙ altMap2 (negsuc zero)) --- helper : (n : ℕ) → intLoop* (negsuc n) ∙ (sym loop*) ≡ altMap2 (negsuc (suc n)) --- helper zero = (λ i → altMapneg1 (~ i) ∙ altMapneg1 (~ i)) ∙ sym (anotherHelper zero) --- helper (suc n) = (λ i → (helper n i) ∙ altMapneg1 (~ i) ) ∙ --- sym (anotherHelper (suc n)) - - --- {- We have that the transported map is an equivalence -} --- mapIsEquiv : isEquiv altMap2 --- mapIsEquiv = compEquiv (intLoop , (isoToIsEquiv (iso intLoop winding (decodeEncode base) windingIntLoop))) ((λ x i → transport S¹≡SuspBool (x i)) , helper) .snd --- where --- helper : isEquiv {A = ΩS¹} {B = _≡_ {A = Susp Bool} north north} (λ x i → transport S¹≡SuspBool (x i)) --- helper = congEquiv (transport S¹≡SuspBool , helper3) .snd --- where --- helper2 : transport S¹≡SuspBool ≡ S¹→SuspBool --- helper2 = funExt λ z → transportRefl (S¹→SuspBool z) --- helper3 : isEquiv (transport S¹≡SuspBool ) --- helper3 = transport (λ i → isEquiv (helper2 (~ i))) (S¹≃SuspBool .snd) - --- {- From this we get that intLoop* is an equivalence-} --- intLoop*Equiv : isEquiv intLoop* --- intLoop*Equiv = transport (λ i → isEquiv (funExt (altMap≡intLoop*2) (~ i))) mapIsEquiv - --- {- We now only need to work with Susp Bool and S₊ 1. We transport intLoop* along a path S1≡SuspBool and show that this gives us looper. -} --- SuspBool→S1 : Susp Bool → S₊ 1 --- SuspBool→S1 north = north --- SuspBool→S1 south = south --- SuspBool→S1 (merid false i) = merid north i --- SuspBool→S1 (merid true i) = merid south i - --- {- We need to spell out the trivial equivalence S1≃SuspBool. Of course it's important to make sure that we choose the right version of it. -} --- S1≃SuspBool : Susp Bool ≃ S₊ 1 --- S1≃SuspBool = isoToEquiv (iso SuspBool→S1 S1→SuspBool retrHelper sectHelper) --- where --- S1→SuspBool : S₊ 1 → Susp Bool --- S1→SuspBool north = north --- S1→SuspBool south = south --- S1→SuspBool (merid north i) = merid false i --- S1→SuspBool (merid south i) = merid true i - --- sectHelper : section S1→SuspBool SuspBool→S1 --- sectHelper north = refl --- sectHelper south = refl --- sectHelper (merid false i) = refl --- sectHelper (merid true i) = refl - --- retrHelper : retract S1→SuspBool SuspBool→S1 --- retrHelper north = refl --- retrHelper south = refl --- retrHelper (merid north i) = refl --- retrHelper (merid south i) = refl - --- {- We show that transporting intLoop* along (ua S1≃SuspBool) gives us looper (again, modulo function extensionality) -} --- looperIntoBool : (x : Int) → looper x ≡ λ i → transport ((ua (S1≃SuspBool))) (intLoop* x i) --- looperIntoBool (pos zero) = refl --- looperIntoBool (pos (suc n)) = (λ j → looperIntoBool (pos n) j ∙ merid north ∙ sym (merid south)) ∙ --- (λ j → (λ i → transportRefl (SuspBool→S1 (intLoop* (pos n) i)) j ) ∙ merid north ∙ sym (merid south)) ∙ --- (λ j → cong SuspBool→S1 (intLoop* (pos n)) ∙ congFunct SuspBool→S1 (merid false) (sym (merid true)) (~ j)) ∙ --- sym (congFunct SuspBool→S1 (intLoop* (pos n)) loop*) ∙ --- λ j i → transportRefl (SuspBool→S1 ((intLoop* (pos n) ∙ loop*) i)) (~ j) --- looperIntoBool (negsuc zero) = symDistr (merid north) (sym (merid south)) ∙ --- sym (congFunct SuspBool→S1 (merid true) (sym (merid false))) ∙ --- (λ j → cong SuspBool→S1 (merid true ∙ sym (merid false))) ∙ --- (λ j → cong SuspBool→S1 (symDistr (merid false) (sym (merid true)) (~ j))) ∙ --- λ j i → transportRefl (SuspBool→S1 (loop* (~ i))) (~ j) --- looperIntoBool (negsuc (suc n)) = (λ i → looperIntoBool (negsuc n) i ∙ sym ((merid north) ∙ (sym (merid south))) ) ∙ --- (λ i → looperIntoBool (negsuc n) i1 ∙ symDistr (merid north) (sym (merid south)) i) ∙ --- (λ j → (λ i → transportRefl (SuspBool→S1 (intLoop* (negsuc n) i)) j) ∙ merid south ∙ sym (merid north)) ∙ --- (λ j → cong SuspBool→S1 (intLoop* (negsuc n)) ∙ congFunct SuspBool→S1 (merid true) (sym (merid false)) (~ j)) ∙ --- ((λ j → cong SuspBool→S1 (intLoop* (negsuc n)) ∙ cong SuspBool→S1 (symDistr (merid false) (sym (merid true)) (~ j)))) ∙ --- sym (congFunct SuspBool→S1 (intLoop* (negsuc n)) (sym loop*)) ∙ --- λ j i → transportRefl (SuspBool→S1 ((intLoop* (negsuc n) ∙ sym loop*) i)) (~ j) - --- {- From this, we can finally deduce that looper is an equivalence-} --- isEquivLooper : isEquiv looper --- isEquivLooper = transport (λ i → isEquiv (funExt (looperIntoBool) (~ i))) isEquivTranspIntLoop --- where --- isEquivTranspIntLoop : isEquiv λ x → cong (transport ((ua (S1≃SuspBool)))) (intLoop* x) --- isEquivTranspIntLoop = compEquiv (intLoop* , intLoop*Equiv) --- (cong (transport (ua S1≃SuspBool)) , --- congEquiv (transport (ua S1≃SuspBool) , --- transportEquiv (ua S1≃SuspBool) .snd) .snd) .snd - --- isSetS1 : isSet (_≡_ {A = S₊ 1} north north) --- isSetS1 = transport (λ i → isSet (helper i)) isSetInt --- where --- helper : Int ≡ (_≡_ {A = S₊ 1} north north) --- helper = sym ΩS¹≡Int ∙ --- (λ i → typ (Ω (S¹≡SuspBool i , transport (λ j → S¹≡SuspBool (j ∧ i)) base))) ∙ --- (λ i → typ (Ω (ua S1≃SuspBool i , transport (λ j → ua S1≃SuspBool (i ∧ j)) north))) - --- isEquivHelper2 : isOfHLevel 3 A → isEquiv {B = ∥ A ∥ 1} ∣_∣ --- isEquivHelper2 ofHlevl = --- isoToIsEquiv (iso ∣_∣ --- (trElim (λ _ → ofHlevl) (λ a → a)) --- (trElim {B = λ b → ∣ trElim (λ _ → ofHlevl) (λ a₁ → a₁) b ∣ ≡ b} (λ b → isOfHLevelSuc 3 (isOfHLevel∥∥ 1) ∣ trElim (λ _ → ofHlevl) (λ a₁ → a₁) b ∣ b) λ a → refl) --- λ b → refl) - --- isEquivHelper : {a b : A} → isOfHLevel 3 A → isEquiv {B = _≡_ {A = ∥ A ∥ 1} ∣ a ∣ ∣ b ∣ } (cong ∣_∣) --- isEquivHelper {A = A} {a = a} {b = b} ofHlevl = congEquiv (∣_∣ , isEquivHelper2 ofHlevl) .snd - - --- S1≡S¹ : S₊ 1 ≡ S¹ --- S1≡S¹ = {!!} - - --- d-map : typ (Ω ((Susp S¹) , north)) → S¹ --- d-map p = subst HopfSuspS¹ p base - --- d-mapId : (r : S¹) → d-map (ϕ base r) ≡ r --- d-mapId r = (λ i → subst HopfSuspS¹ (□≡∙ (merid r) (sym (merid base)) (~ i)) base) ∙ --- (substComposite-□ HopfSuspS¹ (merid r) (sym (merid base)) base) ∙ --- test r --- where --- test : (r : S¹) → rot r base ≡ r --- test base = refl --- test (loop i) = refl - - --- testProp : (B : A → Type ℓ') → {!x !} --- testProp = {!!} - - --- S³≡SuspSuspS¹ : S³ ≡ Susp (Susp S¹) --- S³≡SuspSuspS¹ = S³≡SuspS² ∙ λ i → Susp (S²≡SuspS¹ i) - --- d-mapComp : fiber d-map base ≡ (_≡_ {A = Susp (Susp S¹)} north north) --- d-mapComp = sym (pathSigma≡sigmaPath {B = HopfSuspS¹} _ _) ∙ helper --- where --- helper : (_≡_ {A = Σ (Susp S¹) λ x → HopfSuspS¹ x} (north , base) (north , base)) ≡ (_≡_ {A = Susp (Susp S¹)} north north) --- helper = (λ i → (_≡_ {A = S³≡TotalHopf (~ i)} --- (transp (λ j → S³≡TotalHopf (~ i ∨ ~ j)) (~ i) (north , base)) --- ((transp (λ j → S³≡TotalHopf (~ i ∨ ~ j)) (~ i) (north , base))))) ∙ --- (λ i → _≡_ {A = S³} base base) ∙ --- (λ i → _≡_ {A = S³≡SuspSuspS¹ i} (transp (λ j → S³≡SuspSuspS¹ (i ∧ j)) (~ i) base) ((transp (λ j → S³≡SuspSuspS¹ (i ∧ j)) (~ i) base))) - --- n=1Equiv : isEquiv {A = S₊ 1} {B = typ (Ω ((S₊ 2) , north))} (ϕ north) --- n=1Equiv = isoToIsEquiv (iso (ϕ north) --- (λ x → {!(cong (transport (λ i → Susp (sym (S¹≡SuspBool) i))) (cong (transport (λ i → Susp ((sym (ua S1≃SuspBool)) i))) x))!}) --- {!Hopf!} --- {!HopfSuspS¹ !}) --- where --- helper : transport (λ i → Susp ((S¹≡SuspBool) (~ i))) (transport (λ i → Susp (ua S1≃SuspBool (~ i))) north) ≡ north --- helper = refl -- fantastic - --- is1Connected-dmap : is- 1 -Connected d-map --- is1Connected-dmap base = transport (λ j → isContr (∥ d-mapComp (~ j) ∥ ℕ→ℕ₋₂ 1)) {!!} --- is1Connected-dmap (loop i) = {!!} - --- isOfHLevelLemma : (n : ℕ₋₂) → isOfHLevel (2+ n) A → isOfHLevel (suc (2+ n)) (Susp A) --- isOfHLevelLemma neg2 hlvl = λ x y → {!!} --- isOfHLevelLemma neg1 hlvl = {!!} --- isOfHLevelLemma (ℕ→ℕ₋₂ n) hlvl = {!!} --- where --- helper : (n : ℕ) → isOfHLevel n A → (x y : Susp A) → isOfHLevel (suc n) (x ≡ y) --- helper zero hlvl north north = {!s!} --- helper zero hlvl north south = {!!} --- helper zero hlvl north (merid a i) = {!!} --- helper zero hlvl south north = {!!} --- helper zero hlvl south south = {!!} --- helper zero hlvl south (merid a i) = {!!} --- helper zero hlvl (merid a i) y = {!!} --- helper (suc n) hlvl x y = {!!} - --- -- Kn→ΩKn+1Sucn : (n : ℕ) → Kn→ΩKn+1 (suc n) ≡ λ x → truncEquivΩ (ℕ→ℕ₋₂ (suc n)) .fst (trElim (λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (suc n))) (λ a → ∣ ϕ north a ∣) x) -- --- -- Kn→ΩKn+1Sucn n = funExt (trElim (λ x → isOfHLevelSuc (suc (suc n)) --- -- ((isOfHLevel∥∥ (ℕ→ℕ₋₂ (suc (suc n))) ∣ north ∣ ∣ north ∣) --- -- (Kn→ΩKn+1 (suc n) x) --- -- (truncEquivΩ (ℕ→ℕ₋₂ (suc n)) .fst (trElim (λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (suc n))) (λ a → ∣ ϕ north a ∣) x)))) --- -- λ a → refl) - - --- -- isEquivKn→ΩKn+1 : (n : ℕ) → isEquiv (Kn→ΩKn+1 n) --- -- isEquivKn→ΩKn+1 zero = compEquiv (looper , isEquivLooper) (cong ∣_∣ , isEquivHelper hLevl3) .snd --- -- where --- -- hLevl3 : (x y : S₊ 1) (p q : x ≡ y) → isProp (p ≡ q) --- -- hLevl3 x y = J (λ y p → (q : x ≡ y) → isProp (p ≡ q) ) --- -- (transport (λ i → isSet (helper (~ i))) isSetInt refl) --- -- where --- -- helper : (x ≡ x) ≡ Int --- -- helper = (λ i → transp (λ j → ua S1≃SuspBool (~ j ∨ ~ i)) (~ i) x ≡ transp (λ j → ua S1≃SuspBool (~ j ∨ ~ i)) (~ i) x) ∙ --- -- (λ i → transp (λ j → S¹≡SuspBool (~ j ∨ ~ i)) (~ i) (transport (sym (ua S1≃SuspBool)) x) ≡ transp (λ j → S¹≡SuspBool (~ j ∨ ~ i)) (~ i) (transport (sym (ua S1≃SuspBool)) x)) ∙ --- -- basedΩS¹≡Int (transport (sym S¹≡SuspBool) (transport (sym (ua S1≃SuspBool)) x)) --- -- isEquivKn→ΩKn+1 (suc zero) = transport (λ i → isEquiv (Kn→ΩKn+1Sucn zero (~ i))) --- -- (compEquiv (trElim (λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (suc zero))) (λ a → ∣ ϕ north a ∣) , --- -- {!!}) --- -- (truncEquivΩ (ℕ→ℕ₋₂ (suc zero))) .snd) --- -- isEquivKn→ΩKn+1 (suc (suc n)) = {!!} + helper : (n : ℕ) → intLoop* (negsuc n) ∙ (sym loop*) ≡ altMap2 (negsuc (suc n)) + helper zero = (λ i → altMapneg1 (~ i) ∙ altMapneg1 (~ i)) ∙ sym (anotherHelper zero) + helper (suc n) = (λ i → (helper n i) ∙ altMapneg1 (~ i) ) ∙ + sym (anotherHelper (suc n)) + + +{- We have that the transported map is an equivalence -} +mapIsEquiv : isEquiv altMap2 +mapIsEquiv = compEquiv (intLoop , (isoToIsEquiv (iso intLoop winding (decodeEncode base) windingIntLoop))) ((λ x i → transport S¹≡SuspBool (x i)) , helper) .snd + where + helper : isEquiv {A = ΩS¹} {B = _≡_ {A = Susp Bool} north north} (λ x i → transport S¹≡SuspBool (x i)) + helper = congEquiv (transport S¹≡SuspBool , helper3) .snd + where + helper2 : transport S¹≡SuspBool ≡ S¹→SuspBool + helper2 = funExt λ z → transportRefl (S¹→SuspBool z) + helper3 : isEquiv (transport S¹≡SuspBool ) + helper3 = transport (λ i → isEquiv (helper2 (~ i))) (S¹≃SuspBool .snd) + +{- From this we get that intLoop* is an equivalence-} +intLoop*Equiv : isEquiv intLoop* +intLoop*Equiv = transport (λ i → isEquiv (funExt (altMap≡intLoop*2) (~ i))) mapIsEquiv + +{- We now only need to work with Susp Bool and S₊ 1. We transport intLoop* along a path S1≡SuspBool and show that this gives us looper. -} +SuspBool→S1 : Susp Bool → S₊ 1 +SuspBool→S1 north = north +SuspBool→S1 south = south +SuspBool→S1 (merid false i) = merid north i +SuspBool→S1 (merid true i) = merid south i + +S1→SuspBool : S₊ 1 → Susp Bool +S1→SuspBool north = north +S1→SuspBool south = south +S1→SuspBool (merid north i) = merid false i +S1→SuspBool (merid south i) = merid true i + +{- We need to spell out the trivial equivalence S1≃SuspBool. Of course it's important to make sure that we choose the right version of it. -} +S1≃SuspBool : Susp Bool ≃ S₊ 1 +S1≃SuspBool = isoToEquiv (iso SuspBool→S1 S1→SuspBool retrHelper sectHelper) + where + + + sectHelper : section S1→SuspBool SuspBool→S1 + sectHelper north = refl + sectHelper south = refl + sectHelper (merid false i) = refl + sectHelper (merid true i) = refl + + retrHelper : retract S1→SuspBool SuspBool→S1 + retrHelper north = refl + retrHelper south = refl + retrHelper (merid north i) = refl + retrHelper (merid south i) = refl + +{- We show that transporting intLoop* along (ua S1≃SuspBool) gives us looper (again, modulo function extensionality) -} +looperIntoBool : (x : Int) → looper x ≡ λ i → transport ((ua (S1≃SuspBool))) (intLoop* x i) +looperIntoBool (pos zero) = refl +looperIntoBool (pos (suc n)) = (λ j → looperIntoBool (pos n) j ∙ merid north ∙ sym (merid south)) ∙ + (λ j → (λ i → transportRefl (SuspBool→S1 (intLoop* (pos n) i)) j ) ∙ merid north ∙ sym (merid south)) ∙ + (λ j → cong SuspBool→S1 (intLoop* (pos n)) ∙ congFunct SuspBool→S1 (merid false) (sym (merid true)) (~ j)) ∙ + sym (congFunct SuspBool→S1 (intLoop* (pos n)) loop*) ∙ + λ j i → transportRefl (SuspBool→S1 ((intLoop* (pos n) ∙ loop*) i)) (~ j) +looperIntoBool (negsuc zero) = symDistr (merid north) (sym (merid south)) ∙ + sym (congFunct SuspBool→S1 (merid true) (sym (merid false))) ∙ + (λ j → cong SuspBool→S1 (merid true ∙ sym (merid false))) ∙ + (λ j → cong SuspBool→S1 (symDistr (merid false) (sym (merid true)) (~ j))) ∙ + λ j i → transportRefl (SuspBool→S1 (loop* (~ i))) (~ j) +looperIntoBool (negsuc (suc n)) = (λ i → looperIntoBool (negsuc n) i ∙ sym ((merid north) ∙ (sym (merid south))) ) ∙ + (λ i → looperIntoBool (negsuc n) i1 ∙ symDistr (merid north) (sym (merid south)) i) ∙ + (λ j → (λ i → transportRefl (SuspBool→S1 (intLoop* (negsuc n) i)) j) ∙ merid south ∙ sym (merid north)) ∙ + (λ j → cong SuspBool→S1 (intLoop* (negsuc n)) ∙ congFunct SuspBool→S1 (merid true) (sym (merid false)) (~ j)) ∙ + ((λ j → cong SuspBool→S1 (intLoop* (negsuc n)) ∙ cong SuspBool→S1 (symDistr (merid false) (sym (merid true)) (~ j)))) ∙ + sym (congFunct SuspBool→S1 (intLoop* (negsuc n)) (sym loop*)) ∙ + λ j i → transportRefl (SuspBool→S1 ((intLoop* (negsuc n) ∙ sym loop*) i)) (~ j) + +{- From this, we can finally deduce that looper is an equivalence-} +isEquivLooper : isEquiv looper +isEquivLooper = transport (λ i → isEquiv (funExt (looperIntoBool) (~ i))) isEquivTranspIntLoop + where + isEquivTranspIntLoop : isEquiv λ x → cong (transport ((ua (S1≃SuspBool)))) (intLoop* x) + isEquivTranspIntLoop = compEquiv (intLoop* , intLoop*Equiv) + (cong (transport (ua S1≃SuspBool)) , + congEquiv (transport (ua S1≃SuspBool) , + transportEquiv (ua S1≃SuspBool) .snd) .snd) .snd + +isSetS1 : isSet (_≡_ {A = S₊ 1} north north) +isSetS1 = transport (λ i → isSet (helper i)) isSetInt + where + helper : Int ≡ (_≡_ {A = S₊ 1} north north) + helper = sym ΩS¹≡Int ∙ + (λ i → typ (Ω (S¹≡SuspBool i , transport (λ j → S¹≡SuspBool (j ∧ i)) base))) ∙ + (λ i → typ (Ω (ua S1≃SuspBool i , transport (λ j → ua S1≃SuspBool (i ∧ j)) north))) + +isEquivHelper2 : isOfHLevel 3 A → isEquiv {B = ∥ A ∥ 1} ∣_∣ +isEquivHelper2 ofHlevl = + isoToIsEquiv (iso ∣_∣ + (trElim (λ _ → ofHlevl) (λ a → a)) + (trElim {B = λ b → ∣ trElim (λ _ → ofHlevl) (λ a₁ → a₁) b ∣ ≡ b} (λ b → isOfHLevelSuc 3 (isOfHLevel∥∥ 1) ∣ trElim (λ _ → ofHlevl) (λ a₁ → a₁) b ∣ b) λ a → refl) + λ b → refl) + +isEquivHelper : {a b : A} → isOfHLevel 3 A → isEquiv {B = _≡_ {A = ∥ A ∥ 1} ∣ a ∣ ∣ b ∣ } (cong ∣_∣) +isEquivHelper {A = A} {a = a} {b = b} ofHlevl = congEquiv (∣_∣ , isEquivHelper2 ofHlevl) .snd + + +d-map : typ (Ω ((Susp S¹) , north)) → S¹ +d-map p = subst HopfSuspS¹ p base + +d-mapId : (r : S¹) → d-map (ϕ base r) ≡ r +d-mapId r = (λ i → subst HopfSuspS¹ (□≡∙ (merid r) (sym (merid base)) (~ i)) base) ∙ + (substComposite-□ HopfSuspS¹ (merid r) (sym (merid base)) base) ∙ + test r + where + test : (r : S¹) → rot r base ≡ r + test base = refl + test (loop i) = refl + + +S³≡SuspSuspS¹ : S³ ≡ Susp (Susp S¹) +S³≡SuspSuspS¹ = S³≡SuspS² ∙ λ i → Susp (S²≡SuspS¹ i) + +d-mapComp : fiber d-map base ≡ (_≡_ {A = Susp (Susp S¹)} north north) +d-mapComp = sym (pathSigma≡sigmaPath {B = HopfSuspS¹} _ _) ∙ helper + where + helper : (_≡_ {A = Σ (Susp S¹) λ x → HopfSuspS¹ x} (north , base) (north , base)) ≡ (_≡_ {A = Susp (Susp S¹)} north north) + helper = (λ i → (_≡_ {A = S³≡TotalHopf (~ i)} + (transp (λ j → S³≡TotalHopf (~ i ∨ ~ j)) (~ i) (north , base)) + ((transp (λ j → S³≡TotalHopf (~ i ∨ ~ j)) (~ i) (north , base))))) ∙ + (λ i → _≡_ {A = S³} base base) ∙ + (λ i → _≡_ {A = S³≡SuspSuspS¹ i} (transp (λ j → S³≡SuspSuspS¹ (i ∧ j)) (~ i) base) ((transp (λ j → S³≡SuspSuspS¹ (i ∧ j)) (~ i) base))) + + +is1Connected-dmap : is- 1 -Connected d-map +is1Connected-dmap base = transport (λ j → isContr (∥ d-mapComp (~ j) ∥ ℕ→ℕ₋₂ 1)) + (transport (λ i → isContr (PathΩ {A = Susp (Susp S¹)} {a = north} (ℕ→ℕ₋₂ 1) i)) + (refl , isOfHLevelSuc 1 (isOfHLevelSuc 0 Pr242SpecCase) ∣ north ∣ ∣ north ∣ (λ _ → ∣ north ∣))) +is1Connected-dmap (loop j) = + hcomp (λ k → λ{(j = i0) → is1Connected-dmap base ; + (j = i1) → isPropIsOfHLevel 0 (transport (λ i → isContr (∥ fiber d-map (loop i) ∥ ℕ→ℕ₋₂ 1)) + (is1Connected-dmap base)) + (is1Connected-dmap base) + k}) + (transp (λ i → isContr (∥ fiber d-map (loop (i ∧ j)) ∥ ℕ→ℕ₋₂ 1)) (~ j) + (transport (λ j → isContr (∥ d-mapComp (~ j) ∥ ℕ→ℕ₋₂ 1)) + (transport (λ i → isContr (PathΩ {A = Susp (Susp S¹)} {a = north} (ℕ→ℕ₋₂ 1) i)) + (refl , isOfHLevelSuc 1 (isOfHLevelSuc 0 Pr242SpecCase) ∣ north ∣ ∣ north ∣ (λ _ → ∣ north ∣))))) + +d-equiv : isEquiv {A = ∥ typ (Ω (Susp S¹ , north)) ∥ (ℕ→ℕ₋₂ 1)} {B = ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (trElim (λ x → isOfHLevel∥∥ 1) λ x → ∣ d-map x ∣ ) +d-equiv = conEquiv (ℕ→ℕ₋₂ 1) d-map is1Connected-dmap .snd + +d-mapId2 : (λ (x : ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)) → (trElim {n = (ℕ→ℕ₋₂ 1)} {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (λ x → isOfHLevel∥∥ 1) λ x → ∣ d-map x ∣) + (trElim (λ _ → isOfHLevel∥∥ 1) (λ a → ∣ ϕ base a ∣) x)) ≡ λ x → x +d-mapId2 = funExt (trElim (λ x → isOfHLevelSuc 2 (isOfHLevel∥∥ 1 ((trElim {n = (ℕ→ℕ₋₂ 1)} {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (λ x → isOfHLevel∥∥ 1) λ x → ∣ d-map x ∣) (trElim (λ _ → isOfHLevel∥∥ 1) (λ a → ∣ ϕ base a ∣) x)) x)) λ a i → ∣ d-mapId a i ∣) + +isEquiv∣ϕ∣ : isEquiv {A = ∥ S¹ ∥ ℕ→ℕ₋₂ 1} (trElim (λ _ → isOfHLevel∥∥ 1) (λ a → ∣ ϕ base a ∣)) +isEquiv∣ϕ∣ = compToIdEquiv (trElim {n = (ℕ→ℕ₋₂ 1)} {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (λ x → isOfHLevel∥∥ 1) λ x → ∣ d-map x ∣) + (trElim (λ _ → isOfHLevel∥∥ 1) (λ a → ∣ ϕ base a ∣)) + d-mapId2 + d-equiv + +--- + +S¹≡S1 : S₊ 1 ≡ S¹ +S¹≡S1 = (λ i → Susp (ua (Bool≃Susp⊥) (~ i))) ∙ sym S¹≡SuspBool + +S¹→S1 : S¹ → S₊ 1 +S¹→S1 x = SuspBool→S1 (S¹→SuspBool x) + +S1→S¹ : S₊ 1 → S¹ +S1→S¹ x = SuspBool→S¹ (S1→SuspBool x) + +S2→S² : S₊ 2 → S² +S2→S² north = base +S2→S² south = base +S2→S² (merid a i) = SuspS¹→S² (merid (SuspBool→S¹ (S1→SuspBool a)) i) + +S²→S2 : S² → S₊ 2 +S²→S2 base = north +S²→S2 (surf i i₁) = {!!} + +S2→SuspS¹ : S₊ 2 → Susp S¹ +S2→SuspS¹ x = S²→SuspS¹ (S2→S² x) + +S²≡S2 : S₊ 2 ≡ Susp S¹ +S²≡S2 i = Susp (S¹≡S1 i) + +funTEST2 : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ} {C : (A : Type ℓ) (a1 : A) → Type ℓ'} (p : A ≡ B) (a : A) (b : B) → + (transport p a ≡ b) → + (f : (A : Type ℓ) → + (a1 : A) (a2 : ∥ A ∥ 1) → C A a1) → + isEquiv (f A a) → + isEquiv (f B b) +funTEST2 {ℓ = ℓ}{A = A} {B = B} {C = C} = + J (λ B p → (a : A) (b : B) → + (transport p a ≡ b) → + (f : (A : Type ℓ) → + (a1 : A) (a2 : ∥ A ∥ 1) → C A a1) → + isEquiv (f A a) → + isEquiv (f B b)) + λ a b trefl f is → transport (λ i → isEquiv (f A ((sym (transportRefl a) ∙ trefl) i))) is + + +final : isEquiv {A = ∥ S₊ 1 ∥ 1} {B = ∥ typ (Ω (S₊ 2 , north)) ∥ 1} (trElim (λ _ → isOfHLevel∥∥ 1) (λ a → ∣ ϕ north a ∣)) +final = funTEST2 {A = S¹} (λ i → S¹≡S1 (~ i)) base north refl (λ A a1 → trElim (λ _ → isOfHLevel∥∥ 1) (λ a → ∣ ϕ a1 a ∣)) isEquiv∣ϕ∣ + +Kn→ΩKn+1Sucn : (n : ℕ) → Kn→ΩKn+1 (suc n) ≡ λ x → truncEquivΩ (ℕ→ℕ₋₂ (suc n)) .fst (trElim (λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (suc n))) (λ a → ∣ ϕ north a ∣) x) +Kn→ΩKn+1Sucn n = funExt (trElim (λ x → isOfHLevelSuc (suc (suc n)) + ((isOfHLevel∥∥ (ℕ→ℕ₋₂ (suc (suc n))) ∣ north ∣ ∣ north ∣) + (Kn→ΩKn+1 (suc n) x) + (truncEquivΩ (ℕ→ℕ₋₂ (suc n)) .fst (trElim (λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (suc n))) (λ a → ∣ ϕ north a ∣) x)))) + λ a → refl) + + + + +isEquivKn→ΩKn+1 : (n : ℕ) → isEquiv (Kn→ΩKn+1 n) +isEquivKn→ΩKn+1 zero = compEquiv (looper , isEquivLooper) (cong ∣_∣ , isEquivHelper hLevl3) .snd + where + hLevl3 : (x y : S₊ 1) (p q : x ≡ y) → isProp (p ≡ q) + hLevl3 x y = J (λ y p → (q : x ≡ y) → isProp (p ≡ q) ) + (transport (λ i → isSet (helper (~ i))) isSetInt refl) + where + helper : (x ≡ x) ≡ Int + helper = (λ i → transp (λ j → ua S1≃SuspBool (~ j ∨ ~ i)) (~ i) x ≡ transp (λ j → ua S1≃SuspBool (~ j ∨ ~ i)) (~ i) x) ∙ + (λ i → transp (λ j → S¹≡SuspBool (~ j ∨ ~ i)) (~ i) (transport (sym (ua S1≃SuspBool)) x) ≡ transp (λ j → S¹≡SuspBool (~ j ∨ ~ i)) (~ i) (transport (sym (ua S1≃SuspBool)) x)) ∙ + basedΩS¹≡Int (transport (sym S¹≡SuspBool) (transport (sym (ua S1≃SuspBool)) x)) +isEquivKn→ΩKn+1 (suc zero) = transport (λ i → isEquiv (Kn→ΩKn+1Sucn zero (~ i))) + (compEquiv (trElim (λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (suc zero))) (λ a → ∣ ϕ north a ∣) , + final) + (truncEquivΩ (ℕ→ℕ₋₂ (suc zero))) .snd) +isEquivKn→ΩKn+1 (suc (suc n)) = transport (λ i → isEquiv (Kn→ΩKn+1Sucn (suc n) (~ i))) + (compEquiv (conEquiv3 (4 + n) _ (ϕ north) (n , λ i → suc (suc (suc (+-suc n n (~ i))))) (FthalFun-2nConnected (suc n) _ (Pr242 _))) + (truncEquivΩ (ℕ→ℕ₋₂ (suc (suc n)))) .snd) + +Kn≃ΩKn+1 : (n : ℕ) → coHomK n ≃ typ (Ω (coHomK-ptd (suc n))) +Kn≃ΩKn+1 n = Kn→ΩKn+1 n , isEquivKn→ΩKn+1 n + +ΩKn+1→Kn : (n : ℕ) → typ (Ω (coHomK-ptd (suc n))) → coHomK n +ΩKn+1→Kn n a = equiv-proof (isEquivKn→ΩKn+1 n) a .fst .fst From e571a825cfe08228b95bf76b6667318d405caae3 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Wed, 8 Apr 2020 17:33:24 +0200 Subject: [PATCH 15/89] updated Trunc --- .../Connected/FreudenthalProof/7-5-14.agda | 22 +++--- .../Connected/FreudenthalProof/7-5-7.agda | 9 ++- .../Connected/FreudenthalProof/8-6-4.agda | 2 +- .../Connected/FreudenthalProof/8-6-4Code.agda | 40 +++++----- .../Connected/FreudenthalProof/Prelim.agda | 2 +- .../HITs/Truncation/Connected/Properties.agda | 2 +- Cubical/HITs/Truncation/Properties.agda | 18 ++--- Cubical/ZCohomology/S1Loop.agda | 74 +++++++------------ 8 files changed, 74 insertions(+), 95 deletions(-) diff --git a/Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-14.agda b/Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-14.agda index c344d5286..203ac5178 100644 --- a/Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-14.agda +++ b/Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-14.agda @@ -7,6 +7,7 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism open import Cubical.Data.NatMinusTwo.Base +open import Cubical.Data.Nat hiding (elim) open import Cubical.HITs.Nullification hiding (elim) open import Cubical.HITs.Truncation.Base open import Cubical.HITs.Truncation.Properties @@ -17,26 +18,27 @@ private A : Type ℓ B : Type ℓ' + Lemma7-5-14 : (n : ℕ₋₂) (f : A → B) → (is- n -Connected f) → ∥ A ∥ n ≃ ∥ B ∥ n -Lemma7-5-14 {A = A} {B = B} neg2 f c = isoToEquiv (iso (λ _ → fst (isOfHLevel∥∥ neg2)) - (λ _ → fst (isOfHLevel∥∥ neg2)) - (λ b → snd (isOfHLevel∥∥ neg2) b) - (λ b → snd (isOfHLevel∥∥ neg2) b)) +Lemma7-5-14 {A = A} {B = B} neg2 f c = isoToEquiv (iso (λ _ → fst (isOfHLevelTrunc 0)) + (λ _ → fst (isOfHLevelTrunc 0)) + (λ b → snd (isOfHLevelTrunc 0) b) + (λ b → snd (isOfHLevelTrunc 0) b)) Lemma7-5-14 {A = A} {B = B} (-1+ n) f c = isoToEquiv (iso (∥ f ∥-fun (-1+ n)) - (elim (λ _ → isOfHLevel∥∥ (-1+ n)) back) - (elim (λ x → (isOfHLevelSuc (2+ (-1+ n)) (isOfHLevel∥∥ (-1+ n))) _ _) backSection) - (elim (λ x → (isOfHLevelSuc (2+ (-1+ n)) (isOfHLevel∥∥ (-1+ n))) _ _) backRetract)) + (elim (λ _ → isOfHLevelTrunc (suc n)) back) + (elim (λ x → (isOfHLevelSuc (suc n) (isOfHLevelTrunc (suc n))) _ _) backSection) + (elim (λ x → (isOfHLevelSuc (suc n) (isOfHLevelTrunc (suc n))) _ _) backRetract)) where back : B → ∥ A ∥ (-1+ n) back y = ∥ fst ∥-fun (-1+ n) (fst (c y)) backSection : (b : B) → _≡_ {A = ∥ B ∥ (-1+ n)} - (elim (λ _ → isOfHLevel∥∥ (-1+ n)) (λ a → ∣ f a ∣) (elim {n = (-1+ n)} - {B = λ _ → ∥ A ∥ (-1+ n)} (λ _ → isOfHLevel∥∥ (-1+ n)) back ∣ b ∣)) + (elim (λ _ → isOfHLevelTrunc (suc n)) (λ a → ∣ f a ∣) (elim {n = suc n } + {B = λ _ → ∥ A ∥ (-1+ n)} (λ _ → isOfHLevelTrunc (suc n)) back ∣ b ∣)) ∣ b ∣ backSection b = helper {P = λ p → ∥ f ∥-fun (-1+ n) p ≡ ∣ b ∣} - (λ x → (isOfHLevelSuc (2+ (-1+ n)) (isOfHLevel∥∥ (-1+ n))) + (λ x → (isOfHLevelSuc (suc n) (isOfHLevelTrunc (suc n))) (∥ f ∥-fun (-1+ n) (∥ fst ∥-fun (-1+ n) x)) ∣ b ∣) (λ a p → cong (λ x → ∣ x ∣) p) (fst (c b)) diff --git a/Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-7.agda b/Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-7.agda index 57e95f634..8f621cc5d 100644 --- a/Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-7.agda +++ b/Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-7.agda @@ -6,6 +6,7 @@ open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv.Properties open import Cubical.Foundations.Equiv open import Cubical.Foundations.Prelude +open import Cubical.Data.Nat hiding (elim) open import Cubical.Foundations.Isomorphism open import Cubical.Data.NatMinusTwo.Base open import Cubical.HITs.Nullification hiding (elim) @@ -79,11 +80,11 @@ conInd-iii→i {A = A} {B = B} f n P→hasSection = λ b → (c n P→hasSection P n s b = ∥ fiber f b ∥ n c : (n : ℕ₋₂) → (∀ {ℓ} (P : B → HLevel ℓ (2+ n)) → hasSection (indConFun f P)) → (b : B) → ∥ fiber f b ∥ n - c n s = (s λ b → ( ∥ fiber f b ∥ n , isOfHLevel∥∥ _)) .fst λ a → ∣ a , refl ∣ + c n s = (s λ b → ( ∥ fiber f b ∥ n , isOfHLevelTrunc _)) .fst λ a → ∣ a , refl ∣ fun : (n : ℕ₋₂) → (s : (∀ {ℓ} (P : B → HLevel ℓ (2+ n)) → hasSection (indConFun f P))) → (b : B) (w : (∥ fiber f b ∥ n) ) → w ≡ c n s b - fun neg2 s b w = isOfHLevelSuc (2+ neg2) (isOfHLevel∥∥ neg2) w (c neg2 s b) - fun (-1+ n) s b = elim (λ x → (isOfHLevelSuc (2+ (-1+ n)) (isOfHLevel∥∥ {A = (fiber f b)} (-1+ n))) x (c (-1+ n) s b) ) (λ a → witness b (fst a) (snd a)) + fun neg2 s b w = isOfHLevelSuc (2+ neg2) (isOfHLevelTrunc _) w (c neg2 s b) + fun (-1+ n) s b = elim (λ x → (isOfHLevelSuc (2+ (-1+ n)) (isOfHLevelTrunc {A = (fiber f b)} (suc n))) x (c (-1+ n) s b) ) (λ a → witness b (fst a) (snd a)) where eqtyp : ((a : A) → ∣ (a , refl {x = f a}) ∣ ≡ c (-1+ n) s (f a)) ≡ ((b : B) (a : A) (p : f (a) ≡ b) → ∣ (a , p) ∣ ≡ c (-1+ n) s b) eqtyp = isoToPath (iso @@ -96,7 +97,7 @@ conInd-iii→i {A = A} {B = B} f n P→hasSection = λ b → (c n P→hasSection λ h → funExt λ a → JRefl (λ b₁ p → ∣ a , p ∣ ≡ c (-1+ n) s b₁) (h a)) c* : ((a : A) → ∣ (a , refl {x = f a}) ∣ ≡ c (-1+ n) s (f a)) - c* = λ a → sym (cong (λ x → x a) ((s λ b → ( ∥ fiber f b ∥ (-1+ n) , isOfHLevel∥∥ _)) .snd λ a → ∣ a , refl ∣)) + c* = λ a → sym (cong (λ x → x a) ((s λ b → ( ∥ fiber f b ∥ (-1+ n) , isOfHLevelTrunc _)) .snd λ a → ∣ a , refl ∣)) witness : ((b : B) (a : A) (p : f (a) ≡ b) → ∣ (a , p) ∣ ≡ c (-1+ n) s b) witness = transport (λ i → eqtyp i) c* diff --git a/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4.agda b/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4.agda index 0d81d1c9b..ae79a7bbc 100644 --- a/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4.agda +++ b/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4.agda @@ -1337,7 +1337,7 @@ Thm8-6-4 {ℓ} {A} n a iscon {y = y} = J {ℓ} {Susp A} {north} (λ y p → (isC where mainId : (p : north ≡ north) (d : CODE a n iscon north p) → d ≡ c n a iscon north p - mainId p = elim (λ x → isOfHLevelSuc (suc (n + n)) (isOfHLevel∥∥ {A = fiber (λ y → σ a y) p} (ℕ→ℕ₋₂ (n + n)) x (c n a iscon north p))) + mainId p = elim (λ x → isOfHLevelSuc (suc (n + n)) (isOfHLevelTrunc {A = fiber (λ y → σ a y) p} (2 + (n + n)) x (c n a iscon north p))) base where base : (x : fiber (λ y → σ a y) p) → ∣ x ∣ ≡ c n a iscon north p diff --git a/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4Code.agda b/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4Code.agda index 33e42704b..bf994ab59 100644 --- a/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4Code.agda +++ b/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4Code.agda @@ -36,28 +36,28 @@ private {- We first need to show that (a variant of) the canceller function defined in FreudenthalProof.Prelim is an equivalence -} abstract - isEquivCancel : {a : A} (n : ℕ) (q : north ≡ south) → + isEquivCancel : ∀ {ℓ} {A : Type ℓ}{a : A} (n : ℕ) (q : north ≡ south) → isEquiv {A = ∥ fiber (λ y → σ a y) (q ∙ sym (merid a)) ∥ ℕ→ℕ₋₂ (n + n)} {B = ∥ fiber merid q ∥ ℕ→ℕ₋₂ (n + n)} - (elim (λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (n + n))) + (elim (λ _ → isOfHLevelTrunc (2 + (n + n) )) λ b → ∣ (fst b) , canceller (sym (merid a)) (merid (fst b)) q (snd b) ∣) isEquivCancel {a = a} n q = isoToEquiv (iso - ((elim (λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (n + n))) λ b → ∣ (fst b) , canceller (sym (merid a)) (merid (fst b)) q (snd b) ∣)) - (elim (λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (n + n))) λ s → ∣ (fst s) , cancellerInv (λ i → (merid a) (~ i)) (merid (fst s)) q (snd s) ∣) - (λ b → elim {B = λ b → ((elim (λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (n + n))) + ((elim (λ _ → isOfHLevelTrunc (2 + (n + n))) λ b → ∣ (fst b) , canceller (sym (merid a)) (merid (fst b)) q (snd b) ∣)) + (elim (λ _ → isOfHLevelTrunc (2 + (n + n))) λ s → ∣ (fst s) , cancellerInv (λ i → (merid a) (~ i)) (merid (fst s)) q (snd s) ∣) + (λ b → elim {B = λ b → ((elim (λ _ → isOfHLevelTrunc (2 + (n + n))) λ b → ∣ (fst b) , canceller (sym (merid a)) (merid (fst b)) q (snd b) ∣)) - ((elim (λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (n + n))) + ((elim (λ _ → isOfHLevelTrunc (2 + (n + n))) λ s → ∣ (fst s) , cancellerInv (sym (merid a)) (merid (fst s)) q (snd s) ∣) b) ≡ b} - (λ x → isOfHLevelSuc (suc (n + n)) (isOfHLevel∥∥ (ℕ→ℕ₋₂ (n + n)) _ x)) + (λ x → isOfHLevelSuc _ (isOfHLevelTrunc (2 + (n + n)) _ x)) (λ b i → ∣ fst b , cancellerSection (sym (merid a)) (merid (fst b)) q (snd b) i ∣) b) - (λ b → elim {B = λ b → ((elim (λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (n + n))) + (λ b → elim {B = λ b → ((elim (λ _ → isOfHLevelTrunc (2 + (n + n))) λ b → ∣ (fst b) , cancellerInv (sym (merid a)) (merid (fst b)) q (snd b) ∣)) - ((elim (λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (n + n))) + ((elim (λ _ → isOfHLevelTrunc (2 + (n + n))) λ s → ∣ (fst s) , canceller (sym (merid a)) (merid (fst s)) q (snd s) ∣) b) ≡ b} - (λ x → isOfHLevelSuc (suc (n + n)) (isOfHLevel∥∥ (ℕ→ℕ₋₂ (n + n)) _ x)) + (λ x → isOfHLevelSuc (suc (n + n)) (isOfHLevelTrunc (2 + (n + n)) _ x)) (λ b i → ∣ fst b , cancellerRetract (sym (merid a)) (merid (fst b)) q (snd b) i ∣) b)) .snd {- CODE will be defined by means of univalence applied to an equivalence @@ -70,7 +70,7 @@ sufMap : (n : ℕ) (c a x₂ : A) ≡ (q ∙ sym (merid c)) → ∥ Σ A (λ x → merid x ≡ q) ∥ (ℕ→ℕ₋₂ (n + n)) sufMap {A = A} n c a x₂ iscon q = Lemma8-6-2 (A , a) (A , a) n n iscon iscon (λ x₂ c → ((merid x₂ ∙ sym (merid a)) ≡ (q ∙ sym (merid c)) → ∥ Σ A (λ x → merid x ≡ q) ∥ (ℕ→ℕ₋₂ (n + n))) , - isOfHLevelPi (2+ ((ℕ→ℕ₋₂ (n + n)))) λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (n + n))) + isOfHLevelPi (2+ ((ℕ→ℕ₋₂ (n + n)))) λ _ → isOfHLevelTrunc (2 + (n + n))) (λ x r → ∣ x , canceller (sym (merid a)) (merid x) q r ∣) (λ x r → ∣ x , switcher (merid a) q (merid x) r ∣) (funExt (λ x i → ∣ (refl i) , ((switcherCancellerIdGeneral (merid a) q @@ -81,7 +81,7 @@ sufMap {A = A} n c a x₂ iscon q = Lemma8-6-2 (A , a) (A , a) n n iscon iscon RlFun : (a c : A) (n : ℕ) → (iscon : is- (ℕ→ℕ₋₂ n) -ConnectedType A) → (q : north ≡ south) → (∥ fiber (λ y → σ a y) (q ∙ sym (merid c)) ∥ (ℕ→ℕ₋₂ (n + n))) → ∥ fiber merid q ∥ (ℕ→ℕ₋₂ (n + n)) -RlFun a c n iscon q = elim (λ x → isOfHLevel∥∥ ((ℕ→ℕ₋₂ (n + n)))) +RlFun a c n iscon q = elim (λ x → isOfHLevelTrunc ((2 + (n + n)))) λ r → sufMap n c a (fst r) iscon q (snd r) ------------- @@ -96,7 +96,7 @@ RlFunId {A = A} a n iscon q b = cong (λ x → x (snd b)) (proj₁ (((Lemma8-6-2 (A , a) (A , a) n n iscon iscon (λ x₂ c → (((merid x₂ ∙ sym (merid a)) ≡ (q ∙ sym (merid c)) → ∥ Σ A (λ x → merid x ≡ q) ∥ (ℕ→ℕ₋₂ (n + n)) ) , - isOfHLevelPi (2+ ((ℕ→ℕ₋₂ (n + n)))) λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (n + n)))) + isOfHLevelPi (2+ ((ℕ→ℕ₋₂ (n + n)))) λ _ → isOfHLevelTrunc (2 + (n + n)))) (λ x r → ∣ x , canceller (sym (merid a)) (merid x) q r ∣) (λ b s → ∣ b , switcher (merid a) q (merid b) s ∣) (funExt (λ x → λ j → ∣ (refl j) , (switcherCancellerIdGeneral (merid a) q @@ -111,17 +111,17 @@ baseCase : (a : A) (n : ℕ) (q : north ≡ south) → isEquiv (RlFun a a n iscon q) baseCase {A = A} a n iscon q = transport (λ i → isEquiv (helper (~ i))) - (isEquivCancel n q) + (isEquivCancel n q ) where helper : RlFun a a n iscon q - ≡ elim (λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (n + n))) + ≡ elim (λ _ → isOfHLevelTrunc (2 + (n + n))) λ b → ∣ (fst b) , canceller (sym (merid a)) (merid (fst b)) q (snd b) ∣ helper = funExt (elim {B = λ y → RlFun a a n iscon q y - ≡ (elim (λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (n + n))) + ≡ (elim (λ _ → isOfHLevelTrunc (2 + (n + n))) λ b → ∣ (fst b) , canceller (sym (merid a)) (merid (fst b)) q (snd b) ∣) y } - ((λ z y p → (isOfHLevelSuc (suc (n + n)) ((isOfHLevel∥∥ {A = (fiber merid q)} (ℕ→ℕ₋₂ (n + n))) + ((λ z y p → (isOfHLevelSuc (suc (n + n)) ((isOfHLevelTrunc {A = (fiber merid q)} (2 + (n + n))) (RlFun a a n iscon q z) - ((elim (λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (n + n))) + ((elim (λ _ → isOfHLevelTrunc (2 + (n + n))) λ b → ∣ (fst b) , canceller (sym (merid a)) (merid (fst b)) q (snd b) ∣) z) )) y p)) (RlFunId a n iscon q)) @@ -211,7 +211,7 @@ sufMapId {A = A} n a x1 iscon = (proj₂ (Lemma8-6-2 (A , a) (A , a) n n iscon i (λ x₂ c → ((merid x₂ ∙ sym (merid a)) ≡ ((merid x1) ∙ sym (merid c)) → ∥ Σ A (λ x → merid x ≡ (merid x1)) ∥ (ℕ→ℕ₋₂ (n + n))) , - isOfHLevelPi (2+ ((ℕ→ℕ₋₂ (n + n)))) λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (n + n))) + isOfHLevelPi (2+ ((ℕ→ℕ₋₂ (n + n)))) λ _ → isOfHLevelTrunc (2 + (n + n))) (λ x r → ∣ x , canceller (sym (merid a)) (merid x) (merid x1) r ∣) (λ x r → ∣ x , switcher (merid a) (merid x1) (merid x) r ∣) (funExt (λ x i → ∣ (refl i) , ((switcherCancellerIdGeneral (merid a) (merid x1) @@ -224,7 +224,7 @@ sufMapId2 : (n : ℕ) (a x1 : A) sufMapId2 {A = A} n a x1 iscon i = (proj₁ (Lemma8-6-2 (A , a) (A , a) n n iscon iscon (λ x₂ c → ((merid x₂ ∙ sym (merid a)) ≡ ((merid x1) ∙ sym (merid c)) → ∥ Σ A (λ x → merid x ≡ (merid x1)) ∥ (ℕ→ℕ₋₂ (n + n))) , - isOfHLevelPi (2+ ((ℕ→ℕ₋₂ (n + n)))) λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (n + n))) + isOfHLevelPi (2+ ((ℕ→ℕ₋₂ (n + n)))) λ _ → isOfHLevelTrunc (2 + (n + n))) (λ x r → ∣ x , canceller (sym (merid a)) (merid x) (merid x1) r ∣) (λ x r → ∣ x , switcher (merid a) (merid x1) (merid x) r ∣) (funExt (λ x i → ∣ (refl i) , ((switcherCancellerIdGeneral (merid a) (merid x1) diff --git a/Cubical/HITs/Truncation/Connected/FreudenthalProof/Prelim.agda b/Cubical/HITs/Truncation/Connected/FreudenthalProof/Prelim.agda index 983b345db..1715b8669 100644 --- a/Cubical/HITs/Truncation/Connected/FreudenthalProof/Prelim.agda +++ b/Cubical/HITs/Truncation/Connected/FreudenthalProof/Prelim.agda @@ -328,7 +328,7 @@ pairLemmaReflRefl = pairLemmaRefl refl ∙ λ i → pairLemma2Refl i ∙ lUnit r substComposite-∙ : ∀ {ℓ ℓ′} {A : Type ℓ} → (B : A → Type ℓ′) → {x y z : A} (p : x ≡ y) (q : y ≡ z) (u : B x) → subst B (p ∙ q) u ≡ subst B q (subst B p u) -substComposite-∙ {A = A} B p q u = transport (λ i → subst B (□≡∙ p q i) u ≡ subst B q (subst B p u)) (substComposite-□ B p q u) +substComposite-∙ = substComposite inv-rUnit : ∀ {ℓ} {A : Type ℓ} (x : A) → (λ i → rUnit (rUnit (λ _ → x) (~ i)) i ) ≡ refl inv-rUnit x = transport (λ i → PathP (λ j → (lCancel (λ k → (λ _ → x) ∙ (λ _ → x) ≡ rUnit (λ _ → x) k) i) j) diff --git a/Cubical/HITs/Truncation/Connected/Properties.agda b/Cubical/HITs/Truncation/Connected/Properties.agda index 40d366c15..95c1be4f8 100644 --- a/Cubical/HITs/Truncation/Connected/Properties.agda +++ b/Cubical/HITs/Truncation/Connected/Properties.agda @@ -34,7 +34,7 @@ private isConnectedSubtr : (n m : ℕ) (f : A → B) → is- (-2+ (n + m)) -Connected f → is- (-2+ n) -Connected f isConnectedSubtr n m f iscon b = transport (λ i → isContr (ua (truncOfTruncEq {A = fiber f b} n m) (~ i) )) (∣ iscon b .fst ∣ , - trElim (λ x → isOfHLevelPath n (isOfHLevel∥∥ (-2+ n)) _ _) + trElim (λ x → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → cong ∣_∣ (iscon b .snd a)) -- diff --git a/Cubical/HITs/Truncation/Properties.agda b/Cubical/HITs/Truncation/Properties.agda index d2f0280e8..e760b4764 100644 --- a/Cubical/HITs/Truncation/Properties.agda +++ b/Cubical/HITs/Truncation/Properties.agda @@ -164,8 +164,8 @@ elim3 {n = n} hB g = (λ a b → elim (λ _ → hB _ _ _) (λ c → g a b c)) --- ∥_∥-fun : ∀ {ℓ'} {B : Type ℓ'} (f : A → B) (n : ℕ₋₂) → ∥ A ∥ n → ∥ B ∥ n --- ∥ f ∥-fun n = elim (λ _ → isOfHLevel∥∥ n) λ a → ∣ f a ∣ +∥_∥-fun : ∀ {ℓ'} {B : Type ℓ'} (f : A → B) (n : ℕ₋₂) → ∥ A ∥ n → ∥ B ∥ n +∥ f ∥-fun n = elim (λ _ → isOfHLevelTrunc (2+ n)) λ a → ∣ f a ∣ HLevelTruncModality : ∀ {ℓ} (n : ℕ) → Modality ℓ isModal (HLevelTruncModality n) = isOfHLevel n @@ -195,7 +195,7 @@ module univTrunc where retr : ∀ {ℓ} (n : ℕ₋₂) {B : HLevel ℓ (2+ n)} → retract {A = (∥ A ∥ n → (fst B))} {B = (A → (fst B))} (fun n {B}) (inv n {B}) retr neg2 {B , lev} b = funExt λ x → sym ((snd lev) (elim (λ _ → lev) (λ a → b ∣ a ∣) x)) ∙ (snd lev) (b x) - retr (-1+ n) {B , lev} b = funExt (elim (λ x → (isOfHLevelSuc (2+ (-1+ n) ) lev) ((elim (λ _ → lev) (λ a → b ∣ a ∣) x)) (b x)) λ a → refl) + retr (-2+ suc n) {B , lev} b = funExt (elim (λ x → (isOfHLevelSuc (2+ (-2+ (suc n)) ) lev) ((elim (λ _ → lev) (λ a → b ∣ a ∣) x)) (b x)) λ a → refl) univTrunc : ∀ {ℓ} (n : ℕ₋₂) {B : HLevel ℓ (2+ n)} → (∥ A ∥ n → (fst B)) ≃ (A → (fst B)) univTrunc n {B} = isoToEquiv (iso (fun n {B}) (inv n {B}) (sect n {B}) (retr n {B})) @@ -372,19 +372,19 @@ truncOfTruncEq : (n m : ℕ) → ∥ A ∥ (-2+ n) ≃ ∥ ∥ A ∥ (-2+ (n + m truncOfTruncEq {A = A} n m = isoToEquiv (iso fun funInv sect retr) where fun : ∥ A ∥ (-2+ n) → ∥ ∥ A ∥ (-2+ (n + m)) ∥ (-2+ n) - fun = elim (λ _ → isOfHLevel∥∥ (-2+ n)) λ a → ∣ ∣ a ∣ ∣ + fun = elim (λ _ → isOfHLevelTrunc n) λ a → ∣ ∣ a ∣ ∣ funInv : ∥ ∥ A ∥ (-2+ (n + m)) ∥ (-2+ n) → ∥ A ∥ (-2+ n) - funInv = elim (λ _ → isOfHLevel∥∥ (-2+ n)) + funInv = elim (λ _ → isOfHLevelTrunc n) (elim (λ _ → transport (λ i → isOfHLevel (+-comm n m (~ i)) (∥ A ∥ (-2+ n))) - (isOfHLevelPlus m (isOfHLevel∥∥ (-2+ n)))) + (isOfHLevelPlus m (isOfHLevelTrunc n ))) λ a → ∣ a ∣) sect : section fun funInv - sect = elim (λ x → isOfHLevelPath n (isOfHLevel∥∥ (-2+ n)) _ _ ) + sect = elim (λ x → isOfHLevelPath n (isOfHLevelTrunc n) _ _ ) (elim (λ x → isOfHLevelPath (n + m) (transport (λ i → isOfHLevel (+-comm n m (~ i)) (∥ ∥ A ∥ (-2+ (n + m)) ∥ (-2+ n))) - (isOfHLevelPlus m (isOfHLevel∥∥ (-2+ n)))) _ _ ) + (isOfHLevelPlus m (isOfHLevelTrunc n))) _ _ ) λ a → refl) retr : retract fun funInv - retr = elim (λ x → isOfHLevelPath n (isOfHLevel∥∥ (-2+ n)) _ _) λ a → refl + retr = elim (λ x → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → refl diff --git a/Cubical/ZCohomology/S1Loop.agda b/Cubical/ZCohomology/S1Loop.agda index 9c5e0f845..2b9921db9 100644 --- a/Cubical/ZCohomology/S1Loop.agda +++ b/Cubical/ZCohomology/S1Loop.agda @@ -99,7 +99,7 @@ Pr2310 {A = A} {B = B} {C = C} n f g iscon = elim.3→1 inr n λ P → (λ k → Pr242 : (n : ℕ) → is- (-1+ n) -ConnectedType (S₊ n) -Pr242 zero = ∣ north ∣ , (isOfHLevel∥∥ neg1 ∣ north ∣) +Pr242 zero = ∣ north ∣ , (isOfHLevelTrunc 1 ∣ north ∣) Pr242 (suc n) = transport (λ i → is- ℕ→ℕ₋₂ n -ConnectedType (Susp≡Push {A = S₊ n} (~ i))) (trivFunCon← {A = Pushout {A = S₊ n} (λ x → tt) λ x → tt} {a = inr tt } _ (transport (λ i → is- (-1+ n) -Connected (mapsAgree (~ i))) @@ -148,7 +148,7 @@ looper (negsuc (suc n)) = looper (negsuc n) ∙ sym (ϕ south north) {- The map of Kn≃ΩKn+1 is given as follows. -} Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) Kn→ΩKn+1 zero x = cong ∣_∣ (looper x) -Kn→ΩKn+1 (suc n) = trElim (λ x → (isOfHLevel∥∥ (ℕ→ℕ₋₂ (suc (suc n))) ∣ north ∣ ∣ north ∣)) +Kn→ΩKn+1 (suc n) = trElim (λ x → (isOfHLevelTrunc (2 + (suc (suc n))) ∣ north ∣ ∣ north ∣)) λ a → cong ∣_∣ ((merid a) ∙ (sym (merid north))) {- @@ -305,7 +305,7 @@ isEquivHelper2 : isOfHLevel 3 A → isEquiv {B = ∥ A ∥ 1} ∣_∣ isEquivHelper2 ofHlevl = isoToIsEquiv (iso ∣_∣ (trElim (λ _ → ofHlevl) (λ a → a)) - (trElim {B = λ b → ∣ trElim (λ _ → ofHlevl) (λ a₁ → a₁) b ∣ ≡ b} (λ b → isOfHLevelSuc 3 (isOfHLevel∥∥ 1) ∣ trElim (λ _ → ofHlevl) (λ a₁ → a₁) b ∣ b) λ a → refl) + (trElim {B = λ b → ∣ trElim (λ _ → ofHlevl) (λ a₁ → a₁) b ∣ ≡ b} (λ b → isOfHLevelSuc 3 (isOfHLevelTrunc 3) ∣ trElim (λ _ → ofHlevl) (λ a₁ → a₁) b ∣ b) λ a → refl) λ b → refl) isEquivHelper : {a b : A} → isOfHLevel 3 A → isEquiv {B = _≡_ {A = ∥ A ∥ 1} ∣ a ∣ ∣ b ∣ } (cong ∣_∣) @@ -316,14 +316,12 @@ d-map : typ (Ω ((Susp S¹) , north)) → S¹ d-map p = subst HopfSuspS¹ p base d-mapId : (r : S¹) → d-map (ϕ base r) ≡ r -d-mapId r = (λ i → subst HopfSuspS¹ (□≡∙ (merid r) (sym (merid base)) (~ i)) base) ∙ - (substComposite-□ HopfSuspS¹ (merid r) (sym (merid base)) base) ∙ - test r +d-mapId r = substComposite HopfSuspS¹ (merid r) (sym (merid base)) base ∙ + rotLemma r where - test : (r : S¹) → rot r base ≡ r - test base = refl - test (loop i) = refl - + rotLemma : (r : S¹) → rot r base ≡ r + rotLemma base = refl + rotLemma (loop i) = refl S³≡SuspSuspS¹ : S³ ≡ Susp (Susp S¹) S³≡SuspSuspS¹ = S³≡SuspS² ∙ λ i → Susp (S²≡SuspS¹ i) @@ -354,45 +352,22 @@ is1Connected-dmap (loop j) = (transport (λ i → isContr (PathΩ {A = Susp (Susp S¹)} {a = north} (ℕ→ℕ₋₂ 1) i)) (refl , isOfHLevelSuc 1 (isOfHLevelSuc 0 Pr242SpecCase) ∣ north ∣ ∣ north ∣ (λ _ → ∣ north ∣))))) -d-equiv : isEquiv {A = ∥ typ (Ω (Susp S¹ , north)) ∥ (ℕ→ℕ₋₂ 1)} {B = ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (trElim (λ x → isOfHLevel∥∥ 1) λ x → ∣ d-map x ∣ ) +d-equiv : isEquiv {A = ∥ typ (Ω (Susp S¹ , north)) ∥ (ℕ→ℕ₋₂ 1)} {B = ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (trElim (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣ ) d-equiv = conEquiv (ℕ→ℕ₋₂ 1) d-map is1Connected-dmap .snd -d-mapId2 : (λ (x : ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)) → (trElim {n = (ℕ→ℕ₋₂ 1)} {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (λ x → isOfHLevel∥∥ 1) λ x → ∣ d-map x ∣) - (trElim (λ _ → isOfHLevel∥∥ 1) (λ a → ∣ ϕ base a ∣) x)) ≡ λ x → x -d-mapId2 = funExt (trElim (λ x → isOfHLevelSuc 2 (isOfHLevel∥∥ 1 ((trElim {n = (ℕ→ℕ₋₂ 1)} {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (λ x → isOfHLevel∥∥ 1) λ x → ∣ d-map x ∣) (trElim (λ _ → isOfHLevel∥∥ 1) (λ a → ∣ ϕ base a ∣) x)) x)) λ a i → ∣ d-mapId a i ∣) +d-mapId2 : (λ (x : ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)) → (trElim {n = 3} {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) + (trElim (λ _ → isOfHLevelTrunc 3) (λ a → ∣ ϕ base a ∣) x)) ≡ λ x → x +d-mapId2 = funExt (trElim (λ x → isOfHLevelSuc 2 (isOfHLevelTrunc 3 ((trElim {n = 3} {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) (trElim (λ _ → isOfHLevelTrunc 3) (λ a → ∣ ϕ base a ∣) x)) x)) λ a i → ∣ d-mapId a i ∣) -isEquiv∣ϕ∣ : isEquiv {A = ∥ S¹ ∥ ℕ→ℕ₋₂ 1} (trElim (λ _ → isOfHLevel∥∥ 1) (λ a → ∣ ϕ base a ∣)) -isEquiv∣ϕ∣ = compToIdEquiv (trElim {n = (ℕ→ℕ₋₂ 1)} {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (λ x → isOfHLevel∥∥ 1) λ x → ∣ d-map x ∣) - (trElim (λ _ → isOfHLevel∥∥ 1) (λ a → ∣ ϕ base a ∣)) +isEquiv∣ϕ∣ : isEquiv {A = ∥ S¹ ∥ ℕ→ℕ₋₂ 1} (trElim (λ _ → isOfHLevelTrunc 3) (λ a → ∣ ϕ base a ∣)) +isEquiv∣ϕ∣ = compToIdEquiv (trElim {n = 3} {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) + (trElim (λ _ → isOfHLevelTrunc 3) (λ a → ∣ ϕ base a ∣)) d-mapId2 d-equiv ---- - -S¹≡S1 : S₊ 1 ≡ S¹ -S¹≡S1 = (λ i → Susp (ua (Bool≃Susp⊥) (~ i))) ∙ sym S¹≡SuspBool - -S¹→S1 : S¹ → S₊ 1 -S¹→S1 x = SuspBool→S1 (S¹→SuspBool x) - -S1→S¹ : S₊ 1 → S¹ -S1→S¹ x = SuspBool→S¹ (S1→SuspBool x) - -S2→S² : S₊ 2 → S² -S2→S² north = base -S2→S² south = base -S2→S² (merid a i) = SuspS¹→S² (merid (SuspBool→S¹ (S1→SuspBool a)) i) - -S²→S2 : S² → S₊ 2 -S²→S2 base = north -S²→S2 (surf i i₁) = {!!} - -S2→SuspS¹ : S₊ 2 → Susp S¹ -S2→SuspS¹ x = S²→SuspS¹ (S2→S² x) - -S²≡S2 : S₊ 2 ≡ Susp S¹ -S²≡S2 i = Susp (S¹≡S1 i) - +--------------------------------- +-- We cheat when n = 1 and use J to prove the following lemmma. There is an obvious dependent path between ϕ base and ϕ north. Since the first one is an equivalence, so is the other. +-- funTEST2 : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ} {C : (A : Type ℓ) (a1 : A) → Type ℓ'} (p : A ≡ B) (a : A) (b : B) → (transport p a ≡ b) → (f : (A : Type ℓ) → @@ -408,15 +383,16 @@ funTEST2 {ℓ = ℓ}{A = A} {B = B} {C = C} = isEquiv (f B b)) λ a b trefl f is → transport (λ i → isEquiv (f A ((sym (transportRefl a) ∙ trefl) i))) is +----------------------------------------------------- -final : isEquiv {A = ∥ S₊ 1 ∥ 1} {B = ∥ typ (Ω (S₊ 2 , north)) ∥ 1} (trElim (λ _ → isOfHLevel∥∥ 1) (λ a → ∣ ϕ north a ∣)) -final = funTEST2 {A = S¹} (λ i → S¹≡S1 (~ i)) base north refl (λ A a1 → trElim (λ _ → isOfHLevel∥∥ 1) (λ a → ∣ ϕ a1 a ∣)) isEquiv∣ϕ∣ +final : isEquiv {A = ∥ S₊ 1 ∥ 1} {B = ∥ typ (Ω (S₊ 2 , north)) ∥ 1} (trElim (λ _ → isOfHLevelTrunc 3) (λ a → ∣ ϕ north a ∣)) +final = funTEST2 {A = S¹} (λ i → S¹≡S1 (~ i)) base north refl (λ A a1 → trElim (λ _ → isOfHLevelTrunc 3) (λ a → ∣ ϕ a1 a ∣)) isEquiv∣ϕ∣ -Kn→ΩKn+1Sucn : (n : ℕ) → Kn→ΩKn+1 (suc n) ≡ λ x → truncEquivΩ (ℕ→ℕ₋₂ (suc n)) .fst (trElim (λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (suc n))) (λ a → ∣ ϕ north a ∣) x) +Kn→ΩKn+1Sucn : (n : ℕ) → Kn→ΩKn+1 (suc n) ≡ λ x → truncEquivΩ (ℕ→ℕ₋₂ (suc n)) .fst (trElim (λ _ → isOfHLevelTrunc (3 + n)) (λ a → ∣ ϕ north a ∣) x) Kn→ΩKn+1Sucn n = funExt (trElim (λ x → isOfHLevelSuc (suc (suc n)) - ((isOfHLevel∥∥ (ℕ→ℕ₋₂ (suc (suc n))) ∣ north ∣ ∣ north ∣) + ((isOfHLevelTrunc ( 2 + (suc (suc n))) ∣ north ∣ ∣ north ∣) (Kn→ΩKn+1 (suc n) x) - (truncEquivΩ (ℕ→ℕ₋₂ (suc n)) .fst (trElim (λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (suc n))) (λ a → ∣ ϕ north a ∣) x)))) + (truncEquivΩ (ℕ→ℕ₋₂ (suc n)) .fst (trElim (λ _ → isOfHLevelTrunc (2 + (suc n))) (λ a → ∣ ϕ north a ∣) x)))) λ a → refl) @@ -434,7 +410,7 @@ isEquivKn→ΩKn+1 zero = compEquiv (looper , isEquivLooper) (cong ∣_∣ , isE (λ i → transp (λ j → S¹≡SuspBool (~ j ∨ ~ i)) (~ i) (transport (sym (ua S1≃SuspBool)) x) ≡ transp (λ j → S¹≡SuspBool (~ j ∨ ~ i)) (~ i) (transport (sym (ua S1≃SuspBool)) x)) ∙ basedΩS¹≡Int (transport (sym S¹≡SuspBool) (transport (sym (ua S1≃SuspBool)) x)) isEquivKn→ΩKn+1 (suc zero) = transport (λ i → isEquiv (Kn→ΩKn+1Sucn zero (~ i))) - (compEquiv (trElim (λ _ → isOfHLevel∥∥ (ℕ→ℕ₋₂ (suc zero))) (λ a → ∣ ϕ north a ∣) , + (compEquiv (trElim (λ _ → isOfHLevelTrunc (2 + (suc zero))) (λ a → ∣ ϕ north a ∣) , final) (truncEquivΩ (ℕ→ℕ₋₂ (suc zero))) .snd) isEquivKn→ΩKn+1 (suc (suc n)) = transport (λ i → isEquiv (Kn→ΩKn+1Sucn (suc n) (~ i))) From 31fe2850714c2bb992e631ff9a2794ae0cdd42aa Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 14 Apr 2020 19:04:27 +0200 Subject: [PATCH 16/89] move some things into ZCohomology.Properties + comments --- Cubical/Foundations/Equiv.agda | 9 + Cubical/Foundations/Path.agda | 1 + .../Connected/FreudenthalProof/8-6-1.agda | 52 +++ .../Connected/FreudenthalProof/8-6-4.agda | 307 +++++++++--------- .../Connected/FreudenthalProof/8-6-4Code.agda | 127 ++++---- .../HITs/Truncation/Connected/Properties.agda | 68 ++++ Cubical/ZCohomology/Properties.agda | 46 ++- .../{S1Loop.agda => cupProdPrelims.agda} | 208 ++++-------- 8 files changed, 461 insertions(+), 357 deletions(-) rename Cubical/ZCohomology/{S1Loop.agda => cupProdPrelims.agda} (67%) diff --git a/Cubical/Foundations/Equiv.agda b/Cubical/Foundations/Equiv.agda index 13bd40439..8686aa4a4 100644 --- a/Cubical/Foundations/Equiv.agda +++ b/Cubical/Foundations/Equiv.agda @@ -155,3 +155,12 @@ _■ = idEquiv infixr 0 _≃⟨_⟩_ infix 1 _■ + +composesToId→Equiv : (f : A → B) (g : B → A) → f ∘ g ≡ idfun B → isEquiv f → isEquiv g +composesToId→Equiv f g id iseqf = + isoToIsEquiv (iso g + f + (λ b → (λ i → (equiv-proof iseqf (f b) .snd (g (f b) , cong (λ h → h (f b)) id) (~ i)) .fst ) ∙ + cong (λ x → (equiv-proof iseqf (f b) .fst .fst )) id ∙ + λ i → (equiv-proof iseqf (f b) .snd) (b , refl) i .fst ) + λ a → cong (λ f → f a) id) diff --git a/Cubical/Foundations/Path.agda b/Cubical/Foundations/Path.agda index 597cd5687..0e47e242e 100644 --- a/Cubical/Foundations/Path.agda +++ b/Cubical/Foundations/Path.agda @@ -6,6 +6,7 @@ open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence private variable diff --git a/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-1.agda b/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-1.agda index 9699a89c5..7e4b9f933 100644 --- a/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-1.agda +++ b/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-1.agda @@ -29,7 +29,10 @@ private A : Type ℓ B : Type ℓ' +{- This file contains lemma 8.6.1. from the HoTT book -} + private +{- The following two lemmas transform the expression as in the HoTT book -} Lemma861-fibId : ∀{ℓ} (n : ℕ₋₂) (k : ℕ) (f : A → B) → (is- n -Connected f) → (P : B → HLevel ℓ (((suc k) + (2+ n)))) @@ -87,6 +90,7 @@ private ≡ Σ ((b : B) → fst (P b)) λ g → ((a : A) → g (f a) ≡ l a) Lemma861-fibId2 n k f isCon P l = isoToPath (iso (λ p → fst p , funExt⁻ (snd p)) (λ p → fst p , funExt (snd p)) (λ b → refl) λ b → refl) +{- Lemma861 : ∀{ℓ} (n : ℕ₋₂) (k : ℕ) (f : A → B) → (is- n -Connected f) → (P : B → HLevel ℓ (k + (2+ n))) → @@ -131,6 +135,54 @@ private ((λ x → ((fst a) x ≡ (fst b) x ) , (snd (P x)) (fst a x) (fst b x) ))) λ z → ((snd a z) ∙ sym (snd b z))) +-} + + Lemma861 : ∀{ℓ} (n : ℕ₋₂) (k : ℕ) (f : A → B) → + (is- n -Connected f) → + (P : B → HLevel ℓ (k + (2+ n))) → + is- (-2+ k) -Truncated (indConFun {k = -2+ (k + (2+ n))} f P) + Lemma861 {B = B} {ℓ = ℓ} neg2 zero f iscon P = equiv-proof (conInd-i→ii f neg2 iscon P) + Lemma861 {B = B} {ℓ = ℓ} neg2 (suc zero) f iscon P l = transport (λ i → isOfHLevel (suc (zero)) + (Lemma861-fibId2 neg2 (zero) f iscon P l (~ i))) + λ a b → transport (λ i → isOfHLevel (zero) + (Lemma861-fibId neg2 (zero) f iscon P l a b + (λ x → ((snd (P x)) (fst a x) (fst b x)) , + (λ y → isOfHLevelSuc 1 (snd (P x)) + (fst a x) (fst b x) _ _)) (~ i))) + ((Lemma861 neg2 (zero) f iscon + ((λ x → ((fst a) x ≡ (fst b) x ) , + ((snd (P x)) (fst a x) (fst b x) , + (λ y → isOfHLevelSuc 1 (snd (P x)) + (fst a x) (fst b x) _ _)) ))) + λ z → (snd a z) ∙ sym (snd b z)) .fst + Lemma861 {B = B} {ℓ = ℓ} neg2 (suc (suc k)) f iscon P l = transport (λ i → isOfHLevel (suc (suc k)) + (Lemma861-fibId2 neg2 (suc k) f iscon P l (~ i))) + λ a b → transport (λ i → isOfHLevel (suc k) + (Lemma861-fibId neg2 (suc k) f iscon P l a b + (λ x → (snd (P x)) (fst a x) (fst b x)) (~ i))) + ((Lemma861 neg2 (suc k) f iscon + ((λ x → ((fst a) x ≡ (fst b) x ) , + (snd (P x)) (fst a x) (fst b x) ))) + λ z → ((snd a z) ∙ sym (snd b z))) + Lemma861 {B = B} {ℓ = ℓ} (-1+ n) zero f iscon P = equiv-proof (conInd-i→ii f (-1+ n) iscon P) + Lemma861 {B = B} {ℓ = ℓ} (-1+ n) (suc zero) f iscon P l = transport (λ i → isOfHLevel (suc (zero)) + (Lemma861-fibId2 (-1+ n) (zero) f iscon P l (~ i))) + λ a b → transport (λ i → isOfHLevel (zero) + (Lemma861-fibId (-1+ n) (zero) f iscon P l a b + (λ x → snd (P x) (fst a x) (fst b x)) (~ i))) + ((Lemma861 (-1+ n) (zero) f iscon + ((λ x → ((fst a) x ≡ (fst b) x ) , + (snd (P x) (fst a x) (fst b x)) ))) + λ z → (snd a z) ∙ sym (snd b z)) .fst + Lemma861 {B = B} {ℓ = ℓ} (-1+ n) (suc (suc k)) f iscon P l = transport (λ i → isOfHLevel (suc (suc k)) + (Lemma861-fibId2 (-1+ n) (suc k) f iscon P l (~ i))) + λ a b → transport (λ i → isOfHLevel (suc k) + (Lemma861-fibId (-1+ n) (suc k) f iscon P l a b + (λ x → (snd (P x)) (fst a x) (fst b x)) (~ i))) + ((Lemma861 (-1+ n) (suc k) f iscon + ((λ x → ((fst a) x ≡ (fst b) x ) , + (snd (P x)) (fst a x) (fst b x) ))) + λ z → ((snd a z) ∙ sym (snd b z))) {- The following more general versions are more useful -} Lemma861Gen : ∀{ℓ} (n : ℕ₋₂) (k : ℕ) (expr : ℕ₋₂ → ℕ → ℕ) → diff --git a/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4.agda b/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4.agda index ae79a7bbc..54ceac0f1 100644 --- a/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4.agda +++ b/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4.agda @@ -30,139 +30,140 @@ private {- this file contains a proof that σ is 2n-connected -} --- defining pair⁼ as in the HoTT book. Using J to be able to have easy access to the refl case ---- -abstract - pair⁼ : ∀ {ℓ'} {B : A → Type ℓ'} {x y : Σ A (λ a → B a)} (p : fst x ≡ fst y) → - transport (λ i → B (p i)) (snd x) ≡ (snd y) → - x ≡ y - pair⁼ {B = B}{x = x1 , x2} {y = y1 , y2} p = J (λ y1 p → ((y2 : B y1) → (transport (λ i → B (p i)) x2 ≡ y2) → (x1 , x2) ≡ (y1 , y2))) - (λ y2 p2 i → x1 , ((sym (transportRefl x2)) ∙ p2) i) p y2 - - pair⁼Refl : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} {a : A} {x y : B a} (p : transport (λ i → B a) x ≡ y) → - pair⁼ {B = B} {x = (a , x)} {y = (a , y)} refl p - ≡ - λ i → a , ((sym (transportRefl x)) ∙ p) i - pair⁼Refl {A = A} {B = B} {a = a} {x = x} {y = y} p = cong (λ f → f y p) - (JRefl (λ y1 p → ((y2 : B y1) → (transport (λ i → B (p i)) x ≡ y2) → - _≡_ {A = Σ A (λ a → B a)} (a , x) (y1 , y2))) - (λ y2 p2 i → a , ((sym (transportRefl x)) ∙ p2) i)) - - pair⁼Sym : ∀ {ℓ'} {B : A → Type ℓ'} {x y : Σ A (λ a → B a)} - (p : fst x ≡ fst y) - (q : transport (λ i → B (p i)) (snd x) ≡ (snd y)) → - sym (pair⁼ {x = x} {y = y} p q) - ≡ pair⁼ (sym p) (transportLemma {B = B} p (snd x) (snd y) q ) - pair⁼Sym {ℓ} {A = A} {B = B} {x = x} {y = y} p = J {ℓ} {A} {fst x} (λ fsty p → (sndy : B (fsty)) (q : transport (λ i → B (p i)) (snd x) ≡ (sndy)) → - sym (pair⁼ p q) - ≡ pair⁼ (sym p) (transportLemma {B = B} p (snd x) (sndy) q )) - (λ sndy → J (λ sndy q → sym (pair⁼ (λ _ → fst x) q) - ≡ pair⁼ refl (transportLemma {B = B} refl (snd x) sndy q)) - ((λ j → (sym (pair⁼Refl refl j) )) ∙ - (λ k i → fst x , ((rUnit (sym (transportRefl (snd x))) (~ k)) (~ i))) ∙ - (λ k i → fst x , helper (~ k) i) ∙ - sym (pair⁼Refl (transportLemma {B = B} (λ _ → fst x) (snd x) - (transport (λ i → B (fst x)) (snd x)) refl)))) - p (snd y) - where - helper : (sym (transportRefl (transport (λ i₁ → B (fst x)) (snd x)))) ∙ - (transportLemma {B = B} (λ _ → fst x) (snd x) - (transport (λ i₂ → B (fst x)) (snd x)) - (λ _ → transport (λ i₂ → B (fst x)) (snd x))) - ≡ (transportRefl (snd x)) - helper = (λ i → sym (transportRefl (transport (λ i₁ → B (fst x)) (snd x))) ∙ - (transportLemmaRefl {B = B} (snd x) (snd x) i) (snd x) - (transport (λ i₂ → B (fst x)) (snd x)) - (λ _ → transport (λ i₂ → B (fst x)) (snd x))) ∙ - (λ i → sym (transportRefl (transport (λ i₁ → B (fst x)) (snd x))) ∙ - transportRefl (transport (λ i₂ → B (fst x)) (snd x)) ∙ - (λ _ → transport (λ i₂ → B (fst x)) (snd x)) ∙ - transportRefl (snd x)) ∙ - (λ i → sym (transportRefl (transport (λ i₁ → B (fst x)) (snd x))) ∙ - transportRefl (transport (λ i₂ → B (fst x)) (snd x)) ∙ - lUnit (transportRefl (snd x)) (~ i)) ∙ - assoc (sym (transportRefl (transport (λ i₁ → B (fst x)) (snd x)))) - (transportRefl (transport (λ i₁ → B (fst x)) (snd x))) - (transportRefl (snd x)) ∙ - (λ i → (lCancel (transportRefl (transport (λ i₁ → B (fst x)) (snd x))) i) ∙ - transportRefl (snd x)) ∙ - sym (lUnit (transportRefl (snd x))) - - {- We need something slightly stronger than functoriality of transport. - Since we will be transporting over dependent pairs, we also need to split the second component -} - functTransp2 : ∀ {ℓ ℓ'} {A : Type ℓ} {a1 b : A} {C : Σ A (λ b → a1 ≡ b) → Type ℓ' } (p : a1 ≡ b) (q : b ≡ a1) → - transport (λ i → C (pair⁼ (p ∙ q) (pairLemma2 (p ∙ q)) i) ) - ≡ λ x → transport (λ i → C (pair⁼ q ((pairLemma p q)) i)) - (transport (λ i → C (pair⁼ p (pairLemma2 p) i)) - x) - functTransp2 {ℓ} {ℓ'} {A = A} {a1 = a1} {b = b} {C = C} = - J (λ b p → (q : b ≡ a1) → - transport (λ i → C (pair⁼ {B = λ a → a1 ≡ a} {x = (a1 , refl {x = a1})} - (p ∙ q) (pairLemma2 (p ∙ q)) i) ) - ≡ λ x → transport (λ i → C (pair⁼ q ((pairLemma p q)) i)) - (transport (λ i → C (pair⁼ p (pairLemma2 p) i)) - x)) - λ q → (λ j → subst C (hej a1 q (~ j))) ∙ - (λ j → subst C (pair⁼ (λ _ → a1) (pairLemma2 (λ _ → a1)) ∙ - pair⁼ q (pairLemmaRefl q (~ j)))) ∙ - funExt λ x → (substComposite-∙ C (pair⁼ refl (pairLemma2 refl)) (pair⁼ q (pairLemma refl q)) x) +pair⁼ : ∀ {ℓ'} {B : A → Type ℓ'} {x y : Σ A (λ a → B a)} (p : fst x ≡ fst y) → + transport (λ i → B (p i)) (snd x) ≡ (snd y) → + x ≡ y +pair⁼ {B = B}{x = x1 , x2} {y = y1 , y2} p = J (λ y1 p → ((y2 : B y1) → (transport (λ i → B (p i)) x2 ≡ y2) → (x1 , x2) ≡ (y1 , y2))) + (λ y2 p2 i → x1 , ((sym (transportRefl x2)) ∙ p2) i) p y2 + +pair⁼Refl : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} {a : A} {x y : B a} (p : transport (λ i → B a) x ≡ y) → + pair⁼ {B = B} {x = (a , x)} {y = (a , y)} refl p + ≡ + λ i → a , ((sym (transportRefl x)) ∙ p) i +pair⁼Refl {A = A} {B = B} {a = a} {x = x} {y = y} p = cong (λ f → f y p) + (JRefl (λ y1 p → ((y2 : B y1) → (transport (λ i → B (p i)) x ≡ y2) → + _≡_ {A = Σ A (λ a → B a)} (a , x) (y1 , y2))) + (λ y2 p2 i → a , ((sym (transportRefl x)) ∙ p2) i)) + + +{- We will also need the following idenitity for sym (pair⁼ p q). -} +pair⁼Sym : ∀ {ℓ'} {B : A → Type ℓ'} {x y : Σ A (λ a → B a)} + (p : fst x ≡ fst y) + (q : transport (λ i → B (p i)) (snd x) ≡ (snd y)) → + sym (pair⁼ {x = x} {y = y} p q) + ≡ pair⁼ (sym p) (transportLemma {B = B} p (snd x) (snd y) q ) +pair⁼Sym {ℓ} {A = A} {B = B} {x = x} {y = y} p = J {ℓ} {A} {fst x} (λ fsty p → (sndy : B (fsty)) (q : transport (λ i → B (p i)) (snd x) ≡ (sndy)) → + sym (pair⁼ p q) + ≡ pair⁼ (sym p) (transportLemma {B = B} p (snd x) (sndy) q )) + (λ sndy → J (λ sndy q → sym (pair⁼ (λ _ → fst x) q) + ≡ pair⁼ refl (transportLemma {B = B} refl (snd x) sndy q)) + ((λ j → (sym (pair⁼Refl refl j) )) ∙ + (λ k i → fst x , ((rUnit (sym (transportRefl (snd x))) (~ k)) (~ i))) ∙ + (λ k i → fst x , helper (~ k) i) ∙ + sym (pair⁼Refl (transportLemma {B = B} (λ _ → fst x) (snd x) + (transport (λ i → B (fst x)) (snd x)) refl)))) + p (snd y) + where + helper : (sym (transportRefl (transport (λ i₁ → B (fst x)) (snd x)))) ∙ + (transportLemma {B = B} (λ _ → fst x) (snd x) + (transport (λ i₂ → B (fst x)) (snd x)) + (λ _ → transport (λ i₂ → B (fst x)) (snd x))) + ≡ (transportRefl (snd x)) + helper = (λ i → sym (transportRefl (transport (λ i₁ → B (fst x)) (snd x))) ∙ + (transportLemmaRefl {B = B} (snd x) (snd x) i) (snd x) + (transport (λ i₂ → B (fst x)) (snd x)) + (λ _ → transport (λ i₂ → B (fst x)) (snd x))) ∙ + (λ i → sym (transportRefl (transport (λ i₁ → B (fst x)) (snd x))) ∙ + transportRefl (transport (λ i₂ → B (fst x)) (snd x)) ∙ + (λ _ → transport (λ i₂ → B (fst x)) (snd x)) ∙ + transportRefl (snd x)) ∙ + (λ i → sym (transportRefl (transport (λ i₁ → B (fst x)) (snd x))) ∙ + transportRefl (transport (λ i₂ → B (fst x)) (snd x)) ∙ + lUnit (transportRefl (snd x)) (~ i)) ∙ + assoc (sym (transportRefl (transport (λ i₁ → B (fst x)) (snd x)))) + (transportRefl (transport (λ i₁ → B (fst x)) (snd x))) + (transportRefl (snd x)) ∙ + (λ i → (lCancel (transportRefl (transport (λ i₁ → B (fst x)) (snd x))) i) ∙ + transportRefl (snd x)) ∙ + sym (lUnit (transportRefl (snd x))) + +{- We need something slightly stronger than functoriality of transport. + Since we will be transporting over dependent pairs, we also need to split the second component -} +functTransp2 : ∀ {ℓ ℓ'} {A : Type ℓ} {a1 b : A} {C : Σ A (λ b → a1 ≡ b) → Type ℓ' } (p : a1 ≡ b) (q : b ≡ a1) → + transport (λ i → C (pair⁼ (p ∙ q) (pairLemma2 (p ∙ q)) i) ) + ≡ λ x → transport (λ i → C (pair⁼ q ((pairLemma p q)) i)) + (transport (λ i → C (pair⁼ p (pairLemma2 p) i)) + x) +functTransp2 {ℓ} {ℓ'} {A = A} {a1 = a1} {b = b} {C = C} = + J (λ b p → (q : b ≡ a1) → + transport (λ i → C (pair⁼ {B = λ a → a1 ≡ a} {x = (a1 , refl {x = a1})} + (p ∙ q) (pairLemma2 (p ∙ q)) i) ) + ≡ λ x → transport (λ i → C (pair⁼ q ((pairLemma p q)) i)) + (transport (λ i → C (pair⁼ p (pairLemma2 p) i)) + x)) + λ q → (λ j → subst C (hej a1 q (~ j))) ∙ + (λ j → subst C (pair⁼ (λ _ → a1) (pairLemma2 (λ _ → a1)) ∙ + pair⁼ q (pairLemmaRefl q (~ j)))) ∙ + funExt λ x → (substComposite-∙ C (pair⁼ refl (pairLemma2 refl)) (pair⁼ q (pairLemma refl q)) x) + where + hej : (b : A) (q : a1 ≡ b) → (pair⁼ {A = A} {B = λ a → a1 ≡ a} + {x = (a1 , λ _ → a1)} {y = (a1 , λ _ → a1)} + (λ _ → a1) (pairLemma2 {a = a1} {b = a1} (λ _ → a1)) ∙ + pair⁼ q ((pairLemma2 q) ∙ lUnit q)) + ≡ pair⁼ (refl ∙ q) (pairLemma2 (refl {x = a1} ∙ q)) + hej b = J (λ b q → pair⁼ (λ _ → a1) (pairLemma2 (λ _ → a1)) ∙ + pair⁼ q (pairLemma2 q ∙ lUnit q) + ≡ pair⁼ (refl ∙ q) (pairLemma2 (refl ∙ q))) + ((λ i → (helper2 i) ∙ pair⁼ refl (pairLemma2 refl ∙ lUnit refl)) ∙ + sym (lUnit (pair⁼ (λ _ → a1) (pairLemma2 (λ _ → a1) ∙ lUnit (λ _ → a1)))) ∙ + desired≡) where - hej : (b : A) (q : a1 ≡ b) → (pair⁼ {A = A} {B = λ a → a1 ≡ a} - {x = (a1 , λ _ → a1)} {y = (a1 , λ _ → a1)} - (λ _ → a1) (pairLemma2 {a = a1} {b = a1} (λ _ → a1)) ∙ - pair⁼ q ((pairLemma2 q) ∙ lUnit q)) - ≡ pair⁼ (refl ∙ q) (pairLemma2 (refl {x = a1} ∙ q)) - hej b = J (λ b q → pair⁼ (λ _ → a1) (pairLemma2 (λ _ → a1)) ∙ - pair⁼ q (pairLemma2 q ∙ lUnit q) - ≡ pair⁼ (refl ∙ q) (pairLemma2 (refl ∙ q))) - ((λ i → (helper2 i) ∙ pair⁼ refl (pairLemma2 refl ∙ lUnit refl)) ∙ - sym (lUnit (pair⁼ (λ _ → a1) (pairLemma2 (λ _ → a1) ∙ lUnit (λ _ → a1)))) ∙ - desired≡) - where - helper2 : (pair⁼ {A = A} {B = λ a → a1 ≡ a} - {x = (a1 , λ _ → a1)} {y = (a1 , λ _ → a1)} - (λ _ → a1) (pairLemma2 {a = a1} {b = a1} (λ _ → a1))) - ≡ refl - helper2 = (λ i → (JRefl (λ y1 p → ((y2 : a1 ≡ y1) → (transport (λ i → a1 ≡ (p i)) refl ≡ y2) → - _≡_ {A = Σ A (λ a → a1 ≡ a)} (a1 , refl) (y1 , y2))) - (λ y2 p2 i → refl {x = a1} i , ((sym (transportRefl refl)) ∙ p2) i) i) - (λ _ → a1) - (pairLemma2 {a = a1} {b = a1} (λ _ → a1))) ∙ - (λ j → λ i → a1 , ((sym (transportRefl refl)) ∙ - JRefl (λ b p → transport (λ i → a1 ≡ p i) refl ≡ p) - (transportRefl refl) j ) i) ∙ - λ j i → a1 , (lCancel (transportRefl refl) j) i - - PathP1 : PathP (λ j → _≡_ {A = Σ A (λ a → a1 ≡ a)} - (a1 , (λ _ → a1)) - (a1 , lUnit (λ _ → a1) (~ j))) - (pair⁼ (λ _ → a1) (pairLemma2 (λ _ → a1) ∙ lUnit (λ _ → a1))) - refl - PathP1 j = hcomp (λ k → λ{(j = i0) → pair⁼ (λ _ → a1) (pairLemma2 (λ _ → a1) ∙ - lUnit (λ _ → a1)) - ; (j = i1) → ((λ i → pair⁼ (λ _ → a1) (rUnit (pairLemma2 (λ _ → a1)) (~ i)) ) ∙ - helper2) k}) - (pair⁼ refl ((pairLemma2 (λ _ → a1)) ∙ - (λ i → lUnit refl (i ∧ (~ j))))) - - PathP2 : PathP (λ j → _≡_ {A = Σ A (λ a → a1 ≡ a)} - (a1 , (λ _ → a1)) - (a1 , lUnit (λ _ → a1) j)) - refl - (pair⁼ ((λ _ → a1) ∙ refl) (pairLemma2 ((λ _ → a1) ∙ refl))) - PathP2 j = hcomp (λ k → λ{(j = i0) → helper2 k - ; (j = i1) → (pair⁼ ((λ _ → a1) ∙ refl) (pairLemma2 ((λ _ → a1) ∙ refl)))}) - (pair⁼ (lUnit (λ _ → a1) j) (pairLemma2 (lUnit (λ _ → a1) j))) - - pathsCancel : (λ j → _≡_ {A = Σ A (λ a → a1 ≡ a)} (a1 , (λ _ → a1)) (a1 , lUnit (λ _ → a1) (~ j))) ∙ - ((λ j → _≡_ {A = Σ A (λ a → a1 ≡ a)} (a1 , (λ _ → a1)) (a1 , lUnit (λ _ → a1) j))) - ≡ refl - pathsCancel = lCancel (λ j → _≡_ {A = Σ A (λ a → a1 ≡ a)} (a1 , (λ _ → a1)) (a1 , lUnit (λ _ → a1) j)) - - desired≡ : pair⁼ (λ _ → a1) (λ i i₁ → (pairLemma2 (λ _ → a1) ∙ lUnit (λ _ → a1)) i i₁) - ≡ pair⁼ ((λ _ → a1) ∙ refl) (pairLemma2 ((λ _ → a1) ∙ refl)) - desired≡ = transport (λ i → PathP (λ j → pathsCancel i j) - (pair⁼ (λ _ → a1) (pairLemma2 (λ _ → a1) ∙ lUnit (λ _ → a1))) - (pair⁼ ((λ _ → a1) ∙ refl) (pairLemma2 ((λ _ → a1) ∙ refl)))) - (compPathP PathP1 PathP2) + helper2 : (pair⁼ {A = A} {B = λ a → a1 ≡ a} + {x = (a1 , λ _ → a1)} {y = (a1 , λ _ → a1)} + (λ _ → a1) (pairLemma2 {a = a1} {b = a1} (λ _ → a1))) + ≡ refl + helper2 = (λ i → (JRefl (λ y1 p → ((y2 : a1 ≡ y1) → (transport (λ i → a1 ≡ (p i)) refl ≡ y2) → + _≡_ {A = Σ A (λ a → a1 ≡ a)} (a1 , refl) (y1 , y2))) + (λ y2 p2 i → refl {x = a1} i , ((sym (transportRefl refl)) ∙ p2) i) i) + (λ _ → a1) + (pairLemma2 {a = a1} {b = a1} (λ _ → a1))) ∙ + (λ j → λ i → a1 , ((sym (transportRefl refl)) ∙ + JRefl (λ b p → transport (λ i → a1 ≡ p i) refl ≡ p) + (transportRefl refl) j ) i) ∙ + λ j i → a1 , (lCancel (transportRefl refl) j) i + + PathP1 : PathP (λ j → _≡_ {A = Σ A (λ a → a1 ≡ a)} + (a1 , (λ _ → a1)) + (a1 , lUnit (λ _ → a1) (~ j))) + (pair⁼ (λ _ → a1) (pairLemma2 (λ _ → a1) ∙ lUnit (λ _ → a1))) + refl + PathP1 j = hcomp (λ k → λ{(j = i0) → pair⁼ (λ _ → a1) (pairLemma2 (λ _ → a1) ∙ + lUnit (λ _ → a1)) + ; (j = i1) → ((λ i → pair⁼ (λ _ → a1) (rUnit (pairLemma2 (λ _ → a1)) (~ i)) ) ∙ + helper2) k}) + (pair⁼ refl ((pairLemma2 (λ _ → a1)) ∙ + (λ i → lUnit refl (i ∧ (~ j))))) + + PathP2 : PathP (λ j → _≡_ {A = Σ A (λ a → a1 ≡ a)} + (a1 , (λ _ → a1)) + (a1 , lUnit (λ _ → a1) j)) + refl + (pair⁼ ((λ _ → a1) ∙ refl) (pairLemma2 ((λ _ → a1) ∙ refl))) + PathP2 j = hcomp (λ k → λ{(j = i0) → helper2 k + ; (j = i1) → (pair⁼ ((λ _ → a1) ∙ refl) (pairLemma2 ((λ _ → a1) ∙ refl)))}) + (pair⁼ (lUnit (λ _ → a1) j) (pairLemma2 (lUnit (λ _ → a1) j))) + + pathsCancel : (λ j → _≡_ {A = Σ A (λ a → a1 ≡ a)} (a1 , (λ _ → a1)) (a1 , lUnit (λ _ → a1) (~ j))) ∙ + ((λ j → _≡_ {A = Σ A (λ a → a1 ≡ a)} (a1 , (λ _ → a1)) (a1 , lUnit (λ _ → a1) j))) + ≡ refl + pathsCancel = lCancel (λ j → _≡_ {A = Σ A (λ a → a1 ≡ a)} (a1 , (λ _ → a1)) (a1 , lUnit (λ _ → a1) j)) + + desired≡ : pair⁼ (λ _ → a1) (λ i i₁ → (pairLemma2 (λ _ → a1) ∙ lUnit (λ _ → a1)) i i₁) + ≡ pair⁼ ((λ _ → a1) ∙ refl) (pairLemma2 ((λ _ → a1) ∙ refl)) + desired≡ = transport (λ i → PathP (λ j → pathsCancel i j) + (pair⁼ (λ _ → a1) (pairLemma2 (λ _ → a1) ∙ lUnit (λ _ → a1))) + (pair⁼ ((λ _ → a1) ∙ refl) (pairLemma2 ((λ _ → a1) ∙ refl)))) + (compPathP PathP1 PathP2) ------------------------------------------------------------------------------------------------------------------------ @@ -456,9 +457,11 @@ codeTranspId2 {ℓ} {X = X} {a = a} = J (λ b q → (p : a ≡ b) (A : (a ≡ a λ k → (transportRefl (A (subst (_≡_ a) (λ i → a) p)) ∙ (λ i → A (pairLemmaId p (λ i₁ → a ) (~ k) i))) ∙ f p where - transpLemma2 : ∀ {ℓ ℓ'} {X : Type ℓ} {x : X} {A : x ≡ x → Type ℓ'} (p : x ≡ x) → (λ i → transportRefl A i p) ≡ (transportRefl (A (transport (λ i → x ≡ x) p)) ∙ (λ i → A ((transportRefl p) i))) + transpLemma2 : ∀ {ℓ ℓ'} {X : Type ℓ} {x : X} {A : x ≡ x → Type ℓ'} (p : x ≡ x) → + (λ i → transportRefl A i p) + ≡ (transportRefl (A (transport (λ i → x ≡ x) p)) ∙ (λ i → A ((transportRefl p) i))) transpLemma2 {ℓ' = ℓ'}{x = x} {A = A} p j = hcomp (λ k → λ{(j = i0) → (sym (lUnit (λ i → transportRefl (A ((transportRefl p) i)) i))) k - ; (j = i1) → (transportRefl (A (transport (λ i → x ≡ x) p)) ∙ + ; (j = i1) → (transportRefl (A (transport (λ i → x ≡ x) p)) ∙ (λ i → A ((transportRefl p) i)))}) ((λ i → transportRefl (A (transport (λ i → x ≡ x) p )) (i ∧ j)) ∙ (λ i → transportRefl (A ((transportRefl p) i)) (i ∨ j))) @@ -564,8 +567,13 @@ codeTranspId3 n a x1 iscon = (RlFun a x1 n iscon (merid x1)) (outer (inner guy)) {- We now need to show that the outer transport applied to ∣ x1 , refl ∣ gives ∣ x1 , refl ∣. We transform the outer transport-} codeTranspId4 : (n : ℕ) (a x1 : A) (iscon : is- (ℕ→ℕ₋₂ n) -ConnectedType A) → - transport (λ i → uncurry (CODE a n iscon) (pair⁼ (sym (merid a)) (pairLemma (merid x1) (sym (merid a)) ) i)) ∣ x1 , refl ∣ - ≡ transport (λ i → uncurry (CODE a n iscon) (pair⁼ (λ i₁ → merid a (~ i₁)) (transpRCancel (λ i → north ≡ (merid a (~ i))) (merid x1 ∙ (sym (merid a)))) i)) ∣ x1 , rUnit (merid x1) ∙ sym (cong (λ x → merid x1 ∙ x) (lCancel (merid a))) ∙ assocJ (merid x1) (sym (merid a)) (merid a) ∙ sym (pairLemma {a1 = north} (merid x1 ∙ (sym (merid a))) (merid a)) ∣ + transport (λ i → uncurry (CODE a n iscon) (pair⁼ (sym (merid a)) (pairLemma (merid x1) (sym (merid a)) ) i)) + ∣ x1 , refl ∣ + ≡ transport (λ i → uncurry (CODE a n iscon) (pair⁼ (λ i₁ → merid a (~ i₁)) (transpRCancel (λ i → north ≡ (merid a (~ i))) (merid x1 ∙ (sym (merid a)))) i)) + ∣ x1 , rUnit (merid x1) ∙ + sym (cong (λ x → merid x1 ∙ x) (lCancel (merid a))) ∙ + assocJ (merid x1) (sym (merid a)) (merid a) ∙ + sym (pairLemma {a1 = north} (merid x1 ∙ (sym (merid a))) (merid a)) ∣ codeTranspId4 {ℓ} {A = A} n a x1 iscon = sym (helper north (sym (merid a))) where helper : (y : Susp A) → (q : south ≡ y) → transport (λ i → uncurry (CODE a n iscon) (pair⁼ q (transpRCancel (λ i → north ≡ q i) ((merid x1) ∙ q)) i)) @@ -573,16 +581,16 @@ codeTranspId4 {ℓ} {A = A} n a x1 iscon = sym (helper north (sym (merid a))) sym (cong (λ x → merid x1 ∙ x) (lCancel (sym q))) ∙ assocJ (merid x1) q (sym q) ∙ sym ((pairLemma {a1 = north} (merid x1 ∙ q) (sym q))) ∣ ≡ transport (λ i → uncurry (CODE a n iscon) (pair⁼ q (pairLemma (merid x1) q ) i)) ∣ x1 , refl ∣ - helper y = J (λ y q → transport (λ i → uncurry (CODE a n iscon) (pair⁼ q (transpRCancel (λ i → north ≡ q i) ((merid x1) ∙ q)) i)) ∣ x1 , rUnit (merid x1) ∙ sym (cong (λ x → merid x1 ∙ x) (lCancel (sym q))) ∙ assocJ (merid x1) q (sym q) ∙ sym ((pairLemma {a1 = north} (merid x1 ∙ q) (sym q))) ∣ ≡ transport (λ i → uncurry (CODE a n iscon) (pair⁼ q (pairLemma (merid x1) q ) i)) ∣ x1 , refl ∣) - (transport - (λ i → - uncurry (CODE a n iscon) - (pair⁼ refl - (transpRCancel (λ i₁ → north ≡ refl i₁) (merid x1 ∙ refl)) i)) ∣ x1 , originalPath ∣ + helper y = J (λ y q → transport (λ i → uncurry (CODE a n iscon) (pair⁼ q (transpRCancel (λ i → north ≡ q i) ((merid x1) ∙ q)) i)) + ∣ x1 , rUnit (merid x1) ∙ + sym (cong (λ x → merid x1 ∙ x) (lCancel (sym q))) ∙ + assocJ (merid x1) q (sym q) ∙ sym ((pairLemma {a1 = north} (merid x1 ∙ q) (sym q))) ∣ + ≡ transport (λ i → uncurry (CODE a n iscon) (pair⁼ q (pairLemma (merid x1) q ) i)) ∣ x1 , refl ∣) + (transport (λ i → uncurry (CODE a n iscon) (pair⁼ refl (transpRCancel (λ i₁ → north ≡ refl i₁) (merid x1 ∙ refl)) i)) ∣ x1 , originalPath ∣ ≡⟨ (λ j → (transport (λ i → uncurry (CODE a n iscon) (pair⁼Refl (transpRCancel (λ i₁ → north ≡ south) (merid x1 ∙ refl)) j i))) ∣ x1 , originalPath ∣) ⟩ (transport (λ i → uncurry (CODE a n iscon) - (south , ((sym (transportRefl (transport (λ i → _≡_ {A = Susp A} north south) (merid x1 ∙ (λ _ → south))))) ∙ ( - transpRCancel {A = north ≡ south} {B = north ≡ south} (λ _ → _≡_ {A = Susp A} north south) + (south , ((sym (transportRefl (transport (λ i → _≡_ {A = Susp A} north south) (merid x1 ∙ (λ _ → south))))) ∙ + (transpRCancel {A = north ≡ south} {B = north ≡ south} (λ _ → _≡_ {A = Susp A} north south) (merid x1 ∙ (refl {x = south})))) i) )) ∣ x1 , originalPath ∣ ≡⟨ (λ j → (transport (λ i → uncurry (CODE a n iscon) @@ -1170,10 +1178,13 @@ codeTranspId6 {ℓ}{A = A} n a x1 iscon = transportLemma {B = uncurry (CODE a n (λ k → transportRefl (p ∙ (λ _ → north)) ∙ (rUnit (sym (transportRefl (p ∙ (λ _ → north)))) (~ k))) ∙ (rCancel (transportRefl (p ∙ (λ _ → north))) ) - finalStep : transport (cong (λ q → ∥ fiber merid q ∥ ℕ→ℕ₋₂ (n + n)) (sym (pairLemma ((merid x1) ∙ sym (merid a)) (merid a) ∙ sym (assocJ (merid x1) (sym (merid a)) (merid a)) ∙ (λ i → (merid x1) ∙ lCancel (merid a) i) ∙ sym (rUnit (merid x1))))) ∣ x1 , refl ∣ - ≡ ∣ x1 , rUnit (merid x1) ∙ sym (cong (λ x → merid x1 ∙ x) (lCancel (merid a))) ∙ - assocJ (merid x1) (sym (merid a)) (merid a) ∙ - sym (pairLemma {a1 = north} (merid x1 ∙ (sym (merid a))) (merid a)) ∣ + finalStep : transport (cong (λ q → ∥ fiber merid q ∥ ℕ→ℕ₋₂ (n + n)) (sym (pairLemma ((merid x1) ∙ sym (merid a)) (merid a) ∙ + sym (assocJ (merid x1) (sym (merid a)) (merid a)) ∙ + (λ i → (merid x1) ∙ lCancel (merid a) i) ∙ sym (rUnit (merid x1))))) + ∣ x1 , refl ∣ + ≡ ∣ x1 , rUnit (merid x1) ∙ sym (cong (λ x → merid x1 ∙ x) (lCancel (merid a))) ∙ + assocJ (merid x1) (sym (merid a)) (merid a) ∙ + sym (pairLemma {a1 = north} (merid x1 ∙ (sym (merid a))) (merid a)) ∣ finalStep = (λ k → transport (cong (λ q → ∥ fiber merid q ∥ ℕ→ℕ₋₂ (n + n)) (λ i → (sym (pairLemma ((merid x1) ∙ sym (merid a)) (merid a) ∙ sym (assocJ (merid x1) (sym (merid a)) (merid a)) ∙ diff --git a/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4Code.agda b/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4Code.agda index bf994ab59..1afa8f10d 100644 --- a/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4Code.agda +++ b/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4Code.agda @@ -35,30 +35,30 @@ private {- We first need to show that (a variant of) the canceller function defined in FreudenthalProof.Prelim is an equivalence -} -abstract - isEquivCancel : ∀ {ℓ} {A : Type ℓ}{a : A} (n : ℕ) (q : north ≡ south) → - isEquiv {A = ∥ fiber (λ y → σ a y) - (q ∙ sym (merid a)) ∥ ℕ→ℕ₋₂ (n + n)} - {B = ∥ fiber merid q ∥ ℕ→ℕ₋₂ (n + n)} - (elim (λ _ → isOfHLevelTrunc (2 + (n + n) )) - λ b → ∣ (fst b) , canceller (sym (merid a)) (merid (fst b)) q (snd b) ∣) - isEquivCancel {a = a} n q = isoToEquiv (iso - ((elim (λ _ → isOfHLevelTrunc (2 + (n + n))) λ b → ∣ (fst b) , canceller (sym (merid a)) (merid (fst b)) q (snd b) ∣)) - (elim (λ _ → isOfHLevelTrunc (2 + (n + n))) λ s → ∣ (fst s) , cancellerInv (λ i → (merid a) (~ i)) (merid (fst s)) q (snd s) ∣) - (λ b → elim {B = λ b → ((elim (λ _ → isOfHLevelTrunc (2 + (n + n))) - λ b → ∣ (fst b) , canceller (sym (merid a)) (merid (fst b)) q (snd b) ∣)) + +isEquivCancel : ∀ {ℓ} {A : Type ℓ}{a : A} (n : ℕ) (q : north ≡ south) → + isEquiv {A = ∥ fiber (λ y → σ a y) + (q ∙ sym (merid a)) ∥ ℕ→ℕ₋₂ (n + n)} + {B = ∥ fiber merid q ∥ ℕ→ℕ₋₂ (n + n)} + (elim (λ _ → isOfHLevelTrunc (2 + (n + n) )) + λ b → ∣ (fst b) , canceller (sym (merid a)) (merid (fst b)) q (snd b) ∣) +isEquivCancel {a = a} n q = isoToEquiv (iso + ((elim (λ _ → isOfHLevelTrunc (2 + (n + n))) λ b → ∣ (fst b) , canceller (sym (merid a)) (merid (fst b)) q (snd b) ∣)) + (elim (λ _ → isOfHLevelTrunc (2 + (n + n))) λ s → ∣ (fst s) , cancellerInv (λ i → (merid a) (~ i)) (merid (fst s)) q (snd s) ∣) + (λ b → elim {B = λ b → ((elim (λ _ → isOfHLevelTrunc (2 + (n + n))) + λ b → ∣ (fst b) , canceller (sym (merid a)) (merid (fst b)) q (snd b) ∣)) + ((elim (λ _ → isOfHLevelTrunc (2 + (n + n))) + λ s → ∣ (fst s) , cancellerInv (sym (merid a)) (merid (fst s)) q (snd s) ∣) b) + ≡ b} + (λ x → isOfHLevelSuc _ (isOfHLevelTrunc (2 + (n + n)) _ x)) + (λ b i → ∣ fst b , cancellerSection (sym (merid a)) (merid (fst b)) q (snd b) i ∣) b) + (λ b → elim {B = λ b → ((elim (λ _ → isOfHLevelTrunc (2 + (n + n))) + λ b → ∣ (fst b) , cancellerInv (sym (merid a)) (merid (fst b)) q (snd b) ∣)) ((elim (λ _ → isOfHLevelTrunc (2 + (n + n))) - λ s → ∣ (fst s) , cancellerInv (sym (merid a)) (merid (fst s)) q (snd s) ∣) b) - ≡ b} - (λ x → isOfHLevelSuc _ (isOfHLevelTrunc (2 + (n + n)) _ x)) - (λ b i → ∣ fst b , cancellerSection (sym (merid a)) (merid (fst b)) q (snd b) i ∣) b) - (λ b → elim {B = λ b → ((elim (λ _ → isOfHLevelTrunc (2 + (n + n))) - λ b → ∣ (fst b) , cancellerInv (sym (merid a)) (merid (fst b)) q (snd b) ∣)) - ((elim (λ _ → isOfHLevelTrunc (2 + (n + n))) - λ s → ∣ (fst s) , canceller (sym (merid a)) (merid (fst s)) q (snd s) ∣) b) - ≡ b} - (λ x → isOfHLevelSuc (suc (n + n)) (isOfHLevelTrunc (2 + (n + n)) _ x)) - (λ b i → ∣ fst b , cancellerRetract (sym (merid a)) (merid (fst b)) q (snd b) i ∣) b)) .snd + λ s → ∣ (fst s) , canceller (sym (merid a)) (merid (fst s)) q (snd s) ∣) b) + ≡ b} + (λ x → isOfHLevelSuc (suc (n + n)) (isOfHLevelTrunc (2 + (n + n)) _ x)) + (λ b i → ∣ fst b , cancellerRetract (sym (merid a)) (merid (fst b)) q (snd b) i ∣) b)) .snd {- CODE will be defined by means of univalence applied to an equivalence ∥ fiber (λ y → σ a y) (q ∙ sym (merid c)) ∥ (ℕ→ℕ₋₂ (n + n)) ≃ ∥ fiber merid q ∥ (ℕ→ℕ₋₂ (n + n)). @@ -70,7 +70,7 @@ sufMap : (n : ℕ) (c a x₂ : A) ≡ (q ∙ sym (merid c)) → ∥ Σ A (λ x → merid x ≡ q) ∥ (ℕ→ℕ₋₂ (n + n)) sufMap {A = A} n c a x₂ iscon q = Lemma8-6-2 (A , a) (A , a) n n iscon iscon (λ x₂ c → ((merid x₂ ∙ sym (merid a)) ≡ (q ∙ sym (merid c)) → ∥ Σ A (λ x → merid x ≡ q) ∥ (ℕ→ℕ₋₂ (n + n))) , - isOfHLevelPi (2+ ((ℕ→ℕ₋₂ (n + n)))) λ _ → isOfHLevelTrunc (2 + (n + n))) + isOfHLevelΠ (2+ ((ℕ→ℕ₋₂ (n + n)))) λ _ → isOfHLevelTrunc (2 + (n + n))) (λ x r → ∣ x , canceller (sym (merid a)) (merid x) q r ∣) (λ x r → ∣ x , switcher (merid a) q (merid x) r ∣) (funExt (λ x i → ∣ (refl i) , ((switcherCancellerIdGeneral (merid a) q @@ -96,7 +96,7 @@ RlFunId {A = A} a n iscon q b = cong (λ x → x (snd b)) (proj₁ (((Lemma8-6-2 (A , a) (A , a) n n iscon iscon (λ x₂ c → (((merid x₂ ∙ sym (merid a)) ≡ (q ∙ sym (merid c)) → ∥ Σ A (λ x → merid x ≡ q) ∥ (ℕ→ℕ₋₂ (n + n)) ) , - isOfHLevelPi (2+ ((ℕ→ℕ₋₂ (n + n)))) λ _ → isOfHLevelTrunc (2 + (n + n)))) + isOfHLevelΠ (2+ ((ℕ→ℕ₋₂ (n + n)))) λ _ → isOfHLevelTrunc (2 + (n + n)))) (λ x r → ∣ x , canceller (sym (merid a)) (merid x) q r ∣) (λ b s → ∣ b , switcher (merid a) q (merid b) s ∣) (funExt (λ x → λ j → ∣ (refl j) , (switcherCancellerIdGeneral (merid a) q @@ -145,43 +145,44 @@ RlFunEquiv {A = A} a c n iscon q = fst (conInd-i→iii {A = Unit} {B = A} (λ x natHelper (suc n) = cong (λ x → suc x) (natHelper n) {- We will also need the following function to get things on the right form -} -abstract - helperFun : ∀ {ℓ ℓ'} {X : Type ℓ} {x y : X} (p : x ≡ y) {A : x ≡ x → Type ℓ'} {B : x ≡ y → Type ℓ'} → - ((q : x ≡ y) → A (q ∙ (sym p)) ≡ B q) → - (q : x ≡ x) → - transport (λ i₁ → Type ℓ') (A q) - ≡ B (transport (λ i → x ≡ p i) q ) - helperFun {ℓ' = ℓ'} {x = x} = J (λ y p → {A : x ≡ x → Type ℓ'} {B : x ≡ y → Type ℓ'} → - ((q : x ≡ y) → A (q ∙ (sym p)) ≡ B q) → - (q : x ≡ x) → - transport (λ i₁ → Type ℓ') (A q) - ≡ B (transport (λ i → x ≡ p i) q ) ) - λ {A} {B} k q → transportRefl (A q) ∙ - cong (λ x → A x) (rUnit q) ∙ - k q ∙ - cong (λ x → B x) (sym (transportRefl q)) - - helperFunRefl : ∀ {ℓ ℓ'} {X : Type ℓ} {x : X} - {A : x ≡ x → Type ℓ'} - {B : x ≡ x → Type ℓ'} → - helperFun {X = X} (refl {x = x}) {A = A} {B = B} - ≡ λ k q → transportRefl (A q) ∙ - cong (λ x → A x) (rUnit q) ∙ - k q ∙ - cong (λ x → B x) (sym (transportRefl q)) - helperFunRefl {ℓ' = ℓ'} {x = x} {A = A} {B = B} = cong (λ x → x {A} {B}) - (JRefl (λ y p → {A : x ≡ x → Type ℓ'} {B : x ≡ y → Type ℓ'} → - ((q : x ≡ y) → A (q ∙ (sym p)) ≡ B q) → - (q : x ≡ x) → - transport (λ i₁ → Type ℓ') (A q) - ≡ B (transport (λ i → x ≡ p i) q ) ) - λ {A} {B} k q → transportRefl (A q) ∙ - cong (λ x → A x) (rUnit q) ∙ - k q ∙ - cong (λ x → B x) (sym (transportRefl q))) - - {- For now things do not compute properly, so we use an abstract version of univalence-} +helperFun : ∀ {ℓ ℓ'} {X : Type ℓ} {x y : X} (p : x ≡ y) {A : x ≡ x → Type ℓ'} {B : x ≡ y → Type ℓ'} → + ((q : x ≡ y) → A (q ∙ (sym p)) ≡ B q) → + (q : x ≡ x) → + transport (λ i₁ → Type ℓ') (A q) + ≡ B (transport (λ i → x ≡ p i) q ) +helperFun {ℓ' = ℓ'} {x = x} = J (λ y p → {A : x ≡ x → Type ℓ'} {B : x ≡ y → Type ℓ'} → + ((q : x ≡ y) → A (q ∙ (sym p)) ≡ B q) → + (q : x ≡ x) → + transport (λ i₁ → Type ℓ') (A q) + ≡ B (transport (λ i → x ≡ p i) q ) ) + λ {A} {B} k q → transportRefl (A q) ∙ + cong (λ x → A x) (rUnit q) ∙ + k q ∙ + cong (λ x → B x) (sym (transportRefl q)) + +helperFunRefl : ∀ {ℓ ℓ'} {X : Type ℓ} {x : X} + {A : x ≡ x → Type ℓ'} + {B : x ≡ x → Type ℓ'} → + helperFun {X = X} (refl {x = x}) {A = A} {B = B} + ≡ λ k q → transportRefl (A q) ∙ + cong (λ x → A x) (rUnit q) ∙ + k q ∙ + cong (λ x → B x) (sym (transportRefl q)) +helperFunRefl {ℓ' = ℓ'} {x = x} {A = A} {B = B} = cong (λ x → x {A} {B}) + (JRefl (λ y p → {A : x ≡ x → Type ℓ'} {B : x ≡ y → Type ℓ'} → + ((q : x ≡ y) → A (q ∙ (sym p)) ≡ B q) → + (q : x ≡ x) → + transport (λ i₁ → Type ℓ') (A q) + ≡ B (transport (λ i → x ≡ p i) q ) ) + λ {A} {B} k q → transportRefl (A q) ∙ + cong (λ x → A x) (rUnit q) ∙ + k q ∙ + cong (λ x → B x) (sym (transportRefl q))) + +{- For now things do not compute properly, so we use an abstract version of univalence-} + +abstract ua' : ∀ {A B : Type ℓ} → A ≃ B → A ≡ B ua' = ua @@ -211,7 +212,7 @@ sufMapId {A = A} n a x1 iscon = (proj₂ (Lemma8-6-2 (A , a) (A , a) n n iscon i (λ x₂ c → ((merid x₂ ∙ sym (merid a)) ≡ ((merid x1) ∙ sym (merid c)) → ∥ Σ A (λ x → merid x ≡ (merid x1)) ∥ (ℕ→ℕ₋₂ (n + n))) , - isOfHLevelPi (2+ ((ℕ→ℕ₋₂ (n + n)))) λ _ → isOfHLevelTrunc (2 + (n + n))) + isOfHLevelΠ (2+ ((ℕ→ℕ₋₂ (n + n)))) λ _ → isOfHLevelTrunc (2 + (n + n))) (λ x r → ∣ x , canceller (sym (merid a)) (merid x) (merid x1) r ∣) (λ x r → ∣ x , switcher (merid a) (merid x1) (merid x) r ∣) (funExt (λ x i → ∣ (refl i) , ((switcherCancellerIdGeneral (merid a) (merid x1) @@ -224,7 +225,7 @@ sufMapId2 : (n : ℕ) (a x1 : A) sufMapId2 {A = A} n a x1 iscon i = (proj₁ (Lemma8-6-2 (A , a) (A , a) n n iscon iscon (λ x₂ c → ((merid x₂ ∙ sym (merid a)) ≡ ((merid x1) ∙ sym (merid c)) → ∥ Σ A (λ x → merid x ≡ (merid x1)) ∥ (ℕ→ℕ₋₂ (n + n))) , - isOfHLevelPi (2+ ((ℕ→ℕ₋₂ (n + n)))) λ _ → isOfHLevelTrunc (2 + (n + n))) + isOfHLevelΠ (2+ ((ℕ→ℕ₋₂ (n + n)))) λ _ → isOfHLevelTrunc (2 + (n + n))) (λ x r → ∣ x , canceller (sym (merid a)) (merid x) (merid x1) r ∣) (λ x r → ∣ x , switcher (merid a) (merid x1) (merid x) r ∣) (funExt (λ x i → ∣ (refl i) , ((switcherCancellerIdGeneral (merid a) (merid x1) diff --git a/Cubical/HITs/Truncation/Connected/Properties.agda b/Cubical/HITs/Truncation/Connected/Properties.agda index 95c1be4f8..8b7d16f09 100644 --- a/Cubical/HITs/Truncation/Connected/Properties.agda +++ b/Cubical/HITs/Truncation/Connected/Properties.agda @@ -13,8 +13,10 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.Equiv.Properties open import Cubical.Foundations.Function open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Transport open import Cubical.Foundations.Pointed open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Isomorphism open import Cubical.Data.NatMinusTwo.Base open import Cubical.Data.Prod.Base open import Cubical.HITs.Susp @@ -22,6 +24,9 @@ open import Cubical.Data.Nat open import Cubical.HITs.Truncation.Base open import Cubical.HITs.Truncation.Properties renaming (elim to trElim) open import Cubical.HITs.Nullification +open import Cubical.HITs.Pushout.Base +open import Cubical.HITs.Sn.Base +open import Cubical.Data.Unit.Properties open import Cubical.Data.HomotopyGroup @@ -112,3 +117,66 @@ abstract ------------------------------ + +inrConnected : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} (n : ℕ₋₂) + (f : C → A) (g : C → B) → + is- n -Connected f → + is-_-Connected {A = B} {B = Pushout f g} n inr +inrConnected {A = A} {B = B} {C = C} n f g iscon = elim.3→1 inr n λ P → (λ k → helpLemmas.k P k) , λ b → refl + where + module helpLemmas {ℓ : Level} (P : (Pushout f g) → HLevel ℓ (2+ n)) + (h : (b : B) → typ (P (inr b))) + where + Q : A → HLevel _ (2+ n) + Q a = (P (inl a)) + + fun : (c : C) → typ (Q (f c)) + fun c = transport (λ i → typ (P (push c (~ i)))) (h (g c)) + + k : (d : Pushout f g) → typ (P d) + k (inl x) = elim.2→3 f n Q (elim.1→2 f n iscon Q) .fst fun x + k (inr x) = h x + k (push a i) = hcomp (λ k → λ{(i = i0) → elim.2→3 f n Q + (elim.1→2 f n iscon Q) .fst + fun (f a) ; + (i = i1) → transportTransport⁻ (λ j → typ (P (push a j))) (h (g a)) k}) + (transp (λ j → typ (P (push a (i ∧ j)))) + (~ i) + (elim.2→3 f n Q + (elim.1→2 f n iscon Q) .snd fun i a)) + + +sphereConnected : (n : ℕ) → is- (-1+ n) -ConnectedType (S₊ n) +sphereConnected zero = ∣ north ∣ , (isOfHLevelTrunc 1 ∣ north ∣) +sphereConnected (suc n) = transport (λ i → is- ℕ→ℕ₋₂ n -ConnectedType (Susp≡Push {A = S₊ n} (~ i))) + (trivFunCon← {A = Pushout {A = S₊ n} (λ x → tt) λ x → tt} {a = inr tt } _ + (transport (λ i → is- (-1+ n) -Connected (mapsAgree (~ i))) + (inrConnected _ (λ x → tt) (λ x → tt) λ{tt → transport (λ i → isContr (∥ fibUnit (~ i) ∥ (-1+ n))) (sphereConnected n)}))) + where + mapsAgree : Path ((x : Unit) → Pushout {A = S₊ n} (λ x → tt) (λ x → tt)) (λ (x : Unit) → inr tt) inr + mapsAgree = funExt λ {tt → refl} + + fibUnit : fiber (λ (x : S₊ n) → tt) tt ≡ S₊ n + fibUnit = isoToPath (iso (λ b → fst b) (λ a → a , refl) (λ b → refl) λ b i → (fst b) , isProp→isSet isPropUnit tt tt refl (snd b) i) + + Susp≡Push : Susp A ≡ Pushout {A = A} (λ a → tt) λ a → tt + Susp≡Push {A = A} = isoToPath (iso fun inverse sect retr) + where + fun : Susp A → Pushout {A = A} (λ a → tt) (λ a → tt) + fun north = inl tt + fun south = inr tt + fun (merid a i) = push a i + inverse : Pushout {A = A} (λ a → tt) (λ a → tt) → Susp A + inverse (inl tt) = north + inverse (inr tt) = south + inverse (push a i) = merid a i + + sect : section fun inverse + sect (inl tt) = refl + sect (inr tt) = refl + sect (push a i) = refl + + retr : retract fun inverse + retr north = refl + retr south = refl + retr (merid a i) = refl diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index bee12fee4..81226b7e4 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -2,14 +2,15 @@ module Cubical.ZCohomology.Properties where open import Cubical.ZCohomology.Base +open import Cubical.HITs.S1 open import Cubical.HITs.Sn open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Everything open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence open import Cubical.Data.NatMinusTwo.Base open import Cubical.Data.Empty open import Cubical.Data.Sigma @@ -17,14 +18,17 @@ open import Cubical.Data.Prod.Base open import Cubical.HITs.Susp open import Cubical.HITs.SetTruncation open import Cubical.HITs.Nullification -open import Cubical.Data.Int +open import Cubical.Data.Int hiding (_+_) open import Cubical.Data.Nat -open import Cubical.HITs.Truncation +open import Cubical.HITs.Truncation renaming (elim to trElim) open import Cubical.HITs.Pushout open import Cubical.Data.Sum.Base open import Cubical.Data.HomotopyGroup open import Cubical.Data.Bool + +open import Cubical.ZCohomology.cupProdPrelims renaming (Kn→ΩKn+1 to Kn→ΩKn+1') + private variable ℓ ℓ' : Level @@ -60,3 +64,39 @@ coHomRed+1Equiv zero A i = ∥ helpLemma {C = (Int , pos 0)} i ∥₀ helper (inr tt) = sym snd coHomRed+1Equiv (suc n) A i = ∥ coHomRed+1.helpLemma A i {C = ((coHomK (suc n)) , ∣ north ∣)} i ∥₀ + + +------------------- +{- This section contains a proof of Kn≃ΩKn+1, which is needed for the cup product -} + +{- Compiles slowly right now. Possible improvements : +1. Use S¹ instead of S 1 in definition of K. +2. Use copatterns -} + +Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) +Kn→ΩKn+1 = Kn→ΩKn+1' + +isEquivKn→ΩKn+1 : (n : ℕ) → isEquiv (Kn→ΩKn+1 n) +isEquivKn→ΩKn+1 zero = compEquiv (looper , isEquivLooper) (cong ∣_∣ , isEquivHelper hLevl3) .snd + where + hLevl3 : (x y : S₊ 1) (p q : x ≡ y) → isProp (p ≡ q) + hLevl3 x y = J (λ y p → (q : x ≡ y) → isProp (p ≡ q) ) + (transport (λ i → isSet (helper (~ i))) isSetInt refl) + where + helper : (x ≡ x) ≡ Int + helper = (λ i → transp (λ j → ua S1≃SuspBool (~ j ∨ ~ i)) (~ i) x ≡ transp (λ j → ua S1≃SuspBool (~ j ∨ ~ i)) (~ i) x) ∙ + (λ i → transp (λ j → S¹≡SuspBool (~ j ∨ ~ i)) (~ i) (transport (sym (ua S1≃SuspBool)) x) ≡ transp (λ j → S¹≡SuspBool (~ j ∨ ~ i)) (~ i) (transport (sym (ua S1≃SuspBool)) x)) ∙ + basedΩS¹≡Int (transport (sym S¹≡SuspBool) (transport (sym (ua S1≃SuspBool)) x)) +isEquivKn→ΩKn+1 (suc zero) = transport (λ i → isEquiv (Kn→ΩKn+1Sucn zero (~ i))) + (compEquiv (trElim (λ _ → isOfHLevelTrunc (2 + (suc zero))) (λ a → ∣ ϕ north a ∣) , + isEquiv∣ϕ∣) + (truncEquivΩ (ℕ→ℕ₋₂ (suc zero))) .snd) +isEquivKn→ΩKn+1 (suc (suc n)) = transport (λ i → isEquiv (Kn→ΩKn+1Sucn (suc n) (~ i))) + (compEquiv (conEquiv3 (4 + n) _ (ϕ north) (n , λ i → suc (suc (suc (+-suc n n (~ i))))) (FthalFun-2nConnected (suc n) _ (sphereConnected _))) + (truncEquivΩ (ℕ→ℕ₋₂ (suc (suc n)))) .snd) + +Kn≃ΩKn+1 : (n : ℕ) → coHomK n ≃ typ (Ω (coHomK-ptd (suc n))) +Kn≃ΩKn+1 n = Kn→ΩKn+1 n , isEquivKn→ΩKn+1 n + +ΩKn+1→Kn : (n : ℕ) → typ (Ω (coHomK-ptd (suc n))) → coHomK n +ΩKn+1→Kn n a = equiv-proof (isEquivKn→ΩKn+1 n) a .fst .fst diff --git a/Cubical/ZCohomology/S1Loop.agda b/Cubical/ZCohomology/cupProdPrelims.agda similarity index 67% rename from Cubical/ZCohomology/S1Loop.agda rename to Cubical/ZCohomology/cupProdPrelims.agda index 2b9921db9..2adde3478 100644 --- a/Cubical/ZCohomology/S1Loop.agda +++ b/Cubical/ZCohomology/cupProdPrelims.agda @@ -1,5 +1,5 @@ {-# OPTIONS --cubical --safe #-} -module Cubical.ZCohomology.S1Loop where +module Cubical.ZCohomology.cupProdPrelims where open import Cubical.ZCohomology.Base open import Cubical.HITs.S1 @@ -8,19 +8,19 @@ open import Cubical.HITs.S3 open import Cubical.HITs.Sn open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Everything open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed open import Cubical.Foundations.Transport open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.HAEquiv open import Cubical.Data.NatMinusTwo.Base open import Cubical.Data.Empty open import Cubical.Data.Sigma open import Cubical.Data.Prod.Base open import Cubical.HITs.Susp open import Cubical.HITs.Nullification -open import Cubical.Data.Unit open import Cubical.Data.Int renaming (_+_ to +Int) open import Cubical.Data.Nat open import Cubical.HITs.Truncation renaming (elim to trElim) @@ -38,88 +38,7 @@ private A : Type ℓ B : Type ℓ' - -compToIdEquiv : (f : A → B) (g : B → A) → f ∘ g ≡ idfun B → isEquiv f → isEquiv g -compToIdEquiv f g id iseqf = - isoToIsEquiv (iso g - f - (λ b → (λ i → (equiv-proof iseqf (f b) .snd (g (f b) , cong (λ h → h (f b)) id) (~ i)) .fst ) ∙ - cong (λ x → (equiv-proof iseqf (f b) .fst .fst )) id ∙ - λ i → (equiv-proof iseqf (f b) .snd) (b , refl) i .fst ) - λ a → cong (λ f → f a) id) - -Susp≡Push : Susp A ≡ Pushout {A = A} (λ a → tt) λ a → tt -Susp≡Push {A = A} = isoToPath (iso fun inverse sect retr) - where - fun : Susp A → Pushout {A = A} (λ a → tt) (λ a → tt) - fun north = inl tt - fun south = inr tt - fun (merid a i) = push a i - inverse : Pushout {A = A} (λ a → tt) (λ a → tt) → Susp A - inverse (inl tt) = north - inverse (inr tt) = south - inverse (push a i) = merid a i - - sect : section fun inverse - sect (inl tt) = refl - sect (inr tt) = refl - sect (push a i) = refl - - retr : retract fun inverse - retr north = refl - retr south = refl - retr (merid a i) = refl - -Pr2310 : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} (n : ℕ₋₂) - (f : C → A) (g : C → B) → - is- n -Connected f → - is-_-Connected {A = B} {B = Pushout f g} n inr -Pr2310 {A = A} {B = B} {C = C} n f g iscon = elim.3→1 inr n λ P → (λ k → helpLemmas.k P k) , λ b → refl - where - module helpLemmas {ℓ : Level} (P : (Pushout f g) → HLevel ℓ (2+ n)) - (h : (b : B) → typ (P (inr b))) - where - Q : A → HLevel _ (2+ n) - Q a = (P (inl a)) - - fun : (c : C) → typ (Q (f c)) - fun c = transport (λ i → typ (P (push c (~ i)))) (h (g c)) - - k : (d : Pushout f g) → typ (P d) - k (inl x) = elim.2→3 f n Q (elim.1→2 f n iscon Q) .fst fun x - k (inr x) = h x - k (push a i) = hcomp (λ k → λ{(i = i0) → elim.2→3 f n Q - (elim.1→2 f n iscon Q) .fst - fun (f a) ; - (i = i1) → transportTransport⁻ (λ j → typ (P (push a j))) (h (g a)) k}) - (transp (λ j → typ (P (push a (i ∧ j)))) - (~ i) - (elim.2→3 f n Q - (elim.1→2 f n iscon Q) .snd fun i a)) - - -Pr242 : (n : ℕ) → is- (-1+ n) -ConnectedType (S₊ n) -Pr242 zero = ∣ north ∣ , (isOfHLevelTrunc 1 ∣ north ∣) -Pr242 (suc n) = transport (λ i → is- ℕ→ℕ₋₂ n -ConnectedType (Susp≡Push {A = S₊ n} (~ i))) - (trivFunCon← {A = Pushout {A = S₊ n} (λ x → tt) λ x → tt} {a = inr tt } _ - (transport (λ i → is- (-1+ n) -Connected (mapsAgree (~ i))) - (Pr2310 _ (λ x → tt) (λ x → tt) λ{tt → transport (λ i → isContr (∥ fibUnit (~ i) ∥ (-1+ n))) (Pr242 n)}))) - where - mapsAgree : Path ((x : Unit) → Pushout {A = S₊ n} (λ x → tt) (λ x → tt)) (λ (x : Unit) → inr tt) inr - mapsAgree = funExt λ {tt → refl} - fibUnit : fiber (λ (x : S₊ n) → tt) tt ≡ S₊ n - fibUnit = isoToPath (iso (λ b → fst b) (λ a → a , refl) (λ b → refl) λ b i → (fst b) , isProp→isSet isPropUnit tt tt refl (snd b) i) - - - -- need to show susp and pushout are same here - -Pr242SpecCase : is- 2 -ConnectedType (Susp (Susp S¹)) -Pr242SpecCase = transport (λ i → is- 2 -ConnectedType (helper i)) (Pr242 3) - where - helper : S₊ 3 ≡ Susp (Susp S¹) - helper = (λ i → Susp (Susp (Susp (ua Bool≃Susp⊥ (~ i))))) ∙ λ i → Susp (Susp (S¹≡SuspBool (~ i))) - - +{- Some useful lemmas -- should probably be moved -} congFunct : {a b c : A} (f : A → B) (p : a ≡ b) (q : b ≡ c) → cong f (p ∙ q) ≡ cong f p ∙ cong f q congFunct f p q i = hcomp (λ j → λ{(i = i0) → rUnit (cong f (p ∙ q)) (~ j) ; (i = i1) → cong f (rUnit p (~ j)) ∙ cong f q}) @@ -130,16 +49,16 @@ symDistr p q i = hcomp (λ j → λ{(i = i0) → rUnit (sym (p ∙ q)) (~ j) ; (i = i1) → sym (lUnit q (~ j)) ∙ sym p}) (sym ((λ k → p (k ∨ i)) ∙ q) ∙ sym λ k → p (i ∧ k)) -{- We want to prove that Kn≃ΩKn+1. For this we need the map ϕ-} +{- We want to prove that Kn≃ΩKn+1. For this we use the map ϕ-} -private - ϕ : (pt a : A) → typ (Ω (Susp A , north)) - ϕ pt a = (merid a) ∙ sym (merid pt) + +ϕ : (pt a : A) → typ (Ω (Susp A , north)) +ϕ pt a = (merid a) ∙ sym (merid pt) {- To define the map for n=0 we use the λ k → loopᵏ map for S₊ 1. The loop is given by ϕ south north -} -looper : Int → _≡_ {A = S₊ 1} north north +looper : Int → Path (S₊ 1) north north looper (pos zero) = refl looper (pos (suc n)) = looper (pos n) ∙ (ϕ south north) looper (negsuc zero) = sym (ϕ south north) @@ -151,28 +70,30 @@ Kn→ΩKn+1 zero x = cong ∣_∣ (looper x) Kn→ΩKn+1 (suc n) = trElim (λ x → (isOfHLevelTrunc (2 + (suc (suc n))) ∣ north ∣ ∣ north ∣)) λ a → cong ∣_∣ ((merid a) ∙ (sym (merid north))) +{- We want to show that this map is an equivalence. n ≥ 2 follows from Freudenthal, and -} + {- We want to show that the function (looper : Int → S₊ 1) defined by λ k → loopᵏ is an equivalece. We already know that the corresponding function (intLoop : Int → S¹ is) an equivalence, so the idea is to show that when intLoop is transported along a suitable path S₊ 1 ≡ S¹ we get looper. Instead of using S₊ 1 straight away, we begin by showing this for the equivalent Susp Bool. -} -- loop for Susp Bool -- -loop* : _≡_ {A = Susp Bool} north north +loop* : Path (Susp Bool) north north loop* = merid false ∙ sym (merid true) -- the loop function -- -intLoop* : Int → _≡_ {A = Susp Bool} north north +intLoop* : Int → Path (Susp Bool) north north intLoop* (pos zero) = refl intLoop* (pos (suc n)) = intLoop* (pos n) ∙ loop* intLoop* (negsuc zero) = sym loop* intLoop* (negsuc (suc n)) = intLoop* (negsuc n) ∙ sym loop* -- we show that the loop spaces agree -- -loopSpId : ΩS¹ ≡ _≡_ {A = Susp Bool} north north +loopSpId : ΩS¹ ≡ Path (Susp Bool) north north loopSpId i = typ (Ω (S¹≡SuspBool i , transp ((λ j → S¹≡SuspBool (j ∧ i))) (~ i) base)) -- the transport map -- -altMap2 : Int → _≡_ {A = Susp Bool} north north +altMap2 : Int → Path (Susp Bool) north north altMap2 n i = transport S¹≡SuspBool (intLoop n i) -- We show that the transporting intLoop over S¹≡SuspBool gives intLoop* (modulo function extensionality) -- @@ -194,7 +115,7 @@ altMap≡intLoop*2 (pos (suc n)) = (λ i → (altMap≡intLoop*2 (pos n) i) ∙ anotherHelper : (n : ℕ) → altMap2 (pos (suc (suc n))) ≡ altMap2 (pos (suc n)) ∙ altMap2 (pos 1) anotherHelper n = ((λ i j → transport S¹≡SuspBool ((intLoop (pos (suc n)) ∙ loop) j))) ∙ rUnit (λ j → transport S¹≡SuspBool ((intLoop (pos (suc n)) ∙ loop) j) ) ∙ - (λ i → (λ j → transport S¹≡SuspBool ((intLoop (pos (suc n)) ∙ λ k → loop (k ∧ (~ i))) j)) ∙ λ j → transport S¹≡SuspBool (loop (j ∨ (~ i))) ) ∙ + (λ i → (λ j → transport S¹≡SuspBool ((intLoop (pos (suc n)) ∙ λ k → loop (k ∧ (~ i))) j)) ∙ λ j → transport S¹≡SuspBool (loop (j ∨ (~ i)))) ∙ (λ i → (λ j → transport S¹≡SuspBool (rUnit (intLoop (pos (suc n))) (~ i) j)) ∙ λ j → transport S¹≡SuspBool ((lUnit loop i) j)) altMap≡intLoop*2 (negsuc zero) = sym ((λ i j → transport S¹≡SuspBool (loop (~ j))) ∙ @@ -208,7 +129,7 @@ altMap≡intLoop*2 (negsuc (suc n)) = helper n anotherHelper n = ((λ i → rUnit (λ j → (transport S¹≡SuspBool ((intLoop (negsuc n) ∙ sym loop) j))) i)) ∙ ((λ i → (λ j → transport S¹≡SuspBool ((intLoop (negsuc n) ∙ (λ k → loop ((~ k) ∨ i))) j)) ∙ λ j → transport S¹≡SuspBool (loop ((~ j) ∧ i)))) ∙ (λ i → ((λ j → transport S¹≡SuspBool (rUnit (intLoop (negsuc n)) (~ i) j))) ∙ altMap2 (negsuc zero)) - + helper : (n : ℕ) → intLoop* (negsuc n) ∙ (sym loop*) ≡ altMap2 (negsuc (suc n)) helper zero = (λ i → altMapneg1 (~ i) ∙ altMapneg1 (~ i)) ∙ sym (anotherHelper zero) helper (suc n) = (λ i → (helper n i) ∙ altMapneg1 (~ i) ) ∙ @@ -248,8 +169,6 @@ S1→SuspBool (merid south i) = merid true i S1≃SuspBool : Susp Bool ≃ S₊ 1 S1≃SuspBool = isoToEquiv (iso SuspBool→S1 S1→SuspBool retrHelper sectHelper) where - - sectHelper : section S1→SuspBool SuspBool→S1 sectHelper north = refl sectHelper south = refl @@ -293,10 +212,26 @@ isEquivLooper = transport (λ i → isEquiv (funExt (looperIntoBool) (~ i))) isE congEquiv (transport (ua S1≃SuspBool) , transportEquiv (ua S1≃SuspBool) .snd) .snd) .snd -isSetS1 : isSet (_≡_ {A = S₊ 1} north north) +----------------------------------- n = 1 ----------------------------------------------------- + +{- We begin by stating some useful lemmas -} + +sphereConnectedSpecCase : is- 2 -ConnectedType (Susp (Susp S¹)) +sphereConnectedSpecCase = transport (λ i → is- 2 -ConnectedType (helper i)) (sphereConnected 3) + where + helper : S₊ 3 ≡ Susp (Susp S¹) + helper = (λ i → Susp (Susp (Susp (ua Bool≃Susp⊥ (~ i))))) ∙ λ i → Susp (Susp (S¹≡SuspBool (~ i))) + +S¹≡S1 : S₊ 1 ≡ S¹ +S¹≡S1 = (λ i → Susp (ua (Bool≃Susp⊥) (~ i))) ∙ sym S¹≡SuspBool + +S³≡SuspSuspS¹ : S³ ≡ Susp (Susp S¹) +S³≡SuspSuspS¹ = S³≡SuspS² ∙ λ i → Susp (S²≡SuspS¹ i) + +isSetS1 : isSet (Path (S₊ 1) north north) isSetS1 = transport (λ i → isSet (helper i)) isSetInt where - helper : Int ≡ (_≡_ {A = S₊ 1} north north) + helper : Int ≡ (Path (S₊ 1) north north) helper = sym ΩS¹≡Int ∙ (λ i → typ (Ω (S¹≡SuspBool i , transport (λ j → S¹≡SuspBool (j ∧ i)) base))) ∙ (λ i → typ (Ω (ua S1≃SuspBool i , transport (λ j → ua S1≃SuspBool (i ∧ j)) north))) @@ -305,12 +240,16 @@ isEquivHelper2 : isOfHLevel 3 A → isEquiv {B = ∥ A ∥ 1} ∣_∣ isEquivHelper2 ofHlevl = isoToIsEquiv (iso ∣_∣ (trElim (λ _ → ofHlevl) (λ a → a)) - (trElim {B = λ b → ∣ trElim (λ _ → ofHlevl) (λ a₁ → a₁) b ∣ ≡ b} (λ b → isOfHLevelSuc 3 (isOfHLevelTrunc 3) ∣ trElim (λ _ → ofHlevl) (λ a₁ → a₁) b ∣ b) λ a → refl) + (trElim {B = λ b → ∣ trElim (λ _ → ofHlevl) (λ a₁ → a₁) b ∣ ≡ b} + (λ b → isOfHLevelSuc 3 (isOfHLevelTrunc 3) + ∣ trElim (λ _ → ofHlevl) (λ a₁ → a₁) b ∣ b) + λ a → refl) λ b → refl) -isEquivHelper : {a b : A} → isOfHLevel 3 A → isEquiv {B = _≡_ {A = ∥ A ∥ 1} ∣ a ∣ ∣ b ∣ } (cong ∣_∣) +isEquivHelper : {a b : A} → isOfHLevel 3 A → isEquiv {B = Path (∥ A ∥ 1) ∣ a ∣ ∣ b ∣ } (cong ∣_∣) isEquivHelper {A = A} {a = a} {b = b} ofHlevl = congEquiv (∣_∣ , isEquivHelper2 ofHlevl) .snd +{- We give the following map and show that it is an equivalence -} d-map : typ (Ω ((Susp S¹) , north)) → S¹ d-map p = subst HopfSuspS¹ p base @@ -323,13 +262,10 @@ d-mapId r = substComposite HopfSuspS¹ (merid r) (sym (merid base)) base ∙ rotLemma base = refl rotLemma (loop i) = refl -S³≡SuspSuspS¹ : S³ ≡ Susp (Susp S¹) -S³≡SuspSuspS¹ = S³≡SuspS² ∙ λ i → Susp (S²≡SuspS¹ i) - -d-mapComp : fiber d-map base ≡ (_≡_ {A = Susp (Susp S¹)} north north) +d-mapComp : fiber d-map base ≡ Path (Susp (Susp S¹)) north north d-mapComp = sym (pathSigma≡sigmaPath {B = HopfSuspS¹} _ _) ∙ helper where - helper : (_≡_ {A = Σ (Susp S¹) λ x → HopfSuspS¹ x} (north , base) (north , base)) ≡ (_≡_ {A = Susp (Susp S¹)} north north) + helper : Path (Σ (Susp S¹) λ x → HopfSuspS¹ x) (north , base) (north , base) ≡ Path (Susp (Susp S¹)) north north helper = (λ i → (_≡_ {A = S³≡TotalHopf (~ i)} (transp (λ j → S³≡TotalHopf (~ i ∨ ~ j)) (~ i) (north , base)) ((transp (λ j → S³≡TotalHopf (~ i ∨ ~ j)) (~ i) (north , base))))) ∙ @@ -340,7 +276,7 @@ d-mapComp = sym (pathSigma≡sigmaPath {B = HopfSuspS¹} _ _) ∙ helper is1Connected-dmap : is- 1 -Connected d-map is1Connected-dmap base = transport (λ j → isContr (∥ d-mapComp (~ j) ∥ ℕ→ℕ₋₂ 1)) (transport (λ i → isContr (PathΩ {A = Susp (Susp S¹)} {a = north} (ℕ→ℕ₋₂ 1) i)) - (refl , isOfHLevelSuc 1 (isOfHLevelSuc 0 Pr242SpecCase) ∣ north ∣ ∣ north ∣ (λ _ → ∣ north ∣))) + (refl , isOfHLevelSuc 1 (isOfHLevelSuc 0 sphereConnectedSpecCase) ∣ north ∣ ∣ north ∣ (λ _ → ∣ north ∣))) is1Connected-dmap (loop j) = hcomp (λ k → λ{(j = i0) → is1Connected-dmap base ; (j = i1) → isPropIsOfHLevel 0 (transport (λ i → isContr (∥ fiber d-map (loop i) ∥ ℕ→ℕ₋₂ 1)) @@ -350,17 +286,25 @@ is1Connected-dmap (loop j) = (transp (λ i → isContr (∥ fiber d-map (loop (i ∧ j)) ∥ ℕ→ℕ₋₂ 1)) (~ j) (transport (λ j → isContr (∥ d-mapComp (~ j) ∥ ℕ→ℕ₋₂ 1)) (transport (λ i → isContr (PathΩ {A = Susp (Susp S¹)} {a = north} (ℕ→ℕ₋₂ 1) i)) - (refl , isOfHLevelSuc 1 (isOfHLevelSuc 0 Pr242SpecCase) ∣ north ∣ ∣ north ∣ (λ _ → ∣ north ∣))))) + (refl , isOfHLevelSuc 1 (isOfHLevelSuc 0 sphereConnectedSpecCase) ∣ north ∣ ∣ north ∣ (λ _ → ∣ north ∣))))) d-equiv : isEquiv {A = ∥ typ (Ω (Susp S¹ , north)) ∥ (ℕ→ℕ₋₂ 1)} {B = ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (trElim (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣ ) d-equiv = conEquiv (ℕ→ℕ₋₂ 1) d-map is1Connected-dmap .snd +{- We show that composing (λ a → ∣ ϕ base a ∣) and (λ x → ∣ d-map x ∣) gives us the identity function. -} + d-mapId2 : (λ (x : ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)) → (trElim {n = 3} {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) (trElim (λ _ → isOfHLevelTrunc 3) (λ a → ∣ ϕ base a ∣) x)) ≡ λ x → x -d-mapId2 = funExt (trElim (λ x → isOfHLevelSuc 2 (isOfHLevelTrunc 3 ((trElim {n = 3} {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) (trElim (λ _ → isOfHLevelTrunc 3) (λ a → ∣ ϕ base a ∣) x)) x)) λ a i → ∣ d-mapId a i ∣) - -isEquiv∣ϕ∣ : isEquiv {A = ∥ S¹ ∥ ℕ→ℕ₋₂ 1} (trElim (λ _ → isOfHLevelTrunc 3) (λ a → ∣ ϕ base a ∣)) -isEquiv∣ϕ∣ = compToIdEquiv (trElim {n = 3} {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) +d-mapId2 = funExt (trElim (λ x → isOfHLevelSuc 2 (isOfHLevelTrunc 3 ((trElim {n = 3} + {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} + (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) + (trElim (λ _ → isOfHLevelTrunc 3) + (λ a → ∣ ϕ base a ∣) x)) x)) + λ a i → ∣ d-mapId a i ∣) + +{- This means that (λ a → ∣ ϕ base a ∣) is an equivalence -} +isEquiv∣ϕ-base∣ : isEquiv {A = ∥ S¹ ∥ ℕ→ℕ₋₂ 1} (trElim (λ _ → isOfHLevelTrunc 3) (λ a → ∣ ϕ base a ∣)) +isEquiv∣ϕ-base∣ = composesToId→Equiv (trElim {n = 3} {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) (trElim (λ _ → isOfHLevelTrunc 3) (λ a → ∣ ϕ base a ∣)) d-mapId2 d-equiv @@ -368,13 +312,13 @@ isEquiv∣ϕ∣ = compToIdEquiv (trElim {n = 3} {B = λ _ → ∥ S¹ ∥ (ℕ --------------------------------- -- We cheat when n = 1 and use J to prove the following lemmma. There is an obvious dependent path between ϕ base and ϕ north. Since the first one is an equivalence, so is the other. -- -funTEST2 : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ} {C : (A : Type ℓ) (a1 : A) → Type ℓ'} (p : A ≡ B) (a : A) (b : B) → +pointFunEquiv : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ} {C : (A : Type ℓ) (a1 : A) → Type ℓ'} (p : A ≡ B) (a : A) (b : B) → (transport p a ≡ b) → (f : (A : Type ℓ) → (a1 : A) (a2 : ∥ A ∥ 1) → C A a1) → isEquiv (f A a) → isEquiv (f B b) -funTEST2 {ℓ = ℓ}{A = A} {B = B} {C = C} = +pointFunEquiv {ℓ = ℓ}{A = A} {B = B} {C = C} = J (λ B p → (a : A) (b : B) → (transport p a ≡ b) → (f : (A : Type ℓ) → @@ -383,11 +327,14 @@ funTEST2 {ℓ = ℓ}{A = A} {B = B} {C = C} = isEquiv (f B b)) λ a b trefl f is → transport (λ i → isEquiv (f A ((sym (transportRefl a) ∙ trefl) i))) is ------------------------------------------------------ +{- By pointFunEquiv, this gives that λ a → ∣ ϕ north a ∣ is an equivalence. -} -final : isEquiv {A = ∥ S₊ 1 ∥ 1} {B = ∥ typ (Ω (S₊ 2 , north)) ∥ 1} (trElim (λ _ → isOfHLevelTrunc 3) (λ a → ∣ ϕ north a ∣)) -final = funTEST2 {A = S¹} (λ i → S¹≡S1 (~ i)) base north refl (λ A a1 → trElim (λ _ → isOfHLevelTrunc 3) (λ a → ∣ ϕ a1 a ∣)) isEquiv∣ϕ∣ +isEquiv∣ϕ∣ : isEquiv {A = ∥ S₊ 1 ∥ 1} {B = ∥ typ (Ω (S₊ 2 , north)) ∥ 1} (trElim (λ _ → isOfHLevelTrunc 3) (λ a → ∣ ϕ north a ∣)) +isEquiv∣ϕ∣ = pointFunEquiv {A = S¹} (λ i → S¹≡S1 (~ i)) base north refl (λ A a1 → trElim (λ _ → isOfHLevelTrunc 3) (λ a → ∣ ϕ a1 a ∣)) isEquiv∣ϕ-base∣ +---------------------------------------------------- Finishing up --------------------------------- + +{- For n ≥ 1, we rewrite our function as the composition below. -} Kn→ΩKn+1Sucn : (n : ℕ) → Kn→ΩKn+1 (suc n) ≡ λ x → truncEquivΩ (ℕ→ℕ₋₂ (suc n)) .fst (trElim (λ _ → isOfHLevelTrunc (3 + n)) (λ a → ∣ ϕ north a ∣) x) Kn→ΩKn+1Sucn n = funExt (trElim (λ x → isOfHLevelSuc (suc (suc n)) ((isOfHLevelTrunc ( 2 + (suc (suc n))) ∣ north ∣ ∣ north ∣) @@ -397,28 +344,3 @@ Kn→ΩKn+1Sucn n = funExt (trElim (λ x → isOfHLevelSuc (suc (suc n)) - -isEquivKn→ΩKn+1 : (n : ℕ) → isEquiv (Kn→ΩKn+1 n) -isEquivKn→ΩKn+1 zero = compEquiv (looper , isEquivLooper) (cong ∣_∣ , isEquivHelper hLevl3) .snd - where - hLevl3 : (x y : S₊ 1) (p q : x ≡ y) → isProp (p ≡ q) - hLevl3 x y = J (λ y p → (q : x ≡ y) → isProp (p ≡ q) ) - (transport (λ i → isSet (helper (~ i))) isSetInt refl) - where - helper : (x ≡ x) ≡ Int - helper = (λ i → transp (λ j → ua S1≃SuspBool (~ j ∨ ~ i)) (~ i) x ≡ transp (λ j → ua S1≃SuspBool (~ j ∨ ~ i)) (~ i) x) ∙ - (λ i → transp (λ j → S¹≡SuspBool (~ j ∨ ~ i)) (~ i) (transport (sym (ua S1≃SuspBool)) x) ≡ transp (λ j → S¹≡SuspBool (~ j ∨ ~ i)) (~ i) (transport (sym (ua S1≃SuspBool)) x)) ∙ - basedΩS¹≡Int (transport (sym S¹≡SuspBool) (transport (sym (ua S1≃SuspBool)) x)) -isEquivKn→ΩKn+1 (suc zero) = transport (λ i → isEquiv (Kn→ΩKn+1Sucn zero (~ i))) - (compEquiv (trElim (λ _ → isOfHLevelTrunc (2 + (suc zero))) (λ a → ∣ ϕ north a ∣) , - final) - (truncEquivΩ (ℕ→ℕ₋₂ (suc zero))) .snd) -isEquivKn→ΩKn+1 (suc (suc n)) = transport (λ i → isEquiv (Kn→ΩKn+1Sucn (suc n) (~ i))) - (compEquiv (conEquiv3 (4 + n) _ (ϕ north) (n , λ i → suc (suc (suc (+-suc n n (~ i))))) (FthalFun-2nConnected (suc n) _ (Pr242 _))) - (truncEquivΩ (ℕ→ℕ₋₂ (suc (suc n)))) .snd) - -Kn≃ΩKn+1 : (n : ℕ) → coHomK n ≃ typ (Ω (coHomK-ptd (suc n))) -Kn≃ΩKn+1 n = Kn→ΩKn+1 n , isEquivKn→ΩKn+1 n - -ΩKn+1→Kn : (n : ℕ) → typ (Ω (coHomK-ptd (suc n))) → coHomK n -ΩKn+1→Kn n a = equiv-proof (isEquivKn→ΩKn+1 n) a .fst .fst From a771acbe6e75c98903d4ba7d0185fa254131839c Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 21 Apr 2020 01:29:09 +0200 Subject: [PATCH 17/89] K-algebra --- Cubical/Foundations/Transport.agda | 2 + Cubical/Homotopy/Connected.agda | 1 + Cubical/Homotopy/Loopspace.agda | 75 +++++++++++++++++++++ Cubical/ZCohomology/Properties.agda | 86 +++++++++++++++++++++---- Cubical/ZCohomology/cupProd.agda | 0 Cubical/ZCohomology/cupProdPrelims.agda | 30 +++++++-- 6 files changed, 176 insertions(+), 18 deletions(-) create mode 100644 Cubical/ZCohomology/cupProd.agda diff --git a/Cubical/Foundations/Transport.agda b/Cubical/Foundations/Transport.agda index 3e03f6980..2ea376f01 100644 --- a/Cubical/Foundations/Transport.agda +++ b/Cubical/Foundations/Transport.agda @@ -75,6 +75,8 @@ substComposite B p q Bx = sym (substRefl {B = B} _) ∙ helper where helper : subst B refl (subst B (p ∙ q) Bx) ≡ subst B q (subst B p Bx) helper i = subst B (λ k → compSq (~ i ∧ ~ k) (~ i ∨ k)) (subst B (λ k → compSq (~ i ∨ ~ k) (~ i ∧ k)) Bx) + + -- substitution commutes with morphisms in slices substCommSlice : ∀ {ℓ ℓ′} {A : Type ℓ} → (B C : A → Type ℓ′) diff --git a/Cubical/Homotopy/Connected.agda b/Cubical/Homotopy/Connected.agda index e89a7677b..88d12bc70 100644 --- a/Cubical/Homotopy/Connected.agda +++ b/Cubical/Homotopy/Connected.agda @@ -249,6 +249,7 @@ connectedTruncIso2 : ∀ {ℓ} {A B : Type ℓ} (n m : ℕ) (f : A → B) connectedTruncIso2 {A = A} {B = B} n m f (x , pf) con = connectedTruncIso n f (isConnectedSubtr n x f (transport (λ i → isHLevelConnectedFun (pf (~ i)) f) con)) + connectedTruncEquiv : ∀ {ℓ} {A B : Type ℓ} (n : ℕ) (f : A → B) → isHLevelConnectedFun n f → hLevelTrunc n A ≃ hLevelTrunc n B diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda index 1ede1968e..d604424e1 100644 --- a/Cubical/Homotopy/Loopspace.agda +++ b/Cubical/Homotopy/Loopspace.agda @@ -8,7 +8,9 @@ open import Cubical.Data.Nat open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed +open import Cubical.Foundations.HLevels open import Cubical.Foundations.GroupoidLaws +open import Cubical.HITs.SetTruncation {- loop space of a pointed type -} Ω : {ℓ : Level} → Pointed ℓ → Pointed ℓ @@ -22,3 +24,76 @@ open import Cubical.Foundations.GroupoidLaws {- loop space map -} Ω→ : ∀ {ℓA ℓB} {A : Pointed ℓA} {B : Pointed ℓB} (f : A →∙ B) → (Ω A →∙ Ω B) Ω→ (f , f∙) = (λ p → (sym f∙ ∙ cong f p) ∙ f∙) , cong (λ q → q ∙ f∙) (sym (rUnit (sym f∙))) ∙ lCancel f∙ + + +-- Eckmann-Hilton -- + +{- Horizontal composition -} +_⋆_ : ∀ {ℓ} {A : Type ℓ} {a b c : A} {p q : a ≡ b} {p' q' : b ≡ c} → (p ≡ q) → (p' ≡ q') + → p ∙ p' ≡ q ∙ q' +_⋆_ α β i = α i ∙ β i + +{- Exchange law-} +exchange : ∀ {ℓ} {A : Type ℓ} {a b c : A} {p q r : a ≡ b} {p' q' r' : b ≡ c} + (α : p ≡ q) (β : q ≡ r) (α' : p' ≡ q') (β' : q' ≡ r') + → (α ∙ β) ⋆ (α' ∙ β') ≡ (α ⋆ α') ∙ (β ⋆ β') +exchange {A = A} {a = a} {c = c} {p = p} {r = r} {p' = p'} {r' = r'} α β α' β' z i = + hcomp (λ k → λ { (i = i0) → p ∙ p' + ; (i = i1) → (β ⋆ β') (k ∨ ~ z) + ; (z = i0) → ((α ∙ β) ⋆ (α' ∙ β')) i }) + (btm-Filler z i) + where + btm-Filler : PathP (λ i → p ∙ p' ≡ (β ⋆ β') (~ i)) ((α ∙ β) ⋆ (α' ∙ β')) (α ⋆ α') + btm-Filler i = hcomp (λ k → λ { (i = i0) → (α ∙ β) ⋆ (α' ∙ β') + ; (i = i1) → rUnit α (~ k) ⋆ rUnit α' (~ k)}) + ((α ∙ λ j → β (~ i ∧ j)) ⋆ (α' ∙ λ j → β' (~ i ∧ j))) + + +Eckmann-Hilton : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (α β : typ ((Ω^ (2 + n)) A)) + → α ∙ β ≡ β ∙ α +Eckmann-Hilton {A = (A , pt)} zero α β = + α ∙ β ≡⟨ (λ i → unit1 α (~ i) ∙ unit2 β (~ i)) ⟩ + (rUnit refl ∙ (refl ⋆ α) ∙ sym (rUnit refl)) ∙ lUnit refl ∙ (β ⋆ refl) ∙ sym (lUnit refl) ≡⟨ helper (rUnit refl) (refl ⋆ α) (β ⋆ refl) ⟩ + rUnit refl ∙ ((refl ⋆ α) ∙ (β ⋆ refl)) ∙ sym (rUnit refl) ≡⟨ (λ i → rUnit refl ∙ exchange refl β α refl (~ i) ∙ sym (rUnit refl)) ⟩ + rUnit refl ∙ ((refl ∙ β) ⋆ (α ∙ refl)) ∙ sym (rUnit refl) ≡⟨ (λ i → rUnit refl ∙ helper2 i ∙ sym (rUnit refl)) ⟩ + rUnit refl ∙ ((β ∙ refl) ⋆ (refl ∙ α)) ∙ sym (rUnit refl) ≡⟨ (λ i → rUnit refl ∙ exchange β refl refl α i ∙ sym (rUnit refl) ) ⟩ + rUnit refl ∙ ((β ⋆ refl) ∙ (refl ⋆ α)) ∙ sym (rUnit refl) ≡⟨ sym (helper (rUnit refl) (β ⋆ refl) (refl ⋆ α)) ⟩ + (lUnit refl ∙ (β ⋆ refl) ∙ sym (lUnit refl)) ∙ rUnit refl ∙ (refl ⋆ α) ∙ sym (rUnit refl) ≡⟨ (λ i → unit2 β i ∙ unit1 α i) ⟩ + β ∙ α ∎ + where + helper : ∀ {ℓ} {A : Type ℓ} {a b : A} (p : a ≡ b) (q q' : b ≡ b) → (p ∙ q ∙ (sym p)) ∙ p ∙ q' ∙ sym p ≡ p ∙ (q ∙ q') ∙ sym p + helper p q q' = (p ∙ q ∙ (sym p)) ∙ p ∙ q' ∙ sym p ≡⟨ (λ i → assoc p q (sym p) i ∙ p ∙ q' ∙ sym p) ⟩ + ((p ∙ q) ∙ sym p) ∙ p ∙ q' ∙ sym p ≡⟨ assoc ((p ∙ q) ∙ sym p) p (q' ∙ sym p) ⟩ + (((p ∙ q) ∙ sym p) ∙ p) ∙ q' ∙ sym p ≡⟨ (λ i → assoc (p ∙ q) (sym p) p (~ i) ∙ q' ∙ sym p) ⟩ + ((p ∙ q) ∙ (sym p ∙ p)) ∙ q' ∙ sym p ≡⟨ (λ i → ((p ∙ q) ∙ lCancel p i) ∙ q' ∙ sym p) ⟩ + ((p ∙ q) ∙ refl) ∙ q' ∙ sym p ≡⟨ (λ i → rUnit (p ∙ q) (~ i) ∙ q' ∙ sym p) ⟩ + (p ∙ q) ∙ q' ∙ sym p ≡⟨ assoc (p ∙ q) q' (sym p) ⟩ + ((p ∙ q) ∙ q') ∙ sym p ≡⟨ (λ i → assoc p q q' (~ i) ∙ sym p) ⟩ + (p ∙ (q ∙ q')) ∙ sym p ≡⟨ sym (assoc p (q ∙ q') (sym p)) ⟩ + p ∙ (q ∙ q') ∙ sym p ∎ + + helper2 : (refl ∙ β) ⋆ (α ∙ refl) ≡ ((β ∙ refl) ⋆ (refl ∙ α)) + helper2 = (λ i → (rUnit ((lUnit β) (~ i)) i) ⋆ (lUnit (rUnit α (~ i)) i)) + + unit1 : {a : A} (γ : Path (a ≡ a) (λ _ → a) λ _ → a) → lUnit refl ∙ (refl ⋆ γ) ∙ sym (lUnit refl) ≡ γ + unit1 {a = a} γ = lUnit refl ∙ (refl ⋆ γ) ∙ sym (lUnit refl) ≡⟨ (λ i → (λ j → lUnit refl (j ∧ ~ i)) ∙ (λ j → lUnit (γ j) (~ i)) ∙ λ j → lUnit refl (~ i ∧ ~ j)) ⟩ + refl ∙ γ ∙ refl ≡⟨ (λ i → lUnit (rUnit γ (~ i)) (~ i)) ⟩ + γ ∎ + + + unit2 : {a : A} (γ : Path (a ≡ a) (λ _ → a) λ _ → a) → rUnit refl ∙ (γ ⋆ refl) ∙ sym (rUnit refl) ≡ γ + unit2 γ = rUnit refl ∙ (γ ⋆ refl) ∙ sym (rUnit refl) ≡⟨ (λ i → (λ j → rUnit refl (j ∧ ~ i)) ∙ (λ j → rUnit (γ j) (~ i)) ∙ λ j → rUnit refl (~ i ∧ ~ j)) ⟩ + refl ∙ γ ∙ refl ≡⟨ (λ i → lUnit (rUnit γ (~ i)) (~ i)) ⟩ + γ ∎ + +Eckmann-Hilton {A = (A , pt)} (suc n) α β = Eckmann-Hilton 0 α β + + +{- Composition in fundamental group -} +π-comp : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → ∥ typ ((Ω^ (suc n)) A) ∥₀ → ∥ typ ((Ω^ (suc n)) A) ∥₀ → ∥ typ ((Ω^ (suc n)) A) ∥₀ +π-comp n = elim2 (λ _ _ → setTruncIsSet) λ p q → ∣ p ∙ q ∣₀ + +Eckmann-Hilton-π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p q : ∥ typ ((Ω^ (2 + n)) A) ∥₀) + → π-comp (1 + n) p q ≡ π-comp (1 + n) q p +Eckmann-Hilton-π {A = (A , pt)} n = elim2 (λ x y → isOfHLevelPath 2 setTruncIsSet _ _) + λ p q → cong ∣_∣₀ (Eckmann-Hilton n p q) diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index 557404585..60e03a0f8 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -8,6 +8,7 @@ open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Transport open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Univalence @@ -19,7 +20,7 @@ open import Cubical.HITs.SetTruncation open import Cubical.HITs.Nullification open import Cubical.Data.Int hiding (_+_) open import Cubical.Data.Nat -open import Cubical.HITs.Truncation renaming (elim to trElim) +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) open import Cubical.Homotopy.Loopspace open import Cubical.Homotopy.Connected open import Cubical.Homotopy.Freudenthal @@ -70,26 +71,22 @@ coHomRed+1Equiv (suc n) A i = ∥ coHomRed+1.helpLemma A i {C = ((coHomK (suc n) ----------- Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) -Kn→ΩKn+1 n = Iso.fun (Iso-Kn-ΩKn+1 n) +Kn→ΩKn+1 n = Iso.fun (Iso2-Kn-ΩKn+1 n) ΩKn+1→Kn : {n : ℕ} → typ (Ω (coHomK-ptd (suc n))) → coHomK n -ΩKn+1→Kn {n = n} = Iso.inv (Iso-Kn-ΩKn+1 n) - -sectionK : {n : ℕ} → section (Kn→ΩKn+1 n) ΩKn+1→Kn -sectionK {n = n} a = {!Iso.rightInv (Iso-Kn-ΩKn+1 n) a!} - +ΩKn+1→Kn {n = n} = Iso.inv (Iso2-Kn-ΩKn+1 n) Kn≃ΩKn+1 : {n : ℕ} → coHomK n ≃ typ (Ω (coHomK-ptd (suc n))) Kn≃ΩKn+1 {n = n} = isoToEquiv (Iso-Kn-ΩKn+1 n) - +---------- Algebra -------- _+ₖ_ : {n : ℕ} → coHomK n → coHomK n → coHomK n _+ₖ_ {n = n} x y = ΩKn+1→Kn (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) -_-ₖ_ : {n : ℕ} → coHomK n → coHomK n -_-ₖ_ {n = n} x = ΩKn+1→Kn (sym (Kn→ΩKn+1 n x)) +-ₖ_ : {n : ℕ} → coHomK n → coHomK n +-ₖ_ {n = n} x = ΩKn+1→Kn (sym (Kn→ΩKn+1 n x)) 0ₖ : {n : ℕ} → coHomK n 0ₖ {n = zero} = pt (coHomK-ptd 0) @@ -97,12 +94,73 @@ _-ₖ_ {n = n} x = ΩKn+1→Kn (sym (Kn→ΩKn+1 n x)) Kn→ΩKn+10ₖ : (n : ℕ) → Kn→ΩKn+1 n 0ₖ ≡ refl Kn→ΩKn+10ₖ zero = refl -Kn→ΩKn+10ₖ (suc n) = {!refl!} ∙ (λ i → cong ∣_∣ (rCancel (merid north) i)) +Kn→ΩKn+10ₖ (suc n) = (λ i → cong ∣_∣ (rCancel (merid north) i)) rUnitₖ : {n : ℕ} (x : coHomK n) → x +ₖ 0ₖ ≡ x -rUnitₖ {n = zero} x = {!!} +rUnitₖ {n = zero} x = cong ΩKn+1→Kn (sym (rUnit (Kn→ΩKn+1 zero x))) ∙ + Iso.leftInv (Iso2-Kn-ΩKn+1 zero) x rUnitₖ {n = suc n} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 (suc n) x ∙ Kn→ΩKn+10ₖ (suc n) i)) ∙ (λ i → ΩKn+1→Kn (rUnit (Kn→ΩKn+1 (suc n) x) (~ i))) ∙ - {!!} - {!Iso.rightInv (Iso-Kn-ΩKn+1 (suc n)) (Kn→ΩKn+1 (suc n) x)!} + Iso.leftInv (Iso2-Kn-ΩKn+1 (suc n)) x + +lUnitₖ : {n : ℕ} (x : coHomK n) → 0ₖ +ₖ x ≡ x +lUnitₖ {n = zero} x = cong ΩKn+1→Kn (sym (lUnit (Kn→ΩKn+1 zero x))) ∙ + Iso.leftInv (Iso2-Kn-ΩKn+1 zero) x +lUnitₖ {n = suc n} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc n) i ∙ Kn→ΩKn+1 (suc n) x)) ∙ + (λ i → ΩKn+1→Kn (lUnit (Kn→ΩKn+1 (suc n) x) (~ i))) ∙ + Iso.leftInv (Iso2-Kn-ΩKn+1 (suc n)) x + +rCancelₖ : {n : ℕ} (x : coHomK n) → x +ₖ (-ₖ x) ≡ 0ₖ +rCancelₖ {n = zero} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 zero x ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 zero) (sym (Kn→ΩKn+1 zero x)) i)) ∙ + cong ΩKn+1→Kn (rCancel (Kn→ΩKn+1 zero x)) +rCancelₖ {n = suc zero} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 1 x ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 1) (sym (Kn→ΩKn+1 1 x)) i)) ∙ + cong ΩKn+1→Kn (rCancel (Kn→ΩKn+1 1 x)) +rCancelₖ {n = suc (suc n)} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 (2 + n) x ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 (2 + n)) (sym (Kn→ΩKn+1 (2 + n) x)) i)) ∙ + cong ΩKn+1→Kn (rCancel (Kn→ΩKn+1 (2 + n) x)) ∙ + (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc (suc n)) (~ i))) ∙ + Iso.leftInv (Iso2-Kn-ΩKn+1 (suc (suc n))) 0ₖ + +lCancelₖ : {n : ℕ} (x : coHomK n) → (-ₖ x) +ₖ x ≡ 0ₖ +lCancelₖ {n = zero} x = (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 zero) (sym (Kn→ΩKn+1 zero x)) i ∙ Kn→ΩKn+1 zero x)) ∙ + cong ΩKn+1→Kn (lCancel (Kn→ΩKn+1 zero x)) +lCancelₖ {n = suc zero} x = ((λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 1) (sym (Kn→ΩKn+1 1 x)) i ∙ Kn→ΩKn+1 1 x))) ∙ + cong ΩKn+1→Kn (lCancel (Kn→ΩKn+1 1 x)) +lCancelₖ {n = suc (suc n)} x = (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 (2 + n)) (sym (Kn→ΩKn+1 (2 + n) x)) i ∙ Kn→ΩKn+1 (2 + n) x)) ∙ + cong ΩKn+1→Kn (lCancel (Kn→ΩKn+1 (2 + n) x)) ∙ + (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc (suc n)) (~ i))) ∙ + Iso.leftInv (Iso2-Kn-ΩKn+1 (suc (suc n))) 0ₖ + + +assocₖ : {n : ℕ} (x y z : coHomK n) → ((x +ₖ y) +ₖ z) ≡ (x +ₖ (y +ₖ z)) +assocₖ {n = n} x y z = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 n (ΩKn+1→Kn (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y)) ∙ Kn→ΩKn+1 n z)) ∙ + (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 n) (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) i ∙ Kn→ΩKn+1 n z)) ∙ + (λ i → ΩKn+1→Kn (assoc (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) (Kn→ΩKn+1 n z) (~ i))) ∙ + (λ i → ΩKn+1→Kn ((Kn→ΩKn+1 n x) ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 n) ((Kn→ΩKn+1 n y ∙ Kn→ΩKn+1 n z)) (~ i))) + +commₖ : {n : ℕ} (x y : coHomK n) → (x +ₖ y) ≡ (y +ₖ x) +commₖ {n = n} x y = λ i → ΩKn+1→Kn (EH-instance (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) i) + where + EH-instance : (p q : typ (Ω ((∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) , ∣ north ∣))) → p ∙ q ≡ q ∙ p + EH-instance p q = (λ i → transport⁻Transport (K-Id n) p (~ i) ∙ transport⁻Transport (K-Id n) q (~ i) ) ∙ + (λ i → transport (sym (K-Id n)) (rUnit (transport (K-Id n) p) i) ∙ transport (sym (K-Id n)) (transport (K-Id n) q)) ∙ + (λ i → transport (sym (K-Id n)) ((transport (K-Id n) p) ∙ {!!}) ∙ {!λ j → (transport (sym (K-Id n)) (transport (K-Id n) q)) ?!}) ∙ + {!!} ∙ + {!!} ∙ + {!!} ∙ + {!!} + where + K-Id : (n : ℕ) → typ (Ω (hLevelTrunc (3 + n) (S₊ (1 + n)) , ∣ north ∣)) ≡ typ ((Ω^ 2) (hLevelTrunc (4 + n) (S₊ (2 + n)) , ∣ north ∣ )) + K-Id n = (λ i → typ (Ω (isoToPath (Iso2-Kn-ΩKn+1 (suc n)) i , transp (λ j → isoToPath (Iso2-Kn-ΩKn+1 (suc n)) (i ∧ j)) (~ i) ∣ north ∣))) ∙ + (λ i → typ (Ω (typ (Ω (hLevelTrunc (4 + n) (S₊ (2 + n)) , ∣ north ∣)) , transportRefl (λ j → ∣ rCancel (merid north) i j ∣) i))) + + K-Iso : (n : ℕ) → Iso (typ (Ω (hLevelTrunc (3 + n) (S₊ (1 + n)) , ∣ north ∣))) (typ ((Ω^ 2) (hLevelTrunc (4 + n) (S₊ (2 + n)) , ∣ north ∣ ))) + K-Iso n = {!!} + + subst∙ : ∀ {ℓ} {A : Type ℓ} {x y : A} (P : (x ≡ x) ≡ (y ≡ y)) (p q : x ≡ x) + → transport P p ∙ transport P q ≡ transport P (p ∙ q) + subst∙ {x = x} {y = y} P p q = (λ i → transport (lUnit P i) (rUnit p i) ∙ transport (lUnit P i) q) ∙ + (λ i → transport ((λ j → x ≡ q (i ∧ (~ j))) ∙ P) (p ∙ (λ j → q (i ∧ j))) ∙ transport ((λ j → x ≡ q ((~ i) ∨ j)) ∙ P) λ j → q ((~ i) ∧ j) ) ∙ + {!!} ∙ + {!transport ? ?!} ∙ + {!!} diff --git a/Cubical/ZCohomology/cupProd.agda b/Cubical/ZCohomology/cupProd.agda new file mode 100644 index 000000000..e69de29bb diff --git a/Cubical/ZCohomology/cupProdPrelims.agda b/Cubical/ZCohomology/cupProdPrelims.agda index 18e48abfe..e67700d04 100644 --- a/Cubical/ZCohomology/cupProdPrelims.agda +++ b/Cubical/ZCohomology/cupProdPrelims.agda @@ -43,12 +43,12 @@ private B : Type ℓ' {- We want to prove that Kn≃ΩKn+1. For this we use the map ϕ-} -private - ϕ : (pt a : A) → typ (Ω (Susp A , north)) - ϕ pt a = (merid a) ∙ sym (merid pt) - {- To define the map for n=0 we use the λ k → loopᵏ map for S₊ 1. The loop is given by ϕ south north -} +ϕ : (pt a : A) → typ (Ω (Susp A , north)) +ϕ pt a = (merid a) ∙ sym (merid pt) + {- To define the map for n=0 we use the λ k → loopᵏ map for S₊ 1. The loop is given by ϕ south north -} +private loop* : Path (S₊ 1) north north loop* = ϕ north south @@ -237,6 +237,22 @@ decode-fun2 : (n : ℕ) → (x : A) → hLevelTrunc n (x ≡ x) → Path (hLevel decode-fun2 zero x = trElim (λ _ → isOfHLevelPath 0 (∣ x ∣ , isOfHLevelTrunc 1 ∣ x ∣) ∣ x ∣ ∣ x ∣) (λ p i → ∣ p i ∣) decode-fun2 (suc n) x = trElim (λ _ → isOfHLevelPath' (suc n) (isOfHLevelTrunc (suc (suc n))) ∣ x ∣ ∣ x ∣) (cong ∣_∣) +-- {- auxiliary function r used to define encode -} + +-- r : {m : ℕ₋₂} (u : ∥ B ∥ (suc₋₂ m)) → P u u +-- r = elim (λ x → hLevelP x x) (λ a → ∣ refl ∣) + +-- {- encode function from x ≡ y to P x y -} +-- encode-fun : {n : ℕ₋₂} (x y : ∥ B ∥ (suc₋₂ n)) → x ≡ y → P x y +-- encode-fun x y p = transport (λ i → P x (p i)) (r x) +{- +encode2 : (n : ℕ) (x : A) → Path (hLevelTrunc (suc n) A) ∣ x ∣ ∣ x ∣ → hLevelTrunc n (x ≡ x) +encode2 {A = A} n x p = transport (λ i → hLevelTrunc n (x ≡ {!p i!})) {!!} + where + r : (m : ℕ) (u : hLevelTrunc (1 + m) A) → hLevelTrunc m (u ≡ u) + r m = trElim (λ x → isOfHLevelSuc m (isOfHLevelTrunc _) ) λ a → ∣ refl ∣ +-} + funsAreSame : (n : ℕ) (x : A) (b : hLevelTrunc n (x ≡ x)) → (decode-fun2 n x b) ≡ (ΩTrunc.decode-fun ∣ x ∣ ∣ x ∣ b) funsAreSame zero x = trElim (λ a → isOfHLevelPath 0 (refl , (isOfHLevelSuc 1 (isOfHLevelTrunc 1) ∣ x ∣ ∣ x ∣ refl)) _ _) λ a → refl funsAreSame (suc n) x = trElim (λ a → isOfHLevelPath _ (isOfHLevelPath' (suc n) (isOfHLevelTrunc (suc (suc n))) ∣ x ∣ ∣ x ∣) _ _) λ a → refl @@ -267,3 +283,9 @@ mapId2 : (n : ℕ) → Kn→ΩKn+1 n ≡ Iso.fun (Iso-Kn-ΩKn+1 n) mapId2 zero = refl mapId2 (suc zero) = funExt (trElim (λ x → isOfHLevelPath 3 (isOfHLevelTrunc 4 ∣ north ∣ ∣ north ∣) _ _) λ a → refl) mapId2 (suc (suc n)) = funExt (trElim (λ x → isOfHLevelPath (4 + n) (isOfHLevelTrunc (5 + n) ∣ north ∣ ∣ north ∣) _ _) λ a → refl) + +Iso2-Kn-ΩKn+1 : (n : ℕ) → Iso (coHomK n) (typ (Ω (coHomK-ptd (suc n)))) +Iso.fun (Iso2-Kn-ΩKn+1 n) = Kn→ΩKn+1 n +Iso.inv (Iso2-Kn-ΩKn+1 n) = Iso.inv (Iso-Kn-ΩKn+1 n) +Iso.rightInv (Iso2-Kn-ΩKn+1 n) a = funExt⁻ (mapId2 n) _ ∙ Iso.rightInv (Iso-Kn-ΩKn+1 n) a +Iso.leftInv (Iso2-Kn-ΩKn+1 n) a = cong (Iso.inv (Iso-Kn-ΩKn+1 n)) (funExt⁻ (mapId2 n) a) ∙ Iso.leftInv (Iso-Kn-ΩKn+1 n) a From 4b33a57da8714d522dfdcfbd3a2924c38eab993d Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 21 Apr 2020 04:38:18 +0200 Subject: [PATCH 18/89] K algebra laws done --- Cubical/ZCohomology/Properties.agda | 30 ++++++++--------------------- 1 file changed, 8 insertions(+), 22 deletions(-) diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index 60e03a0f8..28b425f72 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -38,6 +38,7 @@ private A : Type ℓ B : Type ℓ' + {- Equivalence between cohomology of A and reduced cohomology of (A + 1) -} coHomRed+1Equiv : (n : ℕ) → (A : Set ℓ) → @@ -139,28 +140,13 @@ assocₖ {n = n} x y z = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 n (ΩKn+1→Kn (Kn (λ i → ΩKn+1→Kn ((Kn→ΩKn+1 n x) ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 n) ((Kn→ΩKn+1 n y ∙ Kn→ΩKn+1 n z)) (~ i))) commₖ : {n : ℕ} (x y : coHomK n) → (x +ₖ y) ≡ (y +ₖ x) -commₖ {n = n} x y = λ i → ΩKn+1→Kn (EH-instance (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) i) +commₖ {n = n} x y i = ΩKn+1→Kn (EH-instance (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) i) where EH-instance : (p q : typ (Ω ((∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) , ∣ north ∣))) → p ∙ q ≡ q ∙ p - EH-instance p q = (λ i → transport⁻Transport (K-Id n) p (~ i) ∙ transport⁻Transport (K-Id n) q (~ i) ) ∙ - (λ i → transport (sym (K-Id n)) (rUnit (transport (K-Id n) p) i) ∙ transport (sym (K-Id n)) (transport (K-Id n) q)) ∙ - (λ i → transport (sym (K-Id n)) ((transport (K-Id n) p) ∙ {!!}) ∙ {!λ j → (transport (sym (K-Id n)) (transport (K-Id n) q)) ?!}) ∙ - {!!} ∙ - {!!} ∙ - {!!} ∙ - {!!} + EH-instance = transport (λ i → (p q : K-Id n (~ i)) → p ∙ q ≡ q ∙ p) + λ p q → Eckmann-Hilton 0 p q where - K-Id : (n : ℕ) → typ (Ω (hLevelTrunc (3 + n) (S₊ (1 + n)) , ∣ north ∣)) ≡ typ ((Ω^ 2) (hLevelTrunc (4 + n) (S₊ (2 + n)) , ∣ north ∣ )) - K-Id n = (λ i → typ (Ω (isoToPath (Iso2-Kn-ΩKn+1 (suc n)) i , transp (λ j → isoToPath (Iso2-Kn-ΩKn+1 (suc n)) (i ∧ j)) (~ i) ∣ north ∣))) ∙ - (λ i → typ (Ω (typ (Ω (hLevelTrunc (4 + n) (S₊ (2 + n)) , ∣ north ∣)) , transportRefl (λ j → ∣ rCancel (merid north) i j ∣) i))) - - K-Iso : (n : ℕ) → Iso (typ (Ω (hLevelTrunc (3 + n) (S₊ (1 + n)) , ∣ north ∣))) (typ ((Ω^ 2) (hLevelTrunc (4 + n) (S₊ (2 + n)) , ∣ north ∣ ))) - K-Iso n = {!!} - - subst∙ : ∀ {ℓ} {A : Type ℓ} {x y : A} (P : (x ≡ x) ≡ (y ≡ y)) (p q : x ≡ x) - → transport P p ∙ transport P q ≡ transport P (p ∙ q) - subst∙ {x = x} {y = y} P p q = (λ i → transport (lUnit P i) (rUnit p i) ∙ transport (lUnit P i) q) ∙ - (λ i → transport ((λ j → x ≡ q (i ∧ (~ j))) ∙ P) (p ∙ (λ j → q (i ∧ j))) ∙ transport ((λ j → x ≡ q ((~ i) ∨ j)) ∙ P) λ j → q ((~ i) ∧ j) ) ∙ - {!!} ∙ - {!transport ? ?!} ∙ - {!!} + K-Id : (n : ℕ) → typ (Ω (hLevelTrunc (3 + n) (S₊ (1 + n)) , ∣ north ∣)) ≡ typ ((Ω^ 2) (hLevelTrunc (4 + n) (S₊ (2 + n)) , ∣ north ∣ )) + K-Id n = (λ i → typ (Ω (isoToPath (Iso2-Kn-ΩKn+1 (suc n)) i , hcomp (λ k → λ {(i = i0) → ∣ north ∣ + ; (i = i1) → transportRefl (λ j → ∣ rCancel (merid north) k j ∣) k}) + (transp (λ j → isoToPath (Iso2-Kn-ΩKn+1 (suc n)) (i ∧ j)) (~ i) ∣ north ∣)))) From 972feaf08b1b6200360e6008af90fd00dbfc1c79 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 21 Apr 2020 18:55:17 +0200 Subject: [PATCH 19/89] smash sphere --- Cubical/HITs/Pushout/Base.agda | 16 + Cubical/HITs/Sn.agda | 1 + Cubical/HITs/Sn/Properties.agda | 104 ++++ Cubical/Homotopy/Loopspace.agda | 12 +- Cubical/ZCohomology/Base.agda | 6 + Cubical/ZCohomology/OldDef/Base.agda | 45 -- Cubical/ZCohomology/OldDef/Everything.agda | 7 - Cubical/ZCohomology/OldDef/Properties.agda | 105 ---- .../ZCohomology/OldDef/cupProdPrelims.agda | 510 ------------------ .../ZCohomology/OldDef/cupProdPrelims2.agda | 262 --------- Cubical/ZCohomology/Properties.agda | 92 +++- 11 files changed, 222 insertions(+), 938 deletions(-) delete mode 100644 Cubical/ZCohomology/OldDef/Base.agda delete mode 100644 Cubical/ZCohomology/OldDef/Everything.agda delete mode 100644 Cubical/ZCohomology/OldDef/Properties.agda delete mode 100644 Cubical/ZCohomology/OldDef/cupProdPrelims.agda delete mode 100644 Cubical/ZCohomology/OldDef/cupProdPrelims2.agda diff --git a/Cubical/HITs/Pushout/Base.agda b/Cubical/HITs/Pushout/Base.agda index 38ab84273..252e83b34 100644 --- a/Cubical/HITs/Pushout/Base.agda +++ b/Cubical/HITs/Pushout/Base.agda @@ -4,6 +4,8 @@ module Cubical.HITs.Pushout.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Pointed +open import Cubical.Data.Prod open import Cubical.Data.Unit @@ -16,6 +18,20 @@ data Pushout {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} push : (a : A) → inl (f a) ≡ inr (g a) +--- + +_wedge_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Type (ℓ-max ℓ ℓ') +_wedge_ (A , ptA) (B , ptB) = Pushout {A = Unit} {B = A} {C = B} (λ {tt → ptA}) (λ {tt → ptB}) + +i∧ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → A wedge B → (typ A) × (typ B) +i∧ {A = A , ptA} {B = B , ptB} (inl x) = x , ptB +i∧ {A = A , ptA} {B = B , ptB} (inr x) = ptA , x +i∧ {A = A , ptA} {B = B , ptB} (push tt i) = ptA , ptB + +_smash_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Type (ℓ-max ℓ ℓ') +_smash_ A B = Pushout {A = (A wedge B)} {B = Unit} {C = (typ A) × (typ B)} (λ _ → tt) i∧ + + -- Suspension defined as a pushout PushoutSusp : ∀ {ℓ} (A : Type ℓ) → Type ℓ diff --git a/Cubical/HITs/Sn.agda b/Cubical/HITs/Sn.agda index a85d6f439..327c8823d 100644 --- a/Cubical/HITs/Sn.agda +++ b/Cubical/HITs/Sn.agda @@ -2,3 +2,4 @@ module Cubical.HITs.Sn where open import Cubical.HITs.Sn.Base public +open import Cubical.HITs.Sn.Properties public diff --git a/Cubical/HITs/Sn/Properties.agda b/Cubical/HITs/Sn/Properties.agda index 7d5348af0..4b707849d 100644 --- a/Cubical/HITs/Sn/Properties.agda +++ b/Cubical/HITs/Sn/Properties.agda @@ -3,15 +3,24 @@ module Cubical.HITs.Sn.Properties where open import Cubical.Data.Int open import Cubical.Data.Bool +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Function open import Cubical.Foundations.Equiv open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Univalence open import Cubical.HITs.S1 +open import Cubical.Data.Prod open import Cubical.HITs.Sn.Base open import Cubical.HITs.Susp +open import Cubical.HITs.Pushout +variable + ℓ ℓ' : Level + A : Type ℓ + B : Type ℓ' @@ -51,3 +60,98 @@ SuspBool≃S1 = isoToEquiv iso1 Iso.leftInv iso1 south = refl Iso.leftInv iso1 (merid false i) = refl Iso.leftInv iso1 (merid true i) = refl + + +private + f* : {a : A} → S¹ → A → Susp A + f* base b = north + f* {a = a} (loop i) = funExt (λ x → merid x ∙ sym (merid a)) i + + +proj : {A : Pointed ℓ} {B : Pointed ℓ'} → typ A → typ B → A smash B +proj a b = inr (a , b) + +projᵣ : {A : Pointed ℓ} {B : Pointed ℓ'} (a : typ A) → proj {A = A} a (pt B) ≡ inl tt +projᵣ a = sym (push (inl a)) + +projₗ : {A : Pointed ℓ} {B : Pointed ℓ'} (b : typ B) → proj {B = B} (pt A) b ≡ inl tt +projₗ b = sym (push (inr b)) + +projᵣₗ : {A : Pointed ℓ} {B : Pointed ℓ'} → projᵣ (pt A) ≡ projₗ (pt B) +projᵣₗ {A = A} {B = B} = cong sym (cong push (push tt)) + + +module S1-smash-Iso ((A , pt) : Pointed ℓ) where + fun : (S¹ , base) smash (A , pt) → (Susp A) + fun (inl tt) = north + fun (inr (x , x₁)) = f* {a = pt} x x₁ + fun (push (inl base) i) = north + fun (push (inl (loop i₁)) i) = rCancel (merid pt) (~ i) i₁ + fun (push (inr x) i) = north + fun (push (push a i₁) i) = north + + funInv : Susp A → (S¹ , base) smash (A , pt) + funInv north = inl tt + funInv south = inl tt + funInv (merid a i) = (sym (projₗ a) ∙ + cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ + projₗ a) i + + funSect : (x : Susp A) → fun (funInv x) ≡ x + funSect north = refl + funSect south = merid pt + funSect (merid a i) = hcomp (λ k → λ {(i = i0) → ((λ j → (refl ∙ refl) ∙ refl ∙ refl) ∙ + (λ j → rCancel refl j ∙ refl ∙ refl) ∙ + λ j → lUnit (lUnit refl (~ j)) (~ j)) k ; + (i = i1) → ((λ j → (refl ∙ refl) ∙ sym (refl ∙ (merid a ∙ sym (merid pt)) ∙ refl) ∙ merid a) ∙ + (λ j → lUnit refl (~ j) ∙ sym (lUnit (rUnit (merid a ∙ sym (merid pt)) (~ j)) (~ j)) ∙ merid a) ∙ + (λ j → lUnit (symDistr (merid a) (sym (merid pt)) j ∙ merid a) (~ j)) ∙ + (sym (assoc (merid pt) (sym (merid a)) (merid a))) ∙ + (λ j → merid pt ∙ lCancel (merid a) j) ∙ + sym (rUnit (merid pt)) ) k}) + (helper2 i ∙ filler i) + where + + filler : (i : I) → (refl ∙ (merid a ∙ sym (merid pt)) ∙ refl) i ≡ merid a i + filler i = (λ j → (refl ∙ (merid a ∙ sym (merid pt)) ∙ refl) (i ∧ (~ j))) ∙ λ j → merid a (j ∧ i) + + helper2 : (i : I) → fun ((sym (projₗ a) ∙ + cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ + projₗ a) i) + ≡ ((λ i → fun (sym (projₗ a) i)) ∙ + (λ i → fun (cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop i)) ∙ + λ i → fun (projₗ a i)) i + helper2 i = (λ j → congFunct fun (sym (projₗ a)) (cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ projₗ a) j i) ∙ + (λ j → (cong fun (sym (projₗ a)) ∙ congFunct fun (cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop) (projₗ a) j) i) + + + retrHelper : (x : S¹) (a : A) → funInv (f* {a = pt} x a) ≡ proj x a + retrHelper base a = sym (projₗ a) + retrHelper (loop i) a = hcomp (λ k → λ { (i = i0) → {!!} + ; (i = i1) → {!!} }) + filler + + where + filler : funInv (f* {a = pt} (loop i) a) ≡ proj (loop i) a + filler = (λ j → funInv (funExt (λ x → merid x ∙ sym (merid pt)) (i ∨ j) a)) ∙ sym (projₗ a) ∙ λ j → proj (loop (i ∧ j)) a + + id1 : funInv (f* {a = pt} (loop i) a) ≡ cong funInv (merid a ∙ sym (merid pt)) i + id1 = refl + + id2 : funInv (f* {a = pt} (loop i) a) ≡ ((sym (projₗ a) ∙ cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ (projₗ a)) ∙ sym (projₗ pt) ∙ sym (cong (λ x → proj {A = S¹ , base} {B = A , pt} x pt) loop) ∙ projₗ pt) i + id2 = (λ j → congFunct funInv (merid a) (sym (merid pt)) j i) ∙ + (λ j → (((sym (projₗ a) ∙ cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ (projₗ a))) ∙ symDistr (sym (projₗ pt)) (cong (λ x → proj {A = S¹ , base} {B = A , pt} x pt) loop ∙ projₗ pt) j) i) ∙ + (λ j → (((sym (projₗ a) ∙ cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ (projₗ a))) ∙ symDistr (cong (λ x → proj {A = S¹ , base} {B = A , pt} x pt) loop) (projₗ pt) j ∙ projₗ pt) i) ∙ + λ j → (((sym (projₗ a) ∙ cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ (projₗ a))) ∙ assoc (sym (projₗ pt)) (sym (cong (λ x → proj {A = S¹ , base} {B = A , pt} x pt) loop)) (projₗ pt) (~ j)) i + + id4 : sym (projᵣ base) ∙ cong ((λ x → proj {A = S¹ , base} {B = A , pt} x pt)) loop ∙ projᵣ base ≡ refl + id4 = (λ i → (push (inl (loop i))) ∙ (λ j → inr (loop (i ∨ j) , pt)) ∙ sym (push (inl base))) ∙ + (λ i → push (inl base) ∙ lUnit (sym (push (inl base))) (~ i)) ∙ + rCancel (push (inl base)) + + id3 : funInv ((merid pt) (~ i)) ≡ inl tt + id3 = (λ j → (sym (projₗ pt) ∙ (cong (λ x → proj {A = S¹ , base} {B = A , pt} x pt) loop) ∙ projₗ pt) (~ i)) ∙ + (λ j → (sym ((projᵣₗ (~ j))) ∙ cong (λ x → proj {A = S¹ , base} {B = A , pt} x pt) loop ∙ projᵣₗ (~ j)) (~ i)) ∙ + (λ j → id4 j (~ i)) + + diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda index d604424e1..6b18d4f06 100644 --- a/Cubical/Homotopy/Loopspace.agda +++ b/Cubical/Homotopy/Loopspace.agda @@ -41,12 +41,12 @@ exchange {A = A} {a = a} {c = c} {p = p} {r = r} {p' = p'} {r' = r'} α β α' hcomp (λ k → λ { (i = i0) → p ∙ p' ; (i = i1) → (β ⋆ β') (k ∨ ~ z) ; (z = i0) → ((α ∙ β) ⋆ (α' ∙ β')) i }) - (btm-Filler z i) + (bottom z i) where - btm-Filler : PathP (λ i → p ∙ p' ≡ (β ⋆ β') (~ i)) ((α ∙ β) ⋆ (α' ∙ β')) (α ⋆ α') - btm-Filler i = hcomp (λ k → λ { (i = i0) → (α ∙ β) ⋆ (α' ∙ β') - ; (i = i1) → rUnit α (~ k) ⋆ rUnit α' (~ k)}) - ((α ∙ λ j → β (~ i ∧ j)) ⋆ (α' ∙ λ j → β' (~ i ∧ j))) + bottom : PathP (λ i → p ∙ p' ≡ (β ⋆ β') (~ i)) ((α ∙ β) ⋆ (α' ∙ β')) (α ⋆ α') + bottom i = hcomp (λ k → λ { (i = i0) → (α ∙ β) ⋆ (α' ∙ β') + ; (i = i1) → rUnit α (~ k) ⋆ rUnit α' (~ k)}) + ((α ∙ λ j → β (~ i ∧ j)) ⋆ (α' ∙ λ j → β' (~ i ∧ j))) Eckmann-Hilton : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (α β : typ ((Ω^ (2 + n)) A)) @@ -89,7 +89,7 @@ Eckmann-Hilton {A = (A , pt)} zero α β = Eckmann-Hilton {A = (A , pt)} (suc n) α β = Eckmann-Hilton 0 α β -{- Composition in fundamental group -} +{- Homotopy group version -} π-comp : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → ∥ typ ((Ω^ (suc n)) A) ∥₀ → ∥ typ ((Ω^ (suc n)) A) ∥₀ → ∥ typ ((Ω^ (suc n)) A) ∥₀ π-comp n = elim2 (λ _ _ → setTruncIsSet) λ p q → ∣ p ∙ q ∣₀ diff --git a/Cubical/ZCohomology/Base.agda b/Cubical/ZCohomology/Base.agda index da1337bfe..00713e420 100644 --- a/Cubical/ZCohomology/Base.agda +++ b/Cubical/ZCohomology/Base.agda @@ -43,3 +43,9 @@ coHomK-ptd (suc n) = (coHomK (suc n) , ∣ north ∣) {- Reduced cohomology -} coHomRed : (n : ℕ) → (A : Pointed ℓ) → Type ℓ coHomRed n A = ∥ (A →∙ (coHomK-ptd n)) ∥₀ + + +--- +coHom-pt : (n : ℕ) → coHomK n +coHom-pt zero = pos 0 +coHom-pt (suc n) = ∣ north ∣ diff --git a/Cubical/ZCohomology/OldDef/Base.agda b/Cubical/ZCohomology/OldDef/Base.agda deleted file mode 100644 index b90cf962e..000000000 --- a/Cubical/ZCohomology/OldDef/Base.agda +++ /dev/null @@ -1,45 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.ZCohomology.OldDef.Base where - - -open import Cubical.Data.Int.Base -open import Cubical.Data.Nat.Base -open import Cubical.Data.NatMinusTwo.Base -open import Cubical.Data.Sigma - -open import Cubical.Foundations.Pointed.Base - -open import Cubical.HITs.Nullification.Base -open import Cubical.HITs.SetTruncation.Base -open import Cubical.HITs.Sn.Base -open import Cubical.HITs.Susp.Base -open import Cubical.HITs.Truncation.Base - -private - variable - ℓ : Level - A : Type ℓ - - ---- Cohomology --- - -{- Types Kₙ from Brunerie 2016 -} -coHomK : (n : ℕ) → Type₀ -coHomK zero = Int -coHomK (suc n) = ∥ S₊ (suc n) ∥ (ℕ→ℕ₋₂ (suc n)) - -{- Cohomology -} -coHom : (n : ℕ) → Type ℓ → Type ℓ -coHom n A = ∥ (A → coHomK n) ∥₀ - - ---- Reduced cohomology --- - -{- Pointed version of Kₙ -} -coHomK-ptd : (n : ℕ) → Pointed (ℓ-zero) -coHomK-ptd zero = coHomK zero , (pos 0) -coHomK-ptd (suc n) = (coHomK (suc n) , ∣ north ∣) - -{- Reduced cohomology -} -coHomRed : (n : ℕ) → (A : Pointed ℓ) → Type ℓ -coHomRed n A = ∥ (A →∙ (coHomK-ptd n)) ∥₀ diff --git a/Cubical/ZCohomology/OldDef/Everything.agda b/Cubical/ZCohomology/OldDef/Everything.agda deleted file mode 100644 index 160104dce..000000000 --- a/Cubical/ZCohomology/OldDef/Everything.agda +++ /dev/null @@ -1,7 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.ZCohomology.Everything where - -open import Cubical.ZCohomology.Base -open import Cubical.ZCohomology.Properties - -open import Cubical.ZCohomology.S1.S1 diff --git a/Cubical/ZCohomology/OldDef/Properties.agda b/Cubical/ZCohomology/OldDef/Properties.agda deleted file mode 100644 index acfb034c5..000000000 --- a/Cubical/ZCohomology/OldDef/Properties.agda +++ /dev/null @@ -1,105 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.ZCohomology.OldDef.Properties where - -open import Cubical.ZCohomology.OldDef.Base -open import Cubical.HITs.S1 -open import Cubical.HITs.Sn -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Univalence -open import Cubical.Data.NatMinusTwo.Base -open import Cubical.Data.Empty -open import Cubical.Data.Sigma -open import Cubical.HITs.Susp -open import Cubical.HITs.SetTruncation -open import Cubical.HITs.Nullification -open import Cubical.Data.Int hiding (_+_) -open import Cubical.Data.Nat -open import Cubical.HITs.Truncation renaming (elim to trElim) -open import Cubical.Homotopy.Loopspace -open import Cubical.Homotopy.Connected -open import Cubical.Homotopy.Freudenthal - -open import Cubical.HITs.Pushout -open import Cubical.Data.Sum.Base -open import Cubical.Data.HomotopyGroup -open import Cubical.Data.Bool - -open import Cubical.ZCohomology.OldDef.cupProdPrelims renaming (Kn→ΩKn+1 to Kn→ΩKn+1') - -private - variable - ℓ ℓ' : Level - A : Type ℓ - B : Type ℓ' - -{- Equivalence between cohomology of A and reduced cohomology of (A + 1) -} -coHomRed+1Equiv : (n : ℕ) → - (A : Set ℓ) → - (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) -coHomRed+1Equiv zero A i = ∥ helpLemma {C = (Int , pos 0)} i ∥₀ - module coHomRed+1 where - helpLemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →∙ C))) - helpLemma {C = C} = isoToPath (iso map1 - map2 - (λ b → linvPf b) - (λ _ → refl)) - where - map1 : (A → typ C) → ((((A ⊎ Unit) , inr (tt)) →∙ C)) - map1 f = map1' , refl module helpmap where - map1' : A ⊎ Unit → fst C - map1' (inl x) = f x - map1' (inr x) = pt C - - map2 : ((((A ⊎ Unit) , inr (tt)) →∙ C)) → (A → typ C) - map2 (g , pf) x = g (inl x) - - linvPf : (b :((((A ⊎ Unit) , inr (tt)) →∙ C))) → map1 (map2 b) ≡ b - linvPf (f , snd) i = (λ x → helper x i) , λ j → snd ((~ i) ∨ j) - where - helper : (x : A ⊎ Unit) → ((helpmap.map1') (map2 (f , snd)) x) ≡ f x - helper (inl x) = refl - helper (inr tt) = sym snd - -coHomRed+1Equiv (suc n) A i = ∥ coHomRed+1.helpLemma A i {C = ((coHomK (suc n)) , ∣ north ∣)} i ∥₀ - - - --- ------------------- --- {- This section contains a proof of Kn≃ΩKn+1, which is needed for the cup product -} - --- {- Compiles slowly right now. Possible improvements : --- 1. Use S¹ instead of S 1 in definition of K. --- 2. Use copatterns -} - --- Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) --- Kn→ΩKn+1 = Kn→ΩKn+1' - --- isEquivKn→ΩKn+1 : (n : ℕ) → isEquiv (Kn→ΩKn+1 n) --- isEquivKn→ΩKn+1 zero = compEquiv (looper , isEquivLooper) (cong ∣_∣ , isEquivHelper hLevl3) .snd --- where --- hLevl3 : (x y : S₊ 1) (p q : x ≡ y) → isProp (p ≡ q) --- hLevl3 x y = J (λ y p → (q : x ≡ y) → isProp (p ≡ q) ) --- (transport (λ i → isSet (helper (~ i))) isSetInt refl) --- where --- helper : (x ≡ x) ≡ Int --- helper = (λ i → transp (λ j → ua S1≃SuspBool (~ j ∨ ~ i)) (~ i) x ≡ transp (λ j → ua S1≃SuspBool (~ j ∨ ~ i)) (~ i) x) ∙ --- (λ i → transp (λ j → S¹≡SuspBool (~ j ∨ ~ i)) (~ i) (transport (sym (ua S1≃SuspBool)) x) ≡ transp (λ j → S¹≡SuspBool (~ j ∨ ~ i)) (~ i) (transport (sym (ua S1≃SuspBool)) x)) ∙ --- basedΩS¹≡Int (transport (sym S¹≡SuspBool) (transport (sym (ua S1≃SuspBool)) x)) --- isEquivKn→ΩKn+1 (suc zero) = transport (λ i → isEquiv (Kn→ΩKn+1Sucn zero (~ i))) --- (compEquiv (trElim (λ _ → isOfHLevelTrunc (2 + (suc zero))) (λ a → ∣ ϕ north a ∣) , --- isEquiv∣ϕ∣) --- (truncEquivΩ (ℕ→ℕ₋₂ (suc zero))) .snd) --- isEquivKn→ΩKn+1 (suc (suc n)) = transport (λ i → isEquiv (Kn→ΩKn+1Sucn (suc n) (~ i))) --- (compEquiv (connectedTruncEquiv2 (4 + n) _ (ϕ north) (n , λ i → suc (suc (suc (+-suc n n (~ i))))) ?) --- (truncEquivΩ (ℕ→ℕ₋₂ (suc (suc n)))) .snd) - --- Kn≃ΩKn+1 : (n : ℕ) → coHomK n ≃ typ (Ω (coHomK-ptd (suc n))) --- Kn≃ΩKn+1 n = Kn→ΩKn+1 n , isEquivKn→ΩKn+1 n - --- ΩKn+1→Kn : (n : ℕ) → typ (Ω (coHomK-ptd (suc n))) → coHomK n --- ΩKn+1→Kn n a = equiv-proof (isEquivKn→ΩKn+1 n) a .fst .fst diff --git a/Cubical/ZCohomology/OldDef/cupProdPrelims.agda b/Cubical/ZCohomology/OldDef/cupProdPrelims.agda deleted file mode 100644 index ed5d8fb8d..000000000 --- a/Cubical/ZCohomology/OldDef/cupProdPrelims.agda +++ /dev/null @@ -1,510 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.ZCohomology.OldDef.cupProdPrelims where - -open import Cubical.ZCohomology.OldDef.Base -open import Cubical.HITs.S1 -open import Cubical.HITs.S2 -open import Cubical.HITs.S3 -open import Cubical.HITs.Sn -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Function -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Transport -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.Equiv.HalfAdjoint -open import Cubical.Data.NatMinusTwo.Base -open import Cubical.Data.Empty -open import Cubical.Data.Sigma hiding (_×_) -open import Cubical.Data.Prod.Base -open import Cubical.HITs.Susp -open import Cubical.HITs.Nullification -open import Cubical.Data.Int renaming (_+_ to +Int) -open import Cubical.Data.Nat -open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap) -open import Cubical.HITs.Hopf -open import Cubical.Homotopy.Connected -open import Cubical.Homotopy.Freudenthal -open import Cubical.Homotopy.Loopspace - -open import Cubical.HITs.Pushout -open import Cubical.Data.Sum.Base -open import Cubical.Data.HomotopyGroup -open import Cubical.Data.Bool -private - variable - ℓ ℓ' : Level - A : Type ℓ - B : Type ℓ' - -isIso : (A → B) → Type _ -isIso {A = A} {B = B} f = Σ[ g ∈ (B → A) ] section f g × retract f g - -{- Some useful lemmas -- should probably be moved -} -congFunct : {a b c : A} (f : A → B) (p : a ≡ b) (q : b ≡ c) → cong f (p ∙ q) ≡ cong f p ∙ cong f q -congFunct f p q i = hcomp (λ j → λ{(i = i0) → rUnit (cong f (p ∙ q)) (~ j) ; - (i = i1) → cong f (rUnit p (~ j)) ∙ cong f q}) - (cong f (p ∙ (λ k → q (k ∧ (~ i)))) ∙ cong f λ k → q ((~ i) ∨ k) ) - -- -symDistr : {a b c : A} (p : a ≡ b) (q : b ≡ c) → sym (p ∙ q) ≡ sym q ∙ sym p -symDistr p q i = hcomp (λ j → λ{(i = i0) → rUnit (sym (p ∙ q)) (~ j) ; - (i = i1) → sym (lUnit q (~ j)) ∙ sym p}) - (sym ((λ k → p (k ∨ i)) ∙ q) ∙ sym λ k → p (i ∧ k)) - -{- We want to prove that Kn≃ΩKn+1. For this we use the map ϕ-} - - -ϕ : (pt a : A) → typ (Ω (Susp A , north)) -ϕ pt a = (merid a) ∙ sym (merid pt) - -{- To define the map for n=0 we use the λ k → loopᵏ map for S₊ 1. The loop is given by ϕ south north -} - - -looper : Int → Path (S₊ 1) north north -looper (pos zero) = refl -looper (pos (suc n)) = looper (pos n) ∙ (ϕ south north) -looper (negsuc zero) = sym (ϕ south north) -looper (negsuc (suc n)) = looper (negsuc n) ∙ sym (ϕ south north) - -{- The map of Kn≃ΩKn+1 is given as follows. -} -Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) -Kn→ΩKn+1 zero x = cong ∣_∣ (looper x) -Kn→ΩKn+1 (suc n) = trElim (λ x → (isOfHLevelTrunc (2 + (suc (suc n))) ∣ north ∣ ∣ north ∣)) - λ a → cong ∣_∣ ((merid a) ∙ (sym (merid north))) - -{- We want to show that this map is an equivalence. n ≥ 2 follows from Freudenthal, and -} - -{- -We want to show that the function (looper : Int → S₊ 1) defined by λ k → loopᵏ is an equivalece. We already know that the corresponding function (intLoop : Int → S¹ is) an equivalence, -so the idea is to show that when intLoop is transported along a suitable path S₊ 1 ≡ S¹ we get looper. Instead of using S₊ 1 straight away, we begin by showing this for the equivalent Susp Bool. --} - --- loop for Susp Bool -- -loop* : Path (Susp Bool) north north -loop* = merid false ∙ sym (merid true) - --- the loop function -- -intLoop* : Int → Path (Susp Bool) north north -intLoop* (pos zero) = refl -intLoop* (pos (suc n)) = intLoop* (pos n) ∙ loop* -intLoop* (negsuc zero) = sym loop* -intLoop* (negsuc (suc n)) = intLoop* (negsuc n) ∙ sym loop* - --- we show that the loop spaces agree -- -loopSpId : ΩS¹ ≡ Path (Susp Bool) north north -loopSpId i = typ (Ω (S¹≡SuspBool i , transp ((λ j → S¹≡SuspBool (j ∧ i))) (~ i) base)) - --- the transport map -- -altMap2 : Int → Path (Susp Bool) north north -altMap2 n i = transport S¹≡SuspBool (intLoop n i) - --- We show that the transporting intLoop over S¹≡SuspBool gives intLoop* (modulo function extensionality) -- -altMap≡intLoop*2 : (x : Int) → intLoop* x ≡ altMap2 x -altMap≡intLoop*2 (pos zero) = refl -altMap≡intLoop*2 (pos (suc n)) = (λ i → (altMap≡intLoop*2 (pos n) i) ∙ loop*) ∙ - sym (helper n) - - where - helper : (n : ℕ) → altMap2 (pos (suc n)) ≡ altMap2 (pos n) ∙ loop* - helper zero = (λ i j → transport S¹≡SuspBool (lUnit loop (~ i) j)) ∙ - (λ i j → transport S¹≡SuspBool (loop j)) ∙ - (λ i j → transportRefl ((merid false ∙ (sym (merid true))) j) i ) ∙ - lUnit loop* - helper (suc n) = anotherHelper n ∙ - (λ i → altMap2 (pos (suc n)) ∙ helper zero i) ∙ - (λ i → altMap2 (pos (suc n)) ∙ lUnit loop* (~ i)) - where - anotherHelper : (n : ℕ) → altMap2 (pos (suc (suc n))) ≡ altMap2 (pos (suc n)) ∙ altMap2 (pos 1) - anotherHelper n = ((λ i j → transport S¹≡SuspBool ((intLoop (pos (suc n)) ∙ loop) j))) ∙ - rUnit (λ j → transport S¹≡SuspBool ((intLoop (pos (suc n)) ∙ loop) j) ) ∙ - (λ i → (λ j → transport S¹≡SuspBool ((intLoop (pos (suc n)) ∙ λ k → loop (k ∧ (~ i))) j)) ∙ λ j → transport S¹≡SuspBool (loop (j ∨ (~ i)))) ∙ - (λ i → (λ j → transport S¹≡SuspBool (rUnit (intLoop (pos (suc n))) (~ i) j)) ∙ λ j → transport S¹≡SuspBool ((lUnit loop i) j)) - -altMap≡intLoop*2 (negsuc zero) = sym ((λ i j → transport S¹≡SuspBool (loop (~ j))) ∙ - λ i j → transportRefl (loop* (~ j)) i ) -altMap≡intLoop*2 (negsuc (suc n)) = helper n - where - altMapneg1 : altMap2 (negsuc zero) ≡ sym (loop*) - altMapneg1 i j = transportRefl (loop* (~ j)) i - - anotherHelper : (n : ℕ) → altMap2 (negsuc (suc n)) ≡ altMap2 (negsuc n) ∙ altMap2 (negsuc zero) - anotherHelper n = ((λ i → rUnit (λ j → (transport S¹≡SuspBool ((intLoop (negsuc n) ∙ sym loop) j))) i)) ∙ - ((λ i → (λ j → transport S¹≡SuspBool ((intLoop (negsuc n) ∙ (λ k → loop ((~ k) ∨ i))) j)) ∙ λ j → transport S¹≡SuspBool (loop ((~ j) ∧ i)))) ∙ - (λ i → ((λ j → transport S¹≡SuspBool (rUnit (intLoop (negsuc n)) (~ i) j))) ∙ altMap2 (negsuc zero)) - - helper : (n : ℕ) → intLoop* (negsuc n) ∙ (sym loop*) ≡ altMap2 (negsuc (suc n)) - helper zero = (λ i → altMapneg1 (~ i) ∙ altMapneg1 (~ i)) ∙ sym (anotherHelper zero) - helper (suc n) = (λ i → (helper n i) ∙ altMapneg1 (~ i) ) ∙ - sym (anotherHelper (suc n)) - - -isIsoAltMap2 : Iso Int (Path (Susp Bool) north north) -isIsoAltMap2 = compIso iso1 iso2 - where - iso1 : Iso Int ΩS¹ - Iso.fun iso1 = intLoop - Iso.inv iso1 = winding - Iso.rightInv iso1 = decodeEncode base - Iso.leftInv iso1 = windingIntLoop - iso2 : Iso ΩS¹ (Path (Susp Bool) north north) - Iso.fun iso2 = cong (transport S¹≡SuspBool) -- λ x i → transport S¹≡SuspBool (x i) - Iso.inv iso2 = Iso.inv (congIso S¹≃SuspBool) - Iso.rightInv iso2 b = (λ i → cong (funExt (λ z → transportRefl (S¹→SuspBool z)) i) (Iso.inv (congIso {x = base} {y = base} S¹≃SuspBool) b)) ∙ - Iso.rightInv (congIso {x = base} {y = base} S¹≃SuspBool) b - Iso.leftInv iso2 b = (λ i → Iso.inv (congIso S¹≃SuspBool) (cong (funExt (λ z → transportRefl (S¹→SuspBool z)) i) b)) ∙ - Iso.leftInv (congIso {x = base} {y = base} S¹≃SuspBool) b - -{- We have that the transported map is an equivalence -} -mapIsEquiv : isEquiv altMap2 -mapIsEquiv = compEquiv (intLoop , (isoToIsEquiv (iso intLoop winding (decodeEncode base) windingIntLoop))) ((λ x i → transport S¹≡SuspBool (x i)) , helper) .snd - where - helper : isEquiv {A = ΩS¹} {B = _≡_ {A = Susp Bool} north north} (λ x i → transport S¹≡SuspBool (x i)) - helper = congEquiv (transport S¹≡SuspBool , helper3) .snd - where - helper2 : transport S¹≡SuspBool ≡ S¹→SuspBool - helper2 = funExt λ z → transportRefl (S¹→SuspBool z) - helper3 : isEquiv (transport S¹≡SuspBool ) - helper3 = transport (λ i → isEquiv (helper2 (~ i))) (S¹≃SuspBool .snd) - -{- From this we get that intLoop* is an equivalence-} -intLoop*Equiv : isEquiv intLoop* -intLoop*Equiv = transport (λ i → isEquiv (funExt (altMap≡intLoop*2) (~ i))) mapIsEquiv - -{- We now only need to work with Susp Bool and S₊ 1. We transport intLoop* along a path S1≡SuspBool and show that this gives us looper. -} -SuspBool→S1 : Susp Bool → S₊ 1 -SuspBool→S1 north = north -SuspBool→S1 south = south -SuspBool→S1 (merid false i) = merid north i -SuspBool→S1 (merid true i) = merid south i - -S1→SuspBool : S₊ 1 → Susp Bool -S1→SuspBool north = north -S1→SuspBool south = south -S1→SuspBool (merid north i) = merid false i -S1→SuspBool (merid south i) = merid true i - -{- We need to spell out the trivial equivalence S1≃SuspBool. Of course it's important to make sure that we choose the right version of it. -} -S1≃SuspBool : Susp Bool ≃ S₊ 1 -S1≃SuspBool = isoToEquiv (iso SuspBool→S1 S1→SuspBool retrHelper sectHelper) - where - sectHelper : section S1→SuspBool SuspBool→S1 - sectHelper north = refl - sectHelper south = refl - sectHelper (merid false i) = refl - sectHelper (merid true i) = refl - - retrHelper : retract S1→SuspBool SuspBool→S1 - retrHelper north = refl - retrHelper south = refl - retrHelper (merid north i) = refl - retrHelper (merid south i) = refl - -{- We show that transporting intLoop* along (ua S1≃SuspBool) gives us looper (again, modulo function extensionality) -} -looperIntoBool : (x : Int) → looper x ≡ λ i → transport ((ua (S1≃SuspBool))) (intLoop* x i) -looperIntoBool (pos zero) = refl -looperIntoBool (pos (suc n)) = (λ j → looperIntoBool (pos n) j ∙ merid north ∙ sym (merid south)) ∙ - (λ j → (λ i → transportRefl (SuspBool→S1 (intLoop* (pos n) i)) j ) ∙ merid north ∙ sym (merid south)) ∙ - (λ j → cong SuspBool→S1 (intLoop* (pos n)) ∙ congFunct SuspBool→S1 (merid false) (sym (merid true)) (~ j)) ∙ - sym (congFunct SuspBool→S1 (intLoop* (pos n)) loop*) ∙ - λ j i → transportRefl (SuspBool→S1 ((intLoop* (pos n) ∙ loop*) i)) (~ j) -looperIntoBool (negsuc zero) = symDistr (merid north) (sym (merid south)) ∙ - sym (congFunct SuspBool→S1 (merid true) (sym (merid false))) ∙ - (λ j → cong SuspBool→S1 (merid true ∙ sym (merid false))) ∙ - (λ j → cong SuspBool→S1 (symDistr (merid false) (sym (merid true)) (~ j))) ∙ - λ j i → transportRefl (SuspBool→S1 (loop* (~ i))) (~ j) -looperIntoBool (negsuc (suc n)) = (λ i → looperIntoBool (negsuc n) i ∙ sym ((merid north) ∙ (sym (merid south))) ) ∙ - (λ i → looperIntoBool (negsuc n) i1 ∙ symDistr (merid north) (sym (merid south)) i) ∙ - (λ j → (λ i → transportRefl (SuspBool→S1 (intLoop* (negsuc n) i)) j) ∙ merid south ∙ sym (merid north)) ∙ - (λ j → cong SuspBool→S1 (intLoop* (negsuc n)) ∙ congFunct SuspBool→S1 (merid true) (sym (merid false)) (~ j)) ∙ - ((λ j → cong SuspBool→S1 (intLoop* (negsuc n)) ∙ cong SuspBool→S1 (symDistr (merid false) (sym (merid true)) (~ j)))) ∙ - sym (congFunct SuspBool→S1 (intLoop* (negsuc n)) (sym loop*)) ∙ - λ j i → transportRefl (SuspBool→S1 ((intLoop* (negsuc n) ∙ sym loop*) i)) (~ j) - -{- From this, we can finally deduce that looper is an equivalence-} -isEquivLooper : isEquiv looper -isEquivLooper = transport (λ i → isEquiv (funExt (looperIntoBool) (~ i))) isEquivTranspIntLoop - where - isEquivTranspIntLoop : isEquiv λ x → cong (transport ((ua (S1≃SuspBool)))) (intLoop* x) - isEquivTranspIntLoop = compEquiv (intLoop* , intLoop*Equiv) - (cong (transport (ua S1≃SuspBool)) , - congEquiv (transport (ua S1≃SuspBool) , - transportEquiv (ua S1≃SuspBool) .snd) .snd) .snd - -isIsoLooper : Iso Int (Path (S₊ 1) north north) -isIsoLooper = compIso isIsoAltMap2 (congIso S1≃SuspBool) - where - iso2 : Iso (Path (Susp Bool) north north) (Path (S₊ 1) north north) - iso2 = congIso S1≃SuspBool - -isIsoLooperId : (n : ℕ) → Iso.fun isIsoLooper (pos (suc n)) ≡ Iso.fun isIsoLooper (pos n) ∙ Iso.fun isIsoLooper (pos 1) -isIsoLooperId n = (λ i → Iso.fun (congIso S1≃SuspBool) (congFunct (transport S¹≡SuspBool) (intLoop (pos n)) loop i)) ∙ - congFunct (S1≃SuspBool .fst) (cong (transport S¹≡SuspBool) (intLoop (pos n))) (cong (transport S¹≡SuspBool) loop) ∙ - (λ i → cong (S1≃SuspBool .fst) (cong (transport S¹≡SuspBool) (intLoop (pos n))) ∙ cong (S1≃SuspBool .fst) (cong (transport S¹≡SuspBool) (lUnit loop i))) - -helper : (x : Int) → looper x ≡ Iso.fun isIsoLooper x -helper (pos zero) = refl -helper (pos (suc zero)) = sym (lUnit (ϕ south north)) ∙ (λ i → merid north ∙ sym (merid south)) ∙ {!!} ∙ {!Iso.fun isIsoLooper (pos (suc zero))!} ∙ λ i → Iso.fun (congIso S1≃SuspBool) (cong (transport S¹≡SuspBool) (lUnit loop i)) -helper (pos (suc (suc n))) = {!!} ∙ {!!} -helper (negsuc n) = {!!} - where - helper2 : (x : SuspBool) → SuspBool→S1 ≡ {!!} - helper2 = {!!} - -isIsoLooper2 : Iso Int (Path (S₊ 1) north north) -Iso.fun isIsoLooper2 = looper -Iso.inv isIsoLooper2 = Iso.inv isIsoLooper -Iso.rightInv isIsoLooper2 b = {!SuspBool→S1 - (transp (λ i₂ → SuspBool) i0 (S¹→SuspBool (intLoop x i₁)))!} ∙ Iso.rightInv isIsoLooper b -Iso.leftInv isIsoLooper2 b = {!!} ∙ Iso.leftInv isIsoLooper b - --- ----------------------------------- n = 1 ----------------------------------------------------- - --- {- We begin by stating some useful lemmas -} - --- sphereConnectedSpecCase : isHLevelConnected 4 (Susp (Susp S¹)) -- is- 2 -ConnectedType (Susp (Susp S¹)) --- sphereConnectedSpecCase = transport (λ i → isHLevelConnected 4 (helper i)) (sphereConnected 3) --- where --- helper : S₊ 3 ≡ Susp (Susp S¹) --- helper = (λ i → Susp (Susp (Susp (ua Bool≃Susp⊥ (~ i))))) ∙ λ i → Susp (Susp (S¹≡SuspBool (~ i))) - --- S¹≡S1 : S₊ 1 ≡ S¹ --- S¹≡S1 = (λ i → Susp (ua (Bool≃Susp⊥) (~ i))) ∙ sym S¹≡SuspBool - --- S³≡SuspSuspS¹ : S³ ≡ Susp (Susp S¹) --- S³≡SuspSuspS¹ = S³≡SuspS² ∙ λ i → Susp (S²≡SuspS¹ i) - --- isSetS1 : isSet (Path (S₊ 1) north north) --- isSetS1 = transport (λ i → isSet (helper i)) isSetInt --- where --- helper : Int ≡ (Path (S₊ 1) north north) --- helper = sym ΩS¹≡Int ∙ --- (λ i → typ (Ω (S¹≡SuspBool i , transport (λ j → S¹≡SuspBool (j ∧ i)) base))) ∙ --- (λ i → typ (Ω (ua S1≃SuspBool i , transport (λ j → ua S1≃SuspBool (i ∧ j)) north))) - --- isEquivHelper2 : isOfHLevel 3 A → isEquiv {B = ∥ A ∥ 1} ∣_∣ --- isEquivHelper2 ofHlevl = --- isoToIsEquiv (iso ∣_∣ --- (trElim (λ _ → ofHlevl) (λ a → a)) --- (trElim {B = λ b → ∣ trElim (λ _ → ofHlevl) (λ a₁ → a₁) b ∣ ≡ b} --- (λ b → isOfHLevelSuc 3 (isOfHLevelTrunc 3) --- ∣ trElim (λ _ → ofHlevl) (λ a₁ → a₁) b ∣ b) --- λ a → refl) --- λ b → refl) - --- isEquivHelper : {a b : A} → isOfHLevel 3 A → isEquiv {B = Path (∥ A ∥ 1) ∣ a ∣ ∣ b ∣ } (cong ∣_∣) --- isEquivHelper {A = A} {a = a} {b = b} ofHlevl = congEquiv (∣_∣ , isEquivHelper2 ofHlevl) .snd - --- {- We give the following map and show that it is an equivalence -} - --- d-map : typ (Ω ((Susp S¹) , north)) → S¹ --- d-map p = subst HopfSuspS¹ p base - --- d-mapId : (r : S¹) → d-map (ϕ base r) ≡ r --- d-mapId r = substComposite HopfSuspS¹ (merid r) (sym (merid base)) base ∙ --- rotLemma r --- where --- rotLemma : (r : S¹) → rot r base ≡ r --- rotLemma base = refl --- rotLemma (loop i) = refl - --- d-mapComp : fiber d-map base ≡ Path (Susp (Susp S¹)) north north --- d-mapComp = sym (pathSigma≡sigmaPath {B = HopfSuspS¹} _ _) ∙ helper --- where --- helper : Path (Σ (Susp S¹) λ x → HopfSuspS¹ x) (north , base) (north , base) ≡ Path (Susp (Susp S¹)) north north --- helper = (λ i → (_≡_ {A = S³≡TotalHopf (~ i)} --- (transp (λ j → S³≡TotalHopf (~ i ∨ ~ j)) (~ i) (north , base)) --- ((transp (λ j → S³≡TotalHopf (~ i ∨ ~ j)) (~ i) (north , base))))) ∙ --- (λ i → _≡_ {A = S³} base base) ∙ --- (λ i → _≡_ {A = S³≡SuspSuspS¹ i} (transp (λ j → S³≡SuspSuspS¹ (i ∧ j)) (~ i) base) ((transp (λ j → S³≡SuspSuspS¹ (i ∧ j)) (~ i) base))) - - --- is1Connected-dmap : isHLevelConnectedFun 3 d-map --- is1Connected-dmap = toSetelim (λ s → isPropIsOfHLevel 0) (transport (λ j → isContr (∥ d-mapComp (~ j) ∥ ℕ→ℕ₋₂ 1)) --- (transport (λ i → isContr (PathΩ {A = Susp (Susp S¹)} {a = north} (ℕ→ℕ₋₂ 1) i)) --- (refl , isOfHLevelSuc 1 (isOfHLevelSuc 0 sphereConnectedSpecCase) ∣ north ∣ ∣ north ∣ (λ _ → ∣ north ∣)))) - --- d-equiv : isEquiv {A = ∥ typ (Ω (Susp S¹ , north)) ∥ (ℕ→ℕ₋₂ 1)} {B = ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (trElim (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣ ) --- d-equiv = connectedTruncEquiv _ d-map is1Connected-dmap .snd - --- d-iso2 : Iso (∥ typ (Ω (Susp S¹ , north)) ∥ (ℕ→ℕ₋₂ 1)) (∥ S¹ ∥ (ℕ→ℕ₋₂ 1)) --- Iso.fun d-iso2 = trMap d-map --- Iso.inv d-iso2 = Iso.inv (connectedTruncIso _ d-map is1Connected-dmap) --- Iso.rightInv d-iso2 = Iso.rightInv (connectedTruncIso _ d-map is1Connected-dmap) --- Iso.leftInv d-iso2 = Iso.leftInv (connectedTruncIso _ d-map is1Connected-dmap) - --- d-iso : isIso {A = ∥ typ (Ω (Susp S¹ , north)) ∥ (ℕ→ℕ₋₂ 1)} {B = ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (trElim (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣ ) --- d-iso = (Iso.inv (connectedTruncIso _ d-map is1Connected-dmap)) , (Iso.rightInv (connectedTruncIso _ d-map is1Connected-dmap) --- , Iso.leftInv (connectedTruncIso _ d-map is1Connected-dmap)) - --- {- We show that composing (λ a → ∣ ϕ base a ∣) and (λ x → ∣ d-map x ∣) gives us the identity function. -} - --- d-mapId2 : (λ (x : ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)) → (trElim {n = 3} {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) --- (trElim (λ _ → isOfHLevelTrunc 3) (λ a → ∣ ϕ base a ∣) x)) ≡ λ x → x --- d-mapId2 = funExt (trElim (λ x → isOfHLevelSuc 2 (isOfHLevelTrunc 3 ((trElim {n = 3} --- {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} --- (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) --- (trElim (λ _ → isOfHLevelTrunc 3) --- (λ a → ∣ ϕ base a ∣) x)) x)) --- λ a i → ∣ d-mapId a i ∣) - --- {- This means that (λ a → ∣ ϕ base a ∣) is an equivalence -} --- isEquiv∣ϕ-base∣ : isEquiv {A = ∥ S¹ ∥ ℕ→ℕ₋₂ 1} (trElim (λ _ → isOfHLevelTrunc 3) (λ a → ∣ ϕ base a ∣)) --- isEquiv∣ϕ-base∣ = composesToId→Equiv (trElim {n = 3} {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) --- (trElim (λ _ → isOfHLevelTrunc 3) (λ a → ∣ ϕ base a ∣)) --- d-mapId2 --- d-equiv - --- isIso∣ϕ-base∣ : Iso (∥ S¹ ∥ (ℕ→ℕ₋₂ 1)) (∥ typ (Ω (Susp S¹ , north)) ∥ ℕ→ℕ₋₂ 1) --- isIso∣ϕ-base∣ = composesToId→Iso d-iso2 (trMap (ϕ base)) d-mapId2 - - --- --------------------------------- --- -- We cheat when n = 1 and use J to prove the following lemmma. There is an obvious dependent path between ϕ base and ϕ north. Since the first one is an equivalence, so is the other. --- -- - - --- pointFunIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ} {C : (A : Type ℓ) (a1 : A) → Type ℓ'} (p : A ≡ B) (a : A) (b : B) → --- (transport p a ≡ b) → --- (f : (A : Type ℓ) → --- (a1 : A) (a2 : ∥ A ∥ 1) → C A a1) → --- isIso (f A a) → --- isIso (f B b) --- pointFunIso {ℓ = ℓ}{A = A} {B = B} {C = C} = --- J (λ B p → (a : A) (b : B) → --- (transport p a ≡ b) → --- (f : (A : Type ℓ) → --- (a1 : A) (a2 : ∥ A ∥ 1) → C A a1) → --- isIso (f A a) → --- isIso (f B b)) --- λ a b trefl f is → transport (λ i → isIso (f A ((sym (transportRefl a) ∙ trefl) i))) is - --- {- By pointFunEquiv, this gives that λ a → ∣ ϕ north a ∣ is an equivalence. -} - --- isIso∣ϕ∣ : isIso {A = ∥ S₊ 1 ∥ 1} {B = ∥ typ (Ω (S₊ 2 , north)) ∥ 1} (trMap (ϕ north)) --(trElim (λ _ → isOfHLevelTrunc 3) (λ a → ∣ ϕ north a ∣)) --- isIso∣ϕ∣ = pointFunIso {A = S¹} (λ i → S¹≡S1 (~ i)) base north refl (λ A a1 → trMap (ϕ a1)) ((Iso.inv isIso∣ϕ-base∣) , ((Iso.rightInv isIso∣ϕ-base∣) , (Iso.leftInv isIso∣ϕ-base∣))) - --- Iso∣ϕ∣ : Iso (∥ S₊ 1 ∥ 1) (∥ typ (Ω (S₊ 2 , north)) ∥ 1) --- Iso.fun Iso∣ϕ∣ = trMap (ϕ north) --- Iso.inv Iso∣ϕ∣ = isIso∣ϕ∣ .fst --- Iso.rightInv Iso∣ϕ∣ = proj₁ (isIso∣ϕ∣ .snd) --- Iso.leftInv Iso∣ϕ∣ = proj₂ (isIso∣ϕ∣ .snd) - --- -- isEquiv∣ϕ∣ : isEquiv {A = ∥ S₊ 1 ∥ 1} {B = ∥ typ (Ω (S₊ 2 , north)) ∥ 1} (trElim (λ _ → isOfHLevelTrunc 3) (λ a → ∣ ϕ north a ∣)) --- -- isEquiv∣ϕ∣ = pointFunEquiv {A = S¹} (λ i → S¹≡S1 (~ i)) base north refl (λ A a1 → trElim (λ _ → isOfHLevelTrunc 3) (λ a → ∣ ϕ a1 a ∣)) isEquiv∣ϕ-base∣ - --- -- ---------------------------------------------------- Finishing up --------------------------------- - - - --- decode-fun2 : (n : ℕ) → (x : A) → hLevelTrunc n (x ≡ x) → Path (hLevelTrunc (suc n) A) ∣ x ∣ ∣ x ∣ --- decode-fun2 zero x = trElim (λ _ → isOfHLevelPath 0 (∣ x ∣ , isOfHLevelTrunc 1 ∣ x ∣) ∣ x ∣ ∣ x ∣) (λ p i → ∣ p i ∣) --- decode-fun2 (suc n) x = trElim (λ _ → isOfHLevelPath' (suc n) (isOfHLevelTrunc (suc (suc n))) ∣ x ∣ ∣ x ∣) (cong ∣_∣) - --- r : {n : ℕ} (x : A) → (Path (hLevelTrunc (suc n) A) ∣ x ∣ ∣ x ∣) -- (u : ∥ A ∥ (suc₋₂ m)) → P u u --- r x = (λ _ → ∣ x ∣) - --- -- r : {m : ℕ₋₂} (u : ∥ B ∥ (suc₋₂ m)) → P u u --- -- r = elim (λ x → hLevelP x x) (λ a → ∣ refl ∣) - --- funsAreSame : (n : ℕ) (x : A) (b : hLevelTrunc n (x ≡ x)) → (decode-fun2 n x b) ≡ (ΩTrunc.decode-fun ∣ x ∣ ∣ x ∣ b) --- funsAreSame zero x = trElim (λ a → isOfHLevelPath 0 (refl , (isOfHLevelSuc 1 (isOfHLevelTrunc 1) ∣ x ∣ ∣ x ∣ refl)) _ _) λ a → refl --- funsAreSame (suc n) x = trElim (λ a → isOfHLevelPath _ (isOfHLevelPath' (suc n) (isOfHLevelTrunc (suc (suc n))) ∣ x ∣ ∣ x ∣) _ _) λ a → refl - --- decodeIso : (n : ℕ) → (x : A) → Iso (hLevelTrunc n (x ≡ x)) (Path (hLevelTrunc (suc n) A) ∣ x ∣ ∣ x ∣) --- Iso.fun (decodeIso n x) = decode-fun2 n x --- Iso.inv (decodeIso n x) = ΩTrunc.encode-fun ∣ x ∣ ∣ x ∣ --- Iso.rightInv (decodeIso n x) b = funExt⁻ (funExt (funsAreSame n x)) (ΩTrunc.encode-fun ∣ x ∣ ∣ x ∣ b) ∙ ΩTrunc.P-rinv ∣ x ∣ ∣ x ∣ b --- Iso.leftInv (decodeIso n x) b = cong (ΩTrunc.encode-fun ∣ x ∣ ∣ x ∣) (funsAreSame n x b) ∙ ΩTrunc.P-linv ∣ x ∣ ∣ x ∣ b - --- encode-fun2 : (n : ℕ) → (x y : A) → Path (hLevelTrunc (suc n) A) ∣ x ∣ ∣ y ∣ → hLevelTrunc n (x ≡ y) --- encode-fun2 x y p = transport (λ i → {!!}) {!!} - --- -- encode-fun2 : {n : ℕ₋₂} (x y : ∥ B ∥ (suc₋₂ n)) → x ≡ y → P x y --- -- encode-fun2 x y p = transport (λ i → P x (p i)) (r x) - - - --- Kn→ΩKn+1Sucn : (n : ℕ) → Kn→ΩKn+1 (suc n) ≡ λ x → decode-fun2 _ north (trElim (λ _ → isOfHLevelTrunc (3 + n)) (λ a → ∣ ϕ north a ∣) x) --- Kn→ΩKn+1Sucn n = funExt (trElim (λ x → isOfHLevelSuc (2 + n) (isOfHLevelTrunc (4 + n) ∣ north ∣ ∣ north ∣ _ _)) --- λ a → refl) - - --- Kn≃ΩKn+1 : (n : ℕ) → Iso (coHomK n) (typ (Ω (coHomK-ptd (suc n)))) --- -- n = 0 --- Kn≃ΩKn+1 zero = compIso {!!} (congIso (isoToEquiv (iso {!!} {!!} {!!} {!!}))) --- --- n = 1 --- Kn≃ΩKn+1 (suc zero) = compIso Iso∣ϕ∣ (decodeIso _ north) --- --- n = 2 --- Kn≃ΩKn+1 (suc (suc n)) = compIso (connectedTruncIso2 (4 + n) _ (ϕ north) (n , (λ i → suc (suc (suc (+-suc n n (~ i))))) ∙ --- (λ i → suc (suc (+-suc n (suc n) ((~ i)))))) --- (isConnectedσ (suc n) (sphereConnected _))) --- (decodeIso _ north) - - - --- -- {- For n ≥ 1, we rewrite our function as the composition below. -} --- -- Kn→ΩKn+1Sucn : (n : ℕ) → Kn→ΩKn+1 (suc n) ≡ λ x → truncEquivΩ (ℕ→ℕ₋₂ (suc n)) .fst (trElim (λ _ → isOfHLevelTrunc (3 + n)) (λ a → ∣ ϕ north a ∣) x) --- -- Kn→ΩKn+1Sucn n = funExt (trElim (λ x → isOfHLevelSuc (suc (suc n)) --- -- ((isOfHLevelTrunc ( 2 + (suc (suc n))) ∣ north ∣ ∣ north ∣) --- -- (Kn→ΩKn+1 (suc n) x) --- -- (truncEquivΩ (ℕ→ℕ₋₂ (suc n)) .fst (trElim (λ _ → isOfHLevelTrunc (2 + (suc n))) (λ a → ∣ ϕ north a ∣) x)))) --- -- λ a → refl) - - - - --- -- -- {- For n ≥ 1, we rewrite our function as the composition below. -} --- -- -- Kn→ΩKn+1Sucn : (n : ℕ) → Kn→ΩKn+1 (suc n) ≡ λ x → truncEquivΩ (ℕ→ℕ₋₂ (suc n)) .fst (trElim (λ _ → isOfHLevelTrunc (3 + n)) (λ a → ∣ ϕ north a ∣) x) --- -- -- Kn→ΩKn+1Sucn n = funExt (trElim (λ x → isOfHLevelSuc (suc (suc n)) --- -- -- ((isOfHLevelTrunc ( 2 + (suc (suc n))) ∣ north ∣ ∣ north ∣) --- -- -- (Kn→ΩKn+1 (suc n) x) --- -- -- (truncEquivΩ (ℕ→ℕ₋₂ (suc n)) .fst (trElim (λ _ → isOfHLevelTrunc (2 + (suc n))) (λ a → ∣ ϕ north a ∣) x)))) --- -- -- λ a → refl) - - --- -- isEquivKn→ΩKn+1 : (n : ℕ) → isEquiv (Kn→ΩKn+1 n) --- -- isEquivKn→ΩKn+1 zero = compEquiv (looper , isEquivLooper) ((cong ∣_∣) , isEquivHelper hLev3) .snd --- -- where --- -- hLev3 : (x y : S₊ 1) (p q : x ≡ y) → isProp (p ≡ q) --- -- hLev3 x y = J (λ y p → (q : x ≡ y) → isProp (p ≡ q) ) --- -- (transport (λ i → isSet ({!!} (~ i))) isSetInt refl) - --- -- where --- -- helper : (x ≡ x) ≡ Int --- -- helper = (λ i → transp (λ j → ua S1≃SuspBool (~ j ∨ ~ i)) (~ i) x ≡ transp (λ j → ua S1≃SuspBool (~ j ∨ ~ i)) (~ i) x) ∙ --- -- (λ i → transp (λ j → S¹≡SuspBool (~ j ∨ ~ i)) (~ i) (transport (sym (ua S1≃SuspBool)) x) ≡ transp (λ j → S¹≡SuspBool (~ j ∨ ~ i)) (~ i) (transport (sym (ua S1≃SuspBool)) x)) ∙ --- -- basedΩS¹≡Int (transport (sym S¹≡SuspBool) (transport (sym (ua S1≃SuspBool)) x)) --- -- isEquivKn→ΩKn+1 (suc zero) = transport (λ i → isEquiv (Kn→ΩKn+1Sucn zero (~ i))) --- -- (compEquiv ((trMap (ϕ north)) , isEquiv∣ϕ∣) --- -- (isoToEquiv {!decodeIso ? ?!}) --- -- .snd) --- -- isEquivKn→ΩKn+1 (suc (suc n)) = {!!} - - --- -- -- isEquivKn→ΩKn+1 : (n : ℕ) → isEquiv (Kn→ΩKn+1 n) --- -- -- isEquivKn→ΩKn+1 zero = compEquiv (looper , isEquivLooper) (cong ∣_∣ , isEquivHelper hLevl3) .snd --- -- -- where --- -- -- hLevl3 : (x y : S₊ 1) (p q : x ≡ y) → isProp (p ≡ q) --- -- -- hLevl3 x y = J (λ y p → (q : x ≡ y) → isProp (p ≡ q) ) --- -- -- (transport (λ i → isSet (helper (~ i))) isSetInt refl) --- -- -- where --- -- -- helper : (x ≡ x) ≡ Int --- -- -- helper = (λ i → transp (λ j → ua S1≃SuspBool (~ j ∨ ~ i)) (~ i) x ≡ transp (λ j → ua S1≃SuspBool (~ j ∨ ~ i)) (~ i) x) ∙ --- -- -- (λ i → transp (λ j → S¹≡SuspBool (~ j ∨ ~ i)) (~ i) (transport (sym (ua S1≃SuspBool)) x) ≡ transp (λ j → S¹≡SuspBool (~ j ∨ ~ i)) (~ i) (transport (sym (ua S1≃SuspBool)) x)) ∙ --- -- -- basedΩS¹≡Int (transport (sym S¹≡SuspBool) (transport (sym (ua S1≃SuspBool)) x)) --- -- -- isEquivKn→ΩKn+1 (suc zero) = transport (λ i → isEquiv (Kn→ΩKn+1Sucn zero (~ i))) --- -- -- (compEquiv (trElim (λ _ → isOfHLevelTrunc (2 + (suc zero))) (λ a → ∣ ϕ north a ∣) , --- -- -- isEquiv∣ϕ∣) --- -- -- (truncEquivΩ (ℕ→ℕ₋₂ (suc zero))) .snd) --- -- -- isEquivKn→ΩKn+1 (suc (suc n)) = transport (λ i → isEquiv (Kn→ΩKn+1Sucn (suc n) (~ i))) --- -- -- (compEquiv (connectedTruncEquiv2 (4 + n) _ (ϕ north) (n , λ i → suc (suc (suc (+-suc n n (~ i))))) ?) --- -- -- (truncEquivΩ (ℕ→ℕ₋₂ (suc (suc n)))) .snd) - --- -- -- Kn≃ΩKn+1 : (n : ℕ) → coHomK n ≃ typ (Ω (coHomK-ptd (suc n))) --- -- -- Kn≃ΩKn+1 n = Kn→ΩKn+1 n , isEquivKn→ΩKn+1 n - --- -- -- ΩKn+1→Kn : (n : ℕ) → typ (Ω (coHomK-ptd (suc n))) → coHomK n --- -- -- ΩKn+1→Kn n a = equiv-proof (isEquivKn→ΩKn+1 n) a .fst .fst diff --git a/Cubical/ZCohomology/OldDef/cupProdPrelims2.agda b/Cubical/ZCohomology/OldDef/cupProdPrelims2.agda deleted file mode 100644 index e0e58bb86..000000000 --- a/Cubical/ZCohomology/OldDef/cupProdPrelims2.agda +++ /dev/null @@ -1,262 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.ZCohomology.OldDef.cupProdPrelims2 where - -open import Cubical.ZCohomology.OldDef.Base -open import Cubical.HITs.S1 -open import Cubical.HITs.S2 -open import Cubical.HITs.S3 -open import Cubical.HITs.Sn -open import Cubical.HITs.Sn.Properties -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Function -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Transport -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.Equiv.HalfAdjoint -open import Cubical.Data.NatMinusTwo.Base -open import Cubical.Data.Empty -open import Cubical.Data.Sigma hiding (_×_) -open import Cubical.Data.Prod.Base -open import Cubical.HITs.Susp -open import Cubical.HITs.Nullification -open import Cubical.Data.Int renaming (_+_ to +Int) -open import Cubical.Data.Nat -open import Cubical.HITs.Truncation renaming (elim to trElim ; rec to trRec ; map to trMap) -open import Cubical.HITs.Hopf -open import Cubical.Homotopy.Connected -open import Cubical.Homotopy.Freudenthal hiding (encode) -open import Cubical.Homotopy.Loopspace - -open import Cubical.HITs.Pushout -open import Cubical.Data.Sum.Base -open import Cubical.Data.HomotopyGroup -open import Cubical.Data.Bool -private - variable - ℓ ℓ' : Level - A : Type ℓ - B : Type ℓ' - -{- We want to prove that Kn≃ΩKn+1. For this we use the map ϕ-} -private - ϕ : (pt a : A) → typ (Ω (Susp A , north)) - ϕ pt a = (merid a) ∙ sym (merid pt) - - {- To define the map for n=0 we use the λ k → loopᵏ map for S₊ 1. The loop is given by ϕ south north -} - - loop* : Path (S₊ 1) north north - loop* = ϕ north south - -looper : Int → Path (S₊ 1) north north -looper (pos zero) = refl -looper (pos (suc n)) = looper (pos n) ∙ loop* -looper (negsuc zero) = sym loop* -looper (negsuc (suc n)) = looper (negsuc n) ∙ sym loop* - -{- The map of Kn≃ΩKn+1 is given as follows. -} -Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) -Kn→ΩKn+1 zero x = cong ∣_∣ (looper x) -Kn→ΩKn+1 (suc n) = trRec (isOfHLevelTrunc (2 + (suc (suc n))) ∣ north ∣ ∣ north ∣) - λ a → cong ∣_∣ ((merid a) ∙ (sym (merid north))) - -{- We show that looper is a composition of intLoop with two other maps, all three being isos -} - -private - sndcomp : ΩS¹ → Path (Susp Bool) north north - sndcomp = cong S¹→SuspBool - - thrdcomp : Path (Susp Bool) north north → Path (S₊ 1) north north - thrdcomp = cong SuspBool→S1 - - - looper2 : Int → Path (S₊ 1) north north - looper2 a = thrdcomp (sndcomp (intLoop a)) - - looper≡looper2 : (x : Int) → looper x ≡ looper2 x - looper≡looper2 (pos zero) = refl - looper≡looper2 (pos (suc n)) = - looper (pos n) ∙ loop* ≡⟨ (λ i → looper≡looper2 (pos n) i ∙ congFunct SuspBool→S1 (merid false) (sym (merid true)) (~ i)) ⟩ - looper2 (pos n) ∙ cong SuspBool→S1 (ϕ true false) ≡⟨ sym (congFunct SuspBool→S1 (sndcomp (intLoop (pos n))) (ϕ true false)) ⟩ - cong SuspBool→S1 (sndcomp (intLoop (pos n)) ∙ ϕ true false) ≡⟨ cong thrdcomp (sym (congFunct S¹→SuspBool (intLoop (pos n)) loop)) ⟩ - looper2 (pos (suc n)) ∎ - looper≡looper2 (negsuc zero) = - sym loop* ≡⟨ symDistr (merid south) (sym (merid north)) ⟩ - merid north ∙ sym (merid south) ≡⟨ sym (congFunct SuspBool→S1 (merid true) (sym (merid false))) ⟩ - cong SuspBool→S1 (merid true ∙ sym (merid false)) ≡⟨ cong thrdcomp (sym (symDistr (merid false) (sym (merid true)))) ⟩ - looper2 (negsuc zero) ∎ - looper≡looper2 (negsuc (suc n)) = - looper (negsuc n) ∙ sym loop* ≡⟨ ((λ i → looper≡looper2 (negsuc n) i ∙ symDistr (merid south) (sym (merid north)) i)) ⟩ - looper2 (negsuc n) ∙ merid north ∙ sym (merid south) ≡⟨ cong (λ x → looper2 (negsuc n) ∙ x) (sym (congFunct SuspBool→S1 (merid true) (sym (merid false)))) ⟩ - looper2 (negsuc n) ∙ cong SuspBool→S1 (ϕ false true) ≡⟨ cong (λ x → looper2 (negsuc n) ∙ x) (cong thrdcomp (sym (symDistr (merid false) (sym (merid true))))) ⟩ - looper2 (negsuc n) ∙ cong SuspBool→S1 (sym (ϕ true false)) ≡⟨ sym (congFunct SuspBool→S1 (sndcomp (intLoop (negsuc n))) (sym (ϕ true false))) ⟩ - thrdcomp (cong S¹→SuspBool (intLoop (negsuc n)) ∙ cong S¹→SuspBool (sym loop)) ≡⟨ cong thrdcomp (sym (congFunct S¹→SuspBool (intLoop (negsuc n)) (sym loop))) ⟩ - looper2 (negsuc (suc n)) ∎ - - - isolooper2 : Iso Int (Path (S₊ 1) north north) - isolooper2 = compIso (iso intLoop winding (decodeEncode base) windingIntLoop) - (compIso iso2 - iso1) - where - iso1 : Iso (Path (Susp Bool) north north) (Path (S₊ 1) north north) - iso1 = congIso SuspBool≃S1 - - iso2 : Iso ΩS¹ (Path (Susp Bool) north north) - iso2 = congIso (isoToEquiv (iso S¹→SuspBool SuspBool→S¹ SuspBool→S¹→SuspBool S¹→SuspBool→S¹)) - - isolooper : Iso Int (Path (S₊ 1) north north) - Iso.fun isolooper = looper - Iso.inv isolooper = Iso.inv isolooper2 - Iso.rightInv isolooper a = (looper≡looper2 (Iso.inv isolooper2 a)) ∙ Iso.rightInv isolooper2 a - Iso.leftInv isolooper a = cong (Iso.inv isolooper2) (looper≡looper2 a) ∙ Iso.leftInv isolooper2 a - - - {- We want to show that this map is an equivalence. n ≥ 2 follows from Freudenthal, and -} - - {- - We want to show that the function (looper : Int → S₊ 1) defined by λ k → loopᵏ is an equivalece. We already know that the corresponding function (intLoop : Int → S¹ is) an equivalence, - so the idea is to show that when intLoop is transported along a suitable path S₊ 1 ≡ S¹ we get looper. Instead of using S₊ 1 straight away, we begin by showing this for the equivalent Susp Bool. - -} - - - - ----------------------------------- n = 1 ----------------------------------------------------- - - {- We begin by stating some useful lemmas -} - - - - S³≡SuspSuspS¹ : S³ ≡ Susp (Susp S¹) - S³≡SuspSuspS¹ = S³≡SuspS² ∙ λ i → Susp (S²≡SuspS¹ i) - - S3≡SuspSuspS¹ : S₊ 3 ≡ Susp (Susp S¹) - S3≡SuspSuspS¹ = (λ i → Susp (Susp (Susp (ua Bool≃Susp⊥ (~ i))))) ∙ λ i → Susp (Susp (S¹≡SuspBool (~ i))) - - sphereConnectedSpecCase : isHLevelConnected 4 (Susp (Susp S¹)) -- is- 2 -ConnectedType (Susp (Susp S¹)) - sphereConnectedSpecCase = transport (λ i → isHLevelConnected 4 (S3≡SuspSuspS¹ i)) (sphereConnected 3) - - - - {- We give the following map and show that it is an equivalence -} - - d-map : typ (Ω ((Susp S¹) , north)) → S¹ - d-map p = subst HopfSuspS¹ p base - - d-mapId : (r : S¹) → d-map (ϕ base r) ≡ r - d-mapId r = substComposite HopfSuspS¹ (merid r) (sym (merid base)) base ∙ - rotLemma r - where - rotLemma : (r : S¹) → rot r base ≡ r - rotLemma base = refl - rotLemma (loop i) = refl - - d-mapComp : fiber d-map base ≡ Path (Susp (Susp S¹)) north north - d-mapComp = sym (pathSigma≡sigmaPath {B = HopfSuspS¹} _ _) ∙ helper - where - helper : Path (Σ (Susp S¹) λ x → HopfSuspS¹ x) (north , base) (north , base) ≡ Path (Susp (Susp S¹)) north north - helper = (λ i → (Path (S³≡TotalHopf (~ i)) - (transp (λ j → S³≡TotalHopf (~ i ∨ ~ j)) (~ i) (north , base)) - ((transp (λ j → S³≡TotalHopf (~ i ∨ ~ j)) (~ i) (north , base))))) ∙ - (λ i → Path (S³≡SuspSuspS¹ i) (transp (λ j → S³≡SuspSuspS¹ (i ∧ j)) (~ i) base) ((transp (λ j → S³≡SuspSuspS¹ (i ∧ j)) (~ i) base))) - - - is1Connected-dmap : isHLevelConnectedFun 3 d-map - is1Connected-dmap = toPropElim (λ s → isPropIsOfHLevel 0) (transport (λ j → isContr (∥ d-mapComp (~ j) ∥ ℕ→ℕ₋₂ 1)) - (transport (λ i → isContr (PathΩ {A = Susp (Susp S¹)} {a = north} (ℕ→ℕ₋₂ 1) i)) - (refl , isOfHLevelSuc 1 (isOfHLevelSuc 0 sphereConnectedSpecCase) ∣ north ∣ ∣ north ∣ (λ _ → ∣ north ∣)))) - - d-iso2 : Iso (hLevelTrunc 3 (typ (Ω (Susp S¹ , north)))) (hLevelTrunc 3 S¹) - Iso.fun d-iso2 = trMap d-map - Iso.inv d-iso2 = Iso.inv (connectedTruncIso _ d-map is1Connected-dmap) - Iso.rightInv d-iso2 = Iso.rightInv (connectedTruncIso _ d-map is1Connected-dmap) - Iso.leftInv d-iso2 = Iso.leftInv (connectedTruncIso _ d-map is1Connected-dmap) - - d-iso : isIso {A = ∥ typ (Ω (Susp S¹ , north)) ∥ (ℕ→ℕ₋₂ 1)} {B = ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (trElim (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣ ) - d-iso = (Iso.inv (connectedTruncIso _ d-map is1Connected-dmap)) , (Iso.rightInv (connectedTruncIso _ d-map is1Connected-dmap) - , Iso.leftInv (connectedTruncIso _ d-map is1Connected-dmap)) - - {- We show that composing (λ a → ∣ ϕ base a ∣) and (λ x → ∣ d-map x ∣) gives us the identity function. -} - - d-mapId2 : (λ (x : hLevelTrunc 3 S¹) → (trRec {n = 3} {B = hLevelTrunc 3 S¹} (isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) - (trRec (isOfHLevelTrunc 3) (λ a → ∣ ϕ base a ∣) x)) ≡ λ x → x - d-mapId2 = funExt (trElim (λ x → isOfHLevelSuc 2 (isOfHLevelTrunc 3 ((trElim {n = 3} - {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} - (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) - (trElim (λ _ → isOfHLevelTrunc 3) - (λ a → ∣ ϕ base a ∣) x)) x)) - λ a i → ∣ d-mapId a i ∣) - - {- This means that (λ a → ∣ ϕ base a ∣) is an equivalence -} - - - Iso∣ϕ-base∣ : Iso (hLevelTrunc 3 S¹) (hLevelTrunc 3 (typ (Ω (Susp S¹ , north)))) - Iso∣ϕ-base∣ = composesToId→Iso d-iso2 (trMap (ϕ base)) d-mapId2 - - - --------------------------------- - -- We cheat when n = 1 and use J to prove the following lemmma. There is an obvious dependent path between ϕ base and ϕ north. Since the first one is an iso, so is the other. - -- - - - pointFunIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ} {C : (A : Type ℓ) (a1 : A) → Type ℓ'} (p : A ≡ B) (a : A) (b : B) → - (transport p a ≡ b) → - (f : (A : Type ℓ) → - (a1 : A) (a2 : hLevelTrunc 3 A) → C A a1) → - isIso (f A a) → - isIso (f B b) - pointFunIso {ℓ = ℓ}{A = A} {B = B} {C = C} = - J (λ B p → (a : A) (b : B) → - (transport p a ≡ b) → - (f : (A : Type ℓ) → - (a1 : A) (a2 : hLevelTrunc 3 A) → C A a1) → - isIso (f A a) → - isIso (f B b)) - λ a b trefl f is → transport (λ i → isIso (f A ((sym (transportRefl a) ∙ trefl) i))) is - - {- By pointFunEquiv, this gives that λ a → ∣ ϕ north a ∣ is an iso -} - - isIso∣ϕ∣ : isIso {A = hLevelTrunc 3 (S₊ 1)} {B = hLevelTrunc 3 (typ (Ω (S₊ 2 , north)))} (trMap (ϕ north)) - isIso∣ϕ∣ = pointFunIso {A = S¹} (λ i → S¹≡S1 (~ i)) base north refl (λ A a1 → trMap (ϕ a1)) ((Iso.inv Iso∣ϕ-base∣) , ((Iso.rightInv Iso∣ϕ-base∣) , (Iso.leftInv Iso∣ϕ-base∣))) - - Iso∣ϕ∣ : Iso (hLevelTrunc 3 (S₊ 1)) (hLevelTrunc 3 (typ (Ω (S₊ 2 , north)))) - Iso.fun Iso∣ϕ∣ = trMap (ϕ north) - Iso.inv Iso∣ϕ∣ = isIso∣ϕ∣ .fst - Iso.rightInv Iso∣ϕ∣ = isIso∣ϕ∣ .snd .fst - Iso.leftInv Iso∣ϕ∣ = isIso∣ϕ∣ .snd .snd - --- ---------------------------------------------------- Finishing up --------------------------------- - --- We need ΩTrunc. It computes horribly, so its better to restate the function involved for this particular case -- - -decode-fun2 : (n : ℕ) → (x : A) → hLevelTrunc n (x ≡ x) → Path (hLevelTrunc (suc n) A) ∣ x ∣ ∣ x ∣ -decode-fun2 zero x = trElim (λ _ → isOfHLevelPath 0 (∣ x ∣ , isOfHLevelTrunc 1 ∣ x ∣) ∣ x ∣ ∣ x ∣) (λ p i → ∣ p i ∣) -decode-fun2 (suc n) x = trElim (λ _ → isOfHLevelPath' (suc n) (isOfHLevelTrunc (suc (suc n))) ∣ x ∣ ∣ x ∣) (cong ∣_∣) - -funsAreSame : (n : ℕ) (x : A) (b : hLevelTrunc n (x ≡ x)) → (decode-fun2 n x b) ≡ (ΩTrunc.decode-fun ∣ x ∣ ∣ x ∣ b) -funsAreSame zero x = trElim (λ a → isOfHLevelPath 0 (refl , (isOfHLevelSuc 1 (isOfHLevelTrunc 1) ∣ x ∣ ∣ x ∣ refl)) _ _) λ a → refl -funsAreSame (suc n) x = trElim (λ a → isOfHLevelPath _ (isOfHLevelPath' (suc n) (isOfHLevelTrunc (suc (suc n))) ∣ x ∣ ∣ x ∣) _ _) λ a → refl - -decodeIso : (n : ℕ) → (x : A) → Iso (hLevelTrunc n (x ≡ x)) (Path (hLevelTrunc (suc n) A) ∣ x ∣ ∣ x ∣) -Iso.fun (decodeIso n x) = decode-fun2 n x -Iso.inv (decodeIso n x) = ΩTrunc.encode-fun ∣ x ∣ ∣ x ∣ -Iso.rightInv (decodeIso n x) b = funExt⁻ (funExt (funsAreSame n x)) (ΩTrunc.encode-fun ∣ x ∣ ∣ x ∣ b) ∙ ΩTrunc.P-rinv ∣ x ∣ ∣ x ∣ b -Iso.leftInv (decodeIso n x) b = cong (ΩTrunc.encode-fun ∣ x ∣ ∣ x ∣) (funsAreSame n x b) ∙ ΩTrunc.P-linv ∣ x ∣ ∣ x ∣ b - - -Kn→ΩKn+1Sucn : (n : ℕ) → Kn→ΩKn+1 (suc n) ≡ λ x → decode-fun2 _ north (trElim (λ _ → isOfHLevelTrunc (3 + n)) (λ a → ∣ ϕ north a ∣) x) -Kn→ΩKn+1Sucn n = funExt (trElim (λ x → isOfHLevelSuc (2 + n) (isOfHLevelTrunc (4 + n) ∣ north ∣ ∣ north ∣ _ _)) - λ a → refl) - - -Iso-Kn-ΩKn+1 : (n : ℕ) → Iso (coHomK n) (typ (Ω (coHomK-ptd (suc n)))) -Iso-Kn-ΩKn+1 zero = compIso isolooper (congIso (truncIdempotent≃ _ isOfHLevelS1)) -Iso-Kn-ΩKn+1 (suc zero) = compIso Iso∣ϕ∣ (decodeIso _ north) -Iso-Kn-ΩKn+1 (suc (suc n)) = compIso (connectedTruncIso2 (4 + n) _ (ϕ north) (n , (λ i → suc (suc (suc (+-suc n n (~ i))))) ∙ - (λ i → suc (suc (+-suc n (suc n) ((~ i)))))) - (isConnectedσ (suc n) (sphereConnected _))) - (decodeIso _ north) diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index 28b425f72..d6294fe49 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -14,21 +14,23 @@ open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Univalence open import Cubical.Data.NatMinusTwo.Base open import Cubical.Data.Empty -open import Cubical.Data.Sigma +open import Cubical.Data.Sigma hiding (_×_) open import Cubical.HITs.Susp -open import Cubical.HITs.SetTruncation +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) open import Cubical.HITs.Nullification open import Cubical.Data.Int hiding (_+_) open import Cubical.Data.Nat +open import Cubical.Data.Prod open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) open import Cubical.Homotopy.Loopspace open import Cubical.Homotopy.Connected open import Cubical.Homotopy.Freudenthal + open import Cubical.HITs.Pushout open import Cubical.Data.Sum.Base open import Cubical.Data.HomotopyGroup -open import Cubical.Data.Bool +open import Cubical.Data.Bool hiding (_⊕_) open import Cubical.ZCohomology.cupProdPrelims @@ -150,3 +152,87 @@ commₖ {n = n} x y i = ΩKn+1→Kn (EH-instance (Kn→ΩKn+1 n x) (Kn→ΩKn+1 K-Id n = (λ i → typ (Ω (isoToPath (Iso2-Kn-ΩKn+1 (suc n)) i , hcomp (λ k → λ {(i = i0) → ∣ north ∣ ; (i = i1) → transportRefl (λ j → ∣ rCancel (merid north) k j ∣) k}) (transp (λ j → isoToPath (Iso2-Kn-ΩKn+1 (suc n)) (i ∧ j)) (~ i) ∣ north ∣)))) + + +---- Algebra with cohomology groups --- + +private + § : isSet ∥ A ∥₀ + § = setTruncIsSet + +_+ₕ_ : {n : ℕ} → coHom n A → coHom n A → coHom n A +_+ₕ_ = sElim2 (λ _ _ → §) λ a b → ∣ (λ x → a x +ₖ b x) ∣₀ + +-ₕ_ : {n : ℕ} → coHom n A → coHom n A +-ₕ_ = sRec § λ a → ∣ (λ x → -ₖ a x) ∣₀ + +0ₕ : {n : ℕ} → coHom n A +0ₕ = ∣ (λ _ → 0ₖ) ∣₀ + +rUnitₕ : {n : ℕ} (x : coHom n A) → x +ₕ 0ₕ ≡ x +rUnitₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → rUnitₖ (a x)) i ∣₀ + +lUnitₕ : {n : ℕ} (x : coHom n A) → 0ₕ +ₕ x ≡ x +lUnitₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → lUnitₖ (a x)) i ∣₀ + +rCancelₕ : {n : ℕ} (x : coHom n A) → x +ₕ (-ₕ x) ≡ 0ₕ +rCancelₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → rCancelₖ (a x)) i ∣₀ + +lCancelₕ : {n : ℕ} (x : coHom n A) → (-ₕ x) +ₕ x ≡ 0ₕ +lCancelₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → lCancelₖ (a x)) i ∣₀ + + + +assocₕ : {n : ℕ} (x y z : coHom n A) → ((x +ₕ y) +ₕ z) ≡ (x +ₕ (y +ₕ z)) +assocₕ = elim3 (λ _ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b c i → ∣ funExt (λ x → assocₖ (a x) (b x) (c x)) i ∣₀ + + +commₕ : {n : ℕ} (x y : coHom n A) → (x +ₕ y) ≡ (y +ₕ x) +commₕ {n = n} = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ funExt (λ x → commₖ (a x) (b x)) i ∣₀ + + +-- Cup-prouct -- + +_∧-map_ : (n m : ℕ) → (S₊ (suc n) , north) wedge (S₊ (suc m) , north) → S₊ (suc n + suc m) +_∧-map_ n m = {!!} + where + S' : (n : ℕ) → Pointed ℓ-zero + S' zero = Unit , tt + S' (suc zero) = S₊ 1 , north + S' (suc (suc n)) = ((S₊ 1 , north) wedge S' (suc n)) , inl north + + S'≡S : (n : ℕ) → typ (S' (suc n)) ≡ S₊ (suc n) + helper : (n : ℕ) → transp (λ j → S'≡S n j) i0 (snd (S' (suc n))) ≡ north + S'≡S zero = refl + S'≡S (suc n) = (λ i → (S₊ 1 , north) wedge S' (suc n)) ∙ + (λ i → (S₊ 1 , north) wedge ((S'≡S n i) , transp (λ j → S'≡S n (i ∧ j)) (~ i) (pt (S' (suc n))))) ∙ + (λ i → (S₊ 1 , north) wedge (S₊ (suc n) , helper n i)) ∙ + {!!} ∙ + {!!} + helper zero = refl + helper (suc n) = {!!} + where + helper2 : (n : ℕ) → Iso ((S₊ 1 , north) wedge (S₊ (suc n) , north)) (S₊ (suc (suc n))) + Iso.fun (helper2 n) (inl x) = PushoutSusp→Susp {!!} + Iso.fun (helper2 n) (inr x) = PushoutSusp→Susp {!!} + Iso.fun (helper2 n) (push tt i) = {!!} + Iso.inv (helper2 n) = {!!} + Iso.rightInv (helper2 n) = {!!} + Iso.leftInv (helper2 n) = {!!} + +smashₕ : (A : Type ℓ) (n m : ℕ) → Type ℓ +smashₕ A n m = (coHom n A , ∣ (λ a → coHom-pt n) ∣₀) smash (coHom n A , ∣ (λ a → coHom-pt n) ∣₀) + +⌣' : (n m : ℕ) → (coHomK (suc m) , coHom-pt (suc m)) smash (coHomK (suc m) , coHom-pt (suc m)) + → coHomK ((suc n) + (suc m)) +⌣' n m = equiv-proof (elim.isEquivPrecompose {A = ((S₊ (suc n)) , north) wedge ((S₊ (suc m)) , north)} + (λ a → {!smash!}) + {!!} + {!!} + {!!}) {!!} .fst .fst From 57d9c23901025dfdb64770891d36b9d233fe9f42 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Sat, 25 Apr 2020 10:43:16 +0200 Subject: [PATCH 20/89] push --- Cubical/HITs/Pushout/Base.agda | 14 - Cubical/HITs/SmashProduct/Base.agda | 67 +++++ Cubical/HITs/Sn/Properties.agda | 387 +++++++++++++++++++++------- Cubical/Homotopy/Connected.agda | 2 +- 4 files changed, 363 insertions(+), 107 deletions(-) diff --git a/Cubical/HITs/Pushout/Base.agda b/Cubical/HITs/Pushout/Base.agda index 252e83b34..9c297e49a 100644 --- a/Cubical/HITs/Pushout/Base.agda +++ b/Cubical/HITs/Pushout/Base.agda @@ -18,20 +18,6 @@ data Pushout {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} push : (a : A) → inl (f a) ≡ inr (g a) ---- - -_wedge_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Type (ℓ-max ℓ ℓ') -_wedge_ (A , ptA) (B , ptB) = Pushout {A = Unit} {B = A} {C = B} (λ {tt → ptA}) (λ {tt → ptB}) - -i∧ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → A wedge B → (typ A) × (typ B) -i∧ {A = A , ptA} {B = B , ptB} (inl x) = x , ptB -i∧ {A = A , ptA} {B = B , ptB} (inr x) = ptA , x -i∧ {A = A , ptA} {B = B , ptB} (push tt i) = ptA , ptB - -_smash_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Type (ℓ-max ℓ ℓ') -_smash_ A B = Pushout {A = (A wedge B)} {B = Unit} {C = (typ A) × (typ B)} (λ _ → tt) i∧ - - -- Suspension defined as a pushout PushoutSusp : ∀ {ℓ} (A : Type ℓ) → Type ℓ diff --git a/Cubical/HITs/SmashProduct/Base.agda b/Cubical/HITs/SmashProduct/Base.agda index 15e4515e7..172c4fb5f 100644 --- a/Cubical/HITs/SmashProduct/Base.agda +++ b/Cubical/HITs/SmashProduct/Base.agda @@ -3,6 +3,11 @@ module Cubical.HITs.SmashProduct.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.HITs.Pushout.Base +open import Cubical.Data.Unit +open import Cubical.Data.Prod +open import Cubical.Foundations.GroupoidLaws data Smash (A B : Pointed₀) : Type₀ where basel : Smash A B @@ -43,3 +48,65 @@ SmashPt A B = (Smash A B , basel) -- rearrange (gluel a i) = {!!} -- rearrange (gluer b i) = {!gluel ? i!} + + +--- Alternative definition + +_⋁_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Type (ℓ-max ℓ ℓ') +_⋁_ (A , ptA) (B , ptB) = Pushout {A = Unit} {B = A} {C = B} (λ {tt → ptA}) (λ {tt → ptB}) + +i∧ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → A ⋁ B → (typ A) × (typ B) +i∧ {A = A , ptA} {B = B , ptB} (inl x) = x , ptB +i∧ {A = A , ptA} {B = B , ptB} (inr x) = ptA , x +i∧ {A = A , ptA} {B = B , ptB} (push tt i) = ptA , ptB + +_⋀_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Type (ℓ-max ℓ ℓ') +_⋀_ A B = Pushout {A = (A ⋁ B)} {B = Unit} {C = (typ A) × (typ B)} (λ _ → tt) i∧ + + +_⋀⃗_ : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} {D : Pointed ℓ'''} (f : A →∙ C) (g : B →∙ D) → A ⋀ B → C ⋀ D +(f ⋀⃗ g) (inl tt) = inl tt +((f , fpt) ⋀⃗ (g , gpt)) (inr (x , x₁)) = inr (f x , g x₁) +_⋀⃗_ {B = B} {D = D} (f , fpt) (b , gpt) (push (inl x) i) = (push (inl (f x)) ∙ (λ i → inr (f x , gpt (~ i)))) i +_⋀⃗_ {A = A} {B = B} {C = C} {D = D} (f , fpt) (g , gpt) (push (inr x) i) = (push (inr (g x)) ∙ (λ i → inr (fpt (~ i) , g x))) i +_⋀⃗_ {A = A} {B = B} {C = C} {D = D} (f , fpt) (g , gpt) (push (push tt j) i) = helper (~ j) i + where + helper : Path (Path (C ⋀ D) (inl tt) (inr ((f (pt A)) , (g (pt B))))) + (push (inr (g (pt B))) ∙ λ i → inr ((fpt (~ i)) , (g (pt B)))) + (push (inl (f (pt A))) ∙ λ i → inr ((f (pt A)) , (gpt (~ i)))) + helper = (λ j → rUnit (push (inr (g (pt B)))) j ∙ λ i → inr ((fpt (~ i)) , (g (pt B))) ) ∙ + (λ j → (push (inr (gpt j)) ∙ λ i → inr ((pt C) , (gpt ((~ i) ∧ j)))) ∙ λ i → inr ((fpt (~ i)) , (g (pt B)))) ∙ + (λ j → (cong (push) (push tt) (~ j) ∙ λ i → inr (pt C , gpt (~ i))) ∙ λ i → inr ((fpt (~ i)) , (g (pt B)))) ∙ + (λ j → (push (inl (fpt (~ j))) ∙ λ i → inr (fpt ((~ j) ∨ i) , (gpt (~ i)))) ∙ λ i → inr ((fpt (~ i)) , (g (pt B)))) ∙ + (λ j → (push (inl (f (pt A))) ∙ λ i → inr (fpt (i ∧ (~ j)) , gpt ((~ i) ∨ j))) ∙ λ i → inr ((fpt ((~ i) ∧ (~ j))) , (gpt ((~ i) ∧ j)))) ∙ + (λ j → (rUnit (push (inl (f (pt A)))) (~ j)) ∙ λ i → inr ((f (pt A)) , (gpt (~ i)))) + + +⋀-commuter : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → A ⋀ B → B ⋀ A +⋀-commuter (inl tt) = inl tt +⋀-commuter (inr (x , x₁)) = inr (x₁ , x) +⋀-commuter (push (inl x) i) = push (inr x) i +⋀-commuter (push (inr x) i) = push (inl x) i +⋀-commuter (push (push a i₁) i) = push (push (tt) (~ i₁)) i + + +⋀-comm : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → Iso (A ⋀ B) (B ⋀ A) +Iso.fun ⋀-comm = ⋀-commuter +Iso.inv ⋀-comm = ⋀-commuter +Iso.rightInv ⋀-comm (inl tt) = refl +Iso.rightInv ⋀-comm (inr (x , x₁)) = refl +Iso.rightInv ⋀-comm (push (inl x) i) = refl +Iso.rightInv ⋀-comm (push (inr x) i) = refl +Iso.rightInv ⋀-comm (push (push a i₁) i) = refl +Iso.leftInv ⋀-comm (inl tt) = refl +Iso.leftInv ⋀-comm (inr (x , x₁)) = refl +Iso.leftInv ⋀-comm (push (inl x) i) = refl +Iso.leftInv ⋀-comm (push (inr x) i) = refl +Iso.leftInv ⋀-comm (push (push a i₁) i) = refl + +rearrange-proj : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} → A ⋀ B → (C ⋀ B , inl tt) ⋀ A +rearrange-proj (inl x) = inl tt +rearrange-proj {C = C} (inr (x , x₁)) = inr (inr ((snd C) , x₁) , x) +rearrange-proj {A = A}{B = B} {C = C} (push (inl x) i) = (push (inr x) ∙ (λ i → inr (push (inr (snd B)) i , x))) i +rearrange-proj {A = A}{B = B} {C = C} (push (inr x) i) = (push (inr (snd A)) ∙ (λ i → inr (push (inr x) i , snd A))) i +rearrange-proj {A = A}{B = B} {C = C} (push (push a i₁) i) = (push (inr (snd A)) ∙ (λ i → inr (push (inr (snd B)) i , (snd A)))) i diff --git a/Cubical/HITs/Sn/Properties.agda b/Cubical/HITs/Sn/Properties.agda index 4b707849d..866ed2e61 100644 --- a/Cubical/HITs/Sn/Properties.agda +++ b/Cubical/HITs/Sn/Properties.agda @@ -1,7 +1,7 @@ {-# OPTIONS --cubical --safe #-} module Cubical.HITs.Sn.Properties where -open import Cubical.Data.Int +open import Cubical.Data.Int hiding (_+_) open import Cubical.Data.Bool open import Cubical.Foundations.Pointed open import Cubical.Foundations.Function @@ -12,15 +12,20 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Univalence open import Cubical.HITs.S1 +open import Cubical.Data.Nat open import Cubical.Data.Prod open import Cubical.HITs.Sn.Base open import Cubical.HITs.Susp +open import Cubical.Data.Unit +open import Cubical.HITs.Join open import Cubical.HITs.Pushout +open import Cubical.HITs.SmashProduct -variable - ℓ ℓ' : Level - A : Type ℓ - B : Type ℓ' +private + variable + ℓ ℓ' : Level + A : Type ℓ + B : Type ℓ' @@ -62,96 +67,294 @@ SuspBool≃S1 = isoToEquiv iso1 Iso.leftInv iso1 (merid true i) = refl -private - f* : {a : A} → S¹ → A → Susp A - f* base b = north - f* {a = a} (loop i) = funExt (λ x → merid x ∙ sym (merid a)) i - - -proj : {A : Pointed ℓ} {B : Pointed ℓ'} → typ A → typ B → A smash B -proj a b = inr (a , b) +-- map between S₊ -projᵣ : {A : Pointed ℓ} {B : Pointed ℓ'} (a : typ A) → proj {A = A} a (pt B) ≡ inl tt -projᵣ a = sym (push (inl a)) +f' : {a : A} → A → S₊ 1 → Susp A +f' {a = pt} A north = north +f' {a = pt} A south = north +f' {a = pt} a (merid p i) = ((merid a) ∙ sym (merid pt)) i -projₗ : {A : Pointed ℓ} {B : Pointed ℓ'} (b : typ B) → proj {B = B} (pt A) b ≡ inl tt -projₗ b = sym (push (inr b)) +proj' : {A : Pointed ℓ} {B : Pointed ℓ'} → typ A → typ B → A ⋀ B +proj' a b = inr (a , b) -projᵣₗ : {A : Pointed ℓ} {B : Pointed ℓ'} → projᵣ (pt A) ≡ projₗ (pt B) -projᵣₗ {A = A} {B = B} = cong sym (cong push (push tt)) - - -module S1-smash-Iso ((A , pt) : Pointed ℓ) where - fun : (S¹ , base) smash (A , pt) → (Susp A) +module smashS1→susp {(A , pt) : Pointed ℓ} where + fun : (S₊ 1 , north) ⋀ (A , pt) → (Susp A) fun (inl tt) = north - fun (inr (x , x₁)) = f* {a = pt} x x₁ - fun (push (inl base) i) = north - fun (push (inl (loop i₁)) i) = rCancel (merid pt) (~ i) i₁ + fun (inr (x , x₁)) = f' {a = pt} x₁ x + fun (push (inl north) i) = north + fun (push (inl south) i) = north + fun (push (inl (merid a i₁)) i) = rCancel (merid pt) (~ i) i₁ fun (push (inr x) i) = north - fun (push (push a i₁) i) = north - - funInv : Susp A → (S¹ , base) smash (A , pt) - funInv north = inl tt - funInv south = inl tt - funInv (merid a i) = (sym (projₗ a) ∙ - cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ - projₗ a) i - - funSect : (x : Susp A) → fun (funInv x) ≡ x - funSect north = refl - funSect south = merid pt - funSect (merid a i) = hcomp (λ k → λ {(i = i0) → ((λ j → (refl ∙ refl) ∙ refl ∙ refl) ∙ - (λ j → rCancel refl j ∙ refl ∙ refl) ∙ - λ j → lUnit (lUnit refl (~ j)) (~ j)) k ; - (i = i1) → ((λ j → (refl ∙ refl) ∙ sym (refl ∙ (merid a ∙ sym (merid pt)) ∙ refl) ∙ merid a) ∙ - (λ j → lUnit refl (~ j) ∙ sym (lUnit (rUnit (merid a ∙ sym (merid pt)) (~ j)) (~ j)) ∙ merid a) ∙ - (λ j → lUnit (symDistr (merid a) (sym (merid pt)) j ∙ merid a) (~ j)) ∙ - (sym (assoc (merid pt) (sym (merid a)) (merid a))) ∙ - (λ j → merid pt ∙ lCancel (merid a) j) ∙ - sym (rUnit (merid pt)) ) k}) - (helper2 i ∙ filler i) - where - - filler : (i : I) → (refl ∙ (merid a ∙ sym (merid pt)) ∙ refl) i ≡ merid a i - filler i = (λ j → (refl ∙ (merid a ∙ sym (merid pt)) ∙ refl) (i ∧ (~ j))) ∙ λ j → merid a (j ∧ i) - - helper2 : (i : I) → fun ((sym (projₗ a) ∙ - cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ - projₗ a) i) - ≡ ((λ i → fun (sym (projₗ a) i)) ∙ - (λ i → fun (cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop i)) ∙ - λ i → fun (projₗ a i)) i - helper2 i = (λ j → congFunct fun (sym (projₗ a)) (cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ projₗ a) j i) ∙ - (λ j → (cong fun (sym (projₗ a)) ∙ congFunct fun (cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop) (projₗ a) j) i) - - - retrHelper : (x : S¹) (a : A) → funInv (f* {a = pt} x a) ≡ proj x a - retrHelper base a = sym (projₗ a) - retrHelper (loop i) a = hcomp (λ k → λ { (i = i0) → {!!} - ; (i = i1) → {!!} }) - filler - - where - filler : funInv (f* {a = pt} (loop i) a) ≡ proj (loop i) a - filler = (λ j → funInv (funExt (λ x → merid x ∙ sym (merid pt)) (i ∨ j) a)) ∙ sym (projₗ a) ∙ λ j → proj (loop (i ∧ j)) a - - id1 : funInv (f* {a = pt} (loop i) a) ≡ cong funInv (merid a ∙ sym (merid pt)) i - id1 = refl - - id2 : funInv (f* {a = pt} (loop i) a) ≡ ((sym (projₗ a) ∙ cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ (projₗ a)) ∙ sym (projₗ pt) ∙ sym (cong (λ x → proj {A = S¹ , base} {B = A , pt} x pt) loop) ∙ projₗ pt) i - id2 = (λ j → congFunct funInv (merid a) (sym (merid pt)) j i) ∙ - (λ j → (((sym (projₗ a) ∙ cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ (projₗ a))) ∙ symDistr (sym (projₗ pt)) (cong (λ x → proj {A = S¹ , base} {B = A , pt} x pt) loop ∙ projₗ pt) j) i) ∙ - (λ j → (((sym (projₗ a) ∙ cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ (projₗ a))) ∙ symDistr (cong (λ x → proj {A = S¹ , base} {B = A , pt} x pt) loop) (projₗ pt) j ∙ projₗ pt) i) ∙ - λ j → (((sym (projₗ a) ∙ cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ (projₗ a))) ∙ assoc (sym (projₗ pt)) (sym (cong (λ x → proj {A = S¹ , base} {B = A , pt} x pt) loop)) (projₗ pt) (~ j)) i - - id4 : sym (projᵣ base) ∙ cong ((λ x → proj {A = S¹ , base} {B = A , pt} x pt)) loop ∙ projᵣ base ≡ refl - id4 = (λ i → (push (inl (loop i))) ∙ (λ j → inr (loop (i ∨ j) , pt)) ∙ sym (push (inl base))) ∙ - (λ i → push (inl base) ∙ lUnit (sym (push (inl base))) (~ i)) ∙ - rCancel (push (inl base)) - - id3 : funInv ((merid pt) (~ i)) ≡ inl tt - id3 = (λ j → (sym (projₗ pt) ∙ (cong (λ x → proj {A = S¹ , base} {B = A , pt} x pt) loop) ∙ projₗ pt) (~ i)) ∙ - (λ j → (sym ((projᵣₗ (~ j))) ∙ cong (λ x → proj {A = S¹ , base} {B = A , pt} x pt) loop ∙ projᵣₗ (~ j)) (~ i)) ∙ - (λ j → id4 j (~ i)) - + fun (push (push tt i₁) i) = north + + fun⁻ : Susp A → (S₊ 1 , north) ⋀ (A , pt) + fun⁻ north = inl tt + fun⁻ south = inl tt + fun⁻ (merid a i) = (push (inr a) ∙∙ cong (λ x → proj' {A = S₊ 1 , north} {B = A , pt} x a) (merid south ∙ sym (merid north)) ∙∙ sym (push (inr a))) i +smashMap : {!(A : Pointed ℓ) → ?!} +smashMap = {!!} + + +S' : ℕ → Type₀ +S' zero = Unit +S' (suc zero) = S₊ 1 +S' (suc (suc n)) = (S₊ 1 , north) ⋀ (S₊ (suc n) , north) + +hej : (m : ℕ) → S₊ (suc m) ≡ Susp (S₊ m) +hej m = refl + +S'-map-pt1 : (n m : ℕ) → (S₊ (suc (suc n)) , north) ⋀ (S₊ (suc m) , north) → ((S₊ 1 , north) ⋀ (S₊ (suc n) , north) , inl tt) ⋀ (S₊ (suc m) , north) +S'-map-pt1 n m = (smashS1→susp.fun⁻ , refl) ⋀⃗ ((λ x → x) , refl) + +S'-map : (n m : ℕ) → (S₊ (suc n) , north) ⋀ (S₊ (suc m) , north) → S' (2 + n + m) +S'-map zero n a = a +S'-map (suc n) m = {!!} ∘ {!!} ∘ S'-map-pt1 n m + +sphereSmashMap : (n m : ℕ) → (S₊ (suc n) , north) ⋀ (S₊ (suc m) , north) → S₊ (2 + n + m) +sphereSmashMap n m = {!!} + +-- private +-- f* : {a : A} → S¹ → A → Susp A +-- f* base b = north +-- f* {a = a} (loop i) = funExt (λ x → merid x ∙ sym (merid a)) i + +-- f' : {a : A} → Susp Bool → A → Susp A +-- f' north a = north +-- f' south a = north +-- f' {a = a} (merid p i) = funExt (λ x → merid x ∙ sym (merid a)) i + + +-- -- proj : {A : Pointed ℓ} {B : Pointed ℓ'} → typ A → typ B → A smash B +-- -- proj a b = inr (a , b) + +-- -- projᵣ : {A : Pointed ℓ} {B : Pointed ℓ'} (a : typ A) → proj {A = A} a (pt B) ≡ inl tt +-- -- projᵣ a = sym (push (inl a)) + +-- -- projₗ : {A : Pointed ℓ} {B : Pointed ℓ'} (b : typ B) → proj {B = B} (pt A) b ≡ inl tt +-- -- projₗ b = sym (push (inr b)) + +-- -- projᵣₗ : {A : Pointed ℓ} {B : Pointed ℓ'} → projᵣ (pt A) ≡ projₗ (pt B) +-- -- projᵣₗ {A = A} {B = B} i j = push (push tt i) (~ j) -- cong sym (cong push (push tt)) + + +-- -- compLrFiller : {a b c d : A} (p : a ≡ b) (q : b ≡ c) (s : c ≡ d) → PathP (λ i → p i ≡ s (~ i)) (p ∙ q ∙ s) q +-- -- compLrFiller {A = A} {a = a} {b = b} {c = c} {d = d} p q s i j = filler j i1 i +-- -- where +-- -- rhs-filler : I → I → I → A +-- -- rhs-filler i j z = hfill (λ k → λ {(i = i0) → b ; +-- -- (i = i1) → s (~ z ∧ k) ; +-- -- (z = i1) → q i}) +-- -- (inS (q i)) j +-- -- filler : I → I → I → A +-- -- filler i j z = hfill (λ k → λ {(i = i0) → p z ; +-- -- (i = i1) → rhs-filler k i1 z ; +-- -- (z = i1) → q (k ∧ i)}) +-- -- (inS (p (i ∨ z))) +-- -- j + + +-- -- module S1-smash-Iso ((A , pt) : Pointed ℓ) where + +-- -- funInv : Susp A → (S¹ , base) smash (A , pt) +-- -- funInv north = inl tt +-- -- funInv south = inl tt +-- -- funInv (merid a i) = (sym (projₗ a) ∙ +-- -- cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ +-- -- projₗ a) i + + + +-- -- fun : (S¹ , base) smash (A , pt) → (Susp A) +-- -- fun (inl tt) = north +-- -- fun (inr (x , x₁)) = f* {a = pt} x x₁ +-- -- fun (push (inl base) i) = north +-- -- fun (push (inl (loop i₁)) i) = rCancel (merid pt) (~ i) i₁ -- rCancel (merid pt) (~ i) i₁ +-- -- fun (push (inr x) i) = north +-- -- fun (push (push a i₁) i) = north + +-- -- test18 : cong (cong funInv) (rCancel (merid pt)) ≡ congFunct funInv (merid pt) (sym (merid pt)) ∙ rCancel (cong funInv (merid pt)) -- congFunct {!cong!} {!!} {!!} ∙ rCancel {!!} +-- -- test18 = {!!} + +-- -- funSect : (x : Susp A) → fun (funInv x) ≡ x +-- -- funSect north = refl +-- -- funSect south = merid pt +-- -- funSect (merid a i) = hcomp (λ k → λ {(i = i0) → ((λ j → rCancel refl j ∙ refl ∙ refl) ∙ +-- -- λ j → lUnit (lUnit refl (~ j)) (~ j)) k ; +-- -- (i = i1) → ((λ j → lUnit refl (~ j) ∙ sym (lUnit (rUnit (merid a ∙ sym (merid pt)) (~ j)) (~ j)) ∙ merid a) ∙ +-- -- (λ j → lUnit (symDistr (merid a) (sym (merid pt)) j ∙ merid a) (~ j)) ∙ +-- -- (sym (assoc (merid pt) (sym (merid a)) (merid a))) ∙ +-- -- (λ j → merid pt ∙ lCancel (merid a) j) ∙ +-- -- sym (rUnit (merid pt)) ) k}) +-- -- (helper2 i ∙ filler i) +-- -- where + +-- -- filler : (i : I) → (refl ∙ (merid a ∙ sym (merid pt)) ∙ refl) i ≡ merid a i +-- -- filler i = (λ j → (refl ∙ (merid a ∙ sym (merid pt)) ∙ refl) (i ∧ (~ j))) ∙ λ j → merid a (j ∧ i) + +-- -- helper2 : (i : I) → fun ((sym (projₗ a) ∙ +-- -- cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ +-- -- projₗ a) i) +-- -- ≡ ((λ i → fun (sym (projₗ a) i)) ∙ +-- -- (λ i → fun (cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop i)) ∙ +-- -- λ i → fun (projₗ a i)) i +-- -- helper2 i = (λ j → congFunct fun (sym (projₗ a)) (cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ projₗ a) j i) ∙ +-- -- (λ j → (cong fun (sym (projₗ a)) ∙ congFunct fun (cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop) (projₗ a) j) i) + +-- -- retrHelper : (x : S¹) (a : A) → funInv (f* {a = pt} x a) ≡ proj x a +-- -- retrHelper base a = sym (projₗ a) +-- -- retrHelper (loop i) a j = hcomp (λ k → λ { (i = i0) → push (inr a) j +-- -- ; (i = i1) → push (inr a) j +-- -- ; (j = i0) → id3 (~ k) i +-- -- ; (j = i1) → proj (loop i) a}) +-- -- (compLrFiller (push (inr a)) (λ i → proj (loop i) a) (sym (push (inr a))) j i) + + +-- -- {- hcomp (λ k → λ { (i = i0) → ((λ j → id3 j ∙ sym (projₗ a) ∙ λ j → proj (loop (~ j)) a) ∙ +-- -- invCancelHelper _ _) k -- j +-- -- ; (i = i1) → lUnit (rUnit (sym (projₗ a)) (~ k)) (~ k) -- j +-- -- -- ; (j = i0) → funInv (f* {a = pt} (loop i) a) +-- -- -- ; (j = i1) → proj (loop i) a +-- -- } ) +-- -- (filler) -} + +-- -- module _ where + +-- -- invCancelHelper : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) (q : y ≡ y) → (p ∙ q ∙ sym p) ∙ p ∙ sym q ≡ p +-- -- invCancelHelper p q = (p ∙ q ∙ sym p) ∙ p ∙ sym q ≡⟨ (λ i → assoc p q (sym p) i ∙ p ∙ sym q) ⟩ +-- -- ((p ∙ q) ∙ sym p) ∙ p ∙ sym q ≡⟨ assoc ((p ∙ q) ∙ sym p) p (sym q) ⟩ +-- -- (((p ∙ q) ∙ sym p) ∙ p) ∙ sym q ≡⟨ (λ i → assoc (p ∙ q) (sym p) p (~ i) ∙ sym q) ⟩ +-- -- ((p ∙ q) ∙ sym p ∙ p) ∙ sym q ≡⟨ (λ i → ((p ∙ q) ∙ lCancel p i) ∙ sym q) ⟩ +-- -- ((p ∙ q) ∙ refl) ∙ sym q ≡⟨ (λ i → rUnit (p ∙ q) (~ i) ∙ sym q) ⟩ +-- -- (p ∙ q) ∙ sym q ≡⟨ sym (assoc p q (sym q)) ⟩ +-- -- p ∙ q ∙ sym q ≡⟨ (λ i → p ∙ rCancel q i) ⟩ +-- -- p ∙ refl ≡⟨ sym (rUnit p) ⟩ +-- -- p ∎ + +-- -- filler : funInv (f* {a = pt} (loop i) a) ≡ proj (loop i) a +-- -- filler = (λ j → funInv (funExt (λ x → merid x ∙ sym (merid pt)) (i ∨ j) a)) ∙ sym (projₗ a) ∙ λ j → proj (loop (i ∨ (~ j))) a + +-- -- id1 : sym (projᵣ base) ∙ cong ((λ x → proj {A = S¹ , base} {B = A , pt} x pt)) loop ∙ projᵣ base ≡ refl +-- -- id1 = (λ i → (push (inl (loop i))) ∙ (λ j → inr (loop (i ∨ j) , pt)) ∙ sym (push (inl base))) ∙ +-- -- (λ i → push (inl base) ∙ lUnit (sym (push (inl base))) (~ i)) ∙ +-- -- rCancel (push (inl base)) + +-- -- id2 : cong funInv (sym (merid pt)) ≡ (λ _ → inl tt) +-- -- id2 = (λ i j → (sym (projₗ pt) ∙ (cong (λ x → proj {A = S¹ , base} {B = A , pt} x pt) loop) ∙ projₗ pt) (~ j)) ∙ +-- -- (λ i j → (sym ((projᵣₗ (~ i))) ∙ cong (λ x → proj {A = S¹ , base} {B = A , pt} x pt) loop ∙ projᵣₗ (~ i)) (~ j)) ∙ +-- -- (λ i j → id1 i (~ j)) + +-- -- id3 : (λ i → funInv (f* {a = pt} (loop i) a)) ≡ ((sym (projₗ a) ∙ cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ (projₗ a))) +-- -- id3 = (λ j i → congFunct funInv (merid a) (sym (merid pt)) j i ) ∙ +-- -- (λ j i → (cong funInv (merid a) ∙ id2 j) i) ∙ +-- -- (λ j i → rUnit (cong funInv (merid a)) (~ j) i) + + + +-- -- test5 : ((sym (projₗ a) ∙ cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ (projₗ a))) ≡ λ j → funInv (merid a j) +-- -- test5 = refl + +-- -- test6 : (ia ib : I) → (x : S¹) → retrHelper (loop ia) pt ≡ {!? ∙ sym (cong (projₗ) loop) ∙ ?!} +-- -- test6 = {!!} +-- -- where +-- -- test7 : ((sym (projₗ pt) ∙ cong (λ x → proj {A = S¹ , base} {B = A , pt} x pt) loop ∙ (projₗ pt))) ≡ refl +-- -- test7 = (λ i j → funInv (merid pt j)) ∙ λ i → {!!} +-- -- {- +-- -- test2 : (i : I) → (retrHelper (loop i) pt) ≡ (λ j → funInv (rCancel (merid pt) j i)) ∙ sym (projᵣ (loop i)) +-- -- test2 i = (λ z → hfill ((λ k → λ { (i = i0) → ((λ j → id3 i pt j ∙ sym (projₗ pt) ∙ λ j → proj (loop (~ j)) pt) ∙ invCancelHelper i pt _ _) k +-- -- ; (i = i1) → lUnit (rUnit (sym (projₗ pt)) (~ k)) (~ k) })) +-- -- (inS ((λ j → funInv (funExt (λ x → merid x ∙ sym (merid pt)) (i ∨ j) pt)) ∙ sym (projₗ pt) ∙ λ j → proj (loop (i ∨ (~ j))) pt)) (~ z)) ∙ +-- -- (λ z → (λ j → funInv (funExt (λ x → merid x ∙ sym (merid pt)) (i ∨ j) pt)) ∙ sym (projₗ pt) ∙ λ j → proj (loop (i ∨ (~ j))) pt) ∙ +-- -- (λ z → (λ j → funInv {!funExt (λ x → merid x ∙ sym (merid pt)) (i ∨ j) pt -- rCancel (merid pt) z (i ∨ j)!}) ∙ {!!} ∙ {!!}) ∙ +-- -- {!!} +-- -- -} +-- -- isoHelperPre : (x : S¹ × A) (v : (S¹ , base) wedge (A , pt)) → {!!} +-- -- isoHelperPre = {!!} + +-- -- isoHelper : (x : S¹ × A) (v : (S¹ , base) wedge (A , pt)) (q : inr x ≡ inl tt) (i : I) → (funInv (fun (q i)) ≡ q i) ≡ (funInv (fun ((q ∙ push v) i)) ≡ (q ∙ push v) i) +-- -- isoHelper x v q i j = funInv (fun (test2 j)) ≡ test2 j +-- -- where +-- -- test2 : q i ≡ (q ∙ push v) i +-- -- test2 = (λ j → q (i ∧ (~ j))) ∙ λ j → (q ∙ push v) (i ∧ j) + +-- -- idIsEquiv2 : ∀ {ℓ} {A B : Type ℓ} → A ≡ B → A ≃ B +-- -- idIsEquiv2 = J (λ B p → _ ≃ B) (idEquiv _) + +-- -- helper2 : (x : S¹ × A) → funInv (fun (inr x)) ≡ inr x +-- -- helper2 (x , x₁) = retrHelper x x₁ +-- -- funRetr2 : (x : (S¹ , base) smash (A , pt)) → funInv (fun x) ≡ x +-- -- funRetr2 (inl tt) = {!!} +-- -- funRetr2 (inr x) = {!!} +-- -- funRetr2 (push a i) = ElimR.elimL (λ b path → funInv (fun (path (~ i))) ≡ path (~ i)) (λ b path → funInv (fun ((path) (~ i))) ≡ (path) (~ i)) +-- -- (helper2 (refl i)) -- (helper2 (refl i)) +-- -- (λ z q → idIsEquiv2 {!!})-- isoToEquiv (isoHelper _ a q (~ i))) +-- -- tt (sym (push a)) +-- -- where +-- -- test3 : (a₁ : (S¹ , base) wedge (A , pt)) → (q : inr (i∧ a) ≡ inl tt) → Iso ((funInv (fun (q (~ i)))) ≡ q (~ i)) (funInv (fun ((q ∙ push a₁) (~ i))) ≡ (q ∙ push a₁) (~ i)) +-- -- test3 (inl x) q = {!!} +-- -- test3 (inr x) q = {!!} +-- -- test3 (push a i) q = {!!} + +-- -- test2 : (x : S¹ × A) (v : (S¹ , base) wedge (A , pt)) (q : Path ((S¹ , base) smash (A , pt)) (inr x) (inl tt)) → q (~ i) ≡ (q ∙ push v) (~ i) +-- -- test2 x v q = (λ j → q ((~ i) ∧ (~ j))) ∙ λ j → (q ∙ push v) ((~ i) ∧ j) + +-- -- funRetr : (x : (S¹ , base) smash (A , pt)) → funInv (fun x) ≡ x +-- -- funRetr (inl tt) = refl +-- -- funRetr (inr (x , a)) = retrHelper x a +-- -- funRetr (push (inl base) i) j = hcomp (λ k → λ { (i = i0) → inl tt +-- -- ; (i = i1) → sym (projᵣₗ k) j +-- -- ; (j = i0) → inl tt +-- -- ; (j = i1) → projᵣ base (~ i)}) +-- -- (projᵣ base ((~ j) ∨ (~ i))) +-- -- funRetr (push (inl (loop z)) i) j = {!(projᵣ base (~ j ∨ ~ i))!} +-- -- {- hcomp (λ k → λ {(i₁ = i0) → ((λ l → hcomp (λ r → λ {(i = i0) → (rCancel (λ _ → inl tt)) r +-- -- ; (i = i1) → (lUnit ((λ j → push (inl base) (i ∧ j))) (~ l)) }) +-- -- (lUnit ((λ j → push (inl base) (i ∧ j))) (~ i)))) k +-- -- ; (i₁ = i1) → ((λ l → hcomp (λ r → λ {(i = i0) → (rCancel (λ _ → inl tt)) r +-- -- ; (i = i1) → (lUnit ((λ j → push (inl base) (i ∧ j))) (~ l)) }) +-- -- (lUnit ((λ j → push (inl base) (i ∧ j))) (~ i)))) k +-- -- ; (i = i0) → ({!!} ∙ {!!} ∙ {!push inl !}) k -- rCancel refl k +-- -- ; (i = i1) → {!!}}) +-- -- ((hcomp (λ k → λ {(i = i0) → (rCancel (λ _ → inl tt)) k +-- -- ; (i = i1) → _ }) ((λ j → funInv (rCancel (merid pt) (~ i ∨ j) i₁)) ∙ +-- -- λ j → push (inl (loop i₁)) (i ∧ j)))) -} +-- -- where + +-- -- bottom : funInv (rCancel (merid pt) (~ i) z) ≡ push (inl (loop z)) i +-- -- bottom = (λ j → funInv(rCancel (merid pt) (~ i) (z ∨ j))) ∙ (λ j → push (inl (loop z)) (i ∧ j)) + +-- -- i=i1 : (λ j → funInv(rCancel (merid pt) i0 (z ∨ j))) ∙ (λ j → push (inl (loop z)) j) ≡ retrHelper (loop z) pt +-- -- i=i1 = {!-- rCancel (merid pt) i1!} + +-- -- frontCubeFiller : (i j z : I) → (S¹ , base) smash (A , pt) -- bottom is i j , z is height +-- -- frontCubeFiller i j z = hfill (λ k → λ { (i = i0) → inl tt +-- -- ; (i = i1) → sym (projᵣₗ k) j +-- -- ; (j = i0) → inl tt +-- -- ; (j = i1) → projᵣ base (~ i)}) (inS (push (inl base) (i ∧ j))) z + +-- -- bottomCubeFiller : (i j z : I) → (S¹ , base) smash (A , pt) +-- -- bottomCubeFiller i j z = funInv (rCancel (merid pt) (~ i) (z ∧ j)) + + +-- -- theHComp : (i j z : I) → (S¹ , base) smash (A , pt) +-- -- theHComp i j z = {!!} + + + + +-- -- funRetr (push (inr x) i) = λ j → (projₗ x ((~ i) ∨ (~ j))) +-- -- funRetr (push (push tt i₁) i) = {! + +-- -- i₁ = i0 ⊢ hcomp +-- -- (λ { k (i = i0) → λ _ → inl tt +-- -- ; k (i = i1) → λ i₂ → projᵣₗ k (~ i₂) +-- -- }) +-- -- (λ i₂ → projᵣ base (~ i₂ ∨ ~ i)) +-- -- i₁ = i1 ⊢ hcomp +-- -- (λ { k (i = i0) → λ _ → inl tt +-- -- ; k (i = i1) → λ i₂ → projᵣₗ k (~ i₂) +-- -- }) +-- -- (λ i₂ → projᵣ base (~ i₂ ∨ ~ i)) +-- -- i = i0 ⊢ λ _ → inl tt + +-- -- i = i1 ⊢ !} diff --git a/Cubical/Homotopy/Connected.agda b/Cubical/Homotopy/Connected.agda index 88d12bc70..4674f55c9 100644 --- a/Cubical/Homotopy/Connected.agda +++ b/Cubical/Homotopy/Connected.agda @@ -18,7 +18,7 @@ open import Cubical.HITs.Susp open import Cubical.HITs.Truncation as Trunc open import Cubical.Homotopy.Loopspace open import Cubical.HITs.Pushout -open import Cubical.HITs.Sn +open import Cubical.HITs.Sn.Base open import Cubical.Data.Unit open import Cubical.Data.NatMinusTwo.Base From ce9ce2ceac4b868a1af0c2f6e9d186a67d426a08 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Sun, 26 Apr 2020 23:42:37 +0200 Subject: [PATCH 21/89] smash-assoc --- Cubical/HITs/SmashProduct/Base.agda | 88 +++++++++++++++++++++++++++-- 1 file changed, 82 insertions(+), 6 deletions(-) diff --git a/Cubical/HITs/SmashProduct/Base.agda b/Cubical/HITs/SmashProduct/Base.agda index 172c4fb5f..c1bb06f08 100644 --- a/Cubical/HITs/SmashProduct/Base.agda +++ b/Cubical/HITs/SmashProduct/Base.agda @@ -63,6 +63,9 @@ i∧ {A = A , ptA} {B = B , ptB} (push tt i) = ptA , ptB _⋀_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Type (ℓ-max ℓ ℓ') _⋀_ A B = Pushout {A = (A ⋁ B)} {B = Unit} {C = (typ A) × (typ B)} (λ _ → tt) i∧ +_⋀∙_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Pointed (ℓ-max ℓ ℓ') +A ⋀∙ B = Pushout {A = (A ⋁ B)} {B = Unit} {C = (typ A) × (typ B)} (λ _ → tt) i∧ , (inl tt) + _⋀⃗_ : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} {D : Pointed ℓ'''} (f : A →∙ C) (g : B →∙ D) → A ⋀ B → C ⋀ D (f ⋀⃗ g) (inl tt) = inl tt @@ -104,9 +107,82 @@ Iso.leftInv ⋀-comm (push (inl x) i) = refl Iso.leftInv ⋀-comm (push (inr x) i) = refl Iso.leftInv ⋀-comm (push (push a i₁) i) = refl -rearrange-proj : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} → A ⋀ B → (C ⋀ B , inl tt) ⋀ A -rearrange-proj (inl x) = inl tt -rearrange-proj {C = C} (inr (x , x₁)) = inr (inr ((snd C) , x₁) , x) -rearrange-proj {A = A}{B = B} {C = C} (push (inl x) i) = (push (inr x) ∙ (λ i → inr (push (inr (snd B)) i , x))) i -rearrange-proj {A = A}{B = B} {C = C} (push (inr x) i) = (push (inr (snd A)) ∙ (λ i → inr (push (inr x) i , snd A))) i -rearrange-proj {A = A}{B = B} {C = C} (push (push a i₁) i) = (push (inr (snd A)) ∙ (λ i → inr (push (inr (snd B)) i , (snd A)))) i +⋀-associate : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} → A ⋀ (B ⋀∙ C) → (C ⋀∙ B) ⋀ A +⋀-associate (inl x) = inl tt +⋀-associate (inr (x , inl x₁)) = inr ((inl tt) , x) +⋀-associate (inr (x , inr (x₁ , x₂))) = inr ((inr (x₂ , x₁)) , x) +⋀-associate (inr (x , push (inl x₁) i)) = inr (push (inr x₁) i , x) +⋀-associate (inr (x , push (inr x₁) i)) = inr (push (inl x₁) i , x) +⋀-associate (inr (x , push (push a i₁) i)) = inr (push (push tt (~ i₁)) i , x) +⋀-associate (push (inl x) i) = push (inr x) i +⋀-associate {A = A} (push (inr (inl tt)) i) = push (inr (snd A)) i +⋀-associate (push (inr (inr (x , x₁))) i) = push (inl (inr (x₁ , x))) i +⋀-associate {A = A} {B = B} {C = C} (push (inr (push (inl x) j)) i) = + hcomp (λ k → λ { (i = i0) → inl tt + ; (i = i1) → (inr (push (inr x) j , snd A)) + ; (j = i0) → push (push tt k) i + ; (j = i1) → push (inl (inr (snd C , x))) i}) + (push (inl (push (inr x) j)) i) +⋀-associate {A = A} {B = B} {C = C} (push (inr (push (inr x) j)) i) = + hcomp ((λ k → λ { (i = i0) → inl tt + ; (i = i1) → inr (push (inl x) j , snd A) + ; (j = i0) → push (push tt k) i + ; (j = i1) → push (inl (inr (x , snd B))) i})) + (push (inl (push (inl x) j)) i) +⋀-associate {A = A} {B = B} {C = C} (push (inr (push (push tt z) j)) i) = + hcomp (λ k → λ { (i = i0) → inl tt + ; (i = i1) → inr ((push (push tt (~ z)) j) , (snd A)) + ; (j = i0) → push (push tt k) i + ; (j = i1) → push (inl (inr (snd C , (snd B)))) i + }) + (push (inl (push (push tt (~ z)) j)) i) +⋀-associate {A = A} (push (push a i₁) i) = push (inr (snd A)) i + +⋀-associate⁻ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} → (C ⋀∙ B) ⋀ A → A ⋀ (B ⋀∙ C) +⋀-associate⁻ (inl x) = inl tt +⋀-associate⁻ (inr (inl x , x₁)) = inr (x₁ , (inl tt)) +⋀-associate⁻ (inr (inr (x , x₂) , x₁)) = inr (x₁ , (inr (x₂ , x))) +⋀-associate⁻ (inr (push (inl x) i , x₁)) = inr (x₁ , push (inr x) i) +⋀-associate⁻ (inr (push (inr x) i , x₁)) = inr (x₁ , push (inl x) i) +⋀-associate⁻ (inr (push (push a i₁) i , x)) = inr (x , push (push tt (~ i₁)) i) +⋀-associate⁻ {A = A} (push (inl (inl x)) i) = push (inl (snd A)) i +⋀-associate⁻ (push (inl (inr (x , x₁))) i) = push (inr (inr (x₁ , x))) i +⋀-associate⁻ {A = A} {B = B} (push (inl (push (inl x) j)) i) = + hcomp (λ k → λ {(i = i0) → inl tt + ; (i = i1) → inr (snd A , push (inr x) j) + ; (j = i0) → push (push tt (~ k)) i + ; (j = i1) → push (inr (inr ((snd B) , x))) i}) + (push (inr (push (inr x) j)) i) +⋀-associate⁻ {A = A} {C = C} (push (inl (push (inr x) j)) i) = + hcomp (λ k → λ {(i = i0) → inl tt + ; (i = i1) → inr (snd A , push (inl x) j) + ; (j = i0) → push (push tt (~ k)) i + ; (j = i1) → push (inr (inr (x , (snd C)))) i}) + (push (inr (push (inl x) j)) i) +⋀-associate⁻ {A = A} {B = B} {C = C} (push (inl (push (push a z) j)) i) = + hcomp (λ k → λ { (i = i0) → inl tt + ; (i = i1) → inr (snd A , (push (push tt (~ z)) j)) + ; (j = i0) → push (push tt (~ k)) i + ; (j = i1) → push (inr (inr (snd B , snd C))) i + }) + (push (inr (push (push tt (~ z)) j)) i) +⋀-associate⁻ (push (inr x) i) = push (inl x) i +⋀-associate⁻ {A = A} (push (push a i₁) i) = push (inl (snd A)) i + +⋀-Iso : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} → Iso (A ⋀ (B ⋀∙ C)) ((C ⋀∙ B) ⋀ A) +Iso.fun ⋀-Iso = ⋀-associate +Iso.inv ⋀-Iso = ⋀-associate⁻ +Iso.rightInv ⋀-Iso (inl x) = refl +Iso.rightInv ⋀-Iso (inr (inl x , x₁)) = refl +Iso.rightInv ⋀-Iso (inr (inr (x , x₂) , x₁)) = refl +Iso.rightInv ⋀-Iso (inr (push (inl x) i , x₁)) z = inr (push (inl x) i , x₁) +Iso.rightInv ⋀-Iso (inr (push (inr x) i , x₁)) z = inr (push (inr x) i , x₁) +Iso.rightInv ⋀-Iso (inr (push (push tt j) i , x₁)) z = inr (push (push tt j) i , x₁) +Iso.rightInv ⋀-Iso (push (inl (inl x)) i) z = push (push tt (~ z)) i +Iso.rightInv ⋀-Iso (push (inl (inr (x , x₁))) i) j = push (inl (inr (x , x₁))) i +Iso.rightInv ⋀-Iso (push (inl (push (inl x) i₁)) i) = {!!} +Iso.rightInv ⋀-Iso (push (inl (push (inr x) i₁)) i) = {!!} +Iso.rightInv ⋀-Iso (push (inl (push (push a i₂) i₁)) i) = {!a!} +Iso.rightInv ⋀-Iso (push (inr x) i) = {!!} +Iso.rightInv ⋀-Iso (push (push a i₁) i) = {!!} +Iso.leftInv ⋀-Iso = {!!} From dbbf691c05c2a7e992939db2065ef4399e424427 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Sun, 3 May 2020 02:10:27 +0200 Subject: [PATCH 22/89] K-grousp --- Cubical/Data/Group/Base.agda | 7 + Cubical/Foundations/Pointed/Base.agda | 3 + Cubical/HITs/SetTruncation/Properties.agda | 25 ++ Cubical/HITs/SmashProduct/Base.agda | 287 +++++++------ Cubical/HITs/Sn/Properties.agda | 71 ++-- Cubical/HITs/Truncation/Properties.agda | 32 +- Cubical/HITs/Wedge.agda | 4 + Cubical/HITs/Wedge/Base.agda | 24 ++ Cubical/Homotopy/Connected.agda | 466 +++++++++++++++++++-- Cubical/Homotopy/Loopspace.agda | 57 +-- Cubical/ZCohomology/MayerVietoris.agda | 212 ++++++++++ Cubical/ZCohomology/Properties.agda | 405 ++++++++++++------ Cubical/ZCohomology/cupProdPrelims.agda | 462 +++++++++++--------- 13 files changed, 1487 insertions(+), 568 deletions(-) create mode 100644 Cubical/HITs/Wedge.agda create mode 100644 Cubical/HITs/Wedge/Base.agda create mode 100644 Cubical/ZCohomology/MayerVietoris.agda diff --git a/Cubical/Data/Group/Base.agda b/Cubical/Data/Group/Base.agda index 6b26fdbef..9ad745dba 100644 --- a/Cubical/Data/Group/Base.agda +++ b/Cubical/Data/Group/Base.agda @@ -7,6 +7,7 @@ open import Cubical.Foundations.Prelude hiding ( comp ) import Cubical.Foundations.Isomorphism as I import Cubical.Foundations.Equiv as E import Cubical.Foundations.Equiv.HalfAdjoint as HAE +open import Cubical.Data.Prod.Base record isGroup {ℓ} (A : Type ℓ) : Type ℓ where constructor group-struct @@ -27,6 +28,12 @@ record Group ℓ : Type (ℓ-suc ℓ) where setStruc : isSet type groupStruc : isGroup type +isAbelianGroup : ∀ {ℓ} → Type (ℓ-suc ℓ) +isAbelianGroup {ℓ} = Σ[ G ∈ Group ℓ ] (∀ a b → isGroup.comp (Group.groupStruc G) a b ≡ isGroup.comp (Group.groupStruc G) b a) + +-- (group type setStruc (group-struct id inv comp lUnit rUnit assoc lCancel rCancel)) = + -- (a b : type) → comp a b ≡ comp b a + isMorph : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (f : (Group.type G → Group.type H)) → Type (ℓ-max ℓ ℓ') isMorph (group G Gset (group-struct _ _ _⊙_ _ _ _ _ _)) (group H Hset (group-struct _ _ _∘_ _ _ _ _ _)) f diff --git a/Cubical/Foundations/Pointed/Base.agda b/Cubical/Foundations/Pointed/Base.agda index 8f056dd16..d363888b9 100644 --- a/Cubical/Foundations/Pointed/Base.agda +++ b/Cubical/Foundations/Pointed/Base.agda @@ -22,3 +22,6 @@ _→∙_ A B = Σ[ f ∈ (typ A → typ B) ] f (pt A) ≡ pt B _→∙_∙ : ∀{ℓ ℓ'} → (A : Pointed ℓ) (B : Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') A →∙ B ∙ = (A →∙ B) , (λ x → pt B) , refl + +idfun∙ : ∀ {ℓ} (A : Pointed ℓ) → A →∙ A +idfun∙ A = (λ x → x) , refl diff --git a/Cubical/HITs/SetTruncation/Properties.agda b/Cubical/HITs/SetTruncation/Properties.agda index 9fd2d4952..3196eca14 100644 --- a/Cubical/HITs/SetTruncation/Properties.agda +++ b/Cubical/HITs/SetTruncation/Properties.agda @@ -16,6 +16,7 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv open import Cubical.Foundations.HLevels open import Cubical.Foundations.Univalence +open import Cubical.Data.Sigma private variable @@ -74,3 +75,27 @@ setTruncIdempotent≃ {A = A} hA = isoToEquiv f setTruncIdempotent : isSet A → ∥ A ∥₀ ≡ A setTruncIdempotent hA = ua (setTruncIdempotent≃ hA) + + +setSigmaIso : ∀ {ℓ} {B : A → Type ℓ} → Iso ∥ Σ A B ∥₀ ∥ Σ A (λ x → ∥ B x ∥₀) ∥₀ +setSigmaIso {A = A} {B = B} = iso fun funinv sect retr + where + {- writing it out explicitly to avoid yellow highlighting -} + fun : ∥ Σ A B ∥₀ → ∥ Σ A (λ x → ∥ B x ∥₀) ∥₀ + fun = rec setTruncIsSet λ {(a , p) → ∣ a , ∣ p ∣₀ ∣₀} + funinv : ∥ Σ A (λ x → ∥ B x ∥₀) ∥₀ → ∥ Σ A B ∥₀ + funinv = rec setTruncIsSet (λ {(a , p) → rec setTruncIsSet (λ p → ∣ a , p ∣₀) p}) + sect : section fun funinv + sect = elim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ { (a , p) → elim {B = λ p → fun (funinv ∣ a , p ∣₀) ≡ ∣ a , p ∣₀} + (λ p → isOfHLevelPath 2 setTruncIsSet _ _) (λ p → refl) p } + retr : retract fun funinv + retr = elim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ { (a , p) → refl } + + +setSigmaIso2 : ∀ {ℓ} {B : A → Type ℓ} (f g : (x : A) → B x) (a : A) (p1 p2 : f a ≡ g a) + → ((x : A) → isSet (B x)) + → Path ∥ Σ[ x ∈ A ] f x ≡ g x ∥₀ ∣ (a , p1) ∣₀ ∣ (a , p2) ∣₀ +setSigmaIso2 f g a p1 p2 set i = ∣ a , set a (f a) (g a) p1 p2 i ∣₀ + diff --git a/Cubical/HITs/SmashProduct/Base.agda b/Cubical/HITs/SmashProduct/Base.agda index c1bb06f08..ab347e972 100644 --- a/Cubical/HITs/SmashProduct/Base.agda +++ b/Cubical/HITs/SmashProduct/Base.agda @@ -7,24 +7,41 @@ open import Cubical.Foundations.Isomorphism open import Cubical.HITs.Pushout.Base open import Cubical.Data.Unit open import Cubical.Data.Prod +open import Cubical.HITs.Wedge +open import Cubical.Foundations.Function open import Cubical.Foundations.GroupoidLaws -data Smash (A B : Pointed₀) : Type₀ where +data Smash {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') : Type (ℓ-max ℓ ℓ') where basel : Smash A B baser : Smash A B proj : (x : typ A) → (y : typ B) → Smash A B gluel : (a : typ A) → proj a (pt B) ≡ basel gluer : (b : typ B) → proj (pt A) b ≡ baser +private + variable + ℓ ℓ' ℓ'' ℓ''' : Level + A : Pointed ℓ + B : Pointed ℓ' + C : Pointed ℓ'' + D : Pointed ℓ''' + +Smash-map : (f : A →∙ C) (g : B →∙ D) → Smash A B → Smash C D +Smash-map f g basel = basel +Smash-map f g baser = baser +Smash-map (f , fpt) (g , gpt) (proj x y) = proj (f x) (g y) +Smash-map (f , fpt) (g , gpt) (gluel a i) = ((λ j → proj (f a) (gpt j)) ∙ gluel (f a)) i +Smash-map (f , fpt) (g , gpt) (gluer b i) = ((λ j → proj (fpt j) (g b)) ∙ gluer (g b)) i + -- Commutativity -comm : {A B : Pointed₀} → Smash A B → Smash B A +comm : Smash A B → Smash B A comm basel = baser comm baser = basel comm (proj x y) = proj y x comm (gluel a i) = gluer a i comm (gluer b i) = gluel b i -commK : ∀ {A B : Pointed₀} → (x : Smash A B) → comm (comm x) ≡ x +commK : (x : Smash A B) → comm (comm x) ≡ x commK basel = refl commK baser = refl commK (proj x y) = refl @@ -33,28 +50,14 @@ commK (gluer b x) = refl -- WIP below -SmashPt : (A B : Pointed₀) → Pointed₀ +SmashPt : (A : Pointed ℓ) (B : Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') SmashPt A B = (Smash A B , basel) --- -- A (B C) = C (B A) = C (A B) = (A B) C --- rearrange : ∀ {A B C : Pointed₀} → Smash A (SmashPt B C) → Smash C (SmashPt B A) --- rearrange basel = baser --- rearrange baser = basel --- rearrange {B = B} {C = C} (proj x basel) = proj (pt C) baser --- rearrange {C = C} (proj x baser) = proj (pt C) basel -- ? --- rearrange (proj x (proj y z)) = proj z (proj y x) --- rearrange {C = C} (proj x (gluel a i)) = proj (pt C) {!!} --- rearrange (proj x (gluer b i)) = {!!} --- rearrange (gluel a i) = {!!} --- rearrange (gluer b i) = {!gluel ? i!} - - +Smash∙ : (A : Pointed ℓ) (B : Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') +Smash∙ A B = (Smash A B) , (proj (snd A) (snd B)) --- Alternative definition -_⋁_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Type (ℓ-max ℓ ℓ') -_⋁_ (A , ptA) (B , ptB) = Pushout {A = Unit} {B = A} {C = B} (λ {tt → ptA}) (λ {tt → ptB}) - i∧ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → A ⋁ B → (typ A) × (typ B) i∧ {A = A , ptA} {B = B , ptB} (inl x) = x , ptB i∧ {A = A , ptA} {B = B , ptB} (inr x) = ptA , x @@ -64,15 +67,15 @@ _⋀_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Type (ℓ-max ℓ _⋀_ A B = Pushout {A = (A ⋁ B)} {B = Unit} {C = (typ A) × (typ B)} (λ _ → tt) i∧ _⋀∙_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Pointed (ℓ-max ℓ ℓ') -A ⋀∙ B = Pushout {A = (A ⋁ B)} {B = Unit} {C = (typ A) × (typ B)} (λ _ → tt) i∧ , (inl tt) +A ⋀∙ B = Pushout {A = (A ⋁ B)} {B = Unit} {C = (typ A) × (typ B)} (λ _ → tt) i∧ , (inl tt) -_⋀⃗_ : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} {D : Pointed ℓ'''} (f : A →∙ C) (g : B →∙ D) → A ⋀ B → C ⋀ D +_⋀⃗_ : (f : A →∙ C) (g : B →∙ D) → A ⋀ B → C ⋀ D (f ⋀⃗ g) (inl tt) = inl tt ((f , fpt) ⋀⃗ (g , gpt)) (inr (x , x₁)) = inr (f x , g x₁) _⋀⃗_ {B = B} {D = D} (f , fpt) (b , gpt) (push (inl x) i) = (push (inl (f x)) ∙ (λ i → inr (f x , gpt (~ i)))) i -_⋀⃗_ {A = A} {B = B} {C = C} {D = D} (f , fpt) (g , gpt) (push (inr x) i) = (push (inr (g x)) ∙ (λ i → inr (fpt (~ i) , g x))) i -_⋀⃗_ {A = A} {B = B} {C = C} {D = D} (f , fpt) (g , gpt) (push (push tt j) i) = helper (~ j) i +_⋀⃗_ (f , fpt) (g , gpt) (push (inr x) i) = (push (inr (g x)) ∙ (λ i → inr (fpt (~ i) , g x))) i +_⋀⃗_ {A = A} {C = C} {B = B} {D = D} (f , fpt) (g , gpt) (push (push tt j) i) = helper (~ j) i where helper : Path (Path (C ⋀ D) (inl tt) (inr ((f (pt A)) , (g (pt B))))) (push (inr (g (pt B))) ∙ λ i → inr ((fpt (~ i)) , (g (pt B)))) @@ -83,106 +86,136 @@ _⋀⃗_ {A = A} {B = B} {C = C} {D = D} (f , fpt) (g , gpt) (push (push tt j) i (λ j → (push (inl (fpt (~ j))) ∙ λ i → inr (fpt ((~ j) ∨ i) , (gpt (~ i)))) ∙ λ i → inr ((fpt (~ i)) , (g (pt B)))) ∙ (λ j → (push (inl (f (pt A))) ∙ λ i → inr (fpt (i ∧ (~ j)) , gpt ((~ i) ∨ j))) ∙ λ i → inr ((fpt ((~ i) ∧ (~ j))) , (gpt ((~ i) ∧ j)))) ∙ (λ j → (rUnit (push (inl (f (pt A)))) (~ j)) ∙ λ i → inr ((f (pt A)) , (gpt (~ i)))) - - -⋀-commuter : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → A ⋀ B → B ⋀ A -⋀-commuter (inl tt) = inl tt -⋀-commuter (inr (x , x₁)) = inr (x₁ , x) -⋀-commuter (push (inl x) i) = push (inr x) i -⋀-commuter (push (inr x) i) = push (inl x) i -⋀-commuter (push (push a i₁) i) = push (push (tt) (~ i₁)) i - - -⋀-comm : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → Iso (A ⋀ B) (B ⋀ A) -Iso.fun ⋀-comm = ⋀-commuter -Iso.inv ⋀-comm = ⋀-commuter -Iso.rightInv ⋀-comm (inl tt) = refl -Iso.rightInv ⋀-comm (inr (x , x₁)) = refl -Iso.rightInv ⋀-comm (push (inl x) i) = refl -Iso.rightInv ⋀-comm (push (inr x) i) = refl -Iso.rightInv ⋀-comm (push (push a i₁) i) = refl -Iso.leftInv ⋀-comm (inl tt) = refl -Iso.leftInv ⋀-comm (inr (x , x₁)) = refl -Iso.leftInv ⋀-comm (push (inl x) i) = refl -Iso.leftInv ⋀-comm (push (inr x) i) = refl -Iso.leftInv ⋀-comm (push (push a i₁) i) = refl - -⋀-associate : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} → A ⋀ (B ⋀∙ C) → (C ⋀∙ B) ⋀ A -⋀-associate (inl x) = inl tt -⋀-associate (inr (x , inl x₁)) = inr ((inl tt) , x) -⋀-associate (inr (x , inr (x₁ , x₂))) = inr ((inr (x₂ , x₁)) , x) -⋀-associate (inr (x , push (inl x₁) i)) = inr (push (inr x₁) i , x) -⋀-associate (inr (x , push (inr x₁) i)) = inr (push (inl x₁) i , x) -⋀-associate (inr (x , push (push a i₁) i)) = inr (push (push tt (~ i₁)) i , x) -⋀-associate (push (inl x) i) = push (inr x) i -⋀-associate {A = A} (push (inr (inl tt)) i) = push (inr (snd A)) i -⋀-associate (push (inr (inr (x , x₁))) i) = push (inl (inr (x₁ , x))) i -⋀-associate {A = A} {B = B} {C = C} (push (inr (push (inl x) j)) i) = - hcomp (λ k → λ { (i = i0) → inl tt - ; (i = i1) → (inr (push (inr x) j , snd A)) - ; (j = i0) → push (push tt k) i - ; (j = i1) → push (inl (inr (snd C , x))) i}) - (push (inl (push (inr x) j)) i) -⋀-associate {A = A} {B = B} {C = C} (push (inr (push (inr x) j)) i) = - hcomp ((λ k → λ { (i = i0) → inl tt - ; (i = i1) → inr (push (inl x) j , snd A) - ; (j = i0) → push (push tt k) i - ; (j = i1) → push (inl (inr (x , snd B))) i})) - (push (inl (push (inl x) j)) i) -⋀-associate {A = A} {B = B} {C = C} (push (inr (push (push tt z) j)) i) = - hcomp (λ k → λ { (i = i0) → inl tt - ; (i = i1) → inr ((push (push tt (~ z)) j) , (snd A)) - ; (j = i0) → push (push tt k) i - ; (j = i1) → push (inl (inr (snd C , (snd B)))) i - }) - (push (inl (push (push tt (~ z)) j)) i) -⋀-associate {A = A} (push (push a i₁) i) = push (inr (snd A)) i - -⋀-associate⁻ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} → (C ⋀∙ B) ⋀ A → A ⋀ (B ⋀∙ C) -⋀-associate⁻ (inl x) = inl tt -⋀-associate⁻ (inr (inl x , x₁)) = inr (x₁ , (inl tt)) -⋀-associate⁻ (inr (inr (x , x₂) , x₁)) = inr (x₁ , (inr (x₂ , x))) -⋀-associate⁻ (inr (push (inl x) i , x₁)) = inr (x₁ , push (inr x) i) -⋀-associate⁻ (inr (push (inr x) i , x₁)) = inr (x₁ , push (inl x) i) -⋀-associate⁻ (inr (push (push a i₁) i , x)) = inr (x , push (push tt (~ i₁)) i) -⋀-associate⁻ {A = A} (push (inl (inl x)) i) = push (inl (snd A)) i -⋀-associate⁻ (push (inl (inr (x , x₁))) i) = push (inr (inr (x₁ , x))) i -⋀-associate⁻ {A = A} {B = B} (push (inl (push (inl x) j)) i) = - hcomp (λ k → λ {(i = i0) → inl tt - ; (i = i1) → inr (snd A , push (inr x) j) - ; (j = i0) → push (push tt (~ k)) i - ; (j = i1) → push (inr (inr ((snd B) , x))) i}) - (push (inr (push (inr x) j)) i) -⋀-associate⁻ {A = A} {C = C} (push (inl (push (inr x) j)) i) = - hcomp (λ k → λ {(i = i0) → inl tt - ; (i = i1) → inr (snd A , push (inl x) j) - ; (j = i0) → push (push tt (~ k)) i - ; (j = i1) → push (inr (inr (x , (snd C)))) i}) - (push (inr (push (inl x) j)) i) -⋀-associate⁻ {A = A} {B = B} {C = C} (push (inl (push (push a z) j)) i) = - hcomp (λ k → λ { (i = i0) → inl tt - ; (i = i1) → inr (snd A , (push (push tt (~ z)) j)) - ; (j = i0) → push (push tt (~ k)) i - ; (j = i1) → push (inr (inr (snd B , snd C))) i - }) - (push (inr (push (push tt (~ z)) j)) i) -⋀-associate⁻ (push (inr x) i) = push (inl x) i -⋀-associate⁻ {A = A} (push (push a i₁) i) = push (inl (snd A)) i - -⋀-Iso : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} → Iso (A ⋀ (B ⋀∙ C)) ((C ⋀∙ B) ⋀ A) -Iso.fun ⋀-Iso = ⋀-associate -Iso.inv ⋀-Iso = ⋀-associate⁻ -Iso.rightInv ⋀-Iso (inl x) = refl -Iso.rightInv ⋀-Iso (inr (inl x , x₁)) = refl -Iso.rightInv ⋀-Iso (inr (inr (x , x₂) , x₁)) = refl -Iso.rightInv ⋀-Iso (inr (push (inl x) i , x₁)) z = inr (push (inl x) i , x₁) -Iso.rightInv ⋀-Iso (inr (push (inr x) i , x₁)) z = inr (push (inr x) i , x₁) -Iso.rightInv ⋀-Iso (inr (push (push tt j) i , x₁)) z = inr (push (push tt j) i , x₁) -Iso.rightInv ⋀-Iso (push (inl (inl x)) i) z = push (push tt (~ z)) i -Iso.rightInv ⋀-Iso (push (inl (inr (x , x₁))) i) j = push (inl (inr (x , x₁))) i -Iso.rightInv ⋀-Iso (push (inl (push (inl x) i₁)) i) = {!!} -Iso.rightInv ⋀-Iso (push (inl (push (inr x) i₁)) i) = {!!} -Iso.rightInv ⋀-Iso (push (inl (push (push a i₂) i₁)) i) = {!a!} -Iso.rightInv ⋀-Iso (push (inr x) i) = {!!} -Iso.rightInv ⋀-Iso (push (push a i₁) i) = {!!} -Iso.leftInv ⋀-Iso = {!!} +{- perhaps rewrite this with hcomps ... -} + +⋀→Smash : A ⋀ B → Smash A B +⋀→Smash (inl x) = basel +⋀→Smash (inr (x , x₁)) = proj x x₁ +⋀→Smash (push (inl x) i) = gluel x (~ i) +⋀→Smash {A = A} {B = B} (push (inr x) i) = (sym (gluel (snd A)) ∙ gluer (snd B) ∙ sym (gluer x)) i +⋀→Smash {A = A} {B = B} (push (push a j) i) = (((λ i → sym (gluel (snd A)) ∙ + rCancel (gluer (snd B)) i) ∙ + sym (rUnit (sym (gluel (snd A))))) (~ j)) i + +Smash→⋀ : Smash A B → A ⋀ B +Smash→⋀ basel = inl tt +Smash→⋀ baser = inl tt +Smash→⋀ (proj x y) = inr (x , y) +Smash→⋀ (gluel a i) = push (inl a) (~ i) +Smash→⋀ (gluer b i) = push (inr b) (~ i) + + +private + pivotl : (b b' : typ B) + → Path (Smash A B) (proj (snd A) b) (proj (snd A) b') + pivotl b b' i = (gluer b ∙ sym (gluer b')) i + + pivotr : (a a' : typ A) + → Path (Smash A B) (proj a (snd B)) (proj a' (snd B)) + pivotr a a' i = (gluel a ∙ sym (gluel a')) i + + pivotlrId : {A : Pointed ℓ} {B : Pointed ℓ'} + → Path (Path (Smash A B) (proj (snd A) (snd B)) (proj (snd A) (snd B))) + (pivotl (snd B) (snd B)) + (pivotr (snd A) (snd A)) + pivotlrId {A = A} {B = B} = rCancel (gluer (snd B)) ∙ sym (rCancel (gluel (snd A))) + + rearrange-proj : (c : fst C) + → (Smash A B) → Smash (Smash∙ C B) A + rearrange-proj c basel = baser + rearrange-proj c baser = basel + rearrange-proj c (proj x y) = proj (proj c y) x + rearrange-proj {C = C} c (gluel a i) = + hcomp (λ k → λ { (i = i0) → proj (pivotr (snd C) c k) a + ; (i = i1) → baser}) + (gluer a i) + rearrange-proj c (gluer b i) = gluel (proj c b) i + + rearrange-gluer : (s : Smash A B) + → Path (Smash (Smash∙ C B) A) basel (rearrange-proj (snd C) s) + rearrange-gluer {A = A} {B = B} {C = C} basel = sym (gluel (proj (snd C) (snd B))) ∙ + gluer (snd A) + rearrange-gluer baser = refl + rearrange-gluer {A = A} {B = B} {C = C} (proj a b) i = + hcomp (λ k → λ { (i = i0) → (sym (gluel (proj (snd C) (snd B))) ∙ + gluer (snd A)) (~ k) + ; (i = i1) → proj (pivotl (snd B) b k) a}) + (gluer a (~ i)) + rearrange-gluer {A = A} {B = B} {C = C} (gluel a i) j = + hcomp (λ k → λ { (i = i1) → ((λ i₁ → gluel (proj (snd C) (snd B)) (~ i₁)) ∙ + gluer (snd A)) (~ k ∨ j) + ; (j = i0) → ((λ i₁ → gluel (proj (snd C) (snd B)) (~ i₁)) ∙ + gluer (snd A)) (~ k) + ; (j = i1) → top-filler i k}) + (gluer a (i ∨ ~ j)) + where + top-filler : I → I → Smash (Smash∙ C B) A + top-filler i j = + hcomp (λ k → λ { (i = i0) → side-filler k j + ; (i = i1) → gluer a (j ∨ k) + ; (j = i0) → gluer a (i ∧ k)}) + (gluer a (i ∧ j)) + where + side-filler : I → I → Smash (Smash∙ C B) A + side-filler i j = + hcomp (λ k → λ { (i = i0) → proj (proj (snd C) (snd B)) a + ; (i = i1) → proj ((rCancel (gluel (snd C)) ∙ sym (rCancel (gluer (snd B)))) k j) a + ; (j = i0) → proj (proj (snd C) (snd B)) a + ; (j = i1) → (proj ((gluel (snd C) ∙ sym (gluel (snd C))) i) a)}) + (proj ((gluel (snd C) ∙ sym (gluel (snd C))) (j ∧ i)) a) + rearrange-gluer {A = A} {B = B} {C = C} (gluer b i) j = + hcomp (λ k → λ {(i = i1) → ((sym (gluel (proj (snd C) (snd B)))) ∙ gluer (snd A)) (~ k) + ; (j = i0) → ((sym (gluel (proj (snd C) (snd B)))) ∙ gluer (snd A)) (~ k) + ; (j = i1) → top-filler1 i k}) + (gluer (snd A) (i ∨ (~ j))) + where + top-filler1 : I → I → Smash (Smash∙ C B) A + top-filler1 i j = + hcomp (λ k → λ { (i = i0) → congFunct (λ x → proj x (snd A)) (gluer (snd B)) (sym (gluer b)) (~ k) j + ; (i = i1) → (sym (gluel (proj (snd C) (snd B))) ∙ gluer (snd A)) (~ j) + ; (j = i0) → gluer (snd A) i + ; (j = i1) → gluel (proj (snd C) b) i}) + (top-filler2 i j) -- (top-filler2 i j) + where + top-filler2 : I → I → Smash (Smash∙ C B) A + top-filler2 i j = + hcomp (λ k → λ { (j = i0) → gluer (snd A) (i ∧ k) + ; (j = i1) → gluel (gluer b (~ k)) i}) + (hcomp (λ k → λ { (j = i0) → gluel (gluer (snd B) i0) (~ k ∧ (~ i)) + ; (j = i1) → gluel (baser) (~ k ∨ i) + ; (i = i0) → gluel (gluer (snd B) j) (~ k) + ; (i = i1) → gluel (proj (snd C) (snd B)) j }) + (gluel (proj (snd C) (snd B)) (j ∨ (~ i)))) + + rearrange : Smash (Smash∙ A B) C → Smash (Smash∙ C B) A + rearrange basel = basel + rearrange baser = baser + rearrange (proj x y) = rearrange-proj y x + rearrange (gluel a i) = rearrange-gluer a (~ i) + rearrange {A = A} {B = B} {C = C} (gluer b i) = ((λ j → proj (pivotr b (snd C) j) (snd A)) ∙ + gluer (snd A)) i + + ⋀∙→Smash∙ : (A ⋀∙ B) →∙ Smash∙ A B + ⋀∙→Smash∙ {A = A} {B = B} = fun , refl + where + fun : (A ⋀ B) → Smash A B + fun (inl x) = proj (snd A) (snd B) + fun (inr (x , x₁)) = proj x x₁ + fun (push (inl x) i) = pivotr (snd A) x i + fun (push (inr x) i) = pivotl (snd B) x i + fun (push (push a j) i) = pivotlrId (~ j) i + + Smash∙→⋀∙ : (Smash∙ A B) →∙ (A ⋀∙ B) + Smash∙→⋀∙ {A = A} {B = B} = Smash→⋀ , sym (push (inr (snd B))) + +SmashAssociate : Smash (Smash∙ A B) C → Smash A (Smash∙ B C) +SmashAssociate = comm ∘ Smash-map (comm , refl) (idfun∙ _) ∘ rearrange + +SmashAssociate⁻ : Smash A (Smash∙ B C) → Smash (Smash∙ A B) C +SmashAssociate⁻ = rearrange ∘ comm ∘ Smash-map (idfun∙ _) (comm , refl) + +⋀-associate : (A ⋀∙ B) ⋀ C → A ⋀ (B ⋀∙ C) +⋀-associate = (idfun∙ _ ⋀⃗ Smash∙→⋀∙) ∘ Smash→⋀ ∘ SmashAssociate ∘ ⋀→Smash ∘ (⋀∙→Smash∙ ⋀⃗ idfun∙ _) + +⋀-associate⁻ : A ⋀ (B ⋀∙ C) → (A ⋀∙ B) ⋀ C +⋀-associate⁻ = (Smash∙→⋀∙ ⋀⃗ idfun∙ _) ∘ Smash→⋀ ∘ SmashAssociate⁻ ∘ ⋀→Smash ∘ (idfun∙ _ ⋀⃗ ⋀∙→Smash∙) diff --git a/Cubical/HITs/Sn/Properties.agda b/Cubical/HITs/Sn/Properties.agda index 866ed2e61..442dc988a 100644 --- a/Cubical/HITs/Sn/Properties.agda +++ b/Cubical/HITs/Sn/Properties.agda @@ -67,15 +67,15 @@ SuspBool≃S1 = isoToEquiv iso1 Iso.leftInv iso1 (merid true i) = refl --- map between S₊ - -f' : {a : A} → A → S₊ 1 → Susp A -f' {a = pt} A north = north -f' {a = pt} A south = north -f' {a = pt} a (merid p i) = ((merid a) ∙ sym (merid pt)) i +-- map between S₊ +private + f' : {a : A} → A → S₊ 1 → Susp A + f' {a = pt} A north = north + f' {a = pt} A south = north + f' {a = pt} a (merid p i) = ((merid a) ∙ sym (merid pt)) i -proj' : {A : Pointed ℓ} {B : Pointed ℓ'} → typ A → typ B → A ⋀ B -proj' a b = inr (a , b) + proj' : {A : Pointed ℓ} {B : Pointed ℓ'} → typ A → typ B → A ⋀ B + proj' a b = inr (a , b) module smashS1→susp {(A , pt) : Pointed ℓ} where fun : (S₊ 1 , north) ⋀ (A , pt) → (Susp A) @@ -86,33 +86,20 @@ module smashS1→susp {(A , pt) : Pointed ℓ} where fun (push (inl (merid a i₁)) i) = rCancel (merid pt) (~ i) i₁ fun (push (inr x) i) = north fun (push (push tt i₁) i) = north - + fun⁻ : Susp A → (S₊ 1 , north) ⋀ (A , pt) fun⁻ north = inl tt fun⁻ south = inl tt - fun⁻ (merid a i) = (push (inr a) ∙∙ cong (λ x → proj' {A = S₊ 1 , north} {B = A , pt} x a) (merid south ∙ sym (merid north)) ∙∙ sym (push (inr a))) i - -smashMap : {!(A : Pointed ℓ) → ?!} -smashMap = {!!} - - -S' : ℕ → Type₀ -S' zero = Unit -S' (suc zero) = S₊ 1 -S' (suc (suc n)) = (S₊ 1 , north) ⋀ (S₊ (suc n) , north) - -hej : (m : ℕ) → S₊ (suc m) ≡ Susp (S₊ m) -hej m = refl - -S'-map-pt1 : (n m : ℕ) → (S₊ (suc (suc n)) , north) ⋀ (S₊ (suc m) , north) → ((S₊ 1 , north) ⋀ (S₊ (suc n) , north) , inl tt) ⋀ (S₊ (suc m) , north) -S'-map-pt1 n m = (smashS1→susp.fun⁻ , refl) ⋀⃗ ((λ x → x) , refl) - -S'-map : (n m : ℕ) → (S₊ (suc n) , north) ⋀ (S₊ (suc m) , north) → S' (2 + n + m) -S'-map zero n a = a -S'-map (suc n) m = {!!} ∘ {!!} ∘ S'-map-pt1 n m + fun⁻ (merid a i) = + (push (inr a) ∙∙ cong (λ x → proj' {A = S₊ 1 , north} {B = A , pt} x a) (merid south ∙ sym (merid north)) ∙∙ sym (push (inr a))) i sphereSmashMap : (n m : ℕ) → (S₊ (suc n) , north) ⋀ (S₊ (suc m) , north) → S₊ (2 + n + m) -sphereSmashMap n m = {!!} +sphereSmashMap zero m = smashS1→susp.fun +sphereSmashMap (suc n) m = + smashS1→susp.fun ∘ + (idfun∙ _ ⋀⃗ (sphereSmashMap n m , refl)) ∘ + ⋀-associate ∘ + ((smashS1→susp.fun⁻ , refl) ⋀⃗ idfun∙ _) -- private -- f* : {a : A} → S¹ → A → Susp A @@ -209,7 +196,7 @@ sphereSmashMap n m = {!!} -- -- ; (j = i1) → proj (loop i) a}) -- -- (compLrFiller (push (inr a)) (λ i → proj (loop i) a) (sym (push (inr a))) j i) - + -- -- {- hcomp (λ k → λ { (i = i0) → ((λ j → id3 j ∙ sym (projₗ a) ∙ λ j → proj (loop (~ j)) a) ∙ -- -- invCancelHelper _ _) k -- j -- -- ; (i = i1) → lUnit (rUnit (sym (projₗ a)) (~ k)) (~ k) -- j @@ -233,7 +220,7 @@ sphereSmashMap n m = {!!} -- -- filler : funInv (f* {a = pt} (loop i) a) ≡ proj (loop i) a -- -- filler = (λ j → funInv (funExt (λ x → merid x ∙ sym (merid pt)) (i ∨ j) a)) ∙ sym (projₗ a) ∙ λ j → proj (loop (i ∨ (~ j))) a - + -- -- id1 : sym (projᵣ base) ∙ cong ((λ x → proj {A = S¹ , base} {B = A , pt} x pt)) loop ∙ projᵣ base ≡ refl -- -- id1 = (λ i → (push (inl (loop i))) ∙ (λ j → inr (loop (i ∨ j) , pt)) ∙ sym (push (inl base))) ∙ -- -- (λ i → push (inl base) ∙ lUnit (sym (push (inl base))) (~ i)) ∙ @@ -249,18 +236,18 @@ sphereSmashMap n m = {!!} -- -- (λ j i → (cong funInv (merid a) ∙ id2 j) i) ∙ -- -- (λ j i → rUnit (cong funInv (merid a)) (~ j) i) - --- -- test5 : ((sym (projₗ a) ∙ cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ (projₗ a))) ≡ λ j → funInv (merid a j) + +-- -- test5 : ((sym (projₗ a) ∙ cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ (projₗ a))) ≡ λ j → funInv (merid a j) -- -- test5 = refl - + -- -- test6 : (ia ib : I) → (x : S¹) → retrHelper (loop ia) pt ≡ {!? ∙ sym (cong (projₗ) loop) ∙ ?!} -- -- test6 = {!!} -- -- where -- -- test7 : ((sym (projₗ pt) ∙ cong (λ x → proj {A = S¹ , base} {B = A , pt} x pt) loop ∙ (projₗ pt))) ≡ refl -- -- test7 = (λ i j → funInv (merid pt j)) ∙ λ i → {!!} -- -- {- --- -- test2 : (i : I) → (retrHelper (loop i) pt) ≡ (λ j → funInv (rCancel (merid pt) j i)) ∙ sym (projᵣ (loop i)) +-- -- test2 : (i : I) → (retrHelper (loop i) pt) ≡ (λ j → funInv (rCancel (merid pt) j i)) ∙ sym (projᵣ (loop i)) -- -- test2 i = (λ z → hfill ((λ k → λ { (i = i0) → ((λ j → id3 i pt j ∙ sym (projₗ pt) ∙ λ j → proj (loop (~ j)) pt) ∙ invCancelHelper i pt _ _) k -- -- ; (i = i1) → lUnit (rUnit (sym (projₗ pt)) (~ k)) (~ k) })) -- -- (inS ((λ j → funInv (funExt (λ x → merid x ∙ sym (merid pt)) (i ∨ j) pt)) ∙ sym (projₗ pt) ∙ λ j → proj (loop (i ∨ (~ j))) pt)) (~ z)) ∙ @@ -279,7 +266,7 @@ sphereSmashMap n m = {!!} -- -- idIsEquiv2 : ∀ {ℓ} {A B : Type ℓ} → A ≡ B → A ≃ B -- -- idIsEquiv2 = J (λ B p → _ ≃ B) (idEquiv _) - + -- -- helper2 : (x : S¹ × A) → funInv (fun (inr x)) ≡ inr x -- -- helper2 (x , x₁) = retrHelper x x₁ -- -- funRetr2 : (x : (S¹ , base) smash (A , pt)) → funInv (fun x) ≡ x @@ -296,7 +283,7 @@ sphereSmashMap n m = {!!} -- -- test3 (push a i) q = {!!} -- -- test2 : (x : S¹ × A) (v : (S¹ , base) wedge (A , pt)) (q : Path ((S¹ , base) smash (A , pt)) (inr x) (inl tt)) → q (~ i) ≡ (q ∙ push v) (~ i) --- -- test2 x v q = (λ j → q ((~ i) ∧ (~ j))) ∙ λ j → (q ∙ push v) ((~ i) ∧ j) +-- -- test2 x v q = (λ j → q ((~ i) ∧ (~ j))) ∙ λ j → (q ∙ push v) ((~ i) ∧ j) -- -- funRetr : (x : (S¹ , base) smash (A , pt)) → funInv (fun x) ≡ x -- -- funRetr (inl tt) = refl @@ -319,20 +306,20 @@ sphereSmashMap n m = {!!} -- -- ; (i = i1) → _ }) ((λ j → funInv (rCancel (merid pt) (~ i ∨ j) i₁)) ∙ -- -- λ j → push (inl (loop i₁)) (i ∧ j)))) -} -- -- where - + -- -- bottom : funInv (rCancel (merid pt) (~ i) z) ≡ push (inl (loop z)) i -- -- bottom = (λ j → funInv(rCancel (merid pt) (~ i) (z ∨ j))) ∙ (λ j → push (inl (loop z)) (i ∧ j)) -- -- i=i1 : (λ j → funInv(rCancel (merid pt) i0 (z ∨ j))) ∙ (λ j → push (inl (loop z)) j) ≡ retrHelper (loop z) pt -- -- i=i1 = {!-- rCancel (merid pt) i1!} - + -- -- frontCubeFiller : (i j z : I) → (S¹ , base) smash (A , pt) -- bottom is i j , z is height -- -- frontCubeFiller i j z = hfill (λ k → λ { (i = i0) → inl tt -- -- ; (i = i1) → sym (projᵣₗ k) j -- -- ; (j = i0) → inl tt -- -- ; (j = i1) → projᵣ base (~ i)}) (inS (push (inl base) (i ∧ j))) z --- -- bottomCubeFiller : (i j z : I) → (S¹ , base) smash (A , pt) +-- -- bottomCubeFiller : (i j z : I) → (S¹ , base) smash (A , pt) -- -- bottomCubeFiller i j z = funInv (rCancel (merid pt) (~ i) (z ∧ j)) @@ -343,7 +330,7 @@ sphereSmashMap n m = {!!} -- -- funRetr (push (inr x) i) = λ j → (projₗ x ((~ i) ∨ (~ j))) --- -- funRetr (push (push tt i₁) i) = {! +-- -- funRetr (push (push tt i₁) i) = {! -- -- i₁ = i0 ⊢ hcomp -- -- (λ { k (i = i0) → λ _ → inl tt diff --git a/Cubical/HITs/Truncation/Properties.agda b/Cubical/HITs/Truncation/Properties.agda index 2312a10ca..3c2112a22 100644 --- a/Cubical/HITs/Truncation/Properties.agda +++ b/Cubical/HITs/Truncation/Properties.agda @@ -259,11 +259,7 @@ groupoidTrunc≡Trunc1 = ua groupoidTrunc≃Trunc1 abstract isOfHLevelHLevel2 : ∀ n → isOfHLevel (suc n) (HLevel ℓ n) - isOfHLevelHLevel2 0 = isPropHContr - isOfHLevelHLevel2 (suc n) x y = subst (isOfHLevel (suc n)) eq (isOfHLevel≡ (suc n) (snd x) (snd y)) - where eq : ∀ {A B : Type ℓ} {hA : isOfHLevel (suc n) A} {hB : isOfHLevel (suc n) B} - → (A ≡ B) ≡ ((A , hA) ≡ (B , hB)) - eq = ua (_ , ΣProp≡-equiv (λ _ → isPropIsOfHLevel (suc n))) + isOfHLevelHLevel2 n = isOfHLevelHLevel n {- Proofs of Theorem 7.3.12. and Corollary 7.3.13. in the HoTT book -} @@ -274,19 +270,11 @@ module ΩTrunc where (λ a b → ∥ a ≡ b ∥ n , isOfHLevelTrunc (2+ n)) x y .fst {- We will need P to be of hLevel n + 3 -} - hLevelP : {n : ℕ₋₂} (a b : ∥ B ∥ (suc₋₂ n)) → isOfHLevel (2+ (suc₋₂ n)) (P a b) - hLevelP {n = n} = - elim2 (λ x y → isProp→isOfHLevelSuc (2+ n) (isPropIsOfHLevel (2+ suc₋₂ n))) - (λ a b → isOfHLevelSuc (2+ n) (isOfHLevelTrunc (2+ n))) - - - decode* : ∀ {n : ℕ₋₂} (u v : B) - → P {n = n} ∣ u ∣ ∣ v ∣ → Path (∥ B ∥ (suc₋₂ n)) ∣ u ∣ ∣ v ∣ - decode* {B = B} {n = neg2} u v = - rec ( isOfHLevelTrunc 1 ∣ u ∣ ∣ v ∣ - , λ _ → isOfHLevelSuc 1 (isOfHLevelTrunc 1) _ _ _ _) (λ p i → ∣ p i ∣) - decode* {n = ℕ₋₂.-1+ n} u v = - rec (isOfHLevelTrunc (suc (suc n)) ∣ u ∣ ∣ v ∣) (λ p i → ∣ p i ∣) + abstract + hLevelP : {n : ℕ₋₂} (a b : ∥ B ∥ (suc₋₂ n)) → isOfHLevel (2+ (suc₋₂ n)) (P a b) + hLevelP {n = n} = + elim2 (λ x y → isProp→isOfHLevelSuc (2+ n) (isPropIsOfHLevel (2+ suc₋₂ n))) + (λ a b → isOfHLevelSuc (2+ n) (isOfHLevelTrunc (2+ n))) {- decode function from P x y to x ≡ y -} decode-fun : {n : ℕ₋₂} (x y : ∥ B ∥ (suc₋₂ n)) → P x y → x ≡ y @@ -295,7 +283,13 @@ module ΩTrunc where (λ _ → isOfHLevelSuc (2+ suc₋₂ n) (isOfHLevelTrunc (2+ suc₋₂ n)) u v)) decode* where - + decode* : ∀ {n : ℕ₋₂} (u v : B) + → P {n = n} ∣ u ∣ ∣ v ∣ → Path (∥ B ∥ (suc₋₂ n)) ∣ u ∣ ∣ v ∣ + decode* {B = B} {n = neg2} u v = + rec ( isOfHLevelTrunc 1 ∣ u ∣ ∣ v ∣ + , λ _ → isOfHLevelSuc 1 (isOfHLevelTrunc 1) _ _ _ _) (cong ∣_∣) + decode* {n = ℕ₋₂.-1+ n} u v = + rec (isOfHLevelTrunc (suc (suc n)) ∣ u ∣ ∣ v ∣) (cong ∣_∣) {- auxiliary function r used to define encode -} r : {m : ℕ₋₂} (u : ∥ B ∥ (suc₋₂ m)) → P u u diff --git a/Cubical/HITs/Wedge.agda b/Cubical/HITs/Wedge.agda new file mode 100644 index 000000000..77174fe87 --- /dev/null +++ b/Cubical/HITs/Wedge.agda @@ -0,0 +1,4 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.HITs.Wedge where + +open import Cubical.HITs.Wedge.Base public diff --git a/Cubical/HITs/Wedge/Base.agda b/Cubical/HITs/Wedge/Base.agda new file mode 100644 index 000000000..e60f94450 --- /dev/null +++ b/Cubical/HITs/Wedge/Base.agda @@ -0,0 +1,24 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.HITs.Wedge.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.HITs.Pushout.Base +open import Cubical.Data.Unit +open import Cubical.Data.Prod +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws + +_⋁_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Type (ℓ-max ℓ ℓ') +_⋁_ (A , ptA) (B , ptB) = Pushout {A = Unit} {B = A} {C = B} (λ {tt → ptA}) (λ {tt → ptB}) + + +-- pointed versions + +_⋁∙ₗ_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Pointed (ℓ-max ℓ ℓ') +A ⋁∙ₗ B = (A ⋁ B) , (inl (snd A)) + +_⋁∙ᵣ_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Pointed (ℓ-max ℓ ℓ') +A ⋁∙ᵣ B = (A ⋁ B) , (inr (snd B)) + diff --git a/Cubical/Homotopy/Connected.agda b/Cubical/Homotopy/Connected.agda index 4674f55c9..e76f5e946 100644 --- a/Cubical/Homotopy/Connected.agda +++ b/Cubical/Homotopy/Connected.agda @@ -12,9 +12,11 @@ open import Cubical.Foundations.Path open import Cubical.Foundations.Univalence open import Cubical.Functions.Fibration open import Cubical.Data.Nat -open import Cubical.Data.Sigma +open import Cubical.Data.Prod hiding (map) +open import Cubical.Data.Sigma hiding (_×_) open import Cubical.HITs.Nullification open import Cubical.HITs.Susp +open import Cubical.HITs.SmashProduct open import Cubical.HITs.Truncation as Trunc open import Cubical.Homotopy.Loopspace open import Cubical.HITs.Pushout @@ -29,7 +31,8 @@ isHLevelConnected n A = isContr (hLevelTrunc n A) isHLevelConnectedFun : ∀ {ℓ ℓ'} (n : ℕ) {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Type (ℓ-max ℓ ℓ') isHLevelConnectedFun n f = ∀ b → isHLevelConnected n (fiber f b) - +isHLevelTruncatedFun : ∀ {ℓ ℓ'} (n : ℕ) {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Type (ℓ-max ℓ ℓ') +isHLevelTruncatedFun n f = ∀ b → isOfHLevel n (fiber f b) isConnectedSubtr : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n m : ℕ) (f : A → B) → isHLevelConnectedFun (n + m) f → isHLevelConnectedFun n f isConnectedSubtr n m f iscon b = transport (λ i → isContr (ua (truncOfTruncEq {A = fiber f b} n m) (~ i) )) @@ -50,11 +53,30 @@ private Iso.leftInv typeToFiberIso a = refl - - - module elim {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} (f : A → B) (n : ℕ) where + isIsoPrecompose : ∀ {ℓ'''} (P : B → HLevel ℓ''' n) + → isHLevelConnectedFun n f + → Iso ((b : B) → P b .fst) ((a : A) → P (f a) .fst) + isIsoPrecompose P fConn = + (iso (_∘ f) + (λ t b → inv t b (fConn b .fst)) + (λ t → funExt λ a → + cong (inv t (f a)) (fConn (f a) .snd ∣ a , refl ∣) + ∙ substRefl {B = fst ∘ P} (t a)) + (λ s → funExt λ b → + Trunc.elim + {B = λ d → inv (s ∘ f) b d ≡ s b} + (λ _ → isOfHLevelPath n (P b .snd) _ _) + (λ {(a , p) i → transp (λ j → P (p (j ∨ i)) .fst) i (s (p i))}) + (fConn b .fst))) + where + inv : ((a : A) → P (f a) .fst) → (b : B) → hLevelTrunc n (fiber f b) → P b .fst + inv t b = + Trunc.rec + (P b .snd) + (λ {(a , p) → subst (fst ∘ P) p (t a)}) + isEquivPrecompose : ∀ {ℓ'''} (P : B → HLevel ℓ''' n) → isHLevelConnectedFun n f → isEquiv λ(s : (b : B) → P b .fst) → s ∘ f @@ -71,7 +93,7 @@ module elim {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} (f : A → B) (n : (λ _ → isOfHLevelPath n (P b .snd) _ _) (λ {(a , p) i → transp (λ j → P (p (j ∨ i)) .fst) i (s (p i))}) (fConn b .fst))) - + where inv : ((a : A) → P (f a) .fst) → (b : B) → hLevelTrunc n (fiber f b) → P b .fst inv t b = @@ -99,7 +121,7 @@ module elim {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} (f : A → B) (n : fun : (n : ℕ) → (s : (∀ {ℓ} (P : B → HLevel ℓ n) → hasSection λ(s : (b : B) → P b .fst) → s ∘ f)) → (b : B) (w : (hLevelTrunc n (fiber f b))) → w ≡ c n s b fun zero s b w = isOfHLevelSuc zero (isOfHLevelTrunc _) w (c zero s b) - fun (suc n) s b = Trunc.elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc _) x (c (suc n) s b)) λ a → transport eqtyp c* b (fst a) (snd a) + fun (suc n) s b = Trunc.elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc _) x (c (suc n) s b)) λ a → transport eqtyp c* b (fst a) (snd a) where eqtyp : ((a : A) → ∣ (a , refl {x = f a}) ∣ ≡ c (suc n) s (f a)) @@ -221,7 +243,6 @@ connectedTruncIso {A = A} {B = B} (suc n) f con = g (map f (map fst x)) ∣ b ∣) (λ a p → cong ∣_∣ p) (fst (con b)) - where helper : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : A → Type ℓ'} (P : hLevelTrunc (suc n) A → Type ℓ'') → ((x : hLevelTrunc (suc n) (Σ A B)) → isOfHLevel (suc n) (P (map fst x))) @@ -237,9 +258,17 @@ connectedTruncIso {A = A} {B = B} (suc n) f con = g g : Iso (hLevelTrunc (suc n) A) (hLevelTrunc (suc n) B) Iso.fun g = map f Iso.inv g = Trunc.rec (isOfHLevelTrunc _) back - Iso.leftInv g = Trunc.elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc _) _ _) - λ a → cong (map fst) (con (f a) .snd ∣ a , refl ∣) - Iso.rightInv g = Trunc.elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc _) _ _) + Iso.leftInv g = helper + where + abstract + helper : _ + helper = Trunc.elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc _) _ _) + λ a → cong (map fst) (con (f a) .snd ∣ a , refl ∣) + Iso.rightInv g = helper + where + abstract + helper : _ + helper = Trunc.elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc _) _ _) backSection connectedTruncIso2 : ∀ {ℓ} {A B : Type ℓ} (n m : ℕ) (f : A → B) @@ -258,7 +287,8 @@ connectedTruncEquiv {A = A} {B = B} n f con = isoToEquiv (connectedTruncIso n f connectedTruncEquiv2 : ∀ {ℓ} {A B : Type ℓ} (n m : ℕ) (f : A → B) → Σ[ x ∈ ℕ ] n + x ≡ m → isHLevelConnectedFun m f → hLevelTrunc n A ≃ hLevelTrunc n B -connectedTruncEquiv2 {A = A} {B = B} n m f (x , pf) con = connectedTruncEquiv n f (isConnectedSubtr n x f (transport (λ i → isHLevelConnectedFun (pf (~ i)) f) con)) +connectedTruncEquiv2 {A = A} {B = B} n m f (x , pf) con = + connectedTruncEquiv n f (isConnectedSubtr n x f (transport (λ i → isHLevelConnectedFun (pf (~ i)) f) con)) ------------ @@ -268,9 +298,9 @@ inrConnected : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ → (f : C → A) (g : C → B) → isHLevelConnectedFun n f → isHLevelConnectedFun n {A = B} {B = Pushout f g} inr -inrConnected {A = A} {B = B} {C = C} n f g iscon = elim.isConnectedPrecompose inr n λ P → (λ k → helpLemmas.k P k) , λ b → refl +inrConnected {A = A} {B = B} {C = C} n f g iscon = elim.isConnectedPrecompose inr n λ P → (λ l → k P l) , λ b → refl where - module helpLemmas {ℓ : Level} (P : (Pushout f g) → HLevel ℓ n) + module _ {ℓ : Level} (P : (Pushout f g) → HLevel ℓ n) (h : (b : B) → typ (P (inr b))) where Q : A → HLevel _ n @@ -282,24 +312,25 @@ inrConnected {A = A} {B = B} {C = C} n f g iscon = elim.isConnectedPrecompose in k : (d : Pushout f g) → typ (P d) k (inl x) = elim.hasSectionPrecompose f n Q (elim.isEquivPrecompose f n Q iscon) .fst fun x k (inr x) = h x - k (push a i) = hcomp (λ k → λ { (i = i0) → (elim.hasSectionPrecompose f n Q (elim.isEquivPrecompose f n Q iscon) .snd fun i a) - ; (i = i1) → transportTransport⁻ (λ j → typ (P (push a j))) (h (g a)) k }) - (transp (λ j → typ (P (push a (i ∧ j)))) - (~ i) - (elim.hasSectionPrecompose f n Q (elim.isEquivPrecompose f n Q iscon) .snd fun i a )) + k (push a i) = + hcomp (λ k → λ { (i = i0) → (elim.hasSectionPrecompose f n Q (elim.isEquivPrecompose f n Q iscon) .snd fun i0 a) + ; (i = i1) → transportTransport⁻ (λ j → typ (P (push a j))) (h (g a)) k }) + (transp (λ j → typ (P (push a (i ∧ j)))) + (~ i) + (elim.hasSectionPrecompose f n Q (elim.isEquivPrecompose f n Q iscon) .snd fun i a )) sphereConnected : (n : ℕ) → isHLevelConnected (suc n) (S₊ n) sphereConnected zero = ∣ north ∣ , isOfHLevelTrunc 1 ∣ north ∣ sphereConnected (suc n) = - transport (λ i → isHLevelConnected (2 + n) ((isoToPath (SuspPushIso {A = S₊ n})) (~ i))) - (isHLevelConnectedPoint2 (suc n) - {A = Pushout {A = S₊ n} (λ _ → tt) λ _ → tt} - (inr tt) - ((transport (λ i → isHLevelConnectedFun (suc n) (mapsAgree (~ i))) - (inrConnected _ (λ _ → tt) (λ _ → tt) - λ {tt → transport (λ i → isContr (hLevelTrunc (suc n) ((isoToPath fibIso) (~ i)))) - (sphereConnected n)})))) + transport (λ i → isHLevelConnected (2 + n) (PushoutSusp≡Susp {A = S₊ n} i)) + (isHLevelConnectedPoint2 (suc n) + {A = Pushout {A = S₊ n} (λ _ → tt) λ _ → tt} + (inr tt) + ((transport (λ i → isHLevelConnectedFun (suc n) (mapsAgree (~ i))) + (inrConnected _ (λ _ → tt) (λ _ → tt) + λ {tt → transport (λ i → isContr (hLevelTrunc (suc n) ((isoToPath fibIso) (~ i)))) + (sphereConnected n)})))) where mapsAgree : Path ((x : Unit) → Pushout {A = S₊ n} (λ x → tt) (λ x → tt)) @@ -307,24 +338,377 @@ sphereConnected (suc n) = inr mapsAgree = funExt λ {tt → refl} - fibIso : Iso (fiber (λ (x : S₊ n) → tt) tt) (S₊ n) Iso.fun (fibIso) = fst Iso.inv (fibIso) a = a , refl Iso.leftInv (fibIso) b = refl Iso.rightInv (fibIso) b = refl +-- private +-- variable +-- ℓ ℓ' ℓ'' : Level +-- A : Type ℓ +-- B : Type ℓ' +-- A' : Pointed ℓ +-- B' : Pointed ℓ' +-- C' : Pointed ℓ'' + +-- {- +-- {- Lemma 8.6.1. in the HoTT book -} + +-- private +-- module helpers (n k : ℕ) (f : A → B) (P : B → HLevel ℓ'' (k + n)) (l : (a : A) → fst (P (f a))) (iscon : isHLevelConnectedFun n f) where +-- F : _ +-- F = λ(s : (b : B) → P b .fst) → s ∘ f + +-- fiberPath1 : Iso (fiber F l) (Σ[ g ∈ ((b : B) → P b .fst) ] ((a : A) → g (f a) ≡ l a)) +-- Iso.fun fiberPath1 p = (p .fst) , (funExt⁻ (p .snd)) +-- Iso.inv fiberPath1 p = (p .fst) , (funExt (p .snd)) +-- Iso.rightInv fiberPath1 p = refl +-- Iso.leftInv fiberPath1 p = refl + +-- throwAboutHelper : ∀{ℓ} {A : Type ℓ} {a b c : A} → (r : b ≡ c) (p : a ≡ b) (q : a ≡ c) → (p ≡ q ∙ (sym r) ) ≡ (q ≡ p ∙ r) +-- throwAboutHelper {a = a} {b = b} r p q = (λ i → rUnit p i ≡ q ∙ sym r) ∙ +-- (λ i → p ∙ (λ j → r (j ∧ i)) ≡ q ∙ λ j → r (i ∨ (~ j))) ∙ +-- (λ i → p ∙ r ≡ rUnit q (~ i)) ∙ symId _ _ +-- where +-- symId : ∀{ℓ} {A : Type ℓ} (a b : A) → (a ≡ b) ≡ (b ≡ a) +-- symId a b = isoToPath symiso +-- where +-- symiso : Iso (a ≡ b) (b ≡ a) +-- Iso.fun symiso = sym +-- Iso.inv symiso = sym +-- Iso.rightInv symiso a = refl +-- Iso.leftInv symiso a = refl + +-- throwAbout : (g h : (Σ[ g ∈ ((b : B) → P b .fst) ] ((a : A) → g (f a) ≡ l a))) +-- → (r : (x : B) → (g .fst) x ≡ (h .fst) x) +-- → Iso (g . snd ≡ (λ x → r (f x) ∙ h .snd x)) +-- ((x : A) → r (f x) ≡ (g .snd x) ∙ sym (h .snd x)) +-- Iso.fun (throwAbout (g , p) (h , q) r) id x = +-- transport (λ i → throwAboutHelper (sym (q x)) (p x) (r (f x)) i) +-- (funExt⁻ id x) +-- Iso.inv (throwAbout (g , p) (h , q) r) id = +-- funExt λ x → transport (λ i → throwAboutHelper (sym (q x)) (p x) (r (f x)) (~ i)) (id x) +-- Iso.rightInv (throwAbout (g , p) (h , q) r) id = +-- funExt λ x → transportTransport⁻ (throwAboutHelper (sym (q x)) (p x) (r (f x))) (id x) +-- Iso.leftInv (throwAbout (g , p) (h , q) r) id j = +-- funExt (λ x → (transport⁻Transport (throwAboutHelper (sym (q x)) (p x) (r (f x))) (λ i → id i x)) j) + +-- helper3 : (g h : (Σ[ g ∈ ((b : B) → P b .fst) ] ((a : A) → g (f a) ≡ l a))) +-- → (r : (x : B) → (g .fst) x ≡ (h .fst) x) → +-- (PathP (λ i → (x : A) → r (f x) i ≡ l x) (g .snd) (h .snd)) +-- ≡ ((x : A) → r (f x) ≡ ((g .snd) x) ∙ sym ((h .snd) x)) +-- helper3 (g , p) (h , q) r j = +-- hcomp (λ k → λ{ (j = i0) → PathP (λ i → (x : A) → r (f x) i ≡ l x) p λ x → lUnit (q x) (~ k) +-- ; (j = i1) → isoToPath (throwAbout (g , p) (h , q) r) k}) +-- (PathP (λ i → (x : A) → r (f x) (i ∧ ~ j) ≡ l x) +-- p +-- λ x → (λ i → r (f x) (~ j ∨ i)) ∙ (q x)) + +-- iso2 : (g h : (Σ[ g ∈ ((b : B) → P b .fst) ] ((a : A) → g (f a) ≡ l a))) +-- → Iso (Σ[ p ∈ g .fst ≡ h .fst ] PathP (λ i → (a : A) → p i (f a) ≡ l a) (g .snd) (h .snd)) +-- (Σ[ r ∈ ((b : B) → g .fst b ≡ h .fst b) ] r ∘ f ≡ λ a → (g .snd a) ∙ (h .snd a) ⁻¹) +-- Iso.fun (iso2 (g , p) (h , q)) (funid , pathp) = funExt⁻ funid , funExt (transport (helper3 (g , p) (h , q) λ x i → funid i x) pathp) + +-- Iso.inv (iso2 (g , p) (h , q)) (r , id) = (funExt r) , transport (sym (helper3 (g , p) (h , q) r)) λ x i → id i x +-- Iso.rightInv (iso2 (g , p) (h , q)) a i = (fst a) , (funExt (transportTransport⁻ (helper3 (g , p) (h , q) (a .fst)) (funExt⁻ (snd a)) i)) +-- Iso.leftInv (iso2 (g , p) (h , q)) a i = fst a , transport⁻Transport (helper3 (g , p) (h , q) (λ x i → a .fst i x )) (a .snd) i + + +-- isConnectedFamily : (n k : ℕ) (f : A → B) (P : B → HLevel ℓ'' (k + n)) +-- → isHLevelConnectedFun n f +-- → isHLevelTruncatedFun k (λ(s : (b : B) → P b .fst) → s ∘ f) +-- isConnectedFamily {A = A} {B = B} n zero f P iscon = equiv-proof (elim.isEquivPrecompose f n P iscon) +-- isConnectedFamily {A = A} {B = B} n (suc zero) f P iscon l = +-- transport (λ i → isOfHLevel 1 (isoToPath (helpers.fiberPath1 n 1 f P l iscon) (~ i))) +-- λ a b → transport (λ i → (sym (ua Σ≡) ∙ (isoToPath (helpers.iso2 n 1 f P l iscon a b))) (~ i)) +-- (isConnectedFamily n 0 f +-- (λ x → (a .fst x ≡ b .fst x) , isOfHLevelPath' n (P x .snd) (a .fst x) (b .fst x)) +-- iscon +-- (λ r → (a .snd r) ∙ (b .snd r) ⁻¹) .fst ) +-- isConnectedFamily {A = A} {B = B} n (suc (suc k)) f P iscon l = +-- transport (λ i → isOfHLevel (2 + k) (isoToPath (helpers.fiberPath1 n (2 + k) f P l iscon) (~ i))) +-- λ a b → transport (λ i → isOfHLevel (suc k) ((sym (ua Σ≡) ∙ (isoToPath (helpers.iso2 n (2 + k) f P l iscon a b))) (~ i))) +-- (isConnectedFamily n (suc k) f +-- (λ x → (a .fst x ≡ b .fst x) , ((P x) .snd (a .fst x) (b .fst x))) +-- iscon +-- λ r → (a .snd r) ∙ (b .snd r) ⁻¹) + +-- -} + + +-- helper : (n : ℕ) {A B : Pointed₀} → hLevelTrunc n (fst A) → hLevelTrunc n (fst B) → hLevelTrunc n (A ⋀ B) +-- helper n = Trunc.rec (isOfHLevelΠ _ (λ _ → isOfHLevelTrunc _)) λ a → Trunc.rec (isOfHLevelTrunc _) λ b → ∣ inr (a , b) ∣ + +-- truncTest : (n : ℕ) {A B : Pointed₀} → (hLevelTrunc n (fst A) , ∣ pt A ∣) ⋀ (hLevelTrunc n (fst B) , ∣ pt B ∣) +-- → hLevelTrunc n (A ⋀ B) +-- truncTest n (inl x) = ∣ inl tt ∣ +-- truncTest n {A = A} {B = B} (inr (x , x₁)) = helper n x x₁ +-- truncTest n {A = A} {B = B} (push (inl x) i) = helper2 x (~ i) +-- where +-- helper2 : (x : hLevelTrunc n (fst A)) → helper n {A = A} {B = B} x ∣ pt B ∣ ≡ ∣ inl tt ∣ +-- helper2 = Trunc.elim (λ x → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a i → ∣ push (inl a) (~ i) ∣ +-- truncTest n {A = A} {B = B} (push (inr x) i) = helper2 x (~ i) +-- where +-- helper2 : (x : hLevelTrunc n (fst B)) → helper n {A = A} {B = B} ∣ pt A ∣ x ≡ ∣ inl tt ∣ +-- helper2 = Trunc.elim (λ x → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a i → ∣ push (inr a) (~ i) ∣ +-- truncTest n {A = A} {B = B} (push (push tt i₁) i) = ∣ push (push tt i₁) i ∣ + +-- hLevelSmash : {A B : Pointed₀} (n : ℕ) → isOfHLevel n ((hLevelTrunc n (fst A) , ∣ pt A ∣) ⋀ (hLevelTrunc n (fst B) , ∣ pt B ∣)) +-- hLevelSmash {A = A} {B = B} zero = (inl tt) , {!!} +-- where +-- helper2 : (y : (hLevelTrunc zero (fst A) , ∣ pt A ∣) ⋀ (hLevelTrunc zero (fst B) , ∣ pt B ∣)) → inl tt ≡ y +-- helper2 (inl tt) = refl +-- helper2 (inr (x , x₁)) = push (inr x₁) ∙ (λ i → inr (isOfHLevelTrunc 0 .snd ∣ snd A ∣ (~ i) , x₁)) ∙ λ i → inr (((isOfHLevelTrunc 0 .snd) x i) , x₁) +-- helper2 (push (inl x) i) = {!!} +-- helper2 (push (inr x) i) = {!!} +-- helper2 (push (push a i₁) i) = {!!} +-- hLevelSmash {A = A} {B = B} (suc n) = {!!} + +-- truncTestInv : {!!} +-- truncTestInv = {!!} + +-- isConnectedSmashIdfun2 : {A B C : Pointed₀} +-- → (f : A →∙ B) (nf nc : ℕ) +-- → isHLevelConnectedFun nf (fst f) +-- → isHLevelConnected (2 + nc) (fst C) +-- → isHLevelConnectedFun (1 + nc + nf) (Smash-map f (idfun∙ C)) +-- isConnectedSmashIdfun2 {A = A , ptA} {B = B , ptB} {C = C , ptC} (f , fpt) nf nc conf conC = λ b → {!!} , {!!} +-- -- elim.isConnectedPrecompose (Smash-map (f , fpt) (idfun∙ (C , ptC))) _ λ {ℓ} P → proof.h {ℓ} P , λ b → funExt (proof.hsection {ℓ} P b) -- elim.isConnectedPrecompose (Smash-map (f , fpt) (idfun∙ (C , ptC))) _ λ {ℓ} P → (λ h → proof.h {ℓ} P h) ,push (inl x)) λ b → funExt λ x → {!!} +-- where +-- module proof {ℓ : Level} (P : (Smash (B , ptB) (C , ptC)) → HLevel ℓ (1 + nc + nf)) +-- (d : (x : Smash (A , ptA) (C , ptC)) → P ((Smash-map (f , fpt) (idfun∙ (C , ptC)) ) x) .fst) +-- where +-- test : (c : C) (b : B) → (hLevelTrunc (suc (nc + nf)) (fiber (Smash-map (f , fpt) (idfun∙ (C , ptC))) (proj b c))) +-- ≡ ((hLevelTrunc _ (fiber f b) , conf b .fst) ⋀ {!!}) +-- test = {!!} +-- -- F : (c : C) → _ +-- -- F c = λ(s : (b : B) → P (proj b c) .fst) → s ∘ f + +-- -- step1 : (c : C) → isHLevelTruncatedFun (1 + nc) (F c) +-- -- step1 c = isOfHLevelPrecomposeConnected (1 + nc) nf ((λ b → P (proj b c) .fst , P (proj b c) .snd)) f conf + +-- -- codomFun : (c : C) (a : A) → P (proj (f a) c) .fst +-- -- codomFun c = d ∘ λ a → proj a c + +-- -- Q : C → HLevel _ (1 + nc) +-- -- Q c = fiber (F c) (codomFun c) , step1 c _ + + +-- -- helper : (a : A) → transport (λ i → P (gluel (f a) (~ i)) .fst) (d (basel)) +-- -- ≡ d (proj a ptC) +-- -- helper a = transport (PathP≡Path (λ i → P (gluel (f a) (~ i)) .fst) (d basel) (d (proj a ptC))) +-- -- (transport (λ j → PathP (λ i → P (lUnit (gluel (f a)) (~ j) (~ i)) .fst) (d basel) (d (proj a ptC))) +-- -- λ i → d (gluel a (~ i))) + +-- -- QptC : Q ptC .fst +-- -- QptC = (λ b → transport (λ i → P (gluel b (~ i)) .fst) (d basel)) , +-- -- funExt helper + + +-- -- Q-constructor : (c : C) → Q c .fst +-- -- Q-constructor c = Iso.inv (elim.isIsoPrecompose (λ (_ : Unit) → ptC) (1 + nc) Q (isHLevelConnectedPoint _ conC ptC)) (λ _ → QptC) c + +-- -- g : (b : B) (c : C) → P (proj b c) .fst +-- -- g b c = Q-constructor c .fst b + +-- -- Q-constructor-β : (b : B) → Q-constructor ptC .fst b ≡ transport (λ i → P (gluel b (~ i)) .fst) (d basel) +-- -- Q-constructor-β b = ((λ i → (Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) (helper3 i)) .fst b)) ∙ +-- -- (λ i → transportRefl (QptC) i .fst b) + +-- -- where +-- -- helper3 : (isHLevelConnectedPoint _ conC ptC) ptC .fst ≡ ∣ tt , refl ∣ +-- -- helper3 = (isHLevelConnectedPoint _ conC ptC) ptC .snd ∣ tt , refl ∣ + +-- -- moveTransport : ∀ {ℓ ℓ'} → {A : Type ℓ} {B : A → Type ℓ'} {x y : A} (p : x ≡ y) (x1 : B x) (y1 : B y) +-- -- → transport (λ i → (B (p i))) x1 ≡ y1 +-- -- → transport⁻ (λ i → (B (p i))) y1 ≡ x1 +-- -- moveTransport {B = B} p x1 y1 transp = ((λ i → transport⁻ (λ i → B (p i)) (transp (~ i))) ∙ (transport⁻Transport (λ i → B (p i)) x1)) + +-- -- otherway : (b : B) → transport (λ i → P (gluel b i) .fst) (g b ptC) ≡ d basel +-- -- otherway b = (λ i → transport (λ i → P (gluel b i) .fst) (Q-constructor-β b i)) ∙ transportTransport⁻ (λ i → P (gluel b i) .fst) (d basel) + +-- -- path1 : (b : B) → PathP (λ i → P (gluel b i) .fst) (g b ptC) (d basel) +-- -- path1 b i = hcomp (λ k → λ {(i = i0) → g b ptC +-- -- ; (i = i1) → otherway b k}) +-- -- (transp (λ j → P (gluel b (i ∧ j)) .fst) (~ i) (g b ptC)) + + +-- -- where +-- -- massive : (c : C) → PathP _ (g ptB c) (d baser) +-- -- massive c = compPathP (λ i → g (fpt (~ i)) c) (compPathP (λ i → Q-constructor c .snd i (ptA)) (λ i → d (gluer c i))) + +-- -- pathPCancel : (c : C) +-- -- → ((λ z → P (proj (fpt (~ z)) c) .fst) ∙ (λ i → ((λ _ → P (proj (f ptA) c) .fst) ∙ (λ i₁ → P (((λ j₁ → proj (fpt j₁) c) ∙ gluer c) i₁) .fst)) i)) +-- -- ≡ (λ i → P (gluer c i) .fst) +-- -- pathPCancel c = (λ i → (λ z → P (proj (fpt (~ z)) c) .fst) ∙ lUnit ((λ i₁ → P (((λ j₁ → proj (fpt j₁) c) ∙ gluer c) i₁) .fst)) (~ i)) +-- -- ∙ (λ i → (λ z → P (proj (fpt (~ z ∨ i)) c) .fst) ∙ λ i₁ → P (((λ j₁ → proj (fpt (j₁ ∨ i)) c) ∙ gluer c) i₁) .fst ) +-- -- ∙ (λ i → lUnit (λ i₁ → P (lUnit (gluer c) (~ i) i₁) .fst) (~ i)) + +-- -- path2 : (c : C) → PathP (λ i → P (gluer c i) .fst) (g ptB c) (d baser) +-- -- path2 c = transport (λ i → PathP (λ j → pathPCancel c i j) (g ptB c) (d baser)) (massive c) + +-- -- h : (x : Smash (B , ptB) (C , ptC)) → P x .fst +-- -- h basel = d basel +-- -- h baser = d baser +-- -- h (proj x y) = g x y +-- -- h (gluel b i) = path1 b i +-- -- h (gluer c i) = path2 c i + +-- -- hsection : (x : Smash (A , ptA) (C , ptC)) → h (Smash-map (f , fpt) (idfun∙ (C , ptC)) x) ≡ d x +-- -- hsection basel = refl +-- -- hsection baser = refl +-- -- hsection (proj x y) i = Q-constructor y .snd i x +-- -- hsection (gluel a i) j = hcomp (λ k → λ { (i = i0) → transportRefl {!!} k +-- -- ; (i = i1) → transportRefl (d basel) k +-- -- ; (j = i0) → {!!} +-- -- ; (j = i1) → {!!}}) +-- -- {!hcomp!} -- hcomp {!!} {!!} +-- -- where + +-- -- helper2 : Path ((P (proj (f a) ptC) .fst) ≡ (P basel .fst)) +-- -- (λ i → P (Smash-map (f , fpt) (idfun∙ (C , ptC)) (gluel a i)) .fst) +-- -- λ i → P (gluel (f a) i) .fst +-- -- helper2 = (λ j i → P ((lUnit (gluel (f a)) (~ j)) i) .fst) +-- -- helper6 : (a : A) → PathP ((λ j → PathP (λ i → P {!Q-constructor ? .snd i ?!} .fst) {!P (proj (f a) ptC) .fst!} {!!})) (λ i → h (Smash-map (f , fpt) (idfun∙ (C , ptC)) (gluel a i))) (path1 (f a)) +-- -- helper6 a = ((λ j i → h ((lUnit (gluel (f a)) (~ j)) i))) + +-- -- helper3 : PathP {!!} (path1 (f a)) (λ i → d (gluel a i)) +-- -- helper3 = compPathP {!!} {!!} +-- -- helper4 : PathP {!!} (path1 (f a)) λ i → d (gluel a i) +-- -- helper4 = {!!} +-- -- hsection (gluer b i) = {!!} + +-- -- -- isConnectedSmashIdfun : (f : A' →∙ B') (nf nc : ℕ) +-- -- -- → isHLevelConnectedFun nf (fst f) +-- -- -- → isHLevelConnected (2 + nc) (fst C') +-- -- -- → isHLevelConnectedFun (1 + nf + nc) (f ⋀⃗ idfun∙ C') +-- -- -- isConnectedSmashIdfun {A' = (A , ptA)} {B' = (B , ptB)} {C' = (C , ptC)} (f , fpt) nf nc conf conC = {!isHLel!} +-- -- -- where +-- -- -- module _ (P : ((B , ptB) ⋀ (C , ptC)) → HLevel (ℓ-max (ℓ-max ℓ ℓ') ℓ'') (1 + nc + nf)) +-- -- -- (d : (x : (A , ptA) ⋀ (C , ptC)) → P (((f , fpt) ⋀⃗ idfun∙ (C , ptC) ) x) .fst) +-- -- -- where + +-- -- -- F : (c : C) → _ +-- -- -- F c = λ(s : (b : B) → P (inr (b , c)) .fst) → s ∘ f + +-- -- -- step1 : (c : C) → isHLevelTruncatedFun (1 + nc) (F c) +-- -- -- step1 c = isOfHLevelPrecomposeConnected (1 + nc) nf ((λ b → P (inr (b , c)) .fst , P (inr (b , c)) .snd)) f conf + +-- -- -- codomFun : (c : C) (a : A) → P (inr ((f a) , c)) .fst +-- -- -- codomFun c = d ∘ λ a → inr (a , c) + +-- -- -- Q : C → HLevel _ (1 + nc) +-- -- -- Q c = fiber (F c) (codomFun c) , step1 c _ + + +-- -- -- helper : (a : A) → transport (λ i → P (push (inl (f a)) i) .fst) (d (inl tt)) +-- -- -- ≡ d (inr (a , ptC)) +-- -- -- helper a = transport (PathP≡Path (λ i → P (push (inl (f a)) i) .fst) (d (inl tt)) (d (inr (a , ptC)))) +-- -- -- (transport (λ j → PathP (λ i → P (rUnit (push (inl (f a))) (~ j) i) .fst) (d (inl tt)) (d (inr (a , ptC)))) +-- -- -- λ i → d (push (inl a) i)) + + +-- -- -- QptC : Q ptC .fst +-- -- -- QptC = (λ b → transport (λ i → P (push (inl b) i) .fst) (d (inl tt))) , +-- -- -- funExt helper +-- -- -- where + + +-- -- -- Q-constructor : (c : C) → Q c .fst +-- -- -- Q-constructor c = Iso.inv (elim.isIsoPrecompose (λ _ → ptC) (1 + nc) Q (isHLevelConnectedPoint _ conC ptC)) (λ ( _ : Unit) → QptC) c + +-- -- -- g : (b : B) (c : C) → P (inr (b , c)) .fst +-- -- -- g b c = Q-constructor c .fst b + +-- -- -- Q-constructor-β : (b : B) → Q-constructor ptC .fst b ≡ transport (λ i → P (push (inl b) i) .fst) (d (inl tt)) +-- -- -- Q-constructor-β b = ((λ i → (Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) (helper3 i)) .fst b)) ∙ +-- -- -- (λ i → transportRefl (QptC) i .fst b) + +-- -- -- where +-- -- -- helper3 : (isHLevelConnectedPoint _ conC ptC) ptC .fst ≡ ∣ tt , refl ∣ +-- -- -- helper3 = (isHLevelConnectedPoint _ conC ptC) ptC .snd ∣ tt , refl ∣ + +-- -- -- test : (b : B) (c : C) → PathP (λ i → P (push {!!} {!(~ i)!}) .fst) (Q-constructor ptC .fst b) (Q-constructor c .fst ptB) +-- -- -- test b = {!!} + +-- -- -- Q-constructor-β2 : (c : C) → Q-constructor c .fst ptB ≡ transport (λ i → P (push (inr c) i) .fst) (d (inl tt)) +-- -- -- Q-constructor-β2 c = ((λ i → (Trunc.rec (Q c .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) ((isHLevelConnectedPoint _ conC ptC) c .fst)) .fst ptB)) ∙ +-- -- -- {!isHLevelConnectedPoint _ conC ptC c .fst!} + +-- -- -- gid1 : (a : A) (c : C) → g (f a) c ≡ d (inr (a , c)) +-- -- -- gid1 a c i = (Q-constructor c .snd) i a + +-- -- -- gid1' : {!!} +-- -- -- gid1' = {!!} + +-- -- -- gid1Path : (c : C) → PathP (λ i → P (push (inr c) i) .fst) (d (inl tt)) (g ptB c) +-- -- -- gid1Path c i = +-- -- -- hcomp (λ k → λ {(i = i0) → {!!} ; (i = i1) → {!!}}) +-- -- -- {!!} + +-- -- -- compPath : (c : C) → PathP _ (d (inl tt)) (g ptB c) +-- -- -- compPath c = compPathP (λ i → d (push (inr c) i)) (compPathP (sym (gid1 ptA c)) (λ i → g (fpt i) c)) + + +-- -- -- compPathTransport : (c : C) → ((λ z → P ((push (inr c) ∙ +-- -- -- (λ i → inr (fpt (~ i) , c))) z) .fst) ∙ +-- -- -- (λ i → ((λ i₁ → P (inr (f ptA , c)) .fst) ∙ (λ i₁ → P (inr (fpt i₁ , c)) .fst)) i)) +-- -- -- ≡ λ i → P (push (inr c) i) .fst +-- -- -- compPathTransport c = (λ k → ((λ z → P ((push (inr c) ∙ +-- -- -- (λ i → inr (fpt (~ i) , c))) z) .fst) ∙ +-- -- -- ((lUnit (λ i₁ → P (inr (fpt i₁ , c)) .fst) (~ k) )))) +-- -- -- ∙ (λ k → ((λ z → P ((push (inr c) ∙ +-- -- -- (λ i → inr (fpt (~ i ∨ k) , c))) z) .fst) ∙ +-- -- -- (λ i₁ → P (inr (fpt (i₁ ∨ k) , c)) .fst) )) +-- -- -- ∙ (λ k → ((λ z → P ((push (inr c) ∙ refl) z) .fst) ∙ refl)) +-- -- -- ∙ (λ k → rUnit ((λ z → P ((rUnit (push (inr c)) (~ k)) z) .fst)) (~ k)) + +-- -- -- compPathTransport-ptC : compPathTransport ptC ≡ {!!} +-- -- -- compPathTransport-ptC = {!!} + +-- -- -- compPathTrue : (c : C) → PathP (λ i → P (push (inr c) i) .fst) (d (inl tt)) (g ptB c) +-- -- -- compPathTrue c = transport (λ i → PathP (λ j → compPathTransport c i j) (d (inl tt)) (g ptB c)) (compPath c) + + +-- -- -- {- +-- -- -- Goal: P (push (inr x) i) .fst +-- -- -- ———— Boundary —————————————————————————————————————————————— +-- -- -- i = i0 ⊢ d (inl tt) +-- -- -- i = i1 ⊢ g ptB x +-- -- -- -} + +-- -- -- gid1transp : (c : C) → PathP {!!} (g ptB c) (transport (λ i → P (push (inr c) i) .fst) (d (inl tt))) +-- -- -- gid1transp c = compPathP (λ i → g (fpt (~ i)) c) (compPathP (gid1 ptA c) (compPathP (λ i → d (push (inr c) (~ i))) λ i → transp (λ j → P (push (inr c) (i ∧ j)) .fst) (~ i) (d (inl tt)))) + + +-- -- -- gid2 : (b : B) → g b ptC ≡ transport (λ i → P (push (inl b) i) .fst) (d (inl tt)) -- (λ i → P (push (inl b) i) .fst) (d (inl tt)) +-- -- -- gid2 b = Q-constructor-β b + + +-- -- -- gid1≡gid2 : (c : C) (a : A) → PathP {!gid1transp c !} (gid1 a ptC) (gid2 (f a)) +-- -- -- gid1≡gid2 c a = compPathP {!λ j → (Q-constructor ptC .snd) i a!} (compPathP {!!} {!!}) + + +-- -- -- gid2Path : (b : B) → PathP (λ i → P (push (inl b) i) .fst) (d (inl tt)) (g b ptC) +-- -- -- gid2Path b i = +-- -- -- hcomp (λ k → λ {(i = i0) → (d (inl tt)) ; (i = i1) → gid2 b (~ k)}) +-- -- -- (transp (λ j → P (push (inl b) (j ∧ i)) .fst) (~ i) (d (inl tt))) + +-- -- -- {- +-- -- -- Goal: P (push (inl b) i) .fst +-- -- -- ———— Boundary —————————————————————————————————————————————— +-- -- -- i = i0 ⊢ d (inl tt) +-- -- -- i = i1 ⊢ g b ptC +-- -- -- -} + +-- -- -- PathP2 : PathP {!!} ((gid2Path ptB)) (compPathTrue ptC) +-- -- -- PathP2 = compPathP {!gid2Path ptB!} (compPathP {!!} {!compPathTrue ptC!}) - SuspPushIso : ∀ {ℓ} {A : Type ℓ} → Iso (Susp A) (Pushout {A = A} (λ a → tt) λ a → tt) - Iso.fun SuspPushIso north = inl tt - Iso.fun SuspPushIso south = inr tt - Iso.fun SuspPushIso (merid a i) = push a i - Iso.inv SuspPushIso (inl tt) = north - Iso.inv SuspPushIso (inr tt) = south - Iso.inv SuspPushIso (push a i) = merid a i - Iso.leftInv SuspPushIso north = refl - Iso.leftInv SuspPushIso south = refl - Iso.leftInv SuspPushIso (merid a i) = refl - Iso.rightInv SuspPushIso (inl tt) = refl - Iso.rightInv SuspPushIso (inr tt) = refl - Iso.rightInv SuspPushIso (push a i) = refl +-- -- -- h : (x : (B , ptB) ⋀ (C , ptC)) → P x .fst +-- -- -- h (inl x) = d (inl tt) +-- -- -- h (inr (b , c)) = g b c +-- -- -- h (push (inl b) i) = gid2Path b i +-- -- -- h (push (inr x) i) = compPathTrue x i +-- -- -- h (push (push a i₁) i) = {!!} diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda index 6b18d4f06..c57b7e1c2 100644 --- a/Cubical/Homotopy/Loopspace.agda +++ b/Cubical/Homotopy/Loopspace.agda @@ -38,39 +38,41 @@ exchange : ∀ {ℓ} {A : Type ℓ} {a b c : A} {p q r : a ≡ b} {p' q' r' : b (α : p ≡ q) (β : q ≡ r) (α' : p' ≡ q') (β' : q' ≡ r') → (α ∙ β) ⋆ (α' ∙ β') ≡ (α ⋆ α') ∙ (β ⋆ β') exchange {A = A} {a = a} {c = c} {p = p} {r = r} {p' = p'} {r' = r'} α β α' β' z i = - hcomp (λ k → λ { (i = i0) → p ∙ p' - ; (i = i1) → (β ⋆ β') (k ∨ ~ z) - ; (z = i0) → ((α ∙ β) ⋆ (α' ∙ β')) i }) - (bottom z i) + hcomp (λ k → λ { (i = i0) → p ∙ p' + ; (i = i1) → (β ⋆ β') (k ∨ ~ z) + ; (z = i0) → ((α ∙ β) ⋆ (α' ∙ β')) i }) + (bottom z i) where bottom : PathP (λ i → p ∙ p' ≡ (β ⋆ β') (~ i)) ((α ∙ β) ⋆ (α' ∙ β')) (α ⋆ α') - bottom i = hcomp (λ k → λ { (i = i0) → (α ∙ β) ⋆ (α' ∙ β') - ; (i = i1) → rUnit α (~ k) ⋆ rUnit α' (~ k)}) - ((α ∙ λ j → β (~ i ∧ j)) ⋆ (α' ∙ λ j → β' (~ i ∧ j))) - + bottom i = + hcomp (λ k → λ { (i = i0) → (α ∙ β) ⋆ (α' ∙ β') + ; (i = i1) → rUnit α (~ k) ⋆ rUnit α' (~ k)}) + ((α ∙ λ j → β (~ i ∧ j)) ⋆ (α' ∙ λ j → β' (~ i ∧ j))) Eckmann-Hilton : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (α β : typ ((Ω^ (2 + n)) A)) → α ∙ β ≡ β ∙ α Eckmann-Hilton {A = (A , pt)} zero α β = - α ∙ β ≡⟨ (λ i → unit1 α (~ i) ∙ unit2 β (~ i)) ⟩ - (rUnit refl ∙ (refl ⋆ α) ∙ sym (rUnit refl)) ∙ lUnit refl ∙ (β ⋆ refl) ∙ sym (lUnit refl) ≡⟨ helper (rUnit refl) (refl ⋆ α) (β ⋆ refl) ⟩ - rUnit refl ∙ ((refl ⋆ α) ∙ (β ⋆ refl)) ∙ sym (rUnit refl) ≡⟨ (λ i → rUnit refl ∙ exchange refl β α refl (~ i) ∙ sym (rUnit refl)) ⟩ - rUnit refl ∙ ((refl ∙ β) ⋆ (α ∙ refl)) ∙ sym (rUnit refl) ≡⟨ (λ i → rUnit refl ∙ helper2 i ∙ sym (rUnit refl)) ⟩ - rUnit refl ∙ ((β ∙ refl) ⋆ (refl ∙ α)) ∙ sym (rUnit refl) ≡⟨ (λ i → rUnit refl ∙ exchange β refl refl α i ∙ sym (rUnit refl) ) ⟩ - rUnit refl ∙ ((β ⋆ refl) ∙ (refl ⋆ α)) ∙ sym (rUnit refl) ≡⟨ sym (helper (rUnit refl) (β ⋆ refl) (refl ⋆ α)) ⟩ - (lUnit refl ∙ (β ⋆ refl) ∙ sym (lUnit refl)) ∙ rUnit refl ∙ (refl ⋆ α) ∙ sym (rUnit refl) ≡⟨ (λ i → unit2 β i ∙ unit1 α i) ⟩ - β ∙ α ∎ + α ∙ β ≡⟨ (λ i → unit1 α (~ i) ∙ unit2 β (~ i)) ⟩ + (rUnit refl ∙ (refl ⋆ α) ∙ sym (rUnit refl)) ∙ lUnit refl ∙ (β ⋆ refl) ∙ sym (lUnit refl) ≡⟨ helper (rUnit refl) (refl ⋆ α) (β ⋆ refl) ⟩ + rUnit refl ∙ ((refl ⋆ α) ∙ (β ⋆ refl)) ∙ sym (rUnit refl) ≡⟨ (λ i → rUnit refl ∙ exchange refl β α refl (~ i) ∙ sym (rUnit refl)) ⟩ + rUnit refl ∙ ((refl ∙ β) ⋆ (α ∙ refl)) ∙ sym (rUnit refl) ≡⟨ (λ i → rUnit refl ∙ helper2 i ∙ sym (rUnit refl)) ⟩ + rUnit refl ∙ ((β ∙ refl) ⋆ (refl ∙ α)) ∙ sym (rUnit refl) ≡⟨ (λ i → rUnit refl ∙ exchange β refl refl α i ∙ sym (rUnit refl) ) ⟩ + rUnit refl ∙ ((β ⋆ refl) ∙ (refl ⋆ α)) ∙ sym (rUnit refl) ≡⟨ sym (helper (rUnit refl) (β ⋆ refl) (refl ⋆ α)) ⟩ + (lUnit refl ∙ (β ⋆ refl) ∙ sym (lUnit refl)) ∙ rUnit refl ∙ (refl ⋆ α) ∙ sym (rUnit refl) ≡⟨ (λ i → unit2 β i ∙ unit1 α i) ⟩ + β ∙ α ∎ where - helper : ∀ {ℓ} {A : Type ℓ} {a b : A} (p : a ≡ b) (q q' : b ≡ b) → (p ∙ q ∙ (sym p)) ∙ p ∙ q' ∙ sym p ≡ p ∙ (q ∙ q') ∙ sym p - helper p q q' = (p ∙ q ∙ (sym p)) ∙ p ∙ q' ∙ sym p ≡⟨ (λ i → assoc p q (sym p) i ∙ p ∙ q' ∙ sym p) ⟩ - ((p ∙ q) ∙ sym p) ∙ p ∙ q' ∙ sym p ≡⟨ assoc ((p ∙ q) ∙ sym p) p (q' ∙ sym p) ⟩ - (((p ∙ q) ∙ sym p) ∙ p) ∙ q' ∙ sym p ≡⟨ (λ i → assoc (p ∙ q) (sym p) p (~ i) ∙ q' ∙ sym p) ⟩ - ((p ∙ q) ∙ (sym p ∙ p)) ∙ q' ∙ sym p ≡⟨ (λ i → ((p ∙ q) ∙ lCancel p i) ∙ q' ∙ sym p) ⟩ - ((p ∙ q) ∙ refl) ∙ q' ∙ sym p ≡⟨ (λ i → rUnit (p ∙ q) (~ i) ∙ q' ∙ sym p) ⟩ - (p ∙ q) ∙ q' ∙ sym p ≡⟨ assoc (p ∙ q) q' (sym p) ⟩ - ((p ∙ q) ∙ q') ∙ sym p ≡⟨ (λ i → assoc p q q' (~ i) ∙ sym p) ⟩ - (p ∙ (q ∙ q')) ∙ sym p ≡⟨ sym (assoc p (q ∙ q') (sym p)) ⟩ - p ∙ (q ∙ q') ∙ sym p ∎ + helper : ∀ {ℓ} {A : Type ℓ} {a b : A} (p : a ≡ b) (q q' : b ≡ b) + → (p ∙ q ∙ (sym p)) ∙ p ∙ q' ∙ sym p ≡ p ∙ (q ∙ q') ∙ sym p + helper p q q' = + (p ∙ q ∙ (sym p)) ∙ p ∙ q' ∙ sym p ≡⟨ (λ i → assoc p q (sym p) i ∙ p ∙ q' ∙ sym p) ⟩ + ((p ∙ q) ∙ sym p) ∙ p ∙ q' ∙ sym p ≡⟨ assoc ((p ∙ q) ∙ sym p) p (q' ∙ sym p) ⟩ + (((p ∙ q) ∙ sym p) ∙ p) ∙ q' ∙ sym p ≡⟨ (λ i → assoc (p ∙ q) (sym p) p (~ i) ∙ q' ∙ sym p) ⟩ + ((p ∙ q) ∙ (sym p ∙ p)) ∙ q' ∙ sym p ≡⟨ (λ i → ((p ∙ q) ∙ lCancel p i) ∙ q' ∙ sym p) ⟩ + ((p ∙ q) ∙ refl) ∙ q' ∙ sym p ≡⟨ (λ i → rUnit (p ∙ q) (~ i) ∙ q' ∙ sym p) ⟩ + (p ∙ q) ∙ q' ∙ sym p ≡⟨ assoc (p ∙ q) q' (sym p) ⟩ + ((p ∙ q) ∙ q') ∙ sym p ≡⟨ (λ i → assoc p q q' (~ i) ∙ sym p) ⟩ + (p ∙ (q ∙ q')) ∙ sym p ≡⟨ sym (assoc p (q ∙ q') (sym p)) ⟩ + p ∙ (q ∙ q') ∙ sym p ∎ helper2 : (refl ∙ β) ⋆ (α ∙ refl) ≡ ((β ∙ refl) ⋆ (refl ∙ α)) helper2 = (λ i → (rUnit ((lUnit β) (~ i)) i) ⋆ (lUnit (rUnit α (~ i)) i)) @@ -90,7 +92,8 @@ Eckmann-Hilton {A = (A , pt)} (suc n) α β = Eckmann-Hilton 0 α β {- Homotopy group version -} -π-comp : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → ∥ typ ((Ω^ (suc n)) A) ∥₀ → ∥ typ ((Ω^ (suc n)) A) ∥₀ → ∥ typ ((Ω^ (suc n)) A) ∥₀ +π-comp : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → ∥ typ ((Ω^ (suc n)) A) ∥₀ + → ∥ typ ((Ω^ (suc n)) A) ∥₀ → ∥ typ ((Ω^ (suc n)) A) ∥₀ π-comp n = elim2 (λ _ _ → setTruncIsSet) λ p q → ∣ p ∙ q ∣₀ Eckmann-Hilton-π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p q : ∥ typ ((Ω^ (2 + n)) A) ∥₀) diff --git a/Cubical/ZCohomology/MayerVietoris.agda b/Cubical/ZCohomology/MayerVietoris.agda new file mode 100644 index 000000000..6bd2ad5cf --- /dev/null +++ b/Cubical/ZCohomology/MayerVietoris.agda @@ -0,0 +1,212 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.MayerVietoris where + +open import Cubical.ZCohomology.Base +open import Cubical.HITs.S1 +open import Cubical.ZCohomology.Properties +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence +open import Cubical.Data.NatMinusTwo.Base +open import Cubical.Data.Empty +open import Cubical.Data.Sigma hiding (_×_) +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.Nullification +open import Cubical.Data.Int hiding (_+_) +open import Cubical.Data.Nat +open import Cubical.Data.Prod +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Freudenthal +open import Cubical.HITs.SmashProduct.Base +open import Cubical.Data.Group.Base + +open import Cubical.ZCohomology.cupProdPrelims + + +open import Cubical.HITs.Pushout +open import Cubical.Data.Sum.Base +open import Cubical.Data.HomotopyGroup +open import Cubical.Data.Bool hiding (_⊕_) +{- +record AbGroup {ℓ} (A : Type ℓ) : Type ℓ where + constructor abgr + field + noll : A + _⊕_ : A → A → A + associate : (a b c : A) → (a ⊕ b) ⊕ c ≡ a ⊕ (b ⊕ c) + commute : (a b : A) → a ⊕ b ≡ (b ⊕ a) + r-unit : (a : A) → a ⊕ noll ≡ a + negate : (a : A) → Σ[ a⁻ ∈ A ] (a ⊕ a⁻ ≡ noll) +-} +{- +grHom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (ϕ : A → B) + → AbGroup A → AbGroup B → Type (ℓ-max ℓ ℓ') +grHom {A = A} {B = B} ϕ (abgr 0A _+A_ as-A comm-A r-A neg-A) (abgr 0B _+B_ as-B comm-B r-B neg-B) = (x y : A) → ϕ (x +A y) ≡ (ϕ x +B ϕ y) + +ImHom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (ϕ : A → B) (A' : AbGroup A) (B' : AbGroup B) + → grHom ϕ A' B' → Type (ℓ-max ℓ ℓ') +ImHom {A = A} {B = B} ϕ Agr Bgr hom = Σ[ b ∈ B ] Σ[ a ∈ A ] ϕ a ≡ b + +KerHom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (ϕ : A → B) (A' : AbGroup A) (B' : AbGroup B) + → grHom ϕ A' B' → Type (ℓ-max ℓ ℓ') +KerHom {A = A} {B = B} ϕ Agr (abgr 0B _+B_ as-B comm-B r-B neg-B) hom = Σ[ a ∈ A ] ϕ a ≡ 0B +-} + + +--- + + +private + variable + ℓ ℓ' ℓ'' : Level + A : Type ℓ + B : Type ℓ' + C : Type ℓ'' + A' : Pointed ℓ + B' : Pointed ℓ' + + +coHomFun : (n : ℕ) (f : A → B) → coHom n B → coHom n A +coHomFun n f = sRec setTruncIsSet λ β → ∣ β ∘ f ∣₀ + +coHomFun∙ : (n : ℕ) (f : A' →∙ B') → coHomRed n B' → coHomRed n A' +coHomFun∙ n (f , fpt) = sRec setTruncIsSet λ { (β , βpt) → ∣ β ∘ f , (λ i → β (fpt i)) ∙ βpt ∣₀} + +coHomFun∙2 : (n : ℕ) (f : A → B' .fst) → coHomRed n B' → coHom n A +coHomFun∙2 zero f = sRec setTruncIsSet λ { (β , βpt) → ∣ β ∘ f ∣₀} +coHomFun∙2 (suc n) f = sRec setTruncIsSet λ { (β , βpt) → ∣ β ∘ f ∣₀} + +module MV {ℓ ℓ' ℓ''} ((A , a₀) : Pointed ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) where + D : Pointed (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + D = Pushout f g , inl a₀ + + -- σ : typ A → typ (Ω (∙Susp (typ A))) + -- σ a = merid a ∙ merid (pt A) ⁻¹ + + i : (n : ℕ) → coHomRed n D → coHomRed n (A , a₀) × coHom n B + i zero = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) + λ { (δ , δpt) → ∣ (λ x → δ (inl x)) , δpt ∣₀ , ∣ (λ x → δ (inr x)) ∣₀ } + i (suc n) = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) + λ { (δ , δpt) → ∣ (λ x → δ (inl x)) , δpt ∣₀ , ∣ (λ x → δ (inr x)) ∣₀ } + + Δ : (n : ℕ) → coHomRed n (A , a₀) × coHom n B → coHom n C + Δ n (α , β) = coHomFun∙2 n f α +ₕ (-ₕ (coHomFun n g β)) + + -- + + d-pre1 : (n : ℕ) → (C → coHomK n) → D .fst → coHomK-ptd (suc n) .fst + d-pre1 zero γ (inl x) = 0ₖ + d-pre1 zero γ (inr x) = 0ₖ + d-pre1 zero γ (push a i) = Kn→ΩKn+1 zero (γ a) i + d-pre1 (suc n) γ (inl x) = 0ₖ + d-pre1 (suc n) γ (inr x) = 0ₖ + d-pre1 (suc n) γ (push a i) = Kn→ΩKn+1 (suc n) (γ a) i + + d-pre1∙ : (n : ℕ) → (γ : (C → coHomK n)) → d-pre1 n γ (snd D) ≡ ∣ north ∣ + d-pre1∙ zero γ = refl + d-pre1∙ (suc n) γ = refl + + d-pre2 : (n : ℕ) → (C → coHomK n) → D →∙ coHomK-ptd (suc n) + d-pre2 n γ = d-pre1 n γ , d-pre1∙ n γ + + d : (n : ℕ) → coHom n C → coHomRed (suc n) D + d n = sRec setTruncIsSet λ a → ∣ d-pre2 n a ∣₀ + + iIsHom : (n : ℕ) (x y : coHomRed n D) → i n (+ₕ∙ n x y) ≡ +prod n (i n x) (i n y) + iIsHom zero = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ { (f , fpt) (g , gpt) → refl} + iIsHom (suc n) = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ { (f , fpt) (g , gpt) → refl} + + prodElim : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : ∥ A ∥₀ × ∥ B ∥₀ → Type ℓ''} + → ((x : ∥ A ∥₀ × ∥ B ∥₀) → isOfHLevel 2 (C x)) + → ((a : A) (b : B) → C (∣ a ∣₀ , ∣ b ∣₀)) + → (x : ∥ A ∥₀ × ∥ B ∥₀) → C x + prodElim {A = A} {B = B} {C = C} hlevel ind (a , b) = schonf a b + where + schonf : (a : ∥ A ∥₀) (b : ∥ B ∥₀) → C (a , b) + schonf = sElim (λ x → isOfHLevelΠ 2 λ y → hlevel (_ , _)) λ a → sElim (λ x → hlevel (_ , _)) + λ b → ind a b + + ΔIsHom : (n : ℕ) (x y : coHomRed n (A , a₀) × coHom n B) → Δ n (+prod n x y) ≡ (Δ n x +ₕ Δ n y) + ΔIsHom zero = prodElim (λ x → isOfHLevelΠ 2 λ y → isOfHLevelPath 2 setTruncIsSet _ _ ) + λ {(f' , fpt) x1 → prodElim (λ x → isOfHLevelPath 2 setTruncIsSet _ _) + λ {(g' , gpt) x2 → (λ i → ∣ (λ x → (f' (f x) +ₖ g' (f x)) +ₖ -distrₖ (x1 (g x)) (x2 (g x)) i) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ g' (f x)) (-ₖ x1 (g x)) (-ₖ x2 (g x)) (~ i)) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x)) (g' (f x)) (-ₖ x1 (g x)) i +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → (f' (f x) +ₖ commₖ (g' (f x)) (-ₖ x1 (g x)) i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x)) (-ₖ x1 (g x)) (g' (f x)) (~ i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ (-ₖ x1 (g x))) (g' (f x)) (-ₖ x2 (g x)) i) ∣₀)}} + ΔIsHom (suc n) = prodElim (λ x → isOfHLevelΠ 2 λ y → isOfHLevelPath 2 setTruncIsSet _ _ ) + λ {(f' , fpt) x1 → prodElim (λ x → isOfHLevelPath 2 setTruncIsSet _ _) + λ {(g' , gpt) x2 → (λ i → ∣ (λ x → (f' (f x) +ₖ g' (f x)) +ₖ -distrₖ (x1 (g x)) (x2 (g x)) i) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ g' (f x)) (-ₖ x1 (g x)) (-ₖ x2 (g x)) (~ i)) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x)) (g' (f x)) (-ₖ x1 (g x)) i +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → (f' (f x) +ₖ commₖ (g' (f x)) (-ₖ x1 (g x)) i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x)) (-ₖ x1 (g x)) (g' (f x)) (~ i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ (-ₖ x1 (g x))) (g' (f x)) (-ₖ x2 (g x)) i) ∣₀)}} + + helper : (h l : C → Int) (x : D .fst) → (d-pre1 zero h x +ₖ d-pre1 zero l x) ≡ d-pre1 zero (λ x → h x +ₖ l x) x + helper h l (inl x) = lUnitₖ 0ₖ + helper h l (inr x) = lUnitₖ 0ₖ + helper h l (push a i) = hcomp {!!} (cong (λ x → x i) helper1) + where + helper1 : cong₂ (λ x y → x +ₖ y) (Kn→ΩKn+1 zero (h a)) (Kn→ΩKn+1 zero (l a)) ≡ Kn→ΩKn+1 zero (h a +ₖ l a) + helper1 = {!!} + + pp : PathP (λ i → lUnitₖ (0ₖ {n = 1}) i ≡ lUnitₖ 0ₖ i) (λ i → Kn→ΩKn+1 zero (h a) i +ₖ Kn→ΩKn+1 zero (l a) i) (Kn→ΩKn+1 zero (h a +ₖ l a)) + pp = toPathP ((λ z → transport (λ i → lUnitₖ (0ₖ {n = 1}) i ≡ lUnitₖ 0ₖ i) (lUnit (rUnit (λ i → Kn→ΩKn+1 zero (h a) i +ₖ Kn→ΩKn+1 zero (l a) i) z) z )) + ∙ (λ z → transp (λ i → lUnitₖ (0ₖ {n = 1}) (i ∨ z) ≡ lUnitₖ 0ₖ (i ∨ z)) z + (((λ i → lUnitₖ (0ₖ {n = 1}) ((~ i) ∧ z)) ∙ (λ i → Kn→ΩKn+1 zero (h a) i +ₖ Kn→ΩKn+1 zero (l a) i) + ∙ λ i → lUnitₖ (0ₖ {n = 1}) (i ∧ z)))) + ∙ (λ z → sym ((λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ 1 i ∙ Kn→ΩKn+1 1 0ₖ)) ∙ + (λ i → ΩKn+1→Kn (lUnit (Kn→ΩKn+1 1 0ₖ) (~ i))) ∙ + Iso.leftInv (Iso2-Kn-ΩKn+1 1) 0ₖ) + ∙ (lUnit (rUnit (λ i → Kn→ΩKn+1 zero (h a) i +ₖ Kn→ΩKn+1 zero (l a) i) z) z) + ∙ ((λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ 1 i ∙ Kn→ΩKn+1 1 0ₖ)) ∙ + (λ i → ΩKn+1→Kn (lUnit (Kn→ΩKn+1 1 0ₖ) (~ i))) ∙ + Iso.leftInv (Iso2-Kn-ΩKn+1 1) 0ₖ)) + ∙ (λ z → sym (lUnitₖ (0ₖ {n = 1})) + ∙ (lUnit (rUnit refl z ) z ∙ (λ i → ΩKn+1→Kn (Kn→ΩKn+1 1 (Kn→ΩKn+1 zero (h a) i) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 zero (l a) i))) ∙ lUnit (rUnit refl z) z) + ∙ lUnitₖ (0ₖ {n = 1})) + ∙ ((λ z → sym (lUnitₖ (0ₖ {n = 1})) + ∙ (({!!} ∙ {!ΩKn+1→Kn (Kn→ΩKn+1 1 (Kn→ΩKn+1 zero (h a) i) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 zero (l a) i)!} ∙ λ i → {!l!}) + ∙ (λ i → Iso.leftInv (Iso2-Kn-ΩKn+1 1) {!!} {!!}) ∙ (λ i → (Iso.leftInv (Iso2-Kn-ΩKn+1 1) 0ₖ (~ i ∧ z))) ∙ {!!} ∙ {!!}) + ∙ lUnitₖ (0ₖ {n = 1}))) + ∙ {!!} + ∙ {!!} + ∙ {!!}) -- compPathP (λ j i → {!!}) (compPathP {!!} (compPathP {!!} {!!})) + + dIsHom : (n : ℕ) (x y : coHom n C) → d n (x +ₕ y) ≡ +ₕ∙ (suc n) (d n x) (d n y) + dIsHom zero = sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ h l → (λ i → ∣ (λ x → helper h l x (~ i)) , lUnit (λ j → lUnitₖ 0ₖ (j ∨ (~ i))) i ∣₀) ∙ + (λ i → ∣ (λ x → d-pre1 zero h x +ₖ d-pre1 zero l x) , refl ∙ lUnitₖ 0ₖ ∣₀) + dIsHom (suc n) = {!d-pre1 zero h x +ₖ d-pre1 zero l x!} + +-- record LES {ℓ} (B : ℕ → Type ℓ) (0' : {n : ℕ} → B n) (f : {n : ℕ} → B n → B (suc n)) (_+'_ : {n : ℕ} → B n → B n → B n) : Type ℓ where +-- constructor les +-- field +-- grp : (n : ℕ) → AbGroup (B n) 0' _+'_ +-- imIsKer : (n : ℕ) → {!!} + + +-- {- +-- record Ring {ℓ} (A : Type ℓ) (_+∙_ :) : Type (ℓ-max ℓ ℓ') where +-- constructor iso +-- field +-- fun : A → B +-- inv1 : B → A +-- rightInv : section fun inv1 +-- leftInv : retract fun inv1 +-- -} diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index d6294fe49..012a7abd4 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -16,6 +16,7 @@ open import Cubical.Data.NatMinusTwo.Base open import Cubical.Data.Empty open import Cubical.Data.Sigma hiding (_×_) open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) open import Cubical.HITs.Nullification open import Cubical.Data.Int hiding (_+_) @@ -25,6 +26,8 @@ open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; re open import Cubical.Homotopy.Loopspace open import Cubical.Homotopy.Connected open import Cubical.Homotopy.Freudenthal +open import Cubical.HITs.SmashProduct.Base +open import Cubical.Data.Group.Base hiding (_≃_ ; Iso) open import Cubical.HITs.Pushout @@ -39,6 +42,7 @@ private ℓ ℓ' : Level A : Type ℓ B : Type ℓ' + A' : Pointed ℓ {- Equivalence between cohomology of A and reduced cohomology of (A + 1) -} @@ -82,8 +86,13 @@ Kn→ΩKn+1 n = Iso.fun (Iso2-Kn-ΩKn+1 n) Kn≃ΩKn+1 : {n : ℕ} → coHomK n ≃ typ (Ω (coHomK-ptd (suc n))) Kn≃ΩKn+1 {n = n} = isoToEquiv (Iso-Kn-ΩKn+1 n) +---------- Algebra/Group stuff -------- ----------- Algebra -------- +0ₖ : {n : ℕ} → coHomK n +0ₖ {n = zero} = pt (coHomK-ptd 0) +0ₖ {n = suc n} = pt (coHomK-ptd (suc n)) + +infixr 35 _+ₖ_ _+ₖ_ : {n : ℕ} → coHomK n → coHomK n → coHomK n _+ₖ_ {n = n} x y = ΩKn+1→Kn (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) @@ -91,15 +100,12 @@ _+ₖ_ {n = n} x y = ΩKn+1→Kn (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) -ₖ_ : {n : ℕ} → coHomK n → coHomK n -ₖ_ {n = n} x = ΩKn+1→Kn (sym (Kn→ΩKn+1 n x)) -0ₖ : {n : ℕ} → coHomK n -0ₖ {n = zero} = pt (coHomK-ptd 0) -0ₖ {n = suc n} = pt (coHomK-ptd (suc n)) + Kn→ΩKn+10ₖ : (n : ℕ) → Kn→ΩKn+1 n 0ₖ ≡ refl Kn→ΩKn+10ₖ zero = refl Kn→ΩKn+10ₖ (suc n) = (λ i → cong ∣_∣ (rCancel (merid north) i)) - rUnitₖ : {n : ℕ} (x : coHomK n) → x +ₖ 0ₖ ≡ x rUnitₖ {n = zero} x = cong ΩKn+1→Kn (sym (rUnit (Kn→ΩKn+1 zero x))) ∙ Iso.leftInv (Iso2-Kn-ΩKn+1 zero) x @@ -114,125 +120,280 @@ lUnitₖ {n = suc n} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc n) i ∙ Kn (λ i → ΩKn+1→Kn (lUnit (Kn→ΩKn+1 (suc n) x) (~ i))) ∙ Iso.leftInv (Iso2-Kn-ΩKn+1 (suc n)) x -rCancelₖ : {n : ℕ} (x : coHomK n) → x +ₖ (-ₖ x) ≡ 0ₖ -rCancelₖ {n = zero} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 zero x ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 zero) (sym (Kn→ΩKn+1 zero x)) i)) ∙ - cong ΩKn+1→Kn (rCancel (Kn→ΩKn+1 zero x)) -rCancelₖ {n = suc zero} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 1 x ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 1) (sym (Kn→ΩKn+1 1 x)) i)) ∙ - cong ΩKn+1→Kn (rCancel (Kn→ΩKn+1 1 x)) -rCancelₖ {n = suc (suc n)} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 (2 + n) x ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 (2 + n)) (sym (Kn→ΩKn+1 (2 + n) x)) i)) ∙ - cong ΩKn+1→Kn (rCancel (Kn→ΩKn+1 (2 + n) x)) ∙ - (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc (suc n)) (~ i))) ∙ - Iso.leftInv (Iso2-Kn-ΩKn+1 (suc (suc n))) 0ₖ - -lCancelₖ : {n : ℕ} (x : coHomK n) → (-ₖ x) +ₖ x ≡ 0ₖ -lCancelₖ {n = zero} x = (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 zero) (sym (Kn→ΩKn+1 zero x)) i ∙ Kn→ΩKn+1 zero x)) ∙ - cong ΩKn+1→Kn (lCancel (Kn→ΩKn+1 zero x)) -lCancelₖ {n = suc zero} x = ((λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 1) (sym (Kn→ΩKn+1 1 x)) i ∙ Kn→ΩKn+1 1 x))) ∙ - cong ΩKn+1→Kn (lCancel (Kn→ΩKn+1 1 x)) -lCancelₖ {n = suc (suc n)} x = (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 (2 + n)) (sym (Kn→ΩKn+1 (2 + n) x)) i ∙ Kn→ΩKn+1 (2 + n) x)) ∙ - cong ΩKn+1→Kn (lCancel (Kn→ΩKn+1 (2 + n) x)) ∙ - (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc (suc n)) (~ i))) ∙ - Iso.leftInv (Iso2-Kn-ΩKn+1 (suc (suc n))) 0ₖ - - -assocₖ : {n : ℕ} (x y z : coHomK n) → ((x +ₖ y) +ₖ z) ≡ (x +ₖ (y +ₖ z)) -assocₖ {n = n} x y z = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 n (ΩKn+1→Kn (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y)) ∙ Kn→ΩKn+1 n z)) ∙ - (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 n) (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) i ∙ Kn→ΩKn+1 n z)) ∙ - (λ i → ΩKn+1→Kn (assoc (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) (Kn→ΩKn+1 n z) (~ i))) ∙ - (λ i → ΩKn+1→Kn ((Kn→ΩKn+1 n x) ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 n) ((Kn→ΩKn+1 n y ∙ Kn→ΩKn+1 n z)) (~ i))) - -commₖ : {n : ℕ} (x y : coHomK n) → (x +ₖ y) ≡ (y +ₖ x) -commₖ {n = n} x y i = ΩKn+1→Kn (EH-instance (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) i) - where - EH-instance : (p q : typ (Ω ((∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) , ∣ north ∣))) → p ∙ q ≡ q ∙ p - EH-instance = transport (λ i → (p q : K-Id n (~ i)) → p ∙ q ≡ q ∙ p) - λ p q → Eckmann-Hilton 0 p q - where - K-Id : (n : ℕ) → typ (Ω (hLevelTrunc (3 + n) (S₊ (1 + n)) , ∣ north ∣)) ≡ typ ((Ω^ 2) (hLevelTrunc (4 + n) (S₊ (2 + n)) , ∣ north ∣ )) - K-Id n = (λ i → typ (Ω (isoToPath (Iso2-Kn-ΩKn+1 (suc n)) i , hcomp (λ k → λ {(i = i0) → ∣ north ∣ - ; (i = i1) → transportRefl (λ j → ∣ rCancel (merid north) k j ∣) k}) - (transp (λ j → isoToPath (Iso2-Kn-ΩKn+1 (suc n)) (i ∧ j)) (~ i) ∣ north ∣)))) - - ----- Algebra with cohomology groups --- - -private - § : isSet ∥ A ∥₀ - § = setTruncIsSet - -_+ₕ_ : {n : ℕ} → coHom n A → coHom n A → coHom n A -_+ₕ_ = sElim2 (λ _ _ → §) λ a b → ∣ (λ x → a x +ₖ b x) ∣₀ +-- lrUnitₖ0 : {n : ℕ} → Path (coHomK n) (0ₖ +ₖ 0ₖ) 0ₖ +-- lrUnitₖ0 {n = zero} = refl +-- lrUnitₖ0 {n = suc zero} = refl +-- lrUnitₖ0 {n = suc (suc n)} = λ i → ? --ₕ_ : {n : ℕ} → coHom n A → coHom n A --ₕ_ = sRec § λ a → ∣ (λ x → -ₖ a x) ∣₀ - -0ₕ : {n : ℕ} → coHom n A -0ₕ = ∣ (λ _ → 0ₖ) ∣₀ - -rUnitₕ : {n : ℕ} (x : coHom n A) → x +ₕ 0ₕ ≡ x -rUnitₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) - λ a i → ∣ funExt (λ x → rUnitₖ (a x)) i ∣₀ - -lUnitₕ : {n : ℕ} (x : coHom n A) → 0ₕ +ₕ x ≡ x -lUnitₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) - λ a i → ∣ funExt (λ x → lUnitₖ (a x)) i ∣₀ - -rCancelₕ : {n : ℕ} (x : coHom n A) → x +ₕ (-ₕ x) ≡ 0ₕ -rCancelₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) - λ a i → ∣ funExt (λ x → rCancelₖ (a x)) i ∣₀ - -lCancelₕ : {n : ℕ} (x : coHom n A) → (-ₕ x) +ₕ x ≡ 0ₕ -lCancelₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) - λ a i → ∣ funExt (λ x → lCancelₖ (a x)) i ∣₀ - - - -assocₕ : {n : ℕ} (x y z : coHom n A) → ((x +ₕ y) +ₕ z) ≡ (x +ₕ (y +ₕ z)) -assocₕ = elim3 (λ _ _ _ → isOfHLevelPath 1 (§ _ _)) - λ a b c i → ∣ funExt (λ x → assocₖ (a x) (b x) (c x)) i ∣₀ - - -commₕ : {n : ℕ} (x y : coHom n A) → (x +ₕ y) ≡ (y +ₕ x) -commₕ {n = n} = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) - λ a b i → ∣ funExt (λ x → commₖ (a x) (b x)) i ∣₀ +-- -- rCancelₖ : {n : ℕ} (x : coHomK n) → x +ₖ (-ₖ x) ≡ 0ₖ +-- -- rCancelₖ {n = zero} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 zero x ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 zero) (sym (Kn→ΩKn+1 zero x)) i)) ∙ +-- -- cong ΩKn+1→Kn (rCancel (Kn→ΩKn+1 zero x)) +-- -- rCancelₖ {n = suc zero} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 1 x ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 1) (sym (Kn→ΩKn+1 1 x)) i)) ∙ +-- -- cong ΩKn+1→Kn (rCancel (Kn→ΩKn+1 1 x)) +-- -- rCancelₖ {n = suc (suc n)} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 (2 + n) x ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 (2 + n)) (sym (Kn→ΩKn+1 (2 + n) x)) i)) ∙ +-- -- cong ΩKn+1→Kn (rCancel (Kn→ΩKn+1 (2 + n) x)) ∙ +-- -- (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc (suc n)) (~ i))) ∙ +-- -- Iso.leftInv (Iso2-Kn-ΩKn+1 (suc (suc n))) 0ₖ + +-- -- lCancelₖ : {n : ℕ} (x : coHomK n) → (-ₖ x) +ₖ x ≡ 0ₖ +-- -- lCancelₖ {n = zero} x = (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 zero) (sym (Kn→ΩKn+1 zero x)) i ∙ Kn→ΩKn+1 zero x)) ∙ +-- -- cong ΩKn+1→Kn (lCancel (Kn→ΩKn+1 zero x)) +-- -- lCancelₖ {n = suc zero} x = ((λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 1) (sym (Kn→ΩKn+1 1 x)) i ∙ Kn→ΩKn+1 1 x))) ∙ +-- -- cong ΩKn+1→Kn (lCancel (Kn→ΩKn+1 1 x)) +-- -- lCancelₖ {n = suc (suc n)} x = (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 (2 + n)) (sym (Kn→ΩKn+1 (2 + n) x)) i ∙ Kn→ΩKn+1 (2 + n) x)) ∙ +-- -- cong ΩKn+1→Kn (lCancel (Kn→ΩKn+1 (2 + n) x)) ∙ +-- -- (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc (suc n)) (~ i))) ∙ +-- -- Iso.leftInv (Iso2-Kn-ΩKn+1 (suc (suc n))) 0ₖ + + +-- -- assocₖ : {n : ℕ} (x y z : coHomK n) → ((x +ₖ y) +ₖ z) ≡ (x +ₖ (y +ₖ z)) +-- -- assocₖ {n = n} x y z = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 n (ΩKn+1→Kn (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y)) ∙ Kn→ΩKn+1 n z)) ∙ +-- -- (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 n) (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) i ∙ Kn→ΩKn+1 n z)) ∙ +-- -- (λ i → ΩKn+1→Kn (assoc (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) (Kn→ΩKn+1 n z) (~ i))) ∙ +-- -- (λ i → ΩKn+1→Kn ((Kn→ΩKn+1 n x) ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 n) ((Kn→ΩKn+1 n y ∙ Kn→ΩKn+1 n z)) (~ i))) + +-- -- commₖ : {n : ℕ} (x y : coHomK n) → (x +ₖ y) ≡ (y +ₖ x) +-- -- commₖ {n = n} x y i = ΩKn+1→Kn (EH-instance (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) i) +-- -- where +-- -- EH-instance : (p q : typ (Ω ((∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) , ∣ north ∣))) → p ∙ q ≡ q ∙ p +-- -- EH-instance = transport (λ i → (p q : K-Id n (~ i)) → p ∙ q ≡ q ∙ p) +-- -- λ p q → Eckmann-Hilton 0 p q +-- -- where +-- -- K-Id : (n : ℕ) → typ (Ω (hLevelTrunc (3 + n) (S₊ (1 + n)) , ∣ north ∣)) ≡ typ ((Ω^ 2) (hLevelTrunc (4 + n) (S₊ (2 + n)) , ∣ north ∣ )) +-- -- K-Id n = (λ i → typ (Ω (isoToPath (Iso2-Kn-ΩKn+1 (suc n)) i , hcomp (λ k → λ { (i = i0) → ∣ north ∣ +-- -- ; (i = i1) → transportRefl (λ j → ∣ rCancel (merid north) k j ∣) k}) +-- -- (transp (λ j → isoToPath (Iso2-Kn-ΩKn+1 (suc n)) (i ∧ j)) (~ i) ∣ north ∣)))) + +-- -- -distrₖ : {n : ℕ} → (x y : coHomK n) → -ₖ (x +ₖ y) ≡ (-ₖ x) +ₖ (-ₖ y) +-- -- -distrₖ {n = n} x y = (λ i → ΩKn+1→Kn (sym (Kn→ΩKn+1 n (ΩKn+1→Kn (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y))))) ∙ +-- -- (λ i → ΩKn+1→Kn (sym (Iso.rightInv (Iso2-Kn-ΩKn+1 n) (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) i))) ∙ +-- -- (λ i → ΩKn+1→Kn (symDistr (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) i)) ∙ +-- -- (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 n) (sym (Kn→ΩKn+1 n y)) (~ i) ∙ (Iso.rightInv (Iso2-Kn-ΩKn+1 n) (sym (Kn→ΩKn+1 n x)) (~ i)))) ∙ +-- -- commₖ (-ₖ y) (-ₖ x) + + + +-- -- ---- Group structure of cohomology groups --- + +-- -- private +-- -- § : isSet ∥ A ∥₀ +-- -- § = setTruncIsSet + +-- -- _+ₕ_ : {n : ℕ} → coHom n A → coHom n A → coHom n A +-- -- _+ₕ_ = sElim2 (λ _ _ → §) λ a b → ∣ (λ x → a x +ₖ b x) ∣₀ + +-- -- -ₕ : {n : ℕ} → coHom n A → coHom n A +-- -- -ₕ = sRec § λ a → ∣ (λ x → -ₖ a x) ∣₀ + +-- -- 0ₕ : {n : ℕ} → coHom n A +-- -- 0ₕ = ∣ (λ _ → 0ₖ) ∣₀ + +-- -- rUnitₕ : {n : ℕ} (x : coHom n A) → x +ₕ 0ₕ ≡ x +-- -- rUnitₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) +-- -- λ a i → ∣ funExt (λ x → rUnitₖ (a x)) i ∣₀ + +-- -- lUnitₕ : {n : ℕ} (x : coHom n A) → 0ₕ +ₕ x ≡ x +-- -- lUnitₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) +-- -- λ a i → ∣ funExt (λ x → lUnitₖ (a x)) i ∣₀ + +-- -- rCancelₕ : {n : ℕ} (x : coHom n A) → x +ₕ (-ₕ x) ≡ 0ₕ +-- -- rCancelₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) +-- -- λ a i → ∣ funExt (λ x → rCancelₖ (a x)) i ∣₀ + +-- -- lCancelₕ : {n : ℕ} (x : coHom n A) → (-ₕ x) +ₕ x ≡ 0ₕ +-- -- lCancelₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) +-- -- λ a i → ∣ funExt (λ x → lCancelₖ (a x)) i ∣₀ + +-- -- assocₕ : {n : ℕ} (x y z : coHom n A) → ((x +ₕ y) +ₕ z) ≡ (x +ₕ (y +ₕ z)) +-- -- assocₕ = elim3 (λ _ _ _ → isOfHLevelPath 1 (§ _ _)) +-- -- λ a b c i → ∣ funExt (λ x → assocₖ (a x) (b x) (c x)) i ∣₀ + +-- -- commₕ : {n : ℕ} (x y : coHom n A) → (x +ₕ y) ≡ (y +ₕ x) +-- -- commₕ {n = n} = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) +-- -- λ a b i → ∣ funExt (λ x → commₖ (a x) (b x)) i ∣₀ + + + +-- -- ---- Group structure of reduced cohomology groups --- + + + +-- -- coHomRedId : {A : Pointed ℓ} {n : ℕ} (x y : (A →∙ (coHomK-ptd n))) +-- -- → (p : Path (fst A → fst (coHomK-ptd n)) (x .fst) (y .fst)) +-- -- → transport (λ j → p j (snd A) ≡ snd (coHomK-ptd n)) (x .snd) ≡ y .snd +-- -- → Path (coHomRed n A) ∣ x ∣₀ ∣ y ∣₀ +-- -- coHomRedId {A = A}{n = n}(x , xpt) (y , ypt) p transpid i = ∣ (p i) , hcomp (λ k → λ {(i = i0) → xpt ; (i = i1) → transpid k}) (transp (λ j → p (i ∧ j) (snd A) ≡ snd (coHomK-ptd n)) (~ i) xpt) ∣₀ + +-- -- coHomRedId2 : {A : Pointed ℓ} {n : ℕ} (x y : (A →∙ (coHomK-ptd n))) +-- -- → (p : Path (fst A → fst (coHomK-ptd n)) (x .fst) (y .fst)) +-- -- → (λ j → p (~ j) (snd A)) ∙ (x .snd) ≡ y .snd +-- -- → Path (coHomRed n A) ∣ x ∣₀ ∣ y ∣₀ +-- -- coHomRedId2 {A = A} {n = n}(x , xpt) (y , ypt) p compid = +-- -- coHomRedId {n = n} ((x , xpt)) (y , ypt) p +-- -- ((λ i → transport (λ j → p j (snd A) ≡ snd (coHomK-ptd n)) (lUnit xpt i)) +-- -- ∙ (λ i → transport (λ j → p (j ∨ i) (snd A) ≡ snd (coHomK-ptd n)) ((λ j → p ((~ j) ∧ i) (snd A)) ∙ xpt)) +-- -- ∙ transportRefl ((λ j → p (~ j) (snd A)) ∙ xpt) +-- -- ∙ compid) + +-- -- -- test : {n : ℕ} → rUnitₖ (0ₖ {n = n}) ∙ sym (lUnitₖ 0ₖ) ≡ refl +-- -- -- test {n = zero} = rCancel (rUnitₖ (0ₖ {n = zero})) +-- -- -- test {n = suc n} = {!!} + +-- -- {- +-- -- inheritsStructure : ∀ {n} → (G : isGroup (coHomK (suc n))) +-- -- → isGroup.id G ≡ coHomK-ptd (suc n) .snd +-- -- → isGroup.inv G (coHomK-ptd (suc n) .snd) ≡ (coHomK-ptd (suc n) .snd) +-- -- → isGroup (coHomRed (suc n) A) +-- -- isGroup.id (inheritsStructure {n = n} (group-struct id inv comp lunit runit asso lcancel rcancel) 0id invid) = +-- -- ∣ (λ _ → id) , 0id ∣₀ + +-- -- isGroup.inv (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = +-- -- sRec § λ { (f , pt) → ∣ (λ a → inv (f a)) , (λ i → inv (pt i)) ∙ invid ∣₀} +-- -- isGroup.comp (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = +-- -- sElim2 (λ _ _ → §) λ {(f , ptf) (g , ptg) → ∣ (λ x → comp (f x) (g x)) , +-- -- (λ i → comp (ptf i) (ptg i)) +-- -- ∙ (λ i → comp (0id (~ i)) (∣ north ∣)) ∙ lunit ∣ north ∣ ∣₀} +-- -- isGroup.lUnit (inheritsStructure {n = n} (group-struct id inv comp lunit runit asso lcancel rcancel) 0id invid) = +-- -- sElim (λ _ → isOfHLevelPath 2 § _ _) λ { (f , pt) → coHomRedId2 {n = (suc n)} _ _ (funExt (λ x → lunit (f x))) +-- -- ((λ k → lUnit (sym (lunit (f (snd A))) ∙ (λ i → comp (0id i) (pt i)) +-- -- ∙ (λ i → comp (0id (~ i)) ∣ north ∣) ∙ (lunit ∣ north ∣)) k) +-- -- ∙ (λ k → (λ i → pt (i ∧ k)) ∙ (sym (lunit (pt k)) +-- -- ∙ (λ i → comp (0id i) (pt (i ∨ k))) ∙ (λ i → comp (0id (~ i)) (∣ north ∣)) +-- -- ∙ (lunit ∣ north ∣))) +-- -- ∙ helper pt (sym (lunit ∣ north ∣)) (λ i → comp (0id i) (∣ north ∣)))} +-- -- where +-- -- helper : ∀ {ℓ} {A : Type ℓ} {x y z w : A} (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) → p ∙ q ∙ r ∙ sym r ∙ sym q ≡ p +-- -- helper p q r = (λ i → p ∙ q ∙ assoc r (sym r) (sym q) i) +-- -- ∙ (λ i → p ∙ q ∙ rCancel r i ∙ sym q) +-- -- ∙ (λ i → p ∙ q ∙ lUnit (sym q) (~ i)) +-- -- ∙ (λ i → p ∙ rCancel q i) +-- -- ∙ sym (rUnit p) +-- -- isGroup.rUnit (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = +-- -- sElim (λ _ → isOfHLevelPath 2 § _ _) λ { (f , pt) → coHomRedId2 {n = (suc n)} _ _ (funExt (λ x → runit (f x))) +-- -- ((λ k → rUnit ((sym (runit (f (snd A)))) ∙ (λ i → comp (pt i) (0id i)) +-- -- ∙ (lUnit (lUnit (lUnit (lUnit (λ i → comp (0id (~ i)) ∣ north ∣) k) k) k) k) ∙ lunit ∣ north ∣) k) +-- -- ∙ (λ k → (sym (runit {!!}) ∙ (λ i → comp (pt i) {!!}) +-- -- ∙ ({!!} ∙ {!!} ∙ {!!} ∙ (λ i → {!!}) ∙ (λ i → comp (0id (~ i)) (pt (~ k)))) ∙ lunit (pt (~ k))) ∙ λ i → pt (i ∨ (~ k))) +-- -- ∙ (λ k → {!!}) +-- -- ∙ {!!} +-- -- ∙ {!!} +-- -- ∙ {!!} +-- -- ∙ {!!})} +-- -- isGroup.assoc (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = {!!} +-- -- isGroup.lCancel (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = {!!} +-- -- isGroup.rCancel (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = {!!} + +-- -- -} + +-- -- +ₕ∙ : {A : Pointed ℓ} (n : ℕ) → coHomRed n A → coHomRed n A → coHomRed n A +-- -- +ₕ∙ zero = sElim2 (λ _ _ → §) λ { (a , pa) (b , pb) → ∣ (λ x → a x +ₖ b x) , (λ i → (pa i +ₖ pb i)) ∣₀ } -- ∣ (λ x → a x +ₖ b x) ∣₀ +-- -- +ₕ∙ (suc n) = sElim2 (λ _ _ → §) λ { (a , pa) (b , pb) → ∣ (λ x → a x +ₖ b x) , (λ i → pa i +ₖ pb i) ∙ lUnitₖ 0ₖ ∣₀ } + +-- -- -ₕ∙ : {A : Pointed ℓ} (n : ℕ) → coHomRed n A → coHomRed n A +-- -- -ₕ∙ zero = sRec § λ {(a , pt) → ∣ (λ x → -ₖ a x ) , (λ i → -ₖ (pt i)) ∣₀} +-- -- -ₕ∙ (suc zero) = sRec § λ {(a , pt) → ∣ (λ x → -ₖ a x ) , (λ i → -ₖ (pt i)) ∙ (λ i → ΩKn+1→Kn (sym (Kn→ΩKn+10ₖ (suc zero) i))) ∣₀} +-- -- -ₕ∙ (suc (suc n)) = sRec § λ {(a , pt) → ∣ (λ x → -ₖ a x ) , (λ i → -ₖ (pt i)) ∙ +-- -- (λ i → ΩKn+1→Kn (sym (Kn→ΩKn+10ₖ (suc (suc n)) i))) ∙ +-- -- (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc (suc n)) (~ i))) ∙ +-- -- Iso.leftInv (Iso2-Kn-ΩKn+1 (suc (suc n))) ∣ north ∣ ∣₀} + +-- -- 0ₕ∙ : {A : Pointed ℓ} (n : ℕ) → coHomRed n A +-- -- 0ₕ∙ zero = ∣ (λ _ → 0ₖ) , refl ∣₀ +-- -- 0ₕ∙ (suc n) = ∣ (λ _ → 0ₖ) , refl ∣₀ + + +-- -- module _ {A : Type ℓ} {A' : Pointed ℓ'} (n : ℕ) where +-- -- 0prod : coHomRed n A' × coHom n A +-- -- 0prod = (0ₕ∙ n) , 0ₕ + +-- -- -prod : coHomRed n A' × coHom n A → coHomRed n A' × coHom n A +-- -- -prod (a , b) = (-ₕ∙ n a) , (-ₕ b) + +-- -- +prod : coHomRed n A' × coHom n A → coHomRed n A' × coHom n A → coHomRed n A' × coHom n A +-- -- +prod (a , b) (c , d) = (+ₕ∙ n a c) , (b +ₕ d) +-- -- -- rUnitₕ∙ : {n : ℕ} (x : coHomRed n A) → Path (coHomRed n A) (_+ₕ∙_ {n = n} x (0ₕ∙ {n = n})) x +-- -- -- rUnitₕ∙ {n = zero} = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) +-- -- -- λ { (a , pt) → {!coHomRedId2 {n = zero} ((λ x → a x +ₖ 0ₖ) , λ i → (_+ₖ_ {n = zero} (pt i) (pos 0))) (a , pt) (λ i → funExt (λ x → rUnitₖ (a x)) i) (helper a pt)!}} +-- -- -- where +-- -- -- helper : (a : fst A → Int) → (pt : a (snd A) ≡ pos 0) → sym (rUnitₖ (a (snd A))) ∙ (λ i → (pt i) +ₖ (pos 0)) ≡ pt +-- -- -- helper a pt = (λ i → (rUnit (sym (rUnitₖ (a (snd A)))) i) ∙ (rUnit (λ i → (pt i) +ₖ (pos 0)) i)) +-- -- -- ∙ (λ j → (sym (rUnitₖ (a (snd A))) ∙ λ i → rUnitₖ (a (snd A)) (i ∨ (~ j))) ∙ (λ i → rUnitₖ (pt (i ∧ (~ j))) j) ∙ {!!}) +-- -- -- ∙ {!!} +-- -- -- ∙ {!!} +-- -- -- ∙ {!!} +-- -- -- rUnitₕ∙ {n = suc n} = {!!} + +-- -- -- -- lUnitₕ : {n : ℕ} (x : coHom n A) → 0ₕ +ₕ x ≡ x +-- -- -- -- lUnitₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) +-- -- -- -- λ a i → ∣ funExt (λ x → lUnitₖ (a x)) i ∣₀ + +-- -- -- -- rCancelₕ : {n : ℕ} (x : coHom n A) → x +ₕ (-ₕ x) ≡ 0ₕ +-- -- -- -- rCancelₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) +-- -- -- -- λ a i → ∣ funExt (λ x → rCancelₖ (a x)) i ∣₀ + +-- -- -- -- lCancelₕ : {n : ℕ} (x : coHom n A) → (-ₕ x) +ₕ x ≡ 0ₕ +-- -- -- -- lCancelₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) +-- -- -- -- λ a i → ∣ funExt (λ x → lCancelₖ (a x)) i ∣₀ + + + +-- -- -- -- assocₕ : {n : ℕ} (x y z : coHom n A) → ((x +ₕ y) +ₕ z) ≡ (x +ₕ (y +ₕ z)) +-- -- -- -- assocₕ = elim3 (λ _ _ _ → isOfHLevelPath 1 (§ _ _)) +-- -- -- -- λ a b c i → ∣ funExt (λ x → assocₖ (a x) (b x) (c x)) i ∣₀ + + +-- -- -- -- commₕ : {n : ℕ} (x y : coHom n A) → (x +ₕ y) ≡ (y +ₕ x) +-- -- -- -- commₕ {n = n} = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) +-- -- -- -- λ a b i → ∣ funExt (λ x → commₖ (a x) (b x)) i ∣₀ + + + + + + +-- -- -- -- ---- Groups for reduced cohom --- + +-- -- -- -- 0ₖ∙ : {n : ℕ} {A : Pointed ℓ} → coHomRed n A +-- -- -- -- 0ₖ∙ {n = zero} {A = A} = ∣ (λ _ → 0ₖ) , refl ∣₀ +-- -- -- -- 0ₖ∙ {n = suc n} {A = A} = ∣ (λ _ → 0ₖ) , refl ∣₀ + + + +-- -- -- -- coHomGroup : (n : ℕ) (A : Type ℓ) → Group ℓ +-- -- -- -- Group.type (coHomGroup n A) = coHom n A +-- -- -- -- Group.setStruc (coHomGroup n A) = § +-- -- -- -- isGroup.id (Group.groupStruc (coHomGroup n A)) = 0ₕ +-- -- -- -- isGroup.inv (Group.groupStruc (coHomGroup n A)) = -ₕ +-- -- -- -- isGroup.comp (Group.groupStruc (coHomGroup n A)) = _+ₕ_ +-- -- -- -- isGroup.lUnit (Group.groupStruc (coHomGroup n A)) = lUnitₕ +-- -- -- -- isGroup.rUnit (Group.groupStruc (coHomGroup n A)) = rUnitₕ +-- -- -- -- isGroup.assoc (Group.groupStruc (coHomGroup n A)) = assocₕ +-- -- -- -- isGroup.lCancel (Group.groupStruc (coHomGroup n A)) = lCancelₕ +-- -- -- -- isGroup.rCancel (Group.groupStruc (coHomGroup n A)) = rCancelₕ +-- -- -- -- coHomPtdGroup : (n : ℕ) (A : Pointed ℓ) → Group ℓ +-- -- -- -- Group.type (coHomPtdGroup n (A , a)) = coHomRed n (A , a) +-- -- -- -- Group.setStruc (coHomPtdGroup n (A , a)) = § +-- -- -- -- isGroup.id (Group.groupStruc (coHomPtdGroup zero (A , a))) = ∣ (λ _ → 0ₖ) , refl ∣₀ +-- -- -- -- isGroup.id (Group.groupStruc (coHomPtdGroup (suc n) (A , a))) = ∣ (λ _ → 0ₖ) , refl ∣₀ +-- -- -- -- isGroup.inv (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} +-- -- -- -- isGroup.comp (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} +-- -- -- -- isGroup.lUnit (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} +-- -- -- -- isGroup.rUnit (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} +-- -- -- -- isGroup.assoc (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} +-- -- -- -- isGroup.lCancel (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} +-- -- -- -- isGroup.rCancel (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} --- Cup-prouct -- +-- -- -- -- -- -- Cup-prouct -- + +-- -- -- -- -- smashₕ : (A : Type ℓ) (n m : ℕ) → Type ℓ +-- -- -- -- -- smashₕ A n m = (coHom n A , ∣ (λ a → coHom-pt n) ∣₀) ⋀ (coHom n A , ∣ (λ a → coHom-pt n) ∣₀) -_∧-map_ : (n m : ℕ) → (S₊ (suc n) , north) wedge (S₊ (suc m) , north) → S₊ (suc n + suc m) -_∧-map_ n m = {!!} - where - S' : (n : ℕ) → Pointed ℓ-zero - S' zero = Unit , tt - S' (suc zero) = S₊ 1 , north - S' (suc (suc n)) = ((S₊ 1 , north) wedge S' (suc n)) , inl north +-- -- -- -- -- ⌣' : (n m : ℕ) → (coHomK (suc m) , coHom-pt (suc m)) ⋀ (coHomK (suc m) , coHom-pt (suc m)) +-- -- -- -- -- → coHomK ((suc n) + (suc m)) +-- -- -- -- -- ⌣' n m = equiv-proof (elim.isEquivPrecompose {A = ((S₊ (suc n)) , north) ⋁ ((S₊ (suc m)) , north)} +-- -- -- -- -- (λ a → {!⋀!}) +-- -- -- -- -- {!!} +-- -- -- -- -- {!!} +-- -- -- -- -- {!!}) {!!} .fst .fst - S'≡S : (n : ℕ) → typ (S' (suc n)) ≡ S₊ (suc n) - helper : (n : ℕ) → transp (λ j → S'≡S n j) i0 (snd (S' (suc n))) ≡ north - S'≡S zero = refl - S'≡S (suc n) = (λ i → (S₊ 1 , north) wedge S' (suc n)) ∙ - (λ i → (S₊ 1 , north) wedge ((S'≡S n i) , transp (λ j → S'≡S n (i ∧ j)) (~ i) (pt (S' (suc n))))) ∙ - (λ i → (S₊ 1 , north) wedge (S₊ (suc n) , helper n i)) ∙ - {!!} ∙ - {!!} - helper zero = refl - helper (suc n) = {!!} - where - helper2 : (n : ℕ) → Iso ((S₊ 1 , north) wedge (S₊ (suc n) , north)) (S₊ (suc (suc n))) - Iso.fun (helper2 n) (inl x) = PushoutSusp→Susp {!!} - Iso.fun (helper2 n) (inr x) = PushoutSusp→Susp {!!} - Iso.fun (helper2 n) (push tt i) = {!!} - Iso.inv (helper2 n) = {!!} - Iso.rightInv (helper2 n) = {!!} - Iso.leftInv (helper2 n) = {!!} - -smashₕ : (A : Type ℓ) (n m : ℕ) → Type ℓ -smashₕ A n m = (coHom n A , ∣ (λ a → coHom-pt n) ∣₀) smash (coHom n A , ∣ (λ a → coHom-pt n) ∣₀) - -⌣' : (n m : ℕ) → (coHomK (suc m) , coHom-pt (suc m)) smash (coHomK (suc m) , coHom-pt (suc m)) - → coHomK ((suc n) + (suc m)) -⌣' n m = equiv-proof (elim.isEquivPrecompose {A = ((S₊ (suc n)) , north) wedge ((S₊ (suc m)) , north)} - (λ a → {!smash!}) - {!!} - {!!} - {!!}) {!!} .fst .fst diff --git a/Cubical/ZCohomology/cupProdPrelims.agda b/Cubical/ZCohomology/cupProdPrelims.agda index e67700d04..693262605 100644 --- a/Cubical/ZCohomology/cupProdPrelims.agda +++ b/Cubical/ZCohomology/cupProdPrelims.agda @@ -31,6 +31,7 @@ open import Cubical.HITs.Hopf open import Cubical.Homotopy.Connected open import Cubical.Homotopy.Freudenthal hiding (encode) open import Cubical.Homotopy.Loopspace +open import Cubical.HITs.Sn open import Cubical.HITs.Pushout open import Cubical.Data.Sum.Base @@ -42,250 +43,331 @@ private A : Type ℓ B : Type ℓ' -{- We want to prove that Kn≃ΩKn+1. For this we use the map ϕ-} - -ϕ : (pt a : A) → typ (Ω (Susp A , north)) -ϕ pt a = (merid a) ∙ sym (merid pt) - {- To define the map for n=0 we use the λ k → loopᵏ map for S₊ 1. The loop is given by ϕ south north -} -private - loop* : Path (S₊ 1) north north - loop* = ϕ north south +testfun : (n : ℕ) → (hLevelTrunc n (typ (Ω (S₊ (suc n) , north)))) → (typ (Ω (hLevelTrunc (suc n) (S₊ (suc n)) , ∣ north ∣))) +testfun n = trRec (isOfHLevelPath' n (isOfHLevelTrunc _) _ _) + λ a → cong ∣_∣ a -looper : Int → Path (S₊ 1) north north -looper (pos zero) = refl -looper (pos (suc n)) = looper (pos n) ∙ loop* -looper (negsuc zero) = sym loop* -looper (negsuc (suc n)) = looper (negsuc n) ∙ sym loop* +code1 : (n : ℕ) (s : S₊ (suc n)) → (hLevelTrunc n (north ≡ s)) → Path (hLevelTrunc (suc n) (S₊ (suc n))) ∣ north ∣ ∣ s ∣ -- (typ (Ω (hLevelTrunc (suc n) (S₊ (suc n)) , ∣ north ∣))) +code1 n s = trRec (isOfHLevelPath' n (isOfHLevelTrunc _) _ _) + λ a → cong ∣_∣ a -private +decode1 : (n : ℕ) (s : S₊ (suc n)) → Path (hLevelTrunc (suc n) (S₊ (suc n))) ∣ north ∣ ∣ s ∣ → hLevelTrunc n (north ≡ s) +decode1 = {!!} - {- The map of Kn≃ΩKn+1 is given as follows. -} - Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) - Kn→ΩKn+1 zero x = cong ∣_∣ (looper x) - Kn→ΩKn+1 (suc n) = trRec (isOfHLevelTrunc (2 + (suc (suc n))) ∣ north ∣ ∣ north ∣) - λ a → cong ∣_∣ ((merid a) ∙ (sym (merid north))) - - {- We show that looper is a composition of intLoop with two other maps, all three being isos -} - sndcomp : ΩS¹ → Path (Susp Bool) north north - sndcomp = cong S¹→SuspBool - - thrdcomp : Path (Susp Bool) north north → Path (S₊ 1) north north - thrdcomp = cong SuspBool→S1 - - - looper2 : Int → Path (S₊ 1) north north - looper2 a = thrdcomp (sndcomp (intLoop a)) - - looper≡looper2 : (x : Int) → looper x ≡ looper2 x - looper≡looper2 (pos zero) = refl - looper≡looper2 (pos (suc n)) = - looper (pos n) ∙ loop* ≡⟨ (λ i → looper≡looper2 (pos n) i ∙ congFunct SuspBool→S1 (merid false) (sym (merid true)) (~ i)) ⟩ - looper2 (pos n) ∙ cong SuspBool→S1 (ϕ true false) ≡⟨ sym (congFunct SuspBool→S1 (sndcomp (intLoop (pos n))) (ϕ true false)) ⟩ - cong SuspBool→S1 (sndcomp (intLoop (pos n)) ∙ ϕ true false) ≡⟨ cong thrdcomp (sym (congFunct S¹→SuspBool (intLoop (pos n)) loop)) ⟩ - looper2 (pos (suc n)) ∎ - looper≡looper2 (negsuc zero) = - sym loop* ≡⟨ symDistr (merid south) (sym (merid north)) ⟩ - merid north ∙ sym (merid south) ≡⟨ sym (congFunct SuspBool→S1 (merid true) (sym (merid false))) ⟩ - cong SuspBool→S1 (merid true ∙ sym (merid false)) ≡⟨ cong thrdcomp (sym (symDistr (merid false) (sym (merid true)))) ⟩ - looper2 (negsuc zero) ∎ - looper≡looper2 (negsuc (suc n)) = - looper (negsuc n) ∙ sym loop* ≡⟨ ((λ i → looper≡looper2 (negsuc n) i ∙ symDistr (merid south) (sym (merid north)) i)) ⟩ - looper2 (negsuc n) ∙ merid north ∙ sym (merid south) ≡⟨ cong (λ x → looper2 (negsuc n) ∙ x) (sym (congFunct SuspBool→S1 (merid true) (sym (merid false)))) ⟩ - looper2 (negsuc n) ∙ cong SuspBool→S1 (ϕ false true) ≡⟨ cong (λ x → looper2 (negsuc n) ∙ x) (cong thrdcomp (sym (symDistr (merid false) (sym (merid true))))) ⟩ - looper2 (negsuc n) ∙ cong SuspBool→S1 (sym (ϕ true false)) ≡⟨ sym (congFunct SuspBool→S1 (sndcomp (intLoop (negsuc n))) (sym (ϕ true false))) ⟩ - thrdcomp (cong S¹→SuspBool (intLoop (negsuc n)) ∙ cong S¹→SuspBool (sym loop)) ≡⟨ cong thrdcomp (sym (congFunct S¹→SuspBool (intLoop (negsuc n)) (sym loop))) ⟩ - looper2 (negsuc (suc n)) ∎ - - - isolooper2 : Iso Int (Path (S₊ 1) north north) - isolooper2 = compIso (iso intLoop winding (decodeEncode base) windingIntLoop) - (compIso iso2 - iso1) +isCon-testfun2 : (n : ℕ) (t : S₊ (suc n)) → isHLevelConnectedFun n (code1 n t) +isCon-testfun2 n t = elim.isConnectedPrecompose (code1 n t) n λ P → (λ f → {!P!}) , {!!} + where + module _ {ℓ'' : Level} (P : Path (hLevelTrunc (suc n) (S₊ (suc n))) ∣ north ∣ ∣ t ∣ → HLevel ℓ'' n) (F : (a : hLevelTrunc n (north ≡ t)) → P (code1 n t a) .fst) where + F' : (y : S₊ (suc n)) → _ + F' y = λ(s : (a : _) → P {!a!} .fst) → s ∘ {!code1 n s!} + + Q : {!a!} + Q = {!!} + +isCon-testfun : (n : ℕ) → isHLevelConnectedFun n (testfun n) +isCon-testfun n = elim.isConnectedPrecompose (testfun n) n λ P → (λ f → {!!}) , {!!} + where + module _ {ℓ'' : Level} (P : typ (Ω (hLevelTrunc (suc n) (S₊ (suc n)) , ∣ north ∣)) → HLevel ℓ'' n) (F : (a : hLevelTrunc n (typ (Ω (S₊ (suc n) , north)))) → P (testfun n a) .fst) where + F' : _ + F' = λ(s : (a : _) → P {!a!} .fst) → s ∘ testfun n + + Q : {!a!} + Q = {!!} + + + +test : (n : ℕ) → Iso (hLevelTrunc n (typ (Ω (S₊ (suc n) , north)))) + (typ (Ω (hLevelTrunc (suc n) (S₊ (suc n)) , ∣ north ∣))) +Iso.fun (test n) = trRec (isOfHLevelPath' n (isOfHLevelTrunc _) _ _) + λ a → cong ∣_∣ a +Iso.inv (test n) = {!!} + where + map2 : hLevelTrunc (suc n) (S₊ (suc n)) → hLevelTrunc n (typ (Ω (S₊ (suc n) , north))) + map2 = trRec {!!} λ a → ∣ {!a!} ∙ {! → + HLevel ℓ''' n) →!} ∣ where - iso1 : Iso (Path (Susp Bool) north north) (Path (S₊ 1) north north) - iso1 = congIso SuspBool≃S1 + map3 : (n : ℕ) → S₊ (suc n) → hLevelTrunc n (typ (Ω (S₊ (suc n) , north))) + map3 zero x = {!!} + map3 (suc n) = {!!} -- trRec ? ? -- (λ p → ∣ cong (λ x → map4 n x) p ∣) (map3 n (map5 _ x)) + where + map4 : (n : ℕ) → S₊ (suc n) → S₊ (suc (suc n)) + map4 n north = north + map4 n south = south + map4 n (merid a i) = merid (merid a i) i + + map6 : S₊ (suc (suc (suc n))) → typ (Ω (S₊ (suc (suc n)) , north)) + map6 north = {!!} + map6 south = {!!} + map6 (merid north i) = merid (merid north i) ∙ sym (merid north) + map6 (merid south i) = merid (merid south i) ∙ sym (merid north) + map6 (merid (merid a i₁) i) = merid {!a!} ∙ sym (merid north) + + map5 : (n : ℕ) → S₊ (suc (suc n)) → S₊ (suc n) + map5 zero north = north + map5 zero south = south + map5 zero (merid a i) = {!a!} + map5 (suc zero) north = north + map5 (suc zero) south = south + map5 (suc zero) (merid a i) = merid (map5 zero a) i + map5 (suc (suc n)) north = north + map5 (suc (suc n)) south = south + map5 (suc (suc n)) (merid a i) = merid (map5 (suc n) a) i +Iso.rightInv (test n) = {!!} +Iso.leftInv (test n) = {!!} - iso2 : Iso ΩS¹ (Path (Susp Bool) north north) - iso2 = congIso (isoToEquiv (iso S¹→SuspBool SuspBool→S¹ SuspBool→S¹→SuspBool S¹→SuspBool→S¹)) - isolooper : Iso Int (Path (S₊ 1) north north) - Iso.fun isolooper = looper - Iso.inv isolooper = Iso.inv isolooper2 - Iso.rightInv isolooper a = (looper≡looper2 (Iso.inv isolooper2 a)) ∙ Iso.rightInv isolooper2 a - Iso.leftInv isolooper a = cong (Iso.inv isolooper2) (looper≡looper2 a) ∙ Iso.leftInv isolooper2 a +-- {- We want to prove that Kn≃ΩKn+1. For this we use the map ϕ-} - {- We want to show that this map is an equivalence. n ≥ 2 follows from Freudenthal, and -} +-- ϕ : (pt a : A) → typ (Ω (Susp A , north)) +-- ϕ pt a = (merid a) ∙ sym (merid pt) - {- - We want to show that the function (looper : Int → S₊ 1) defined by λ k → loopᵏ is an equivalece. We already know that the corresponding function (intLoop : Int → S¹ is) an equivalence, - so the idea is to show that when intLoop is transported along a suitable path S₊ 1 ≡ S¹ we get looper. Instead of using S₊ 1 straight away, we begin by showing this for the equivalent Susp Bool. - -} +-- {- To define the map for n=0 we use the λ k → loopᵏ map for S₊ 1. The loop is given by ϕ south north -} +-- private +-- loop* : Path (S₊ 1) north north +-- loop* = ϕ north south +-- looper : Int → Path (S₊ 1) north north +-- looper (pos zero) = refl +-- looper (pos (suc n)) = looper (pos n) ∙ loop* +-- looper (negsuc zero) = sym loop* +-- looper (negsuc (suc n)) = looper (negsuc n) ∙ sym loop* +-- private - ----------------------------------- n = 1 ----------------------------------------------------- +-- {- The map of Kn≃ΩKn+1 is given as follows. -} +-- Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) +-- Kn→ΩKn+1 zero x = cong ∣_∣ (looper x) +-- Kn→ΩKn+1 (suc n) = trRec (isOfHLevelTrunc (2 + (suc (suc n))) ∣ north ∣ ∣ north ∣) +-- λ a → cong ∣_∣ ((merid a) ∙ (sym (merid north))) - {- We begin by stating some useful lemmas -} +-- {- We show that looper is a composition of intLoop with two other maps, all three being isos -} +-- sndcomp : ΩS¹ → Path (Susp Bool) north north +-- sndcomp = cong S¹→SuspBool +-- thrdcomp : Path (Susp Bool) north north → Path (S₊ 1) north north +-- thrdcomp = cong SuspBool→S1 - S³≡SuspSuspS¹ : S³ ≡ Susp (Susp S¹) - S³≡SuspSuspS¹ = S³≡SuspS² ∙ λ i → Susp (S²≡SuspS¹ i) +-- looper2 : Int → Path (S₊ 1) north north +-- looper2 a = thrdcomp (sndcomp (intLoop a)) - S3≡SuspSuspS¹ : S₊ 3 ≡ Susp (Susp S¹) - S3≡SuspSuspS¹ = (λ i → Susp (Susp (Susp (ua Bool≃Susp⊥ (~ i))))) ∙ λ i → Susp (Susp (S¹≡SuspBool (~ i))) +-- looper≡looper2 : (x : Int) → looper x ≡ looper2 x +-- looper≡looper2 (pos zero) = refl +-- looper≡looper2 (pos (suc n)) = +-- looper (pos n) ∙ loop* ≡⟨ (λ i → looper≡looper2 (pos n) i ∙ congFunct SuspBool→S1 (merid false) (sym (merid true)) (~ i)) ⟩ +-- looper2 (pos n) ∙ cong SuspBool→S1 (ϕ true false) ≡⟨ sym (congFunct SuspBool→S1 (sndcomp (intLoop (pos n))) (ϕ true false)) ⟩ +-- cong SuspBool→S1 (sndcomp (intLoop (pos n)) ∙ ϕ true false) ≡⟨ cong thrdcomp (sym (congFunct S¹→SuspBool (intLoop (pos n)) loop)) ⟩ +-- looper2 (pos (suc n)) ∎ +-- looper≡looper2 (negsuc zero) = +-- sym loop* ≡⟨ symDistr (merid south) (sym (merid north)) ⟩ +-- merid north ∙ sym (merid south) ≡⟨ sym (congFunct SuspBool→S1 (merid true) (sym (merid false))) ⟩ +-- cong SuspBool→S1 (merid true ∙ sym (merid false)) ≡⟨ cong thrdcomp (sym (symDistr (merid false) (sym (merid true)))) ⟩ +-- looper2 (negsuc zero) ∎ +-- looper≡looper2 (negsuc (suc n)) = +-- looper (negsuc n) ∙ sym loop* ≡⟨ ((λ i → looper≡looper2 (negsuc n) i ∙ symDistr (merid south) (sym (merid north)) i)) ⟩ +-- looper2 (negsuc n) ∙ merid north ∙ sym (merid south) ≡⟨ cong (λ x → looper2 (negsuc n) ∙ x) (sym (congFunct SuspBool→S1 (merid true) (sym (merid false)))) ⟩ +-- looper2 (negsuc n) ∙ cong SuspBool→S1 (ϕ false true) ≡⟨ cong (λ x → looper2 (negsuc n) ∙ x) (cong thrdcomp (sym (symDistr (merid false) (sym (merid true))))) ⟩ +-- looper2 (negsuc n) ∙ cong SuspBool→S1 (sym (ϕ true false)) ≡⟨ sym (congFunct SuspBool→S1 (sndcomp (intLoop (negsuc n))) (sym (ϕ true false))) ⟩ +-- thrdcomp (cong S¹→SuspBool (intLoop (negsuc n)) ∙ cong S¹→SuspBool (sym loop)) ≡⟨ cong thrdcomp (sym (congFunct S¹→SuspBool (intLoop (negsuc n)) (sym loop))) ⟩ +-- looper2 (negsuc (suc n)) ∎ - sphereConnectedSpecCase : isHLevelConnected 4 (Susp (Susp S¹)) -- is- 2 -ConnectedType (Susp (Susp S¹)) - sphereConnectedSpecCase = transport (λ i → isHLevelConnected 4 (S3≡SuspSuspS¹ i)) (sphereConnected 3) +-- isolooper2 : Iso Int (Path (S₊ 1) north north) +-- isolooper2 = compIso (iso intLoop winding (decodeEncode base) windingIntLoop) +-- (compIso iso2 +-- iso1) +-- where +-- iso1 : Iso (Path (Susp Bool) north north) (Path (S₊ 1) north north) +-- iso1 = congIso SuspBool≃S1 +-- iso2 : Iso ΩS¹ (Path (Susp Bool) north north) +-- iso2 = congIso (isoToEquiv (iso S¹→SuspBool SuspBool→S¹ SuspBool→S¹→SuspBool S¹→SuspBool→S¹)) - {- We give the following map and show that it is an equivalence -} +-- isolooper : Iso Int (Path (S₊ 1) north north) +-- Iso.fun isolooper = looper +-- Iso.inv isolooper = Iso.inv isolooper2 +-- Iso.rightInv isolooper a = (looper≡looper2 (Iso.inv isolooper2 a)) ∙ Iso.rightInv isolooper2 a +-- Iso.leftInv isolooper a = cong (Iso.inv isolooper2) (looper≡looper2 a) ∙ Iso.leftInv isolooper2 a - d-map : typ (Ω ((Susp S¹) , north)) → S¹ - d-map p = subst HopfSuspS¹ p base - d-mapId : (r : S¹) → d-map (ϕ base r) ≡ r - d-mapId r = substComposite HopfSuspS¹ (merid r) (sym (merid base)) base ∙ - rotLemma r - where - rotLemma : (r : S¹) → rot r base ≡ r - rotLemma base = refl - rotLemma (loop i) = refl +-- {- We want to show that this map is an equivalence. n ≥ 2 follows from Freudenthal, and -} - d-mapComp : fiber d-map base ≡ Path (Susp (Susp S¹)) north north - d-mapComp = sym (pathSigma≡sigmaPath {B = HopfSuspS¹} _ _) ∙ helper - where - helper : Path (Σ (Susp S¹) λ x → HopfSuspS¹ x) (north , base) (north , base) ≡ Path (Susp (Susp S¹)) north north - helper = (λ i → (Path (S³≡TotalHopf (~ i)) - (transp (λ j → S³≡TotalHopf (~ i ∨ ~ j)) (~ i) (north , base)) - ((transp (λ j → S³≡TotalHopf (~ i ∨ ~ j)) (~ i) (north , base))))) ∙ - (λ i → Path (S³≡SuspSuspS¹ i) (transp (λ j → S³≡SuspSuspS¹ (i ∧ j)) (~ i) base) ((transp (λ j → S³≡SuspSuspS¹ (i ∧ j)) (~ i) base))) +-- {- +-- We want to show that the function (looper : Int → S₊ 1) defined by λ k → loopᵏ is an equivalece. We already know that the corresponding function (intLoop : Int → S¹ is) an equivalence, +-- so the idea is to show that when intLoop is transported along a suitable path S₊ 1 ≡ S¹ we get looper. Instead of using S₊ 1 straight away, we begin by showing this for the equivalent Susp Bool. +-- -} - is1Connected-dmap : isHLevelConnectedFun 3 d-map - is1Connected-dmap = toPropElim (λ s → isPropIsOfHLevel 0) (transport (λ j → isContr (∥ d-mapComp (~ j) ∥ ℕ→ℕ₋₂ 1)) - (transport (λ i → isContr (PathΩ {A = Susp (Susp S¹)} {a = north} (ℕ→ℕ₋₂ 1) i)) - (refl , isOfHLevelSuc 1 (isOfHLevelSuc 0 sphereConnectedSpecCase) ∣ north ∣ ∣ north ∣ (λ _ → ∣ north ∣)))) - d-iso2 : Iso (hLevelTrunc 3 (typ (Ω (Susp S¹ , north)))) (hLevelTrunc 3 S¹) - Iso.fun d-iso2 = trMap d-map - Iso.inv d-iso2 = Iso.inv (connectedTruncIso _ d-map is1Connected-dmap) - Iso.rightInv d-iso2 = Iso.rightInv (connectedTruncIso _ d-map is1Connected-dmap) - Iso.leftInv d-iso2 = Iso.leftInv (connectedTruncIso _ d-map is1Connected-dmap) +-- ----------------------------------- n = 1 ----------------------------------------------------- - d-iso : isIso {A = ∥ typ (Ω (Susp S¹ , north)) ∥ (ℕ→ℕ₋₂ 1)} {B = ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (trElim (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣ ) - d-iso = (Iso.inv (connectedTruncIso _ d-map is1Connected-dmap)) , (Iso.rightInv (connectedTruncIso _ d-map is1Connected-dmap) - , Iso.leftInv (connectedTruncIso _ d-map is1Connected-dmap)) +-- {- We begin by stating some useful lemmas -} - {- We show that composing (λ a → ∣ ϕ base a ∣) and (λ x → ∣ d-map x ∣) gives us the identity function. -} - d-mapId2 : (λ (x : hLevelTrunc 3 S¹) → (trRec {n = 3} {B = hLevelTrunc 3 S¹} (isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) - (trRec (isOfHLevelTrunc 3) (λ a → ∣ ϕ base a ∣) x)) ≡ λ x → x - d-mapId2 = funExt (trElim (λ x → isOfHLevelSuc 2 (isOfHLevelTrunc 3 ((trElim {n = 3} - {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} - (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) - (trElim (λ _ → isOfHLevelTrunc 3) - (λ a → ∣ ϕ base a ∣) x)) x)) - λ a i → ∣ d-mapId a i ∣) - {- This means that (λ a → ∣ ϕ base a ∣) is an equivalence -} +-- S³≡SuspSuspS¹ : S³ ≡ Susp (Susp S¹) +-- S³≡SuspSuspS¹ = S³≡SuspS² ∙ λ i → Susp (S²≡SuspS¹ i) +-- S3≡SuspSuspS¹ : S₊ 3 ≡ Susp (Susp S¹) +-- S3≡SuspSuspS¹ = (λ i → Susp (Susp (Susp (ua Bool≃Susp⊥ (~ i))))) ∙ λ i → Susp (Susp (S¹≡SuspBool (~ i))) - Iso∣ϕ-base∣ : Iso (hLevelTrunc 3 S¹) (hLevelTrunc 3 (typ (Ω (Susp S¹ , north)))) - Iso∣ϕ-base∣ = composesToId→Iso d-iso2 (trMap (ϕ base)) d-mapId2 +-- sphereConnectedSpecCase : isHLevelConnected 4 (Susp (Susp S¹)) -- is- 2 -ConnectedType (Susp (Susp S¹)) +-- sphereConnectedSpecCase = transport (λ i → isHLevelConnected 4 (S3≡SuspSuspS¹ i)) (sphereConnected 3) - --------------------------------- - -- We cheat when n = 1 and use J to prove the following lemmma. There is an obvious dependent path between ϕ base and ϕ north. Since the first one is an iso, so is the other. - -- +-- {- We give the following map and show that it is an equivalence -} - pointFunIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ} {C : (A : Type ℓ) (a1 : A) → Type ℓ'} (p : A ≡ B) (a : A) (b : B) → - (transport p a ≡ b) → - (f : (A : Type ℓ) → - (a1 : A) (a2 : hLevelTrunc 3 A) → C A a1) → - isIso (f A a) → - isIso (f B b) - pointFunIso {ℓ = ℓ}{A = A} {B = B} {C = C} = - J (λ B p → (a : A) (b : B) → - (transport p a ≡ b) → - (f : (A : Type ℓ) → - (a1 : A) (a2 : hLevelTrunc 3 A) → C A a1) → - isIso (f A a) → - isIso (f B b)) - λ a b trefl f is → transport (λ i → isIso (f A ((sym (transportRefl a) ∙ trefl) i))) is +-- d-map : typ (Ω ((Susp S¹) , north)) → S¹ +-- d-map p = subst HopfSuspS¹ p base - {- By pointFunEquiv, this gives that λ a → ∣ ϕ north a ∣ is an iso -} +-- d-mapId : (r : S¹) → d-map (ϕ base r) ≡ r +-- d-mapId r = substComposite HopfSuspS¹ (merid r) (sym (merid base)) base ∙ +-- rotLemma r +-- where +-- rotLemma : (r : S¹) → rot r base ≡ r +-- rotLemma base = refl +-- rotLemma (loop i) = refl - isIso∣ϕ∣ : isIso {A = hLevelTrunc 3 (S₊ 1)} {B = hLevelTrunc 3 (typ (Ω (S₊ 2 , north)))} (trMap (ϕ north)) - isIso∣ϕ∣ = pointFunIso {A = S¹} (λ i → S¹≡S1 (~ i)) base north refl (λ A a1 → trMap (ϕ a1)) ((Iso.inv Iso∣ϕ-base∣) , ((Iso.rightInv Iso∣ϕ-base∣) , (Iso.leftInv Iso∣ϕ-base∣))) +-- d-mapComp : fiber d-map base ≡ Path (Susp (Susp S¹)) north north +-- d-mapComp = sym (pathSigma≡sigmaPath {B = HopfSuspS¹} _ _) ∙ helper +-- where +-- helper : Path (Σ (Susp S¹) λ x → HopfSuspS¹ x) (north , base) (north , base) ≡ Path (Susp (Susp S¹)) north north +-- helper = (λ i → (Path (S³≡TotalHopf (~ i)) +-- (transp (λ j → S³≡TotalHopf (~ i ∨ ~ j)) (~ i) (north , base)) +-- ((transp (λ j → S³≡TotalHopf (~ i ∨ ~ j)) (~ i) (north , base))))) ∙ +-- (λ i → Path (S³≡SuspSuspS¹ i) (transp (λ j → S³≡SuspSuspS¹ (i ∧ j)) (~ i) base) ((transp (λ j → S³≡SuspSuspS¹ (i ∧ j)) (~ i) base))) - Iso∣ϕ∣ : Iso (hLevelTrunc 3 (S₊ 1)) (hLevelTrunc 3 (typ (Ω (S₊ 2 , north)))) - Iso.fun Iso∣ϕ∣ = trMap (ϕ north) - Iso.inv Iso∣ϕ∣ = isIso∣ϕ∣ .fst - Iso.rightInv Iso∣ϕ∣ = isIso∣ϕ∣ .snd .fst - Iso.leftInv Iso∣ϕ∣ = isIso∣ϕ∣ .snd .snd --- ---------------------------------------------------- Finishing up --------------------------------- +-- is1Connected-dmap : isHLevelConnectedFun 3 d-map +-- is1Connected-dmap = toPropElim (λ s → isPropIsOfHLevel 0) (transport (λ j → isContr (∥ d-mapComp (~ j) ∥ ℕ→ℕ₋₂ 1)) +-- (transport (λ i → isContr (PathΩ {A = Susp (Susp S¹)} {a = north} (ℕ→ℕ₋₂ 1) i)) +-- (refl , isOfHLevelSuc 1 (isOfHLevelSuc 0 sphereConnectedSpecCase) ∣ north ∣ ∣ north ∣ (λ _ → ∣ north ∣)))) --- We need ΩTrunc. It computes horribly, so its better to restate the function involved for this particular case -- +-- d-iso2 : Iso (hLevelTrunc 3 (typ (Ω (Susp S¹ , north)))) (hLevelTrunc 3 S¹) +-- Iso.fun d-iso2 = trMap d-map +-- Iso.inv d-iso2 = Iso.inv (connectedTruncIso _ d-map is1Connected-dmap) +-- Iso.rightInv d-iso2 = Iso.rightInv (connectedTruncIso _ d-map is1Connected-dmap) +-- Iso.leftInv d-iso2 = Iso.leftInv (connectedTruncIso _ d-map is1Connected-dmap) -decode-fun2 : (n : ℕ) → (x : A) → hLevelTrunc n (x ≡ x) → Path (hLevelTrunc (suc n) A) ∣ x ∣ ∣ x ∣ -decode-fun2 zero x = trElim (λ _ → isOfHLevelPath 0 (∣ x ∣ , isOfHLevelTrunc 1 ∣ x ∣) ∣ x ∣ ∣ x ∣) (λ p i → ∣ p i ∣) -decode-fun2 (suc n) x = trElim (λ _ → isOfHLevelPath' (suc n) (isOfHLevelTrunc (suc (suc n))) ∣ x ∣ ∣ x ∣) (cong ∣_∣) +-- d-iso : isIso {A = ∥ typ (Ω (Susp S¹ , north)) ∥ (ℕ→ℕ₋₂ 1)} {B = ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (trElim (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣ ) +-- d-iso = (Iso.inv (connectedTruncIso _ d-map is1Connected-dmap)) , (Iso.rightInv (connectedTruncIso _ d-map is1Connected-dmap) +-- , Iso.leftInv (connectedTruncIso _ d-map is1Connected-dmap)) --- {- auxiliary function r used to define encode -} +-- {- We show that composing (λ a → ∣ ϕ base a ∣) and (λ x → ∣ d-map x ∣) gives us the identity function. -} --- r : {m : ℕ₋₂} (u : ∥ B ∥ (suc₋₂ m)) → P u u --- r = elim (λ x → hLevelP x x) (λ a → ∣ refl ∣) +-- d-mapId2 : (λ (x : hLevelTrunc 3 S¹) → (trRec {n = 3} {B = hLevelTrunc 3 S¹} (isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) +-- (trRec (isOfHLevelTrunc 3) (λ a → ∣ ϕ base a ∣) x)) ≡ λ x → x +-- d-mapId2 = funExt (trElim (λ x → isOfHLevelSuc 2 (isOfHLevelTrunc 3 ((trElim {n = 3} +-- {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} +-- (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) +-- (trElim (λ _ → isOfHLevelTrunc 3) +-- (λ a → ∣ ϕ base a ∣) x)) x)) +-- λ a i → ∣ d-mapId a i ∣) --- {- encode function from x ≡ y to P x y -} --- encode-fun : {n : ℕ₋₂} (x y : ∥ B ∥ (suc₋₂ n)) → x ≡ y → P x y --- encode-fun x y p = transport (λ i → P x (p i)) (r x) -{- -encode2 : (n : ℕ) (x : A) → Path (hLevelTrunc (suc n) A) ∣ x ∣ ∣ x ∣ → hLevelTrunc n (x ≡ x) -encode2 {A = A} n x p = transport (λ i → hLevelTrunc n (x ≡ {!p i!})) {!!} - where - r : (m : ℕ) (u : hLevelTrunc (1 + m) A) → hLevelTrunc m (u ≡ u) - r m = trElim (λ x → isOfHLevelSuc m (isOfHLevelTrunc _) ) λ a → ∣ refl ∣ --} +-- {- This means that (λ a → ∣ ϕ base a ∣) is an equivalence -} + + +-- Iso∣ϕ-base∣ : Iso (hLevelTrunc 3 S¹) (hLevelTrunc 3 (typ (Ω (Susp S¹ , north)))) +-- Iso∣ϕ-base∣ = composesToId→Iso d-iso2 (trMap (ϕ base)) d-mapId2 + + +-- --------------------------------- +-- -- We cheat when n = 1 and use J to prove the following lemmma. There is an obvious dependent path between ϕ base and ϕ north. Since the first one is an iso, so is the other. +-- -- + + +-- pointFunIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ} {C : (A : Type ℓ) (a1 : A) → Type ℓ'} (p : A ≡ B) (a : A) (b : B) → +-- (transport p a ≡ b) → +-- (f : (A : Type ℓ) → +-- (a1 : A) (a2 : hLevelTrunc 3 A) → C A a1) → +-- isIso (f A a) → +-- isIso (f B b) +-- pointFunIso {ℓ = ℓ}{A = A} {B = B} {C = C} = +-- J (λ B p → (a : A) (b : B) → +-- (transport p a ≡ b) → +-- (f : (A : Type ℓ) → +-- (a1 : A) (a2 : hLevelTrunc 3 A) → C A a1) → +-- isIso (f A a) → +-- isIso (f B b)) +-- λ a b trefl f is → transport (λ i → isIso (f A ((sym (transportRefl a) ∙ trefl) i))) is + +-- {- By pointFunIso, this gives that λ a → ∣ ϕ north a ∣ is an iso -} + +-- isIso∣ϕ∣ : isIso {A = hLevelTrunc 3 (S₊ 1)} {B = hLevelTrunc 3 (typ (Ω (S₊ 2 , north)))} (trMap (ϕ north)) +-- isIso∣ϕ∣ = pointFunIso {A = S¹} (λ i → S¹≡S1 (~ i)) base north refl (λ A a1 → trMap (ϕ a1)) ((Iso.inv Iso∣ϕ-base∣) , ((Iso.rightInv Iso∣ϕ-base∣) , (Iso.leftInv Iso∣ϕ-base∣))) + +-- Iso∣ϕ∣ : Iso (hLevelTrunc 3 (S₊ 1)) (hLevelTrunc 3 (typ (Ω (S₊ 2 , north)))) +-- Iso.fun Iso∣ϕ∣ = trMap (ϕ north) +-- Iso.inv Iso∣ϕ∣ = isIso∣ϕ∣ .fst +-- Iso.rightInv Iso∣ϕ∣ = isIso∣ϕ∣ .snd .fst +-- Iso.leftInv Iso∣ϕ∣ = isIso∣ϕ∣ .snd .snd + +-- -- ---------------------------------------------------- Finishing up --------------------------------- + +-- -- We need ΩTrunc. It computes horribly, so its better to restate the function involved for this particular case -- + +-- decode-fun2 : (n : ℕ) → (x : A) → hLevelTrunc n (x ≡ x) → Path (hLevelTrunc (suc n) A) ∣ x ∣ ∣ x ∣ +-- decode-fun2 zero x = trElim (λ _ → isOfHLevelPath 0 (∣ x ∣ , isOfHLevelTrunc 1 ∣ x ∣) ∣ x ∣ ∣ x ∣) (λ p i → ∣ p i ∣) +-- decode-fun2 (suc n) x = trElim (λ _ → isOfHLevelPath' (suc n) (isOfHLevelTrunc (suc (suc n))) ∣ x ∣ ∣ x ∣) (cong ∣_∣) -funsAreSame : (n : ℕ) (x : A) (b : hLevelTrunc n (x ≡ x)) → (decode-fun2 n x b) ≡ (ΩTrunc.decode-fun ∣ x ∣ ∣ x ∣ b) -funsAreSame zero x = trElim (λ a → isOfHLevelPath 0 (refl , (isOfHLevelSuc 1 (isOfHLevelTrunc 1) ∣ x ∣ ∣ x ∣ refl)) _ _) λ a → refl -funsAreSame (suc n) x = trElim (λ a → isOfHLevelPath _ (isOfHLevelPath' (suc n) (isOfHLevelTrunc (suc (suc n))) ∣ x ∣ ∣ x ∣) _ _) λ a → refl +-- -- {- auxiliary function r used to define encode -} -decodeIso : (n : ℕ) → (x : A) → Iso (hLevelTrunc n (x ≡ x)) (Path (hLevelTrunc (suc n) A) ∣ x ∣ ∣ x ∣) -Iso.fun (decodeIso n x) = decode-fun2 n x -Iso.inv (decodeIso n x) = ΩTrunc.encode-fun ∣ x ∣ ∣ x ∣ -Iso.rightInv (decodeIso n x) b = funExt⁻ (funExt (funsAreSame n x)) (ΩTrunc.encode-fun ∣ x ∣ ∣ x ∣ b) ∙ ΩTrunc.P-rinv ∣ x ∣ ∣ x ∣ b -Iso.leftInv (decodeIso n x) b = cong (ΩTrunc.encode-fun ∣ x ∣ ∣ x ∣) (funsAreSame n x b) ∙ ΩTrunc.P-linv ∣ x ∣ ∣ x ∣ b +-- -- r : {m : ℕ₋₂} (u : ∥ B ∥ (suc₋₂ m)) → P u u +-- -- r = elim (λ x → hLevelP x x) (λ a → ∣ refl ∣) + +-- -- {- encode function from x ≡ y to P x y -} +-- -- encode-fun : {n : ℕ₋₂} (x y : ∥ B ∥ (suc₋₂ n)) → x ≡ y → P x y +-- -- encode-fun x y p = transport (λ i → P x (p i)) (r x) +-- {- +-- encode2 : (n : ℕ) (x : A) → Path (hLevelTrunc (suc n) A) ∣ x ∣ ∣ x ∣ → hLevelTrunc n (x ≡ x) +-- encode2 {A = A} n x p = transport (λ i → hLevelTrunc n (x ≡ {!p i!})) {!!} +-- where +-- r : (m : ℕ) (u : hLevelTrunc (1 + m) A) → hLevelTrunc m (u ≡ u) +-- r m = trElim (λ x → isOfHLevelSuc m (isOfHLevelTrunc _) ) λ a → ∣ refl ∣ +-- -} + +-- funsAreSame : (n : ℕ) (x : A) (b : hLevelTrunc n (x ≡ x)) → (decode-fun2 n x b) ≡ (ΩTrunc.decode-fun ∣ x ∣ ∣ x ∣ b) +-- funsAreSame zero x = trElim (λ a → isOfHLevelPath 0 (refl , (isOfHLevelSuc 1 (isOfHLevelTrunc 1) ∣ x ∣ ∣ x ∣ refl)) _ _) λ a → refl +-- funsAreSame (suc n) x = trElim (λ a → isOfHLevelPath _ (isOfHLevelPath' (suc n) (isOfHLevelTrunc (suc (suc n))) ∣ x ∣ ∣ x ∣) _ _) λ a → refl + +-- decodeIso : (n : ℕ) (x : A) → Iso (hLevelTrunc n (x ≡ x)) (Path (hLevelTrunc (suc n) A) ∣ x ∣ ∣ x ∣) +-- Iso.fun (decodeIso n x) = decode-fun2 n x +-- Iso.inv (decodeIso n x) = ΩTrunc.encode-fun ∣ x ∣ ∣ x ∣ +-- Iso.rightInv (decodeIso n x) b = funExt⁻ (funExt (funsAreSame n x)) (ΩTrunc.encode-fun ∣ x ∣ ∣ x ∣ b) ∙ ΩTrunc.P-rinv ∣ x ∣ ∣ x ∣ b +-- Iso.leftInv (decodeIso n x) b = cong (ΩTrunc.encode-fun ∣ x ∣ ∣ x ∣) (funsAreSame n x b) ∙ ΩTrunc.P-linv ∣ x ∣ ∣ x ∣ b -Kn→ΩKn+1Sucn : (n : ℕ) → Kn→ΩKn+1 (suc n) ≡ λ x → decode-fun2 _ north (trElim (λ _ → isOfHLevelTrunc (3 + n)) (λ a → ∣ ϕ north a ∣) x) -Kn→ΩKn+1Sucn n = funExt (trElim (λ x → isOfHLevelSuc (2 + n) (isOfHLevelTrunc (4 + n) ∣ north ∣ ∣ north ∣ _ _)) - λ a → refl) +-- Iso-Kn-ΩKn+1 : (n : ℕ) → Iso (coHomK n) (typ (Ω (coHomK-ptd (suc n)))) +-- Iso-Kn-ΩKn+1 zero = compIso isolooper (congIso (truncIdempotent≃ _ isOfHLevelS1)) +-- Iso-Kn-ΩKn+1 (suc zero) = compIso Iso∣ϕ∣ (decodeIso _ north) +-- Iso-Kn-ΩKn+1 (suc (suc n)) = compIso (connectedTruncIso2 (4 + n) _ (ϕ north) (n , (λ i → suc (suc (suc (+-suc n n (~ i))))) ∙ +-- (λ i → suc (suc (+-suc n (suc n) ((~ i)))))) +-- (isConnectedσ (suc n) (sphereConnected _))) +-- (decodeIso _ north) + +-- mapId2 : (n : ℕ) → Kn→ΩKn+1 n ≡ Iso.fun (Iso-Kn-ΩKn+1 n) +-- mapId2 zero = refl +-- mapId2 (suc zero) = funExt (trElim (λ x → isOfHLevelPath 3 (isOfHLevelTrunc 4 ∣ north ∣ ∣ north ∣) _ _) λ a → refl) +-- mapId2 (suc (suc n)) = funExt (trElim (λ x → isOfHLevelPath (4 + n) (isOfHLevelTrunc (5 + n) ∣ north ∣ ∣ north ∣) _ _) λ a → refl) +-- Iso2-Kn-ΩKn+1 : (n : ℕ) → Iso (coHomK n) (typ (Ω (coHomK-ptd (suc n)))) +-- Iso.fun (Iso2-Kn-ΩKn+1 n) = Kn→ΩKn+1 n +-- Iso.inv (Iso2-Kn-ΩKn+1 n) = Iso.inv (Iso-Kn-ΩKn+1 n) +-- Iso.rightInv (Iso2-Kn-ΩKn+1 n) a = rinv +-- where +-- abstract +-- rinv : Kn→ΩKn+1 n (Iso.inv (Iso-Kn-ΩKn+1 n) a) ≡ a +-- rinv = funExt⁻ (mapId2 n) _ ∙ Iso.rightInv (Iso-Kn-ΩKn+1 n) a +-- Iso.leftInv (Iso2-Kn-ΩKn+1 n) a = linv +-- where +-- abstract +-- linv : Iso.inv (Iso-Kn-ΩKn+1 n) (Kn→ΩKn+1 n a) ≡ a +-- linv = cong (Iso.inv (Iso-Kn-ΩKn+1 n)) (funExt⁻ (mapId2 n) a) ∙ Iso.leftInv (Iso-Kn-ΩKn+1 n) a -Iso-Kn-ΩKn+1 : (n : ℕ) → Iso (coHomK n) (typ (Ω (coHomK-ptd (suc n)))) -Iso-Kn-ΩKn+1 zero = compIso isolooper (congIso (truncIdempotent≃ _ isOfHLevelS1)) -Iso-Kn-ΩKn+1 (suc zero) = compIso Iso∣ϕ∣ (decodeIso _ north) -Iso-Kn-ΩKn+1 (suc (suc n)) = compIso (connectedTruncIso2 (4 + n) _ (ϕ north) (n , (λ i → suc (suc (suc (+-suc n n (~ i))))) ∙ - (λ i → suc (suc (+-suc n (suc n) ((~ i)))))) - (isConnectedσ (suc n) (sphereConnected _))) - (decodeIso _ north) -mapId2 : (n : ℕ) → Kn→ΩKn+1 n ≡ Iso.fun (Iso-Kn-ΩKn+1 n) -mapId2 zero = refl -mapId2 (suc zero) = funExt (trElim (λ x → isOfHLevelPath 3 (isOfHLevelTrunc 4 ∣ north ∣ ∣ north ∣) _ _) λ a → refl) -mapId2 (suc (suc n)) = funExt (trElim (λ x → isOfHLevelPath (4 + n) (isOfHLevelTrunc (5 + n) ∣ north ∣ ∣ north ∣) _ _) λ a → refl) -Iso2-Kn-ΩKn+1 : (n : ℕ) → Iso (coHomK n) (typ (Ω (coHomK-ptd (suc n)))) -Iso.fun (Iso2-Kn-ΩKn+1 n) = Kn→ΩKn+1 n -Iso.inv (Iso2-Kn-ΩKn+1 n) = Iso.inv (Iso-Kn-ΩKn+1 n) -Iso.rightInv (Iso2-Kn-ΩKn+1 n) a = funExt⁻ (mapId2 n) _ ∙ Iso.rightInv (Iso-Kn-ΩKn+1 n) a -Iso.leftInv (Iso2-Kn-ΩKn+1 n) a = cong (Iso.inv (Iso-Kn-ΩKn+1 n)) (funExt⁻ (mapId2 n) a) ∙ Iso.leftInv (Iso-Kn-ΩKn+1 n) a From fc4b2fe8310e33506ba04091f4fc97734e68eb54 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Sun, 3 May 2020 22:57:28 +0200 Subject: [PATCH 23/89] K-alg --- Cubical/ZCohomology/Properties.agda | 505 ++++++++++++----------- Cubical/ZCohomology/cupProdPrelims.agda | 508 ++++++++++-------------- 2 files changed, 456 insertions(+), 557 deletions(-) diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index 012a7abd4..e02d08261 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -101,7 +101,6 @@ _+ₖ_ {n = n} x y = ΩKn+1→Kn (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) -ₖ_ {n = n} x = ΩKn+1→Kn (sym (Kn→ΩKn+1 n x)) - Kn→ΩKn+10ₖ : (n : ℕ) → Kn→ΩKn+1 n 0ₖ ≡ refl Kn→ΩKn+10ₖ zero = refl Kn→ΩKn+10ₖ (suc n) = (λ i → cong ∣_∣ (rCancel (merid north) i)) @@ -120,280 +119,272 @@ lUnitₖ {n = suc n} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc n) i ∙ Kn (λ i → ΩKn+1→Kn (lUnit (Kn→ΩKn+1 (suc n) x) (~ i))) ∙ Iso.leftInv (Iso2-Kn-ΩKn+1 (suc n)) x --- lrUnitₖ0 : {n : ℕ} → Path (coHomK n) (0ₖ +ₖ 0ₖ) 0ₖ --- lrUnitₖ0 {n = zero} = refl --- lrUnitₖ0 {n = suc zero} = refl --- lrUnitₖ0 {n = suc (suc n)} = λ i → ? - --- -- rCancelₖ : {n : ℕ} (x : coHomK n) → x +ₖ (-ₖ x) ≡ 0ₖ --- -- rCancelₖ {n = zero} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 zero x ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 zero) (sym (Kn→ΩKn+1 zero x)) i)) ∙ --- -- cong ΩKn+1→Kn (rCancel (Kn→ΩKn+1 zero x)) --- -- rCancelₖ {n = suc zero} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 1 x ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 1) (sym (Kn→ΩKn+1 1 x)) i)) ∙ --- -- cong ΩKn+1→Kn (rCancel (Kn→ΩKn+1 1 x)) --- -- rCancelₖ {n = suc (suc n)} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 (2 + n) x ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 (2 + n)) (sym (Kn→ΩKn+1 (2 + n) x)) i)) ∙ --- -- cong ΩKn+1→Kn (rCancel (Kn→ΩKn+1 (2 + n) x)) ∙ --- -- (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc (suc n)) (~ i))) ∙ --- -- Iso.leftInv (Iso2-Kn-ΩKn+1 (suc (suc n))) 0ₖ - --- -- lCancelₖ : {n : ℕ} (x : coHomK n) → (-ₖ x) +ₖ x ≡ 0ₖ --- -- lCancelₖ {n = zero} x = (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 zero) (sym (Kn→ΩKn+1 zero x)) i ∙ Kn→ΩKn+1 zero x)) ∙ --- -- cong ΩKn+1→Kn (lCancel (Kn→ΩKn+1 zero x)) --- -- lCancelₖ {n = suc zero} x = ((λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 1) (sym (Kn→ΩKn+1 1 x)) i ∙ Kn→ΩKn+1 1 x))) ∙ --- -- cong ΩKn+1→Kn (lCancel (Kn→ΩKn+1 1 x)) --- -- lCancelₖ {n = suc (suc n)} x = (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 (2 + n)) (sym (Kn→ΩKn+1 (2 + n) x)) i ∙ Kn→ΩKn+1 (2 + n) x)) ∙ --- -- cong ΩKn+1→Kn (lCancel (Kn→ΩKn+1 (2 + n) x)) ∙ --- -- (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc (suc n)) (~ i))) ∙ --- -- Iso.leftInv (Iso2-Kn-ΩKn+1 (suc (suc n))) 0ₖ - - --- -- assocₖ : {n : ℕ} (x y z : coHomK n) → ((x +ₖ y) +ₖ z) ≡ (x +ₖ (y +ₖ z)) --- -- assocₖ {n = n} x y z = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 n (ΩKn+1→Kn (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y)) ∙ Kn→ΩKn+1 n z)) ∙ --- -- (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 n) (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) i ∙ Kn→ΩKn+1 n z)) ∙ --- -- (λ i → ΩKn+1→Kn (assoc (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) (Kn→ΩKn+1 n z) (~ i))) ∙ --- -- (λ i → ΩKn+1→Kn ((Kn→ΩKn+1 n x) ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 n) ((Kn→ΩKn+1 n y ∙ Kn→ΩKn+1 n z)) (~ i))) - --- -- commₖ : {n : ℕ} (x y : coHomK n) → (x +ₖ y) ≡ (y +ₖ x) --- -- commₖ {n = n} x y i = ΩKn+1→Kn (EH-instance (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) i) --- -- where --- -- EH-instance : (p q : typ (Ω ((∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) , ∣ north ∣))) → p ∙ q ≡ q ∙ p --- -- EH-instance = transport (λ i → (p q : K-Id n (~ i)) → p ∙ q ≡ q ∙ p) --- -- λ p q → Eckmann-Hilton 0 p q --- -- where --- -- K-Id : (n : ℕ) → typ (Ω (hLevelTrunc (3 + n) (S₊ (1 + n)) , ∣ north ∣)) ≡ typ ((Ω^ 2) (hLevelTrunc (4 + n) (S₊ (2 + n)) , ∣ north ∣ )) --- -- K-Id n = (λ i → typ (Ω (isoToPath (Iso2-Kn-ΩKn+1 (suc n)) i , hcomp (λ k → λ { (i = i0) → ∣ north ∣ --- -- ; (i = i1) → transportRefl (λ j → ∣ rCancel (merid north) k j ∣) k}) --- -- (transp (λ j → isoToPath (Iso2-Kn-ΩKn+1 (suc n)) (i ∧ j)) (~ i) ∣ north ∣)))) - --- -- -distrₖ : {n : ℕ} → (x y : coHomK n) → -ₖ (x +ₖ y) ≡ (-ₖ x) +ₖ (-ₖ y) --- -- -distrₖ {n = n} x y = (λ i → ΩKn+1→Kn (sym (Kn→ΩKn+1 n (ΩKn+1→Kn (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y))))) ∙ --- -- (λ i → ΩKn+1→Kn (sym (Iso.rightInv (Iso2-Kn-ΩKn+1 n) (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) i))) ∙ --- -- (λ i → ΩKn+1→Kn (symDistr (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) i)) ∙ --- -- (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 n) (sym (Kn→ΩKn+1 n y)) (~ i) ∙ (Iso.rightInv (Iso2-Kn-ΩKn+1 n) (sym (Kn→ΩKn+1 n x)) (~ i)))) ∙ --- -- commₖ (-ₖ y) (-ₖ x) - - - --- -- ---- Group structure of cohomology groups --- - --- -- private --- -- § : isSet ∥ A ∥₀ --- -- § = setTruncIsSet - --- -- _+ₕ_ : {n : ℕ} → coHom n A → coHom n A → coHom n A --- -- _+ₕ_ = sElim2 (λ _ _ → §) λ a b → ∣ (λ x → a x +ₖ b x) ∣₀ - --- -- -ₕ : {n : ℕ} → coHom n A → coHom n A --- -- -ₕ = sRec § λ a → ∣ (λ x → -ₖ a x) ∣₀ - --- -- 0ₕ : {n : ℕ} → coHom n A --- -- 0ₕ = ∣ (λ _ → 0ₖ) ∣₀ - --- -- rUnitₕ : {n : ℕ} (x : coHom n A) → x +ₕ 0ₕ ≡ x --- -- rUnitₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) --- -- λ a i → ∣ funExt (λ x → rUnitₖ (a x)) i ∣₀ - --- -- lUnitₕ : {n : ℕ} (x : coHom n A) → 0ₕ +ₕ x ≡ x --- -- lUnitₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) --- -- λ a i → ∣ funExt (λ x → lUnitₖ (a x)) i ∣₀ - --- -- rCancelₕ : {n : ℕ} (x : coHom n A) → x +ₕ (-ₕ x) ≡ 0ₕ --- -- rCancelₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) --- -- λ a i → ∣ funExt (λ x → rCancelₖ (a x)) i ∣₀ - --- -- lCancelₕ : {n : ℕ} (x : coHom n A) → (-ₕ x) +ₕ x ≡ 0ₕ --- -- lCancelₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) --- -- λ a i → ∣ funExt (λ x → lCancelₖ (a x)) i ∣₀ +rCancelₖ : {n : ℕ} (x : coHomK n) → x +ₖ (-ₖ x) ≡ 0ₖ +rCancelₖ {n = zero} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 zero x ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 zero) (sym (Kn→ΩKn+1 zero x)) i)) ∙ + cong ΩKn+1→Kn (rCancel (Kn→ΩKn+1 zero x)) +rCancelₖ {n = suc zero} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 1 x ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 1) (sym (Kn→ΩKn+1 1 x)) i)) ∙ + cong ΩKn+1→Kn (rCancel (Kn→ΩKn+1 1 x)) +rCancelₖ {n = suc (suc n)} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 (2 + n) x ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 (2 + n)) (sym (Kn→ΩKn+1 (2 + n) x)) i)) ∙ + cong ΩKn+1→Kn (rCancel (Kn→ΩKn+1 (2 + n) x)) ∙ + (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc (suc n)) (~ i))) ∙ + Iso.leftInv (Iso2-Kn-ΩKn+1 (suc (suc n))) 0ₖ + +lCancelₖ : {n : ℕ} (x : coHomK n) → (-ₖ x) +ₖ x ≡ 0ₖ +lCancelₖ {n = zero} x = (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 zero) (sym (Kn→ΩKn+1 zero x)) i ∙ Kn→ΩKn+1 zero x)) ∙ + cong ΩKn+1→Kn (lCancel (Kn→ΩKn+1 zero x)) +lCancelₖ {n = suc zero} x = ((λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 1) (sym (Kn→ΩKn+1 1 x)) i ∙ Kn→ΩKn+1 1 x))) ∙ + cong ΩKn+1→Kn (lCancel (Kn→ΩKn+1 1 x)) +lCancelₖ {n = suc (suc n)} x = (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 (2 + n)) (sym (Kn→ΩKn+1 (2 + n) x)) i ∙ Kn→ΩKn+1 (2 + n) x)) ∙ + cong ΩKn+1→Kn (lCancel (Kn→ΩKn+1 (2 + n) x)) ∙ + (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc (suc n)) (~ i))) ∙ + Iso.leftInv (Iso2-Kn-ΩKn+1 (suc (suc n))) 0ₖ + + +assocₖ : {n : ℕ} (x y z : coHomK n) → ((x +ₖ y) +ₖ z) ≡ (x +ₖ (y +ₖ z)) +assocₖ {n = n} x y z = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 n (ΩKn+1→Kn (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y)) ∙ Kn→ΩKn+1 n z)) ∙ + (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 n) (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) i ∙ Kn→ΩKn+1 n z)) ∙ + (λ i → ΩKn+1→Kn (assoc (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) (Kn→ΩKn+1 n z) (~ i))) ∙ + (λ i → ΩKn+1→Kn ((Kn→ΩKn+1 n x) ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 n) ((Kn→ΩKn+1 n y ∙ Kn→ΩKn+1 n z)) (~ i))) + +commₖ : {n : ℕ} (x y : coHomK n) → (x +ₖ y) ≡ (y +ₖ x) +commₖ {n = n} x y i = ΩKn+1→Kn (EH-instance (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) i) + where + EH-instance : (p q : typ (Ω ((∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) , ∣ north ∣))) → p ∙ q ≡ q ∙ p + EH-instance = transport (λ i → (p q : K-Id n (~ i)) → p ∙ q ≡ q ∙ p) + λ p q → Eckmann-Hilton 0 p q + where + K-Id : (n : ℕ) → typ (Ω (hLevelTrunc (3 + n) (S₊ (1 + n)) , ∣ north ∣)) ≡ typ ((Ω^ 2) (hLevelTrunc (4 + n) (S₊ (2 + n)) , ∣ north ∣ )) + K-Id n = (λ i → typ (Ω (isoToPath (Iso2-Kn-ΩKn+1 (suc n)) i , hcomp (λ k → λ { (i = i0) → ∣ north ∣ + ; (i = i1) → transportRefl (λ j → ∣ rCancel (merid north) k j ∣) k}) + (transp (λ j → isoToPath (Iso2-Kn-ΩKn+1 (suc n)) (i ∧ j)) (~ i) ∣ north ∣)))) + +-distrₖ : {n : ℕ} → (x y : coHomK n) → -ₖ (x +ₖ y) ≡ (-ₖ x) +ₖ (-ₖ y) +-distrₖ {n = n} x y = (λ i → ΩKn+1→Kn (sym (Kn→ΩKn+1 n (ΩKn+1→Kn (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y))))) ∙ + (λ i → ΩKn+1→Kn (sym (Iso.rightInv (Iso2-Kn-ΩKn+1 n) (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) i))) ∙ + (λ i → ΩKn+1→Kn (symDistr (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) i)) ∙ + (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 n) (sym (Kn→ΩKn+1 n y)) (~ i) ∙ (Iso.rightInv (Iso2-Kn-ΩKn+1 n) (sym (Kn→ΩKn+1 n x)) (~ i)))) ∙ + commₖ (-ₖ y) (-ₖ x) + + +---- Group structure of cohomology groups --- + +private + § : isSet ∥ A ∥₀ + § = setTruncIsSet + +_+ₕ_ : {n : ℕ} → coHom n A → coHom n A → coHom n A +_+ₕ_ = sElim2 (λ _ _ → §) λ a b → ∣ (λ x → a x +ₖ b x) ∣₀ + +-ₕ : {n : ℕ} → coHom n A → coHom n A +-ₕ = sRec § λ a → ∣ (λ x → -ₖ a x) ∣₀ + +0ₕ : {n : ℕ} → coHom n A +0ₕ = ∣ (λ _ → 0ₖ) ∣₀ + +rUnitₕ : {n : ℕ} (x : coHom n A) → x +ₕ 0ₕ ≡ x +rUnitₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → rUnitₖ (a x)) i ∣₀ + +lUnitₕ : {n : ℕ} (x : coHom n A) → 0ₕ +ₕ x ≡ x +lUnitₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → lUnitₖ (a x)) i ∣₀ + +rCancelₕ : {n : ℕ} (x : coHom n A) → x +ₕ (-ₕ x) ≡ 0ₕ +rCancelₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → rCancelₖ (a x)) i ∣₀ + +lCancelₕ : {n : ℕ} (x : coHom n A) → (-ₕ x) +ₕ x ≡ 0ₕ +lCancelₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) + λ a i → ∣ funExt (λ x → lCancelₖ (a x)) i ∣₀ + +assocₕ : {n : ℕ} (x y z : coHom n A) → ((x +ₕ y) +ₕ z) ≡ (x +ₕ (y +ₕ z)) +assocₕ = elim3 (λ _ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b c i → ∣ funExt (λ x → assocₖ (a x) (b x) (c x)) i ∣₀ + +commₕ : {n : ℕ} (x y : coHom n A) → (x +ₕ y) ≡ (y +ₕ x) +commₕ {n = n} = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) + λ a b i → ∣ funExt (λ x → commₖ (a x) (b x)) i ∣₀ + + + +---- Group structure of reduced cohomology groups --- + +coHomRedId : {A : Pointed ℓ} {n : ℕ} (x y : (A →∙ (coHomK-ptd n))) + → (p : Path (fst A → fst (coHomK-ptd n)) (x .fst) (y .fst)) + → transport (λ j → p j (snd A) ≡ snd (coHomK-ptd n)) (x .snd) ≡ y .snd + → Path (coHomRed n A) ∣ x ∣₀ ∣ y ∣₀ +coHomRedId {A = A}{n = n}(x , xpt) (y , ypt) p transpid i = ∣ (p i) , hcomp (λ k → λ {(i = i0) → xpt ; (i = i1) → transpid k}) (transp (λ j → p (i ∧ j) (snd A) ≡ snd (coHomK-ptd n)) (~ i) xpt) ∣₀ + +coHomRedId2 : {A : Pointed ℓ} {n : ℕ} (x y : (A →∙ (coHomK-ptd n))) + → (p : Path (fst A → fst (coHomK-ptd n)) (x .fst) (y .fst)) + → (λ j → p (~ j) (snd A)) ∙ (x .snd) ≡ y .snd + → Path (coHomRed n A) ∣ x ∣₀ ∣ y ∣₀ +coHomRedId2 {A = A} {n = n}(x , xpt) (y , ypt) p compid = + coHomRedId {n = n} ((x , xpt)) (y , ypt) p + ((λ i → transport (λ j → p j (snd A) ≡ snd (coHomK-ptd n)) (lUnit xpt i)) + ∙ (λ i → transport (λ j → p (j ∨ i) (snd A) ≡ snd (coHomK-ptd n)) ((λ j → p ((~ j) ∧ i) (snd A)) ∙ xpt)) + ∙ transportRefl ((λ j → p (~ j) (snd A)) ∙ xpt) + ∙ compid) + + -- test : {n : ℕ} → rUnitₖ (0ₖ {n = n}) ∙ sym (lUnitₖ 0ₖ) ≡ refl + -- test {n = zero} = rCancel (rUnitₖ (0ₖ {n = zero})) + -- test {n = suc n} = {!!} + +{- + inheritsStructure : ∀ {n} → (G : isGroup (coHomK (suc n))) + → isGroup.id G ≡ coHomK-ptd (suc n) .snd + → isGroup.inv G (coHomK-ptd (suc n) .snd) ≡ (coHomK-ptd (suc n) .snd) + → isGroup (coHomRed (suc n) A) + isGroup.id (inheritsStructure {n = n} (group-struct id inv comp lunit runit asso lcancel rcancel) 0id invid) = + ∣ (λ _ → id) , 0id ∣₀ + + isGroup.inv (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = + sRec § λ { (f , pt) → ∣ (λ a → inv (f a)) , (λ i → inv (pt i)) ∙ invid ∣₀} + isGroup.comp (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = + sElim2 (λ _ _ → §) λ {(f , ptf) (g , ptg) → ∣ (λ x → comp (f x) (g x)) , + (λ i → comp (ptf i) (ptg i)) + ∙ (λ i → comp (0id (~ i)) (∣ north ∣)) ∙ lunit ∣ north ∣ ∣₀} + isGroup.lUnit (inheritsStructure {n = n} (group-struct id inv comp lunit runit asso lcancel rcancel) 0id invid) = + sElim (λ _ → isOfHLevelPath 2 § _ _) λ { (f , pt) → coHomRedId2 {n = (suc n)} _ _ (funExt (λ x → lunit (f x))) + ((λ k → lUnit (sym (lunit (f (snd A))) ∙ (λ i → comp (0id i) (pt i)) + ∙ (λ i → comp (0id (~ i)) ∣ north ∣) ∙ (lunit ∣ north ∣)) k) + ∙ (λ k → (λ i → pt (i ∧ k)) ∙ (sym (lunit (pt k)) + ∙ (λ i → comp (0id i) (pt (i ∨ k))) ∙ (λ i → comp (0id (~ i)) (∣ north ∣)) + ∙ (lunit ∣ north ∣))) + ∙ helper pt (sym (lunit ∣ north ∣)) (λ i → comp (0id i) (∣ north ∣)))} + where + helper : ∀ {ℓ} {A : Type ℓ} {x y z w : A} (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) → p ∙ q ∙ r ∙ sym r ∙ sym q ≡ p + helper p q r = (λ i → p ∙ q ∙ assoc r (sym r) (sym q) i) + ∙ (λ i → p ∙ q ∙ rCancel r i ∙ sym q) + ∙ (λ i → p ∙ q ∙ lUnit (sym q) (~ i)) + ∙ (λ i → p ∙ rCancel q i) + ∙ sym (rUnit p) + isGroup.rUnit (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = + sElim (λ _ → isOfHLevelPath 2 § _ _) λ { (f , pt) → coHomRedId2 {n = (suc n)} _ _ (funExt (λ x → runit (f x))) + ((λ k → rUnit ((sym (runit (f (snd A)))) ∙ (λ i → comp (pt i) (0id i)) + ∙ (lUnit (lUnit (lUnit (lUnit (λ i → comp (0id (~ i)) ∣ north ∣) k) k) k) k) ∙ lunit ∣ north ∣) k) + ∙ (λ k → (sym (runit {!!}) ∙ (λ i → comp (pt i) {!!}) + ∙ ({!!} ∙ {!!} ∙ {!!} ∙ (λ i → {!!}) ∙ (λ i → comp (0id (~ i)) (pt (~ k)))) ∙ lunit (pt (~ k))) ∙ λ i → pt (i ∨ (~ k))) + ∙ (λ k → {!!}) + ∙ {!!} + ∙ {!!} + ∙ {!!} + ∙ {!!})} + isGroup.assoc (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = {!!} + isGroup.lCancel (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = {!!} + isGroup.rCancel (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = {!!} + +-} + ++ₕ∙ : {A : Pointed ℓ} (n : ℕ) → coHomRed n A → coHomRed n A → coHomRed n A ++ₕ∙ zero = sElim2 (λ _ _ → §) λ { (a , pa) (b , pb) → ∣ (λ x → a x +ₖ b x) , (λ i → (pa i +ₖ pb i)) ∣₀ } -- ∣ (λ x → a x +ₖ b x) ∣₀ ++ₕ∙ (suc n) = sElim2 (λ _ _ → §) λ { (a , pa) (b , pb) → ∣ (λ x → a x +ₖ b x) , (λ i → pa i +ₖ pb i) ∙ lUnitₖ 0ₖ ∣₀ } + +-ₕ∙ : {A : Pointed ℓ} (n : ℕ) → coHomRed n A → coHomRed n A +-ₕ∙ zero = sRec § λ {(a , pt) → ∣ (λ x → -ₖ a x ) , (λ i → -ₖ (pt i)) ∣₀} +-ₕ∙ (suc zero) = sRec § λ {(a , pt) → ∣ (λ x → -ₖ a x ) , (λ i → -ₖ (pt i)) ∙ (λ i → ΩKn+1→Kn (sym (Kn→ΩKn+10ₖ (suc zero) i))) ∣₀} +-ₕ∙ (suc (suc n)) = sRec § λ {(a , pt) → ∣ (λ x → -ₖ a x ) , (λ i → -ₖ (pt i)) ∙ + (λ i → ΩKn+1→Kn (sym (Kn→ΩKn+10ₖ (suc (suc n)) i))) ∙ + (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc (suc n)) (~ i))) ∙ + Iso.leftInv (Iso2-Kn-ΩKn+1 (suc (suc n))) ∣ north ∣ ∣₀} + +0ₕ∙ : {A : Pointed ℓ} (n : ℕ) → coHomRed n A +0ₕ∙ zero = ∣ (λ _ → 0ₖ) , refl ∣₀ +0ₕ∙ (suc n) = ∣ (λ _ → 0ₖ) , refl ∣₀ + + +module _ {A : Type ℓ} {A' : Pointed ℓ'} (n : ℕ) where + 0prod : coHomRed n A' × coHom n A + 0prod = (0ₕ∙ n) , 0ₕ + + -prod : coHomRed n A' × coHom n A → coHomRed n A' × coHom n A + -prod (a , b) = (-ₕ∙ n a) , (-ₕ b) + + +prod : coHomRed n A' × coHom n A → coHomRed n A' × coHom n A → coHomRed n A' × coHom n A + +prod (a , b) (c , d) = (+ₕ∙ n a c) , (b +ₕ d) +-- rUnitₕ∙ : {n : ℕ} (x : coHomRed n A) → Path (coHomRed n A) (_+ₕ∙_ {n = n} x (0ₕ∙ {n = n})) x +-- rUnitₕ∙ {n = zero} = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) +-- λ { (a , pt) → {!coHomRedId2 {n = zero} ((λ x → a x +ₖ 0ₖ) , λ i → (_+ₖ_ {n = zero} (pt i) (pos 0))) (a , pt) (λ i → funExt (λ x → rUnitₖ (a x)) i) (helper a pt)!}} +-- where +-- helper : (a : fst A → Int) → (pt : a (snd A) ≡ pos 0) → sym (rUnitₖ (a (snd A))) ∙ (λ i → (pt i) +ₖ (pos 0)) ≡ pt +-- helper a pt = (λ i → (rUnit (sym (rUnitₖ (a (snd A)))) i) ∙ (rUnit (λ i → (pt i) +ₖ (pos 0)) i)) +-- ∙ (λ j → (sym (rUnitₖ (a (snd A))) ∙ λ i → rUnitₖ (a (snd A)) (i ∨ (~ j))) ∙ (λ i → rUnitₖ (pt (i ∧ (~ j))) j) ∙ {!!}) +-- ∙ {!!} +-- ∙ {!!} +-- ∙ {!!} +-- rUnitₕ∙ {n = suc n} = {!!} + +-- -- lUnitₕ : {n : ℕ} (x : coHom n A) → 0ₕ +ₕ x ≡ x +-- -- lUnitₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) +-- -- λ a i → ∣ funExt (λ x → lUnitₖ (a x)) i ∣₀ + +-- -- rCancelₕ : {n : ℕ} (x : coHom n A) → x +ₕ (-ₕ x) ≡ 0ₕ +-- -- rCancelₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) +-- -- λ a i → ∣ funExt (λ x → rCancelₖ (a x)) i ∣₀ + +-- -- lCancelₕ : {n : ℕ} (x : coHom n A) → (-ₕ x) +ₕ x ≡ 0ₕ +-- -- lCancelₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) +-- -- λ a i → ∣ funExt (λ x → lCancelₖ (a x)) i ∣₀ + + -- -- assocₕ : {n : ℕ} (x y z : coHom n A) → ((x +ₕ y) +ₕ z) ≡ (x +ₕ (y +ₕ z)) -- -- assocₕ = elim3 (λ _ _ _ → isOfHLevelPath 1 (§ _ _)) -- -- λ a b c i → ∣ funExt (λ x → assocₖ (a x) (b x) (c x)) i ∣₀ + -- -- commₕ : {n : ℕ} (x y : coHom n A) → (x +ₕ y) ≡ (y +ₕ x) -- -- commₕ {n = n} = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) -- -- λ a b i → ∣ funExt (λ x → commₖ (a x) (b x)) i ∣₀ --- -- ---- Group structure of reduced cohomology groups --- --- -- coHomRedId : {A : Pointed ℓ} {n : ℕ} (x y : (A →∙ (coHomK-ptd n))) --- -- → (p : Path (fst A → fst (coHomK-ptd n)) (x .fst) (y .fst)) --- -- → transport (λ j → p j (snd A) ≡ snd (coHomK-ptd n)) (x .snd) ≡ y .snd --- -- → Path (coHomRed n A) ∣ x ∣₀ ∣ y ∣₀ --- -- coHomRedId {A = A}{n = n}(x , xpt) (y , ypt) p transpid i = ∣ (p i) , hcomp (λ k → λ {(i = i0) → xpt ; (i = i1) → transpid k}) (transp (λ j → p (i ∧ j) (snd A) ≡ snd (coHomK-ptd n)) (~ i) xpt) ∣₀ +-- -- ---- Groups for reduced cohom --- --- -- coHomRedId2 : {A : Pointed ℓ} {n : ℕ} (x y : (A →∙ (coHomK-ptd n))) --- -- → (p : Path (fst A → fst (coHomK-ptd n)) (x .fst) (y .fst)) --- -- → (λ j → p (~ j) (snd A)) ∙ (x .snd) ≡ y .snd --- -- → Path (coHomRed n A) ∣ x ∣₀ ∣ y ∣₀ --- -- coHomRedId2 {A = A} {n = n}(x , xpt) (y , ypt) p compid = --- -- coHomRedId {n = n} ((x , xpt)) (y , ypt) p --- -- ((λ i → transport (λ j → p j (snd A) ≡ snd (coHomK-ptd n)) (lUnit xpt i)) --- -- ∙ (λ i → transport (λ j → p (j ∨ i) (snd A) ≡ snd (coHomK-ptd n)) ((λ j → p ((~ j) ∧ i) (snd A)) ∙ xpt)) --- -- ∙ transportRefl ((λ j → p (~ j) (snd A)) ∙ xpt) --- -- ∙ compid) +-- -- 0ₖ∙ : {n : ℕ} {A : Pointed ℓ} → coHomRed n A +-- -- 0ₖ∙ {n = zero} {A = A} = ∣ (λ _ → 0ₖ) , refl ∣₀ +-- -- 0ₖ∙ {n = suc n} {A = A} = ∣ (λ _ → 0ₖ) , refl ∣₀ --- -- -- test : {n : ℕ} → rUnitₖ (0ₖ {n = n}) ∙ sym (lUnitₖ 0ₖ) ≡ refl --- -- -- test {n = zero} = rCancel (rUnitₖ (0ₖ {n = zero})) --- -- -- test {n = suc n} = {!!} --- -- {- --- -- inheritsStructure : ∀ {n} → (G : isGroup (coHomK (suc n))) --- -- → isGroup.id G ≡ coHomK-ptd (suc n) .snd --- -- → isGroup.inv G (coHomK-ptd (suc n) .snd) ≡ (coHomK-ptd (suc n) .snd) --- -- → isGroup (coHomRed (suc n) A) --- -- isGroup.id (inheritsStructure {n = n} (group-struct id inv comp lunit runit asso lcancel rcancel) 0id invid) = --- -- ∣ (λ _ → id) , 0id ∣₀ - --- -- isGroup.inv (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = --- -- sRec § λ { (f , pt) → ∣ (λ a → inv (f a)) , (λ i → inv (pt i)) ∙ invid ∣₀} --- -- isGroup.comp (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = --- -- sElim2 (λ _ _ → §) λ {(f , ptf) (g , ptg) → ∣ (λ x → comp (f x) (g x)) , --- -- (λ i → comp (ptf i) (ptg i)) --- -- ∙ (λ i → comp (0id (~ i)) (∣ north ∣)) ∙ lunit ∣ north ∣ ∣₀} --- -- isGroup.lUnit (inheritsStructure {n = n} (group-struct id inv comp lunit runit asso lcancel rcancel) 0id invid) = --- -- sElim (λ _ → isOfHLevelPath 2 § _ _) λ { (f , pt) → coHomRedId2 {n = (suc n)} _ _ (funExt (λ x → lunit (f x))) --- -- ((λ k → lUnit (sym (lunit (f (snd A))) ∙ (λ i → comp (0id i) (pt i)) --- -- ∙ (λ i → comp (0id (~ i)) ∣ north ∣) ∙ (lunit ∣ north ∣)) k) --- -- ∙ (λ k → (λ i → pt (i ∧ k)) ∙ (sym (lunit (pt k)) --- -- ∙ (λ i → comp (0id i) (pt (i ∨ k))) ∙ (λ i → comp (0id (~ i)) (∣ north ∣)) --- -- ∙ (lunit ∣ north ∣))) --- -- ∙ helper pt (sym (lunit ∣ north ∣)) (λ i → comp (0id i) (∣ north ∣)))} --- -- where --- -- helper : ∀ {ℓ} {A : Type ℓ} {x y z w : A} (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) → p ∙ q ∙ r ∙ sym r ∙ sym q ≡ p --- -- helper p q r = (λ i → p ∙ q ∙ assoc r (sym r) (sym q) i) --- -- ∙ (λ i → p ∙ q ∙ rCancel r i ∙ sym q) --- -- ∙ (λ i → p ∙ q ∙ lUnit (sym q) (~ i)) --- -- ∙ (λ i → p ∙ rCancel q i) --- -- ∙ sym (rUnit p) --- -- isGroup.rUnit (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = --- -- sElim (λ _ → isOfHLevelPath 2 § _ _) λ { (f , pt) → coHomRedId2 {n = (suc n)} _ _ (funExt (λ x → runit (f x))) --- -- ((λ k → rUnit ((sym (runit (f (snd A)))) ∙ (λ i → comp (pt i) (0id i)) --- -- ∙ (lUnit (lUnit (lUnit (lUnit (λ i → comp (0id (~ i)) ∣ north ∣) k) k) k) k) ∙ lunit ∣ north ∣) k) --- -- ∙ (λ k → (sym (runit {!!}) ∙ (λ i → comp (pt i) {!!}) --- -- ∙ ({!!} ∙ {!!} ∙ {!!} ∙ (λ i → {!!}) ∙ (λ i → comp (0id (~ i)) (pt (~ k)))) ∙ lunit (pt (~ k))) ∙ λ i → pt (i ∨ (~ k))) --- -- ∙ (λ k → {!!}) --- -- ∙ {!!} --- -- ∙ {!!} --- -- ∙ {!!} --- -- ∙ {!!})} --- -- isGroup.assoc (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = {!!} --- -- isGroup.lCancel (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = {!!} --- -- isGroup.rCancel (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = {!!} - --- -- -} - --- -- +ₕ∙ : {A : Pointed ℓ} (n : ℕ) → coHomRed n A → coHomRed n A → coHomRed n A --- -- +ₕ∙ zero = sElim2 (λ _ _ → §) λ { (a , pa) (b , pb) → ∣ (λ x → a x +ₖ b x) , (λ i → (pa i +ₖ pb i)) ∣₀ } -- ∣ (λ x → a x +ₖ b x) ∣₀ --- -- +ₕ∙ (suc n) = sElim2 (λ _ _ → §) λ { (a , pa) (b , pb) → ∣ (λ x → a x +ₖ b x) , (λ i → pa i +ₖ pb i) ∙ lUnitₖ 0ₖ ∣₀ } - --- -- -ₕ∙ : {A : Pointed ℓ} (n : ℕ) → coHomRed n A → coHomRed n A --- -- -ₕ∙ zero = sRec § λ {(a , pt) → ∣ (λ x → -ₖ a x ) , (λ i → -ₖ (pt i)) ∣₀} --- -- -ₕ∙ (suc zero) = sRec § λ {(a , pt) → ∣ (λ x → -ₖ a x ) , (λ i → -ₖ (pt i)) ∙ (λ i → ΩKn+1→Kn (sym (Kn→ΩKn+10ₖ (suc zero) i))) ∣₀} --- -- -ₕ∙ (suc (suc n)) = sRec § λ {(a , pt) → ∣ (λ x → -ₖ a x ) , (λ i → -ₖ (pt i)) ∙ --- -- (λ i → ΩKn+1→Kn (sym (Kn→ΩKn+10ₖ (suc (suc n)) i))) ∙ --- -- (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc (suc n)) (~ i))) ∙ --- -- Iso.leftInv (Iso2-Kn-ΩKn+1 (suc (suc n))) ∣ north ∣ ∣₀} - --- -- 0ₕ∙ : {A : Pointed ℓ} (n : ℕ) → coHomRed n A --- -- 0ₕ∙ zero = ∣ (λ _ → 0ₖ) , refl ∣₀ --- -- 0ₕ∙ (suc n) = ∣ (λ _ → 0ₖ) , refl ∣₀ - - --- -- module _ {A : Type ℓ} {A' : Pointed ℓ'} (n : ℕ) where --- -- 0prod : coHomRed n A' × coHom n A --- -- 0prod = (0ₕ∙ n) , 0ₕ - --- -- -prod : coHomRed n A' × coHom n A → coHomRed n A' × coHom n A --- -- -prod (a , b) = (-ₕ∙ n a) , (-ₕ b) - --- -- +prod : coHomRed n A' × coHom n A → coHomRed n A' × coHom n A → coHomRed n A' × coHom n A --- -- +prod (a , b) (c , d) = (+ₕ∙ n a c) , (b +ₕ d) --- -- -- rUnitₕ∙ : {n : ℕ} (x : coHomRed n A) → Path (coHomRed n A) (_+ₕ∙_ {n = n} x (0ₕ∙ {n = n})) x --- -- -- rUnitₕ∙ {n = zero} = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) --- -- -- λ { (a , pt) → {!coHomRedId2 {n = zero} ((λ x → a x +ₖ 0ₖ) , λ i → (_+ₖ_ {n = zero} (pt i) (pos 0))) (a , pt) (λ i → funExt (λ x → rUnitₖ (a x)) i) (helper a pt)!}} --- -- -- where --- -- -- helper : (a : fst A → Int) → (pt : a (snd A) ≡ pos 0) → sym (rUnitₖ (a (snd A))) ∙ (λ i → (pt i) +ₖ (pos 0)) ≡ pt --- -- -- helper a pt = (λ i → (rUnit (sym (rUnitₖ (a (snd A)))) i) ∙ (rUnit (λ i → (pt i) +ₖ (pos 0)) i)) --- -- -- ∙ (λ j → (sym (rUnitₖ (a (snd A))) ∙ λ i → rUnitₖ (a (snd A)) (i ∨ (~ j))) ∙ (λ i → rUnitₖ (pt (i ∧ (~ j))) j) ∙ {!!}) --- -- -- ∙ {!!} --- -- -- ∙ {!!} --- -- -- ∙ {!!} --- -- -- rUnitₕ∙ {n = suc n} = {!!} - --- -- -- -- lUnitₕ : {n : ℕ} (x : coHom n A) → 0ₕ +ₕ x ≡ x --- -- -- -- lUnitₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) --- -- -- -- λ a i → ∣ funExt (λ x → lUnitₖ (a x)) i ∣₀ - --- -- -- -- rCancelₕ : {n : ℕ} (x : coHom n A) → x +ₕ (-ₕ x) ≡ 0ₕ --- -- -- -- rCancelₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) --- -- -- -- λ a i → ∣ funExt (λ x → rCancelₖ (a x)) i ∣₀ - --- -- -- -- lCancelₕ : {n : ℕ} (x : coHom n A) → (-ₕ x) +ₕ x ≡ 0ₕ --- -- -- -- lCancelₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) --- -- -- -- λ a i → ∣ funExt (λ x → lCancelₖ (a x)) i ∣₀ - - - --- -- -- -- assocₕ : {n : ℕ} (x y z : coHom n A) → ((x +ₕ y) +ₕ z) ≡ (x +ₕ (y +ₕ z)) --- -- -- -- assocₕ = elim3 (λ _ _ _ → isOfHLevelPath 1 (§ _ _)) --- -- -- -- λ a b c i → ∣ funExt (λ x → assocₖ (a x) (b x) (c x)) i ∣₀ - - --- -- -- -- commₕ : {n : ℕ} (x y : coHom n A) → (x +ₕ y) ≡ (y +ₕ x) --- -- -- -- commₕ {n = n} = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) --- -- -- -- λ a b i → ∣ funExt (λ x → commₖ (a x) (b x)) i ∣₀ - - - - - - --- -- -- -- ---- Groups for reduced cohom --- - --- -- -- -- 0ₖ∙ : {n : ℕ} {A : Pointed ℓ} → coHomRed n A --- -- -- -- 0ₖ∙ {n = zero} {A = A} = ∣ (λ _ → 0ₖ) , refl ∣₀ --- -- -- -- 0ₖ∙ {n = suc n} {A = A} = ∣ (λ _ → 0ₖ) , refl ∣₀ - - - --- -- -- -- coHomGroup : (n : ℕ) (A : Type ℓ) → Group ℓ --- -- -- -- Group.type (coHomGroup n A) = coHom n A --- -- -- -- Group.setStruc (coHomGroup n A) = § --- -- -- -- isGroup.id (Group.groupStruc (coHomGroup n A)) = 0ₕ --- -- -- -- isGroup.inv (Group.groupStruc (coHomGroup n A)) = -ₕ --- -- -- -- isGroup.comp (Group.groupStruc (coHomGroup n A)) = _+ₕ_ --- -- -- -- isGroup.lUnit (Group.groupStruc (coHomGroup n A)) = lUnitₕ --- -- -- -- isGroup.rUnit (Group.groupStruc (coHomGroup n A)) = rUnitₕ --- -- -- -- isGroup.assoc (Group.groupStruc (coHomGroup n A)) = assocₕ --- -- -- -- isGroup.lCancel (Group.groupStruc (coHomGroup n A)) = lCancelₕ --- -- -- -- isGroup.rCancel (Group.groupStruc (coHomGroup n A)) = rCancelₕ --- -- -- -- coHomPtdGroup : (n : ℕ) (A : Pointed ℓ) → Group ℓ --- -- -- -- Group.type (coHomPtdGroup n (A , a)) = coHomRed n (A , a) --- -- -- -- Group.setStruc (coHomPtdGroup n (A , a)) = § --- -- -- -- isGroup.id (Group.groupStruc (coHomPtdGroup zero (A , a))) = ∣ (λ _ → 0ₖ) , refl ∣₀ --- -- -- -- isGroup.id (Group.groupStruc (coHomPtdGroup (suc n) (A , a))) = ∣ (λ _ → 0ₖ) , refl ∣₀ --- -- -- -- isGroup.inv (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} --- -- -- -- isGroup.comp (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} --- -- -- -- isGroup.lUnit (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} --- -- -- -- isGroup.rUnit (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} --- -- -- -- isGroup.assoc (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} --- -- -- -- isGroup.lCancel (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} --- -- -- -- isGroup.rCancel (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} +-- -- coHomGroup : (n : ℕ) (A : Type ℓ) → Group ℓ +-- -- Group.type (coHomGroup n A) = coHom n A +-- -- Group.setStruc (coHomGroup n A) = § +-- -- isGroup.id (Group.groupStruc (coHomGroup n A)) = 0ₕ +-- -- isGroup.inv (Group.groupStruc (coHomGroup n A)) = -ₕ +-- -- isGroup.comp (Group.groupStruc (coHomGroup n A)) = _+ₕ_ +-- -- isGroup.lUnit (Group.groupStruc (coHomGroup n A)) = lUnitₕ +-- -- isGroup.rUnit (Group.groupStruc (coHomGroup n A)) = rUnitₕ +-- -- isGroup.assoc (Group.groupStruc (coHomGroup n A)) = assocₕ +-- -- isGroup.lCancel (Group.groupStruc (coHomGroup n A)) = lCancelₕ +-- -- isGroup.rCancel (Group.groupStruc (coHomGroup n A)) = rCancelₕ + +-- -- coHomPtdGroup : (n : ℕ) (A : Pointed ℓ) → Group ℓ +-- -- Group.type (coHomPtdGroup n (A , a)) = coHomRed n (A , a) +-- -- Group.setStruc (coHomPtdGroup n (A , a)) = § +-- -- isGroup.id (Group.groupStruc (coHomPtdGroup zero (A , a))) = ∣ (λ _ → 0ₖ) , refl ∣₀ +-- -- isGroup.id (Group.groupStruc (coHomPtdGroup (suc n) (A , a))) = ∣ (λ _ → 0ₖ) , refl ∣₀ +-- -- isGroup.inv (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} +-- -- isGroup.comp (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} +-- -- isGroup.lUnit (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} +-- -- isGroup.rUnit (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} +-- -- isGroup.assoc (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} +-- -- isGroup.lCancel (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} +-- -- isGroup.rCancel (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} --- -- -- -- -- -- Cup-prouct -- +-- -- -- -- Cup-prouct -- --- -- -- -- -- smashₕ : (A : Type ℓ) (n m : ℕ) → Type ℓ --- -- -- -- -- smashₕ A n m = (coHom n A , ∣ (λ a → coHom-pt n) ∣₀) ⋀ (coHom n A , ∣ (λ a → coHom-pt n) ∣₀) - --- -- -- -- -- ⌣' : (n m : ℕ) → (coHomK (suc m) , coHom-pt (suc m)) ⋀ (coHomK (suc m) , coHom-pt (suc m)) --- -- -- -- -- → coHomK ((suc n) + (suc m)) --- -- -- -- -- ⌣' n m = equiv-proof (elim.isEquivPrecompose {A = ((S₊ (suc n)) , north) ⋁ ((S₊ (suc m)) , north)} --- -- -- -- -- (λ a → {!⋀!}) --- -- -- -- -- {!!} --- -- -- -- -- {!!} --- -- -- -- -- {!!}) {!!} .fst .fst +-- -- -- smashₕ : (A : Type ℓ) (n m : ℕ) → Type ℓ +-- -- -- smashₕ A n m = (coHom n A , ∣ (λ a → coHom-pt n) ∣₀) ⋀ (coHom n A , ∣ (λ a → coHom-pt n) ∣₀) + +-- -- -- ⌣' : (n m : ℕ) → (coHomK (suc m) , coHom-pt (suc m)) ⋀ (coHomK (suc m) , coHom-pt (suc m)) +-- -- -- → coHomK ((suc n) + (suc m)) +-- -- -- ⌣' n m = equiv-proof (elim.isEquivPrecompose {A = ((S₊ (suc n)) , north) ⋁ ((S₊ (suc m)) , north)} +-- -- -- (λ a → {!⋀!}) +-- -- -- {!!} +-- -- -- {!!} +-- -- -- {!!}) {!!} .fst .fst diff --git a/Cubical/ZCohomology/cupProdPrelims.agda b/Cubical/ZCohomology/cupProdPrelims.agda index 693262605..ca01e9584 100644 --- a/Cubical/ZCohomology/cupProdPrelims.agda +++ b/Cubical/ZCohomology/cupProdPrelims.agda @@ -44,330 +44,238 @@ private B : Type ℓ' -testfun : (n : ℕ) → (hLevelTrunc n (typ (Ω (S₊ (suc n) , north)))) → (typ (Ω (hLevelTrunc (suc n) (S₊ (suc n)) , ∣ north ∣))) -testfun n = trRec (isOfHLevelPath' n (isOfHLevelTrunc _) _ _) - λ a → cong ∣_∣ a +{- We want to prove that Kn≃ΩKn+1. For this we use the map ϕ-} -code1 : (n : ℕ) (s : S₊ (suc n)) → (hLevelTrunc n (north ≡ s)) → Path (hLevelTrunc (suc n) (S₊ (suc n))) ∣ north ∣ ∣ s ∣ -- (typ (Ω (hLevelTrunc (suc n) (S₊ (suc n)) , ∣ north ∣))) -code1 n s = trRec (isOfHLevelPath' n (isOfHLevelTrunc _) _ _) - λ a → cong ∣_∣ a +ϕ : (pt a : A) → typ (Ω (Susp A , north)) +ϕ pt a = (merid a) ∙ sym (merid pt) -decode1 : (n : ℕ) (s : S₊ (suc n)) → Path (hLevelTrunc (suc n) (S₊ (suc n))) ∣ north ∣ ∣ s ∣ → hLevelTrunc n (north ≡ s) -decode1 = {!!} - -isCon-testfun2 : (n : ℕ) (t : S₊ (suc n)) → isHLevelConnectedFun n (code1 n t) -isCon-testfun2 n t = elim.isConnectedPrecompose (code1 n t) n λ P → (λ f → {!P!}) , {!!} - where - module _ {ℓ'' : Level} (P : Path (hLevelTrunc (suc n) (S₊ (suc n))) ∣ north ∣ ∣ t ∣ → HLevel ℓ'' n) (F : (a : hLevelTrunc n (north ≡ t)) → P (code1 n t a) .fst) where - F' : (y : S₊ (suc n)) → _ - F' y = λ(s : (a : _) → P {!a!} .fst) → s ∘ {!code1 n s!} - - Q : {!a!} - Q = {!!} - -isCon-testfun : (n : ℕ) → isHLevelConnectedFun n (testfun n) -isCon-testfun n = elim.isConnectedPrecompose (testfun n) n λ P → (λ f → {!!}) , {!!} - where - module _ {ℓ'' : Level} (P : typ (Ω (hLevelTrunc (suc n) (S₊ (suc n)) , ∣ north ∣)) → HLevel ℓ'' n) (F : (a : hLevelTrunc n (typ (Ω (S₊ (suc n) , north)))) → P (testfun n a) .fst) where - F' : _ - F' = λ(s : (a : _) → P {!a!} .fst) → s ∘ testfun n - - Q : {!a!} - Q = {!!} - - - -test : (n : ℕ) → Iso (hLevelTrunc n (typ (Ω (S₊ (suc n) , north)))) - (typ (Ω (hLevelTrunc (suc n) (S₊ (suc n)) , ∣ north ∣))) -Iso.fun (test n) = trRec (isOfHLevelPath' n (isOfHLevelTrunc _) _ _) - λ a → cong ∣_∣ a -Iso.inv (test n) = {!!} - where - map2 : hLevelTrunc (suc n) (S₊ (suc n)) → hLevelTrunc n (typ (Ω (S₊ (suc n) , north))) - map2 = trRec {!!} λ a → ∣ {!a!} ∙ {! → - HLevel ℓ''' n) →!} ∣ - where - map3 : (n : ℕ) → S₊ (suc n) → hLevelTrunc n (typ (Ω (S₊ (suc n) , north))) - map3 zero x = {!!} - map3 (suc n) = {!!} -- trRec ? ? -- (λ p → ∣ cong (λ x → map4 n x) p ∣) (map3 n (map5 _ x)) - where - map4 : (n : ℕ) → S₊ (suc n) → S₊ (suc (suc n)) - map4 n north = north - map4 n south = south - map4 n (merid a i) = merid (merid a i) i - - map6 : S₊ (suc (suc (suc n))) → typ (Ω (S₊ (suc (suc n)) , north)) - map6 north = {!!} - map6 south = {!!} - map6 (merid north i) = merid (merid north i) ∙ sym (merid north) - map6 (merid south i) = merid (merid south i) ∙ sym (merid north) - map6 (merid (merid a i₁) i) = merid {!a!} ∙ sym (merid north) - - map5 : (n : ℕ) → S₊ (suc (suc n)) → S₊ (suc n) - map5 zero north = north - map5 zero south = south - map5 zero (merid a i) = {!a!} - map5 (suc zero) north = north - map5 (suc zero) south = south - map5 (suc zero) (merid a i) = merid (map5 zero a) i - map5 (suc (suc n)) north = north - map5 (suc (suc n)) south = south - map5 (suc (suc n)) (merid a i) = merid (map5 (suc n) a) i -Iso.rightInv (test n) = {!!} -Iso.leftInv (test n) = {!!} - - - --- {- We want to prove that Kn≃ΩKn+1. For this we use the map ϕ-} - --- ϕ : (pt a : A) → typ (Ω (Susp A , north)) --- ϕ pt a = (merid a) ∙ sym (merid pt) - --- {- To define the map for n=0 we use the λ k → loopᵏ map for S₊ 1. The loop is given by ϕ south north -} --- private --- loop* : Path (S₊ 1) north north --- loop* = ϕ north south - --- looper : Int → Path (S₊ 1) north north --- looper (pos zero) = refl --- looper (pos (suc n)) = looper (pos n) ∙ loop* --- looper (negsuc zero) = sym loop* --- looper (negsuc (suc n)) = looper (negsuc n) ∙ sym loop* - --- private - --- {- The map of Kn≃ΩKn+1 is given as follows. -} --- Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) --- Kn→ΩKn+1 zero x = cong ∣_∣ (looper x) --- Kn→ΩKn+1 (suc n) = trRec (isOfHLevelTrunc (2 + (suc (suc n))) ∣ north ∣ ∣ north ∣) --- λ a → cong ∣_∣ ((merid a) ∙ (sym (merid north))) - --- {- We show that looper is a composition of intLoop with two other maps, all three being isos -} --- sndcomp : ΩS¹ → Path (Susp Bool) north north --- sndcomp = cong S¹→SuspBool - --- thrdcomp : Path (Susp Bool) north north → Path (S₊ 1) north north --- thrdcomp = cong SuspBool→S1 - - --- looper2 : Int → Path (S₊ 1) north north --- looper2 a = thrdcomp (sndcomp (intLoop a)) - --- looper≡looper2 : (x : Int) → looper x ≡ looper2 x --- looper≡looper2 (pos zero) = refl --- looper≡looper2 (pos (suc n)) = --- looper (pos n) ∙ loop* ≡⟨ (λ i → looper≡looper2 (pos n) i ∙ congFunct SuspBool→S1 (merid false) (sym (merid true)) (~ i)) ⟩ --- looper2 (pos n) ∙ cong SuspBool→S1 (ϕ true false) ≡⟨ sym (congFunct SuspBool→S1 (sndcomp (intLoop (pos n))) (ϕ true false)) ⟩ --- cong SuspBool→S1 (sndcomp (intLoop (pos n)) ∙ ϕ true false) ≡⟨ cong thrdcomp (sym (congFunct S¹→SuspBool (intLoop (pos n)) loop)) ⟩ --- looper2 (pos (suc n)) ∎ --- looper≡looper2 (negsuc zero) = --- sym loop* ≡⟨ symDistr (merid south) (sym (merid north)) ⟩ --- merid north ∙ sym (merid south) ≡⟨ sym (congFunct SuspBool→S1 (merid true) (sym (merid false))) ⟩ --- cong SuspBool→S1 (merid true ∙ sym (merid false)) ≡⟨ cong thrdcomp (sym (symDistr (merid false) (sym (merid true)))) ⟩ --- looper2 (negsuc zero) ∎ --- looper≡looper2 (negsuc (suc n)) = --- looper (negsuc n) ∙ sym loop* ≡⟨ ((λ i → looper≡looper2 (negsuc n) i ∙ symDistr (merid south) (sym (merid north)) i)) ⟩ --- looper2 (negsuc n) ∙ merid north ∙ sym (merid south) ≡⟨ cong (λ x → looper2 (negsuc n) ∙ x) (sym (congFunct SuspBool→S1 (merid true) (sym (merid false)))) ⟩ --- looper2 (negsuc n) ∙ cong SuspBool→S1 (ϕ false true) ≡⟨ cong (λ x → looper2 (negsuc n) ∙ x) (cong thrdcomp (sym (symDistr (merid false) (sym (merid true))))) ⟩ --- looper2 (negsuc n) ∙ cong SuspBool→S1 (sym (ϕ true false)) ≡⟨ sym (congFunct SuspBool→S1 (sndcomp (intLoop (negsuc n))) (sym (ϕ true false))) ⟩ --- thrdcomp (cong S¹→SuspBool (intLoop (negsuc n)) ∙ cong S¹→SuspBool (sym loop)) ≡⟨ cong thrdcomp (sym (congFunct S¹→SuspBool (intLoop (negsuc n)) (sym loop))) ⟩ --- looper2 (negsuc (suc n)) ∎ - - --- isolooper2 : Iso Int (Path (S₊ 1) north north) --- isolooper2 = compIso (iso intLoop winding (decodeEncode base) windingIntLoop) --- (compIso iso2 --- iso1) --- where --- iso1 : Iso (Path (Susp Bool) north north) (Path (S₊ 1) north north) --- iso1 = congIso SuspBool≃S1 - --- iso2 : Iso ΩS¹ (Path (Susp Bool) north north) --- iso2 = congIso (isoToEquiv (iso S¹→SuspBool SuspBool→S¹ SuspBool→S¹→SuspBool S¹→SuspBool→S¹)) - --- isolooper : Iso Int (Path (S₊ 1) north north) --- Iso.fun isolooper = looper --- Iso.inv isolooper = Iso.inv isolooper2 --- Iso.rightInv isolooper a = (looper≡looper2 (Iso.inv isolooper2 a)) ∙ Iso.rightInv isolooper2 a --- Iso.leftInv isolooper a = cong (Iso.inv isolooper2) (looper≡looper2 a) ∙ Iso.leftInv isolooper2 a - - --- {- We want to show that this map is an equivalence. n ≥ 2 follows from Freudenthal, and -} - --- {- --- We want to show that the function (looper : Int → S₊ 1) defined by λ k → loopᵏ is an equivalece. We already know that the corresponding function (intLoop : Int → S¹ is) an equivalence, --- so the idea is to show that when intLoop is transported along a suitable path S₊ 1 ≡ S¹ we get looper. Instead of using S₊ 1 straight away, we begin by showing this for the equivalent Susp Bool. --- -} - - - --- ----------------------------------- n = 1 ----------------------------------------------------- - --- {- We begin by stating some useful lemmas -} - - - --- S³≡SuspSuspS¹ : S³ ≡ Susp (Susp S¹) --- S³≡SuspSuspS¹ = S³≡SuspS² ∙ λ i → Susp (S²≡SuspS¹ i) - --- S3≡SuspSuspS¹ : S₊ 3 ≡ Susp (Susp S¹) --- S3≡SuspSuspS¹ = (λ i → Susp (Susp (Susp (ua Bool≃Susp⊥ (~ i))))) ∙ λ i → Susp (Susp (S¹≡SuspBool (~ i))) - --- sphereConnectedSpecCase : isHLevelConnected 4 (Susp (Susp S¹)) -- is- 2 -ConnectedType (Susp (Susp S¹)) --- sphereConnectedSpecCase = transport (λ i → isHLevelConnected 4 (S3≡SuspSuspS¹ i)) (sphereConnected 3) - - - --- {- We give the following map and show that it is an equivalence -} - --- d-map : typ (Ω ((Susp S¹) , north)) → S¹ --- d-map p = subst HopfSuspS¹ p base - --- d-mapId : (r : S¹) → d-map (ϕ base r) ≡ r --- d-mapId r = substComposite HopfSuspS¹ (merid r) (sym (merid base)) base ∙ --- rotLemma r --- where --- rotLemma : (r : S¹) → rot r base ≡ r --- rotLemma base = refl --- rotLemma (loop i) = refl - --- d-mapComp : fiber d-map base ≡ Path (Susp (Susp S¹)) north north --- d-mapComp = sym (pathSigma≡sigmaPath {B = HopfSuspS¹} _ _) ∙ helper --- where --- helper : Path (Σ (Susp S¹) λ x → HopfSuspS¹ x) (north , base) (north , base) ≡ Path (Susp (Susp S¹)) north north --- helper = (λ i → (Path (S³≡TotalHopf (~ i)) --- (transp (λ j → S³≡TotalHopf (~ i ∨ ~ j)) (~ i) (north , base)) --- ((transp (λ j → S³≡TotalHopf (~ i ∨ ~ j)) (~ i) (north , base))))) ∙ --- (λ i → Path (S³≡SuspSuspS¹ i) (transp (λ j → S³≡SuspSuspS¹ (i ∧ j)) (~ i) base) ((transp (λ j → S³≡SuspSuspS¹ (i ∧ j)) (~ i) base))) - - --- is1Connected-dmap : isHLevelConnectedFun 3 d-map --- is1Connected-dmap = toPropElim (λ s → isPropIsOfHLevel 0) (transport (λ j → isContr (∥ d-mapComp (~ j) ∥ ℕ→ℕ₋₂ 1)) --- (transport (λ i → isContr (PathΩ {A = Susp (Susp S¹)} {a = north} (ℕ→ℕ₋₂ 1) i)) --- (refl , isOfHLevelSuc 1 (isOfHLevelSuc 0 sphereConnectedSpecCase) ∣ north ∣ ∣ north ∣ (λ _ → ∣ north ∣)))) - --- d-iso2 : Iso (hLevelTrunc 3 (typ (Ω (Susp S¹ , north)))) (hLevelTrunc 3 S¹) --- Iso.fun d-iso2 = trMap d-map --- Iso.inv d-iso2 = Iso.inv (connectedTruncIso _ d-map is1Connected-dmap) --- Iso.rightInv d-iso2 = Iso.rightInv (connectedTruncIso _ d-map is1Connected-dmap) --- Iso.leftInv d-iso2 = Iso.leftInv (connectedTruncIso _ d-map is1Connected-dmap) - --- d-iso : isIso {A = ∥ typ (Ω (Susp S¹ , north)) ∥ (ℕ→ℕ₋₂ 1)} {B = ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (trElim (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣ ) --- d-iso = (Iso.inv (connectedTruncIso _ d-map is1Connected-dmap)) , (Iso.rightInv (connectedTruncIso _ d-map is1Connected-dmap) --- , Iso.leftInv (connectedTruncIso _ d-map is1Connected-dmap)) + {- To define the map for n=0 we use the λ k → loopᵏ map for S₊ 1. The loop is given by ϕ south north -} +private + loop* : Path (S₊ 1) north north + loop* = ϕ north south --- {- We show that composing (λ a → ∣ ϕ base a ∣) and (λ x → ∣ d-map x ∣) gives us the identity function. -} +looper : Int → Path (S₊ 1) north north +looper (pos zero) = refl +looper (pos (suc n)) = looper (pos n) ∙ loop* +looper (negsuc zero) = sym loop* +looper (negsuc (suc n)) = looper (negsuc n) ∙ sym loop* --- d-mapId2 : (λ (x : hLevelTrunc 3 S¹) → (trRec {n = 3} {B = hLevelTrunc 3 S¹} (isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) --- (trRec (isOfHLevelTrunc 3) (λ a → ∣ ϕ base a ∣) x)) ≡ λ x → x --- d-mapId2 = funExt (trElim (λ x → isOfHLevelSuc 2 (isOfHLevelTrunc 3 ((trElim {n = 3} --- {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} --- (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) --- (trElim (λ _ → isOfHLevelTrunc 3) --- (λ a → ∣ ϕ base a ∣) x)) x)) --- λ a i → ∣ d-mapId a i ∣) --- {- This means that (λ a → ∣ ϕ base a ∣) is an equivalence -} +private --- Iso∣ϕ-base∣ : Iso (hLevelTrunc 3 S¹) (hLevelTrunc 3 (typ (Ω (Susp S¹ , north)))) --- Iso∣ϕ-base∣ = composesToId→Iso d-iso2 (trMap (ϕ base)) d-mapId2 + {- The map of Kn≃ΩKn+1 is given as follows. -} + Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) + Kn→ΩKn+1 zero x = cong ∣_∣ (looper x) + Kn→ΩKn+1 (suc n) = trRec (isOfHLevelTrunc (2 + (suc (suc n))) ∣ north ∣ ∣ north ∣) + λ a → cong ∣_∣ ((merid a) ∙ (sym (merid north))) + + {- We show that looper is a composition of intLoop with two other maps, all three being isos -} + sndcomp : ΩS¹ → Path (Susp Bool) north north + sndcomp = cong S¹→SuspBool + + thrdcomp : Path (Susp Bool) north north → Path (S₊ 1) north north + thrdcomp = cong SuspBool→S1 + + + looper2 : Int → Path (S₊ 1) north north + looper2 a = thrdcomp (sndcomp (intLoop a)) + + looper≡looper2 : (x : Int) → looper x ≡ looper2 x + looper≡looper2 (pos zero) = refl + looper≡looper2 (pos (suc n)) = + looper (pos n) ∙ loop* ≡⟨ (λ i → looper≡looper2 (pos n) i ∙ congFunct SuspBool→S1 (merid false) (sym (merid true)) (~ i)) ⟩ + looper2 (pos n) ∙ cong SuspBool→S1 (ϕ true false) ≡⟨ sym (congFunct SuspBool→S1 (sndcomp (intLoop (pos n))) (ϕ true false)) ⟩ + cong SuspBool→S1 (sndcomp (intLoop (pos n)) ∙ ϕ true false) ≡⟨ cong thrdcomp (sym (congFunct S¹→SuspBool (intLoop (pos n)) loop)) ⟩ + looper2 (pos (suc n)) ∎ + looper≡looper2 (negsuc zero) = + sym loop* ≡⟨ symDistr (merid south) (sym (merid north)) ⟩ + merid north ∙ sym (merid south) ≡⟨ sym (congFunct SuspBool→S1 (merid true) (sym (merid false))) ⟩ + cong SuspBool→S1 (merid true ∙ sym (merid false)) ≡⟨ cong thrdcomp (sym (symDistr (merid false) (sym (merid true)))) ⟩ + looper2 (negsuc zero) ∎ + looper≡looper2 (negsuc (suc n)) = + looper (negsuc n) ∙ sym loop* ≡⟨ ((λ i → looper≡looper2 (negsuc n) i ∙ symDistr (merid south) (sym (merid north)) i)) ⟩ + looper2 (negsuc n) ∙ merid north ∙ sym (merid south) ≡⟨ cong (λ x → looper2 (negsuc n) ∙ x) (sym (congFunct SuspBool→S1 (merid true) (sym (merid false)))) ⟩ + looper2 (negsuc n) ∙ cong SuspBool→S1 (ϕ false true) ≡⟨ cong (λ x → looper2 (negsuc n) ∙ x) (cong thrdcomp (sym (symDistr (merid false) (sym (merid true))))) ⟩ + looper2 (negsuc n) ∙ cong SuspBool→S1 (sym (ϕ true false)) ≡⟨ sym (congFunct SuspBool→S1 (sndcomp (intLoop (negsuc n))) (sym (ϕ true false))) ⟩ + thrdcomp (cong S¹→SuspBool (intLoop (negsuc n)) ∙ cong S¹→SuspBool (sym loop)) ≡⟨ cong thrdcomp (sym (congFunct S¹→SuspBool (intLoop (negsuc n)) (sym loop))) ⟩ + looper2 (negsuc (suc n)) ∎ + + + isolooper2 : Iso Int (Path (S₊ 1) north north) + isolooper2 = compIso (iso intLoop winding (decodeEncode base) windingIntLoop) + (compIso iso2 + iso1) + where + iso1 : Iso (Path (Susp Bool) north north) (Path (S₊ 1) north north) + iso1 = congIso SuspBool≃S1 + iso2 : Iso ΩS¹ (Path (Susp Bool) north north) + iso2 = congIso (isoToEquiv (iso S¹→SuspBool SuspBool→S¹ SuspBool→S¹→SuspBool S¹→SuspBool→S¹)) --- --------------------------------- --- -- We cheat when n = 1 and use J to prove the following lemmma. There is an obvious dependent path between ϕ base and ϕ north. Since the first one is an iso, so is the other. --- -- + isolooper : Iso Int (Path (S₊ 1) north north) + Iso.fun isolooper = looper + Iso.inv isolooper = Iso.inv isolooper2 + Iso.rightInv isolooper a = (looper≡looper2 (Iso.inv isolooper2 a)) ∙ Iso.rightInv isolooper2 a + Iso.leftInv isolooper a = cong (Iso.inv isolooper2) (looper≡looper2 a) ∙ Iso.leftInv isolooper2 a --- pointFunIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ} {C : (A : Type ℓ) (a1 : A) → Type ℓ'} (p : A ≡ B) (a : A) (b : B) → --- (transport p a ≡ b) → --- (f : (A : Type ℓ) → --- (a1 : A) (a2 : hLevelTrunc 3 A) → C A a1) → --- isIso (f A a) → --- isIso (f B b) --- pointFunIso {ℓ = ℓ}{A = A} {B = B} {C = C} = --- J (λ B p → (a : A) (b : B) → --- (transport p a ≡ b) → --- (f : (A : Type ℓ) → --- (a1 : A) (a2 : hLevelTrunc 3 A) → C A a1) → --- isIso (f A a) → --- isIso (f B b)) --- λ a b trefl f is → transport (λ i → isIso (f A ((sym (transportRefl a) ∙ trefl) i))) is + {- We want to show that this map is an equivalence. n ≥ 2 follows from Freudenthal, and -} --- {- By pointFunIso, this gives that λ a → ∣ ϕ north a ∣ is an iso -} + {- + We want to show that the function (looper : Int → S₊ 1) defined by λ k → loopᵏ is an equivalece. We already know that the corresponding function (intLoop : Int → S¹ is) an equivalence, + so the idea is to show that when intLoop is transported along a suitable path S₊ 1 ≡ S¹ we get looper. Instead of using S₊ 1 straight away, we begin by showing this for the equivalent Susp Bool. + -} --- isIso∣ϕ∣ : isIso {A = hLevelTrunc 3 (S₊ 1)} {B = hLevelTrunc 3 (typ (Ω (S₊ 2 , north)))} (trMap (ϕ north)) --- isIso∣ϕ∣ = pointFunIso {A = S¹} (λ i → S¹≡S1 (~ i)) base north refl (λ A a1 → trMap (ϕ a1)) ((Iso.inv Iso∣ϕ-base∣) , ((Iso.rightInv Iso∣ϕ-base∣) , (Iso.leftInv Iso∣ϕ-base∣))) --- Iso∣ϕ∣ : Iso (hLevelTrunc 3 (S₊ 1)) (hLevelTrunc 3 (typ (Ω (S₊ 2 , north)))) --- Iso.fun Iso∣ϕ∣ = trMap (ϕ north) --- Iso.inv Iso∣ϕ∣ = isIso∣ϕ∣ .fst --- Iso.rightInv Iso∣ϕ∣ = isIso∣ϕ∣ .snd .fst --- Iso.leftInv Iso∣ϕ∣ = isIso∣ϕ∣ .snd .snd --- -- ---------------------------------------------------- Finishing up --------------------------------- + ----------------------------------- n = 1 ----------------------------------------------------- --- -- We need ΩTrunc. It computes horribly, so its better to restate the function involved for this particular case -- + {- We begin by stating some useful lemmas -} --- decode-fun2 : (n : ℕ) → (x : A) → hLevelTrunc n (x ≡ x) → Path (hLevelTrunc (suc n) A) ∣ x ∣ ∣ x ∣ --- decode-fun2 zero x = trElim (λ _ → isOfHLevelPath 0 (∣ x ∣ , isOfHLevelTrunc 1 ∣ x ∣) ∣ x ∣ ∣ x ∣) (λ p i → ∣ p i ∣) --- decode-fun2 (suc n) x = trElim (λ _ → isOfHLevelPath' (suc n) (isOfHLevelTrunc (suc (suc n))) ∣ x ∣ ∣ x ∣) (cong ∣_∣) --- -- {- auxiliary function r used to define encode -} --- -- r : {m : ℕ₋₂} (u : ∥ B ∥ (suc₋₂ m)) → P u u --- -- r = elim (λ x → hLevelP x x) (λ a → ∣ refl ∣) - --- -- {- encode function from x ≡ y to P x y -} --- -- encode-fun : {n : ℕ₋₂} (x y : ∥ B ∥ (suc₋₂ n)) → x ≡ y → P x y --- -- encode-fun x y p = transport (λ i → P x (p i)) (r x) --- {- --- encode2 : (n : ℕ) (x : A) → Path (hLevelTrunc (suc n) A) ∣ x ∣ ∣ x ∣ → hLevelTrunc n (x ≡ x) --- encode2 {A = A} n x p = transport (λ i → hLevelTrunc n (x ≡ {!p i!})) {!!} --- where --- r : (m : ℕ) (u : hLevelTrunc (1 + m) A) → hLevelTrunc m (u ≡ u) --- r m = trElim (λ x → isOfHLevelSuc m (isOfHLevelTrunc _) ) λ a → ∣ refl ∣ --- -} - --- funsAreSame : (n : ℕ) (x : A) (b : hLevelTrunc n (x ≡ x)) → (decode-fun2 n x b) ≡ (ΩTrunc.decode-fun ∣ x ∣ ∣ x ∣ b) --- funsAreSame zero x = trElim (λ a → isOfHLevelPath 0 (refl , (isOfHLevelSuc 1 (isOfHLevelTrunc 1) ∣ x ∣ ∣ x ∣ refl)) _ _) λ a → refl --- funsAreSame (suc n) x = trElim (λ a → isOfHLevelPath _ (isOfHLevelPath' (suc n) (isOfHLevelTrunc (suc (suc n))) ∣ x ∣ ∣ x ∣) _ _) λ a → refl - --- decodeIso : (n : ℕ) (x : A) → Iso (hLevelTrunc n (x ≡ x)) (Path (hLevelTrunc (suc n) A) ∣ x ∣ ∣ x ∣) --- Iso.fun (decodeIso n x) = decode-fun2 n x --- Iso.inv (decodeIso n x) = ΩTrunc.encode-fun ∣ x ∣ ∣ x ∣ --- Iso.rightInv (decodeIso n x) b = funExt⁻ (funExt (funsAreSame n x)) (ΩTrunc.encode-fun ∣ x ∣ ∣ x ∣ b) ∙ ΩTrunc.P-rinv ∣ x ∣ ∣ x ∣ b --- Iso.leftInv (decodeIso n x) b = cong (ΩTrunc.encode-fun ∣ x ∣ ∣ x ∣) (funsAreSame n x b) ∙ ΩTrunc.P-linv ∣ x ∣ ∣ x ∣ b + S³≡SuspSuspS¹ : S³ ≡ Susp (Susp S¹) + S³≡SuspSuspS¹ = S³≡SuspS² ∙ λ i → Susp (S²≡SuspS¹ i) + S3≡SuspSuspS¹ : S₊ 3 ≡ Susp (Susp S¹) + S3≡SuspSuspS¹ = (λ i → Susp (Susp (Susp (ua Bool≃Susp⊥ (~ i))))) ∙ λ i → Susp (Susp (S¹≡SuspBool (~ i))) --- Iso-Kn-ΩKn+1 : (n : ℕ) → Iso (coHomK n) (typ (Ω (coHomK-ptd (suc n)))) --- Iso-Kn-ΩKn+1 zero = compIso isolooper (congIso (truncIdempotent≃ _ isOfHLevelS1)) --- Iso-Kn-ΩKn+1 (suc zero) = compIso Iso∣ϕ∣ (decodeIso _ north) --- Iso-Kn-ΩKn+1 (suc (suc n)) = compIso (connectedTruncIso2 (4 + n) _ (ϕ north) (n , (λ i → suc (suc (suc (+-suc n n (~ i))))) ∙ --- (λ i → suc (suc (+-suc n (suc n) ((~ i)))))) --- (isConnectedσ (suc n) (sphereConnected _))) --- (decodeIso _ north) - --- mapId2 : (n : ℕ) → Kn→ΩKn+1 n ≡ Iso.fun (Iso-Kn-ΩKn+1 n) --- mapId2 zero = refl --- mapId2 (suc zero) = funExt (trElim (λ x → isOfHLevelPath 3 (isOfHLevelTrunc 4 ∣ north ∣ ∣ north ∣) _ _) λ a → refl) --- mapId2 (suc (suc n)) = funExt (trElim (λ x → isOfHLevelPath (4 + n) (isOfHLevelTrunc (5 + n) ∣ north ∣ ∣ north ∣) _ _) λ a → refl) + sphereConnectedSpecCase : isHLevelConnected 4 (Susp (Susp S¹)) -- is- 2 -ConnectedType (Susp (Susp S¹)) + sphereConnectedSpecCase = transport (λ i → isHLevelConnected 4 (S3≡SuspSuspS¹ i)) (sphereConnected 3) --- Iso2-Kn-ΩKn+1 : (n : ℕ) → Iso (coHomK n) (typ (Ω (coHomK-ptd (suc n)))) --- Iso.fun (Iso2-Kn-ΩKn+1 n) = Kn→ΩKn+1 n --- Iso.inv (Iso2-Kn-ΩKn+1 n) = Iso.inv (Iso-Kn-ΩKn+1 n) --- Iso.rightInv (Iso2-Kn-ΩKn+1 n) a = rinv --- where --- abstract --- rinv : Kn→ΩKn+1 n (Iso.inv (Iso-Kn-ΩKn+1 n) a) ≡ a --- rinv = funExt⁻ (mapId2 n) _ ∙ Iso.rightInv (Iso-Kn-ΩKn+1 n) a --- Iso.leftInv (Iso2-Kn-ΩKn+1 n) a = linv --- where --- abstract --- linv : Iso.inv (Iso-Kn-ΩKn+1 n) (Kn→ΩKn+1 n a) ≡ a --- linv = cong (Iso.inv (Iso-Kn-ΩKn+1 n)) (funExt⁻ (mapId2 n) a) ∙ Iso.leftInv (Iso-Kn-ΩKn+1 n) a + {- We give the following map and show that it is an equivalence -} + d-map : typ (Ω ((Susp S¹) , north)) → S¹ + d-map p = subst HopfSuspS¹ p base + d-mapId : (r : S¹) → d-map (ϕ base r) ≡ r + d-mapId r = substComposite HopfSuspS¹ (merid r) (sym (merid base)) base ∙ + rotLemma r + where + rotLemma : (r : S¹) → rot r base ≡ r + rotLemma base = refl + rotLemma (loop i) = refl + d-mapComp : fiber d-map base ≡ Path (Susp (Susp S¹)) north north + d-mapComp = sym (pathSigma≡sigmaPath {B = HopfSuspS¹} _ _) ∙ helper + where + helper : Path (Σ (Susp S¹) λ x → HopfSuspS¹ x) (north , base) (north , base) ≡ Path (Susp (Susp S¹)) north north + helper = (λ i → (Path (S³≡TotalHopf (~ i)) + (transp (λ j → S³≡TotalHopf (~ i ∨ ~ j)) (~ i) (north , base)) + ((transp (λ j → S³≡TotalHopf (~ i ∨ ~ j)) (~ i) (north , base))))) ∙ + (λ i → Path (S³≡SuspSuspS¹ i) (transp (λ j → S³≡SuspSuspS¹ (i ∧ j)) (~ i) base) ((transp (λ j → S³≡SuspSuspS¹ (i ∧ j)) (~ i) base))) + + + is1Connected-dmap : isHLevelConnectedFun 3 d-map + is1Connected-dmap = toPropElim (λ s → isPropIsOfHLevel 0) (transport (λ j → isContr (∥ d-mapComp (~ j) ∥ ℕ→ℕ₋₂ 1)) + (transport (λ i → isContr (PathΩ {A = Susp (Susp S¹)} {a = north} (ℕ→ℕ₋₂ 1) i)) + (refl , isOfHLevelSuc 1 (isOfHLevelSuc 0 sphereConnectedSpecCase) ∣ north ∣ ∣ north ∣ (λ _ → ∣ north ∣)))) + + d-iso2 : Iso (hLevelTrunc 3 (typ (Ω (Susp S¹ , north)))) (hLevelTrunc 3 S¹) + Iso.fun d-iso2 = trMap d-map + Iso.inv d-iso2 = Iso.inv (connectedTruncIso _ d-map is1Connected-dmap) + Iso.rightInv d-iso2 = Iso.rightInv (connectedTruncIso _ d-map is1Connected-dmap) + Iso.leftInv d-iso2 = Iso.leftInv (connectedTruncIso _ d-map is1Connected-dmap) + + d-iso : isIso {A = ∥ typ (Ω (Susp S¹ , north)) ∥ (ℕ→ℕ₋₂ 1)} {B = ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (trElim (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣ ) + d-iso = (Iso.inv (connectedTruncIso _ d-map is1Connected-dmap)) , (Iso.rightInv (connectedTruncIso _ d-map is1Connected-dmap) + , Iso.leftInv (connectedTruncIso _ d-map is1Connected-dmap)) + + {- We show that composing (λ a → ∣ ϕ base a ∣) and (λ x → ∣ d-map x ∣) gives us the identity function. -} + + d-mapId2 : (λ (x : hLevelTrunc 3 S¹) → (trRec {n = 3} {B = hLevelTrunc 3 S¹} (isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) + (trRec (isOfHLevelTrunc 3) (λ a → ∣ ϕ base a ∣) x)) ≡ λ x → x + d-mapId2 = funExt (trElim (λ x → isOfHLevelSuc 2 (isOfHLevelTrunc 3 ((trElim {n = 3} + {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} + (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) + (trElim (λ _ → isOfHLevelTrunc 3) + (λ a → ∣ ϕ base a ∣) x)) x)) + λ a i → ∣ d-mapId a i ∣) + + {- This means that (λ a → ∣ ϕ base a ∣) is an equivalence -} + + + Iso∣ϕ-base∣ : Iso (hLevelTrunc 3 S¹) (hLevelTrunc 3 (typ (Ω (Susp S¹ , north)))) + Iso∣ϕ-base∣ = composesToId→Iso d-iso2 (trMap (ϕ base)) d-mapId2 + + + --------------------------------- + -- We cheat when n = 1 and use J to prove the following lemmma. There is an obvious dependent path between ϕ base and ϕ north. Since the first one is an iso, so is the other. + -- + + + pointFunIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ} {C : (A : Type ℓ) (a1 : A) → Type ℓ'} (p : A ≡ B) (a : A) (b : B) → + (transport p a ≡ b) → + (f : (A : Type ℓ) → + (a1 : A) (a2 : hLevelTrunc 3 A) → C A a1) → + isIso (f A a) → + isIso (f B b) + pointFunIso {ℓ = ℓ}{A = A} {B = B} {C = C} = + J (λ B p → (a : A) (b : B) → + (transport p a ≡ b) → + (f : (A : Type ℓ) → + (a1 : A) (a2 : hLevelTrunc 3 A) → C A a1) → + isIso (f A a) → + isIso (f B b)) + λ a b trefl f is → transport (λ i → isIso (f A ((sym (transportRefl a) ∙ trefl) i))) is + + {- By pointFunIso, this gives that λ a → ∣ ϕ north a ∣ is an iso -} + + isIso∣ϕ∣ : isIso {A = hLevelTrunc 3 (S₊ 1)} {B = hLevelTrunc 3 (typ (Ω (S₊ 2 , north)))} (trMap (ϕ north)) + isIso∣ϕ∣ = pointFunIso {A = S¹} (λ i → S¹≡S1 (~ i)) base north refl (λ A a1 → trMap (ϕ a1)) ((Iso.inv Iso∣ϕ-base∣) , ((Iso.rightInv Iso∣ϕ-base∣) , (Iso.leftInv Iso∣ϕ-base∣))) + + Iso∣ϕ∣ : Iso (hLevelTrunc 3 (S₊ 1)) (hLevelTrunc 3 (typ (Ω (S₊ 2 , north)))) + Iso.fun Iso∣ϕ∣ = trMap (ϕ north) + Iso.inv Iso∣ϕ∣ = isIso∣ϕ∣ .fst + Iso.rightInv Iso∣ϕ∣ = isIso∣ϕ∣ .snd .fst + Iso.leftInv Iso∣ϕ∣ = isIso∣ϕ∣ .snd .snd + +-- ---------------------------------------------------- Finishing up --------------------------------- + +-- We need ΩTrunc. It appears to compute better when restated for this particular case -- + +decode-fun2 : (n : ℕ) → (x : A) → hLevelTrunc n (x ≡ x) → Path (hLevelTrunc (suc n) A) ∣ x ∣ ∣ x ∣ +decode-fun2 zero x = trElim (λ _ → isOfHLevelPath 0 (∣ x ∣ , isOfHLevelTrunc 1 ∣ x ∣) ∣ x ∣ ∣ x ∣) (λ p i → ∣ p i ∣) +decode-fun2 (suc n) x = trElim (λ _ → isOfHLevelPath' (suc n) (isOfHLevelTrunc (suc (suc n))) ∣ x ∣ ∣ x ∣) (cong ∣_∣) + +funsAreSame : (n : ℕ) (x : A) (b : hLevelTrunc n (x ≡ x)) → (decode-fun2 n x b) ≡ (ΩTrunc.decode-fun ∣ x ∣ ∣ x ∣ b) +funsAreSame zero x = trElim (λ a → isOfHLevelPath 0 (refl , (isOfHLevelSuc 1 (isOfHLevelTrunc 1) ∣ x ∣ ∣ x ∣ refl)) _ _) λ a → refl +funsAreSame (suc n) x = trElim (λ a → isOfHLevelPath _ (isOfHLevelPath' (suc n) (isOfHLevelTrunc (suc (suc n))) ∣ x ∣ ∣ x ∣) _ _) λ a → refl + +decodeIso : (n : ℕ) (x : A) → Iso (hLevelTrunc n (x ≡ x)) (Path (hLevelTrunc (suc n) A) ∣ x ∣ ∣ x ∣) +Iso.fun (decodeIso n x) = decode-fun2 n x +Iso.inv (decodeIso n x) = ΩTrunc.encode-fun ∣ x ∣ ∣ x ∣ +Iso.rightInv (decodeIso n x) b = funExt⁻ (funExt (funsAreSame n x)) (ΩTrunc.encode-fun ∣ x ∣ ∣ x ∣ b) ∙ ΩTrunc.P-rinv ∣ x ∣ ∣ x ∣ b +Iso.leftInv (decodeIso n x) b = cong (ΩTrunc.encode-fun ∣ x ∣ ∣ x ∣) (funsAreSame n x b) ∙ ΩTrunc.P-linv ∣ x ∣ ∣ x ∣ b + +Iso-Kn-ΩKn+1 : (n : ℕ) → Iso (coHomK n) (typ (Ω (coHomK-ptd (suc n)))) +Iso-Kn-ΩKn+1 zero = compIso isolooper (congIso (truncIdempotent≃ _ isOfHLevelS1)) +Iso-Kn-ΩKn+1 (suc zero) = compIso Iso∣ϕ∣ (decodeIso _ north) +Iso-Kn-ΩKn+1 (suc (suc n)) = compIso (connectedTruncIso2 (4 + n) _ (ϕ north) (n , (λ i → suc (suc (suc (+-suc n n (~ i))))) ∙ + (λ i → suc (suc (+-suc n (suc n) ((~ i)))))) + (isConnectedσ (suc n) (sphereConnected _))) + (decodeIso _ north) + +mapId2 : (n : ℕ) → Kn→ΩKn+1 n ≡ Iso.fun (Iso-Kn-ΩKn+1 n) +mapId2 zero = refl +mapId2 (suc zero) = funExt (trElim (λ x → isOfHLevelPath 3 (isOfHLevelTrunc 4 ∣ north ∣ ∣ north ∣) _ _) λ a → refl) +mapId2 (suc (suc n)) = funExt (trElim (λ x → isOfHLevelPath (4 + n) (isOfHLevelTrunc (5 + n) ∣ north ∣ ∣ north ∣) _ _) λ a → refl) + +{- This version computes somewhat better -} + +Iso2-Kn-ΩKn+1 : (n : ℕ) → Iso (coHomK n) (typ (Ω (coHomK-ptd (suc n)))) +Iso.fun (Iso2-Kn-ΩKn+1 n) = Kn→ΩKn+1 n +Iso.inv (Iso2-Kn-ΩKn+1 n) = Iso.inv (Iso-Kn-ΩKn+1 n) +Iso.rightInv (Iso2-Kn-ΩKn+1 n) a = rinv + where + abstract + rinv : Kn→ΩKn+1 n (Iso.inv (Iso-Kn-ΩKn+1 n) a) ≡ a + rinv = funExt⁻ (mapId2 n) _ ∙ Iso.rightInv (Iso-Kn-ΩKn+1 n) a +Iso.leftInv (Iso2-Kn-ΩKn+1 n) a = linv + where + abstract + linv : Iso.inv (Iso-Kn-ΩKn+1 n) (Kn→ΩKn+1 n a) ≡ a + linv = cong (Iso.inv (Iso-Kn-ΩKn+1 n)) (funExt⁻ (mapId2 n) a) ∙ Iso.leftInv (Iso-Kn-ΩKn+1 n) a From b75a597fcfa3da484ef803283db857a20d4fb152 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 4 May 2020 00:03:02 +0200 Subject: [PATCH 24/89] whitespace etc --- Cubical/Foundations/Equiv.agda | 20 +- Cubical/Foundations/Equiv/Properties.agda | 7 +- Cubical/Foundations/Isomorphism.agda | 7 + Cubical/HITs/SmashProduct/Base.agda | 14 +- Cubical/HITs/Sn/Properties.agda | 250 +--------- Cubical/Homotopy/Connected.agda | 527 +++------------------- Cubical/Homotopy/Freudenthal.agda | 2 +- Cubical/Homotopy/Loopspace.agda | 6 +- Cubical/ZCohomology/MayerVietoris.agda | 212 --------- Cubical/ZCohomology/Properties.agda | 176 +------- Cubical/ZCohomology/cupProd.agda | 0 11 files changed, 106 insertions(+), 1115 deletions(-) delete mode 100644 Cubical/ZCohomology/MayerVietoris.agda delete mode 100644 Cubical/ZCohomology/cupProd.agda diff --git a/Cubical/Foundations/Equiv.agda b/Cubical/Foundations/Equiv.agda index 07a1a6f2d..b301c969a 100644 --- a/Cubical/Foundations/Equiv.agda +++ b/Cubical/Foundations/Equiv.agda @@ -158,15 +158,11 @@ infix 1 _■ composesToId→Equiv : (f : A → B) (g : B → A) → f ∘ g ≡ idfun B → isEquiv f → isEquiv g composesToId→Equiv f g id iseqf = - isoToIsEquiv (iso g - f - (λ b → (λ i → (equiv-proof iseqf (f b) .snd (g (f b) , cong (λ h → h (f b)) id) (~ i)) .fst ) ∙ - cong (λ x → (equiv-proof iseqf (f b) .fst .fst )) id ∙ - λ i → (equiv-proof iseqf (f b) .snd) (b , refl) i .fst ) - λ a → cong (λ f → f a) id) - -composesToId→Iso : (G : Iso A B) → (g : B → A) → (Iso.fun G) ∘ g ≡ idfun B → Iso B A -Iso.fun (composesToId→Iso (iso fun inv rightInv leftInv) g path) = g -Iso.inv (composesToId→Iso (iso fun inv rightInv leftInv) g path) = fun -Iso.rightInv (composesToId→Iso (iso fun inv rightInv leftInv) g path) b = (sym (leftInv (g (fun b))) ∙ cong (λ x → inv x) (cong (λ f → f (fun b)) path)) ∙ leftInv b -Iso.leftInv (composesToId→Iso (iso fun inv rightInv leftInv) g path) b = cong (λ f → f b) path + isoToIsEquiv + (iso g + f + (λ b → (λ i → (equiv-proof iseqf (f b) .snd (g (f b) , + cong (λ h → h (f b)) id) (~ i)) .fst ) + ∙ cong (λ x → (equiv-proof iseqf (f b) .fst .fst )) id + ∙ λ i → (equiv-proof iseqf (f b) .snd) (b , refl) i .fst ) + λ a → cong (λ f → f a) id) diff --git a/Cubical/Foundations/Equiv/Properties.agda b/Cubical/Foundations/Equiv/Properties.agda index c1229c957..3e61d76d2 100644 --- a/Cubical/Foundations/Equiv/Properties.agda +++ b/Cubical/Foundations/Equiv/Properties.agda @@ -57,13 +57,18 @@ postCompEquiv : ∀ {ℓ ℓ′} {A B : Type ℓ} {C : Type ℓ′} (e : A ≃ B → (C → A) ≃ (C → B) postCompEquiv e = (λ φ x → fst e (φ x)) , isEquivPostComp e - hasSection : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → (A → B) → Type _ hasSection {A = A} {B} f = Σ[ g ∈ (B → A) ] section f g hasRetract : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → (A → B) → Type _ hasRetract {A = A} {B} f = Σ[ g ∈ (B → A) ] retract f g +isEquiv-hasSection : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) + → isEquiv f + → hasSection f +isEquiv-hasSection f isEq = (λ b → equiv-proof isEq b .fst .fst) , + λ b → equiv-proof isEq b .fst .snd + isContr-hasSection : ∀ {ℓ} {A B : Type ℓ} (e : A ≃ B) → isContr (hasSection (fst e)) fst (isContr-hasSection e) = invEq e , retEq e snd (isContr-hasSection e) (f , ε) i = (λ b → fst (p b i)) , (λ b → snd (p b i)) diff --git a/Cubical/Foundations/Isomorphism.agda b/Cubical/Foundations/Isomorphism.agda index a911831be..5b107079d 100644 --- a/Cubical/Foundations/Isomorphism.agda +++ b/Cubical/Foundations/Isomorphism.agda @@ -110,3 +110,10 @@ compIso (iso fun inv rightInv leftInv) (iso fun₁ inv₁ rightInv₁ leftInv₁ = cong fun₁ (rightInv (inv₁ b)) ∙ (rightInv₁ b) compIso (iso fun inv rightInv leftInv) (iso fun₁ inv₁ rightInv₁ leftInv₁) .Iso.leftInv a = cong inv (leftInv₁ (fun a) ) ∙ leftInv a + +composesToId→Iso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (G : Iso A B) → (g : B → A) → (Iso.fun G) ∘ g ≡ idfun B → Iso B A +Iso.fun (composesToId→Iso (iso fun inv rightInv leftInv) g path) = g +Iso.inv (composesToId→Iso (iso fun inv rightInv leftInv) g path) = fun +Iso.rightInv (composesToId→Iso (iso fun inv rightInv leftInv) g path) b = + (sym (leftInv (g (fun b))) ∙ cong (λ x → inv x) (cong (λ f → f (fun b)) path)) ∙ leftInv b +Iso.leftInv (composesToId→Iso (iso fun inv rightInv leftInv) g path) b = cong (λ f → f b) path diff --git a/Cubical/HITs/SmashProduct/Base.agda b/Cubical/HITs/SmashProduct/Base.agda index ab347e972..bee90b783 100644 --- a/Cubical/HITs/SmashProduct/Base.agda +++ b/Cubical/HITs/SmashProduct/Base.agda @@ -104,7 +104,7 @@ Smash→⋀ (proj x y) = inr (x , y) Smash→⋀ (gluel a i) = push (inl a) (~ i) Smash→⋀ (gluer b i) = push (inr b) (~ i) - +{- associativity maps for smash produts. Proof pretty much direcly translated from https://github.com/ecavallo/redtt/blob/master/library/pointed/smash.red -} private pivotl : (b b' : typ B) → Path (Smash A B) (proj (snd A) b) (proj (snd A) b') @@ -180,12 +180,12 @@ private top-filler2 : I → I → Smash (Smash∙ C B) A top-filler2 i j = hcomp (λ k → λ { (j = i0) → gluer (snd A) (i ∧ k) - ; (j = i1) → gluel (gluer b (~ k)) i}) - (hcomp (λ k → λ { (j = i0) → gluel (gluer (snd B) i0) (~ k ∧ (~ i)) - ; (j = i1) → gluel (baser) (~ k ∨ i) - ; (i = i0) → gluel (gluer (snd B) j) (~ k) - ; (i = i1) → gluel (proj (snd C) (snd B)) j }) - (gluel (proj (snd C) (snd B)) (j ∨ (~ i)))) + ; (j = i1) → gluel (gluer b (~ k)) i}) + (hcomp (λ k → λ { (j = i0) → gluel (gluer (snd B) i0) (~ k ∧ (~ i)) + ; (j = i1) → gluel (baser) (~ k ∨ i) + ; (i = i0) → gluel (gluer (snd B) j) (~ k) + ; (i = i1) → gluel (proj (snd C) (snd B)) j }) + (gluel (proj (snd C) (snd B)) (j ∨ (~ i)))) rearrange : Smash (Smash∙ A B) C → Smash (Smash∙ C B) A rearrange basel = basel diff --git a/Cubical/HITs/Sn/Properties.agda b/Cubical/HITs/Sn/Properties.agda index 442dc988a..a8d360420 100644 --- a/Cubical/HITs/Sn/Properties.agda +++ b/Cubical/HITs/Sn/Properties.agda @@ -67,7 +67,7 @@ SuspBool≃S1 = isoToEquiv iso1 Iso.leftInv iso1 (merid true i) = refl --- map between S₊ +-- map between S¹ ∧ A and Susp A. private f' : {a : A} → A → S₊ 1 → Susp A f' {a = pt} A north = north @@ -93,6 +93,9 @@ module smashS1→susp {(A , pt) : Pointed ℓ} where fun⁻ (merid a i) = (push (inr a) ∙∙ cong (λ x → proj' {A = S₊ 1 , north} {B = A , pt} x a) (merid south ∙ sym (merid north)) ∙∙ sym (push (inr a))) i + {- TODO : Prove that they cancel out -} + +{- Map used in definition of cup product. Maybe mover there later -} sphereSmashMap : (n m : ℕ) → (S₊ (suc n) , north) ⋀ (S₊ (suc m) , north) → S₊ (2 + n + m) sphereSmashMap zero m = smashS1→susp.fun sphereSmashMap (suc n) m = @@ -100,248 +103,3 @@ sphereSmashMap (suc n) m = (idfun∙ _ ⋀⃗ (sphereSmashMap n m , refl)) ∘ ⋀-associate ∘ ((smashS1→susp.fun⁻ , refl) ⋀⃗ idfun∙ _) - --- private --- f* : {a : A} → S¹ → A → Susp A --- f* base b = north --- f* {a = a} (loop i) = funExt (λ x → merid x ∙ sym (merid a)) i - --- f' : {a : A} → Susp Bool → A → Susp A --- f' north a = north --- f' south a = north --- f' {a = a} (merid p i) = funExt (λ x → merid x ∙ sym (merid a)) i - - --- -- proj : {A : Pointed ℓ} {B : Pointed ℓ'} → typ A → typ B → A smash B --- -- proj a b = inr (a , b) - --- -- projᵣ : {A : Pointed ℓ} {B : Pointed ℓ'} (a : typ A) → proj {A = A} a (pt B) ≡ inl tt --- -- projᵣ a = sym (push (inl a)) - --- -- projₗ : {A : Pointed ℓ} {B : Pointed ℓ'} (b : typ B) → proj {B = B} (pt A) b ≡ inl tt --- -- projₗ b = sym (push (inr b)) - --- -- projᵣₗ : {A : Pointed ℓ} {B : Pointed ℓ'} → projᵣ (pt A) ≡ projₗ (pt B) --- -- projᵣₗ {A = A} {B = B} i j = push (push tt i) (~ j) -- cong sym (cong push (push tt)) - - --- -- compLrFiller : {a b c d : A} (p : a ≡ b) (q : b ≡ c) (s : c ≡ d) → PathP (λ i → p i ≡ s (~ i)) (p ∙ q ∙ s) q --- -- compLrFiller {A = A} {a = a} {b = b} {c = c} {d = d} p q s i j = filler j i1 i --- -- where --- -- rhs-filler : I → I → I → A --- -- rhs-filler i j z = hfill (λ k → λ {(i = i0) → b ; --- -- (i = i1) → s (~ z ∧ k) ; --- -- (z = i1) → q i}) --- -- (inS (q i)) j --- -- filler : I → I → I → A --- -- filler i j z = hfill (λ k → λ {(i = i0) → p z ; --- -- (i = i1) → rhs-filler k i1 z ; --- -- (z = i1) → q (k ∧ i)}) --- -- (inS (p (i ∨ z))) --- -- j - - --- -- module S1-smash-Iso ((A , pt) : Pointed ℓ) where - --- -- funInv : Susp A → (S¹ , base) smash (A , pt) --- -- funInv north = inl tt --- -- funInv south = inl tt --- -- funInv (merid a i) = (sym (projₗ a) ∙ --- -- cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ --- -- projₗ a) i - - - --- -- fun : (S¹ , base) smash (A , pt) → (Susp A) --- -- fun (inl tt) = north --- -- fun (inr (x , x₁)) = f* {a = pt} x x₁ --- -- fun (push (inl base) i) = north --- -- fun (push (inl (loop i₁)) i) = rCancel (merid pt) (~ i) i₁ -- rCancel (merid pt) (~ i) i₁ --- -- fun (push (inr x) i) = north --- -- fun (push (push a i₁) i) = north - --- -- test18 : cong (cong funInv) (rCancel (merid pt)) ≡ congFunct funInv (merid pt) (sym (merid pt)) ∙ rCancel (cong funInv (merid pt)) -- congFunct {!cong!} {!!} {!!} ∙ rCancel {!!} --- -- test18 = {!!} - --- -- funSect : (x : Susp A) → fun (funInv x) ≡ x --- -- funSect north = refl --- -- funSect south = merid pt --- -- funSect (merid a i) = hcomp (λ k → λ {(i = i0) → ((λ j → rCancel refl j ∙ refl ∙ refl) ∙ --- -- λ j → lUnit (lUnit refl (~ j)) (~ j)) k ; --- -- (i = i1) → ((λ j → lUnit refl (~ j) ∙ sym (lUnit (rUnit (merid a ∙ sym (merid pt)) (~ j)) (~ j)) ∙ merid a) ∙ --- -- (λ j → lUnit (symDistr (merid a) (sym (merid pt)) j ∙ merid a) (~ j)) ∙ --- -- (sym (assoc (merid pt) (sym (merid a)) (merid a))) ∙ --- -- (λ j → merid pt ∙ lCancel (merid a) j) ∙ --- -- sym (rUnit (merid pt)) ) k}) --- -- (helper2 i ∙ filler i) --- -- where - --- -- filler : (i : I) → (refl ∙ (merid a ∙ sym (merid pt)) ∙ refl) i ≡ merid a i --- -- filler i = (λ j → (refl ∙ (merid a ∙ sym (merid pt)) ∙ refl) (i ∧ (~ j))) ∙ λ j → merid a (j ∧ i) - --- -- helper2 : (i : I) → fun ((sym (projₗ a) ∙ --- -- cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ --- -- projₗ a) i) --- -- ≡ ((λ i → fun (sym (projₗ a) i)) ∙ --- -- (λ i → fun (cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop i)) ∙ --- -- λ i → fun (projₗ a i)) i --- -- helper2 i = (λ j → congFunct fun (sym (projₗ a)) (cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ projₗ a) j i) ∙ --- -- (λ j → (cong fun (sym (projₗ a)) ∙ congFunct fun (cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop) (projₗ a) j) i) - --- -- retrHelper : (x : S¹) (a : A) → funInv (f* {a = pt} x a) ≡ proj x a --- -- retrHelper base a = sym (projₗ a) --- -- retrHelper (loop i) a j = hcomp (λ k → λ { (i = i0) → push (inr a) j --- -- ; (i = i1) → push (inr a) j --- -- ; (j = i0) → id3 (~ k) i --- -- ; (j = i1) → proj (loop i) a}) --- -- (compLrFiller (push (inr a)) (λ i → proj (loop i) a) (sym (push (inr a))) j i) - - --- -- {- hcomp (λ k → λ { (i = i0) → ((λ j → id3 j ∙ sym (projₗ a) ∙ λ j → proj (loop (~ j)) a) ∙ --- -- invCancelHelper _ _) k -- j --- -- ; (i = i1) → lUnit (rUnit (sym (projₗ a)) (~ k)) (~ k) -- j --- -- -- ; (j = i0) → funInv (f* {a = pt} (loop i) a) --- -- -- ; (j = i1) → proj (loop i) a --- -- } ) --- -- (filler) -} - --- -- module _ where - --- -- invCancelHelper : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) (q : y ≡ y) → (p ∙ q ∙ sym p) ∙ p ∙ sym q ≡ p --- -- invCancelHelper p q = (p ∙ q ∙ sym p) ∙ p ∙ sym q ≡⟨ (λ i → assoc p q (sym p) i ∙ p ∙ sym q) ⟩ --- -- ((p ∙ q) ∙ sym p) ∙ p ∙ sym q ≡⟨ assoc ((p ∙ q) ∙ sym p) p (sym q) ⟩ --- -- (((p ∙ q) ∙ sym p) ∙ p) ∙ sym q ≡⟨ (λ i → assoc (p ∙ q) (sym p) p (~ i) ∙ sym q) ⟩ --- -- ((p ∙ q) ∙ sym p ∙ p) ∙ sym q ≡⟨ (λ i → ((p ∙ q) ∙ lCancel p i) ∙ sym q) ⟩ --- -- ((p ∙ q) ∙ refl) ∙ sym q ≡⟨ (λ i → rUnit (p ∙ q) (~ i) ∙ sym q) ⟩ --- -- (p ∙ q) ∙ sym q ≡⟨ sym (assoc p q (sym q)) ⟩ --- -- p ∙ q ∙ sym q ≡⟨ (λ i → p ∙ rCancel q i) ⟩ --- -- p ∙ refl ≡⟨ sym (rUnit p) ⟩ --- -- p ∎ - --- -- filler : funInv (f* {a = pt} (loop i) a) ≡ proj (loop i) a --- -- filler = (λ j → funInv (funExt (λ x → merid x ∙ sym (merid pt)) (i ∨ j) a)) ∙ sym (projₗ a) ∙ λ j → proj (loop (i ∨ (~ j))) a - --- -- id1 : sym (projᵣ base) ∙ cong ((λ x → proj {A = S¹ , base} {B = A , pt} x pt)) loop ∙ projᵣ base ≡ refl --- -- id1 = (λ i → (push (inl (loop i))) ∙ (λ j → inr (loop (i ∨ j) , pt)) ∙ sym (push (inl base))) ∙ --- -- (λ i → push (inl base) ∙ lUnit (sym (push (inl base))) (~ i)) ∙ --- -- rCancel (push (inl base)) - --- -- id2 : cong funInv (sym (merid pt)) ≡ (λ _ → inl tt) --- -- id2 = (λ i j → (sym (projₗ pt) ∙ (cong (λ x → proj {A = S¹ , base} {B = A , pt} x pt) loop) ∙ projₗ pt) (~ j)) ∙ --- -- (λ i j → (sym ((projᵣₗ (~ i))) ∙ cong (λ x → proj {A = S¹ , base} {B = A , pt} x pt) loop ∙ projᵣₗ (~ i)) (~ j)) ∙ --- -- (λ i j → id1 i (~ j)) - --- -- id3 : (λ i → funInv (f* {a = pt} (loop i) a)) ≡ ((sym (projₗ a) ∙ cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ (projₗ a))) --- -- id3 = (λ j i → congFunct funInv (merid a) (sym (merid pt)) j i ) ∙ --- -- (λ j i → (cong funInv (merid a) ∙ id2 j) i) ∙ --- -- (λ j i → rUnit (cong funInv (merid a)) (~ j) i) - - - --- -- test5 : ((sym (projₗ a) ∙ cong (λ x → proj {A = S¹ , base} {B = A , pt} x a) loop ∙ (projₗ a))) ≡ λ j → funInv (merid a j) --- -- test5 = refl - --- -- test6 : (ia ib : I) → (x : S¹) → retrHelper (loop ia) pt ≡ {!? ∙ sym (cong (projₗ) loop) ∙ ?!} --- -- test6 = {!!} --- -- where --- -- test7 : ((sym (projₗ pt) ∙ cong (λ x → proj {A = S¹ , base} {B = A , pt} x pt) loop ∙ (projₗ pt))) ≡ refl --- -- test7 = (λ i j → funInv (merid pt j)) ∙ λ i → {!!} --- -- {- --- -- test2 : (i : I) → (retrHelper (loop i) pt) ≡ (λ j → funInv (rCancel (merid pt) j i)) ∙ sym (projᵣ (loop i)) --- -- test2 i = (λ z → hfill ((λ k → λ { (i = i0) → ((λ j → id3 i pt j ∙ sym (projₗ pt) ∙ λ j → proj (loop (~ j)) pt) ∙ invCancelHelper i pt _ _) k --- -- ; (i = i1) → lUnit (rUnit (sym (projₗ pt)) (~ k)) (~ k) })) --- -- (inS ((λ j → funInv (funExt (λ x → merid x ∙ sym (merid pt)) (i ∨ j) pt)) ∙ sym (projₗ pt) ∙ λ j → proj (loop (i ∨ (~ j))) pt)) (~ z)) ∙ --- -- (λ z → (λ j → funInv (funExt (λ x → merid x ∙ sym (merid pt)) (i ∨ j) pt)) ∙ sym (projₗ pt) ∙ λ j → proj (loop (i ∨ (~ j))) pt) ∙ --- -- (λ z → (λ j → funInv {!funExt (λ x → merid x ∙ sym (merid pt)) (i ∨ j) pt -- rCancel (merid pt) z (i ∨ j)!}) ∙ {!!} ∙ {!!}) ∙ --- -- {!!} --- -- -} --- -- isoHelperPre : (x : S¹ × A) (v : (S¹ , base) wedge (A , pt)) → {!!} --- -- isoHelperPre = {!!} - --- -- isoHelper : (x : S¹ × A) (v : (S¹ , base) wedge (A , pt)) (q : inr x ≡ inl tt) (i : I) → (funInv (fun (q i)) ≡ q i) ≡ (funInv (fun ((q ∙ push v) i)) ≡ (q ∙ push v) i) --- -- isoHelper x v q i j = funInv (fun (test2 j)) ≡ test2 j --- -- where --- -- test2 : q i ≡ (q ∙ push v) i --- -- test2 = (λ j → q (i ∧ (~ j))) ∙ λ j → (q ∙ push v) (i ∧ j) - --- -- idIsEquiv2 : ∀ {ℓ} {A B : Type ℓ} → A ≡ B → A ≃ B --- -- idIsEquiv2 = J (λ B p → _ ≃ B) (idEquiv _) - --- -- helper2 : (x : S¹ × A) → funInv (fun (inr x)) ≡ inr x --- -- helper2 (x , x₁) = retrHelper x x₁ --- -- funRetr2 : (x : (S¹ , base) smash (A , pt)) → funInv (fun x) ≡ x --- -- funRetr2 (inl tt) = {!!} --- -- funRetr2 (inr x) = {!!} --- -- funRetr2 (push a i) = ElimR.elimL (λ b path → funInv (fun (path (~ i))) ≡ path (~ i)) (λ b path → funInv (fun ((path) (~ i))) ≡ (path) (~ i)) --- -- (helper2 (refl i)) -- (helper2 (refl i)) --- -- (λ z q → idIsEquiv2 {!!})-- isoToEquiv (isoHelper _ a q (~ i))) --- -- tt (sym (push a)) --- -- where --- -- test3 : (a₁ : (S¹ , base) wedge (A , pt)) → (q : inr (i∧ a) ≡ inl tt) → Iso ((funInv (fun (q (~ i)))) ≡ q (~ i)) (funInv (fun ((q ∙ push a₁) (~ i))) ≡ (q ∙ push a₁) (~ i)) --- -- test3 (inl x) q = {!!} --- -- test3 (inr x) q = {!!} --- -- test3 (push a i) q = {!!} - --- -- test2 : (x : S¹ × A) (v : (S¹ , base) wedge (A , pt)) (q : Path ((S¹ , base) smash (A , pt)) (inr x) (inl tt)) → q (~ i) ≡ (q ∙ push v) (~ i) --- -- test2 x v q = (λ j → q ((~ i) ∧ (~ j))) ∙ λ j → (q ∙ push v) ((~ i) ∧ j) - --- -- funRetr : (x : (S¹ , base) smash (A , pt)) → funInv (fun x) ≡ x --- -- funRetr (inl tt) = refl --- -- funRetr (inr (x , a)) = retrHelper x a --- -- funRetr (push (inl base) i) j = hcomp (λ k → λ { (i = i0) → inl tt --- -- ; (i = i1) → sym (projᵣₗ k) j --- -- ; (j = i0) → inl tt --- -- ; (j = i1) → projᵣ base (~ i)}) --- -- (projᵣ base ((~ j) ∨ (~ i))) --- -- funRetr (push (inl (loop z)) i) j = {!(projᵣ base (~ j ∨ ~ i))!} --- -- {- hcomp (λ k → λ {(i₁ = i0) → ((λ l → hcomp (λ r → λ {(i = i0) → (rCancel (λ _ → inl tt)) r --- -- ; (i = i1) → (lUnit ((λ j → push (inl base) (i ∧ j))) (~ l)) }) --- -- (lUnit ((λ j → push (inl base) (i ∧ j))) (~ i)))) k --- -- ; (i₁ = i1) → ((λ l → hcomp (λ r → λ {(i = i0) → (rCancel (λ _ → inl tt)) r --- -- ; (i = i1) → (lUnit ((λ j → push (inl base) (i ∧ j))) (~ l)) }) --- -- (lUnit ((λ j → push (inl base) (i ∧ j))) (~ i)))) k --- -- ; (i = i0) → ({!!} ∙ {!!} ∙ {!push inl !}) k -- rCancel refl k --- -- ; (i = i1) → {!!}}) --- -- ((hcomp (λ k → λ {(i = i0) → (rCancel (λ _ → inl tt)) k --- -- ; (i = i1) → _ }) ((λ j → funInv (rCancel (merid pt) (~ i ∨ j) i₁)) ∙ --- -- λ j → push (inl (loop i₁)) (i ∧ j)))) -} --- -- where - --- -- bottom : funInv (rCancel (merid pt) (~ i) z) ≡ push (inl (loop z)) i --- -- bottom = (λ j → funInv(rCancel (merid pt) (~ i) (z ∨ j))) ∙ (λ j → push (inl (loop z)) (i ∧ j)) - --- -- i=i1 : (λ j → funInv(rCancel (merid pt) i0 (z ∨ j))) ∙ (λ j → push (inl (loop z)) j) ≡ retrHelper (loop z) pt --- -- i=i1 = {!-- rCancel (merid pt) i1!} - --- -- frontCubeFiller : (i j z : I) → (S¹ , base) smash (A , pt) -- bottom is i j , z is height --- -- frontCubeFiller i j z = hfill (λ k → λ { (i = i0) → inl tt --- -- ; (i = i1) → sym (projᵣₗ k) j --- -- ; (j = i0) → inl tt --- -- ; (j = i1) → projᵣ base (~ i)}) (inS (push (inl base) (i ∧ j))) z - --- -- bottomCubeFiller : (i j z : I) → (S¹ , base) smash (A , pt) --- -- bottomCubeFiller i j z = funInv (rCancel (merid pt) (~ i) (z ∧ j)) - - --- -- theHComp : (i j z : I) → (S¹ , base) smash (A , pt) --- -- theHComp i j z = {!!} - - - - --- -- funRetr (push (inr x) i) = λ j → (projₗ x ((~ i) ∨ (~ j))) --- -- funRetr (push (push tt i₁) i) = {! - --- -- i₁ = i0 ⊢ hcomp --- -- (λ { k (i = i0) → λ _ → inl tt --- -- ; k (i = i1) → λ i₂ → projᵣₗ k (~ i₂) --- -- }) --- -- (λ i₂ → projᵣ base (~ i₂ ∨ ~ i)) --- -- i₁ = i1 ⊢ hcomp --- -- (λ { k (i = i0) → λ _ → inl tt --- -- ; k (i = i1) → λ i₂ → projᵣₗ k (~ i₂) --- -- }) --- -- (λ i₂ → projᵣ base (~ i₂ ∨ ~ i)) --- -- i = i0 ⊢ λ _ → inl tt - --- -- i = i1 ⊢ !} diff --git a/Cubical/Homotopy/Connected.agda b/Cubical/Homotopy/Connected.agda index e76f5e946..bd06687a9 100644 --- a/Cubical/Homotopy/Connected.agda +++ b/Cubical/Homotopy/Connected.agda @@ -34,13 +34,14 @@ isHLevelConnectedFun n f = ∀ b → isHLevelConnected n (fiber f b) isHLevelTruncatedFun : ∀ {ℓ ℓ'} (n : ℕ) {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Type (ℓ-max ℓ ℓ') isHLevelTruncatedFun n f = ∀ b → isOfHLevel n (fiber f b) -isConnectedSubtr : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n m : ℕ) (f : A → B) → isHLevelConnectedFun (n + m) f → isHLevelConnectedFun n f -isConnectedSubtr n m f iscon b = transport (λ i → isContr (ua (truncOfTruncEq {A = fiber f b} n m) (~ i) )) - (∣ iscon b .fst ∣ , - Trunc.elim (λ x → isOfHLevelPath n (isOfHLevelTrunc n) _ _) - λ a → cong ∣_∣ (iscon b .snd a)) - - +isConnectedSubtr : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n m : ℕ) (f : A → B) + → isHLevelConnectedFun (n + m) f + → isHLevelConnectedFun n f +isConnectedSubtr n m f iscon b = + transport (λ i → isContr (ua (truncOfTruncEq {A = fiber f b} n m) (~ i) )) + (∣ iscon b .fst ∣ , + Trunc.elim (λ x → isOfHLevelPath n (isOfHLevelTrunc n) _ _) + λ a → cong ∣_∣ (iscon b .snd a)) private typeToFiber : ∀ {ℓ} (A : Type ℓ) (b : Unit) → A ≡ fiber (λ (x : A) → b) b @@ -80,48 +81,31 @@ module elim {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} (f : A → B) (n : isEquivPrecompose : ∀ {ℓ'''} (P : B → HLevel ℓ''' n) → isHLevelConnectedFun n f → isEquiv λ(s : (b : B) → P b .fst) → s ∘ f - isEquivPrecompose P fConn = - isoToIsEquiv - (iso (_∘ f) - (λ t b → inv t b (fConn b .fst)) - (λ t → funExt λ a → - cong (inv t (f a)) (fConn (f a) .snd ∣ a , refl ∣) - ∙ substRefl {B = fst ∘ P} (t a)) - (λ s → funExt λ b → - Trunc.elim - {B = λ d → inv (s ∘ f) b d ≡ s b} - (λ _ → isOfHLevelPath n (P b .snd) _ _) - (λ {(a , p) i → transp (λ j → P (p (j ∨ i)) .fst) i (s (p i))}) - (fConn b .fst))) - - where - inv : ((a : A) → P (f a) .fst) → (b : B) → hLevelTrunc n (fiber f b) → P b .fst - inv t b = - Trunc.rec - (P b .snd) - (λ {(a , p) → subst (fst ∘ P) p (t a)}) - - hasSectionPrecompose : ∀ {ℓ'''} (P : B → HLevel ℓ''' n) - → isEquiv (λ(s : (b : B) → P b .fst) → s ∘ f) - → hasSection λ(s : (b : B) → P b .fst) → s ∘ f - hasSectionPrecompose P record { equiv-proof = eqPf } = - (λ g → (eqPf g) .fst .fst) , (λ b → (eqPf b) .fst .snd) - + isEquivPrecompose P fConn = isoToIsEquiv (isIsoPrecompose P fConn) isConnectedPrecompose : ((∀ {ℓ'''} (P : B → HLevel ℓ''' n) → hasSection λ(s : (b : B) → P b .fst) → s ∘ f)) → isHLevelConnectedFun n f isConnectedPrecompose P→sect b = (c n P→sect b) , λ y → sym (fun n P→sect b y) where - P : (n : ℕ) → (∀ {ℓ} (P : B → HLevel ℓ n) → hasSection (λ(s : (b : B) → P b .fst) → s ∘ f)) → B → Type _ + P : (n : ℕ) → (∀ {ℓ} (P : B → HLevel ℓ n) + → hasSection (λ(s : (b : B) → P b .fst) → s ∘ f)) + → B → Type _ P n s b = hLevelTrunc n (fiber f b) - c : (n : ℕ) → (∀ {ℓ} (P : B → HLevel ℓ n) → hasSection (λ(s : (b : B) → P b .fst) → s ∘ f)) → (b : B) → hLevelTrunc n (fiber f b) - c n s = (s λ b → (hLevelTrunc n (fiber f b) , isOfHLevelTrunc _)) .fst λ a → ∣ a , refl ∣ + c : (n : ℕ) → (∀ {ℓ} (P : B → HLevel ℓ n) + → hasSection (λ(s : (b : B) → P b .fst) → s ∘ f)) → (b : B) + → hLevelTrunc n (fiber f b) + c n s = (s λ b → (hLevelTrunc n (fiber f b) , isOfHLevelTrunc _)) .fst + λ a → ∣ a , refl ∣ - fun : (n : ℕ) → (s : (∀ {ℓ} (P : B → HLevel ℓ n) → hasSection λ(s : (b : B) → P b .fst) → s ∘ f)) → (b : B) (w : (hLevelTrunc n (fiber f b))) → w ≡ c n s b + fun : (n : ℕ) (s : (∀ {ℓ} (P : B → HLevel ℓ n) + → hasSection λ(s : (b : B) → P b .fst) → s ∘ f)) + → (b : B) (w : (hLevelTrunc n (fiber f b))) + → w ≡ c n s b fun zero s b w = isOfHLevelSuc zero (isOfHLevelTrunc _) w (c zero s b) - fun (suc n) s b = Trunc.elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc _) x (c (suc n) s b)) λ a → transport eqtyp c* b (fst a) (snd a) + fun (suc n) s b = Trunc.elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc _) x (c (suc n) s b)) + λ a → transport eqtyp c* b (fst a) (snd a) where eqtyp : ((a : A) → ∣ (a , refl {x = f a}) ∣ ≡ c (suc n) s (f a)) @@ -131,11 +115,11 @@ module elim {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} (f : A → B) (n : isoToPath (iso (λ g b a → J (λ b p → ∣ a , p ∣ ≡ c (suc n) s b) (g a)) (λ g a → g (f a) a refl) - (λ h → funExt λ b - → funExt λ a - → funExt (J (λ b x → (J (λ b₂ p → ∣ a , p ∣ ≡ c (suc n) s b₂) (h (f a) a (λ _ → f a)) x ≡ h b a x)) - (JRefl (λ b₂ p → ∣ a , p ∣ ≡ c (suc n) s b₂) (h (f a) a (λ _ → f a))))) - λ h → funExt λ a → JRefl (λ b₁ p → ∣ a , p ∣ ≡ c (suc n) s b₁) (h a)) + (λ h i b a p → (J (λ b x → (J (λ b₂ p → ∣ a , p ∣ ≡ c (suc n) s b₂) + (h (f a) a (λ _ → f a)) x ≡ h b a x)) + (JRefl (λ b₂ p → ∣ a , p ∣ ≡ c (suc n) s b₂) (h (f a) a (λ _ → f a)))) + p i) + λ h i a p → JRefl (λ b₁ p → ∣ a , p ∣ ≡ c (suc n) s b₁) (h a) i p) c* : ((a : A) → ∣ (a , refl {x = f a}) ∣ ≡ c (suc n) s (f a)) c* a = sym (cong (λ x → x a) (s (λ b → hLevelTrunc (suc n) (fiber f b) , isOfHLevelTrunc _) .snd λ a → ∣ a , refl ∣)) @@ -158,15 +142,17 @@ isOfHLevelPrecomposeConnected (suc k) n P f fConn t = f fConn (funExt⁻ (p₀ ∙∙ refl ∙∙ sym p₁))))}) -indMapEquiv→conType : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → (∀ {ℓ'} (B : HLevel ℓ' n) → isEquiv (λ (b : (fst B)) → λ (a : A) → b)) → isHLevelConnected n A -indMapEquiv→conType {A = A} n BEq = transport (λ i → isContr (hLevelTrunc n (typeToFiber A tt (~ i)))) - (elim.isConnectedPrecompose (λ _ → tt) n - (λ {ℓ'} P → elim.hasSectionPrecompose (λ _ → tt) n P - (compEquiv ((λ Q → Q tt) , isoToIsEquiv (helper P)) - (_ , BEq (P tt)) .snd)) - tt) +indMapEquiv→conType : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → (∀ {ℓ'} (B : HLevel ℓ' n) + → isEquiv (λ (b : (fst B)) → λ (a : A) → b)) → isHLevelConnected n A +indMapEquiv→conType {A = A} n BEq = + transport (λ i → isContr (hLevelTrunc n (typeToFiber A tt (~ i)))) + (elim.isConnectedPrecompose (λ _ → tt) n + (λ {ℓ'} P → isEquiv-hasSection _ (compEquiv ((λ Q → Q tt) , isoToIsEquiv (helper P)) + (_ , BEq (P tt)) .snd )) + tt) where - helper : ∀ {ℓ'} → (P : Unit → HLevel ℓ' n) → Iso ((b : Unit) → fst (P b)) (fst (P tt)) + helper : ∀ {ℓ'} (P : Unit → HLevel ℓ' n) + → Iso ((b : Unit) → fst (P b)) (fst (P tt)) Iso.fun (helper P) = λ Q → Q tt Iso.inv (helper P) = λ a → λ {tt → a} Iso.rightInv (helper P) b = refl @@ -209,7 +195,8 @@ isHLevelConnectedPoint2 n {A = A} a connMap = indMapEquiv→conType _ λ B → i module _ {ℓ' : Level} (B : HLevel ℓ' (suc n)) where helper : (f : A → fst B) → (a2 : A) → f a2 ≡ f a - helper f = equiv-proof (elim.isEquivPrecompose (λ _ → a) n (λ a2 → (f a2 ≡ f a) , isOfHLevelPath' n (snd B) (f a2) (f a)) connMap) (λ _ → refl) .fst .fst + helper f = equiv-proof (elim.isEquivPrecompose (λ _ → a) n (λ a2 → (f a2 ≡ f a) , + isOfHLevelPath' n (snd B) (f a2) (f a)) connMap) (λ _ → refl) .fst .fst theIso : Iso (fst B) (A → fst B) Iso.fun theIso b a = b @@ -258,67 +245,49 @@ connectedTruncIso {A = A} {B = B} (suc n) f con = g g : Iso (hLevelTrunc (suc n) A) (hLevelTrunc (suc n) B) Iso.fun g = map f Iso.inv g = Trunc.rec (isOfHLevelTrunc _) back - Iso.leftInv g = helper - where - abstract - helper : _ - helper = Trunc.elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc _) _ _) + Iso.leftInv g = Trunc.elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc _) _ _) λ a → cong (map fst) (con (f a) .snd ∣ a , refl ∣) - Iso.rightInv g = helper - where - abstract - helper : _ - helper = Trunc.elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc _) _ _) + Iso.rightInv g = Trunc.elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc _) _ _) backSection connectedTruncIso2 : ∀ {ℓ} {A B : Type ℓ} (n m : ℕ) (f : A → B) → Σ[ x ∈ ℕ ] n + x ≡ m → isHLevelConnectedFun m f → Iso (hLevelTrunc n A) (hLevelTrunc n B) -connectedTruncIso2 {A = A} {B = B} n m f (x , pf) con = connectedTruncIso n f (isConnectedSubtr n x f (transport (λ i → isHLevelConnectedFun (pf (~ i)) f) con)) - - +connectedTruncIso2 {A = A} {B = B} n m f (x , pf) con = + connectedTruncIso n f (isConnectedSubtr n x f (transport (λ i → isHLevelConnectedFun (pf (~ i)) f) con)) connectedTruncEquiv : ∀ {ℓ} {A B : Type ℓ} (n : ℕ) (f : A → B) → isHLevelConnectedFun n f → hLevelTrunc n A ≃ hLevelTrunc n B connectedTruncEquiv {A = A} {B = B} n f con = isoToEquiv (connectedTruncIso n f con) -connectedTruncEquiv2 : ∀ {ℓ} {A B : Type ℓ} (n m : ℕ) (f : A → B) → Σ[ x ∈ ℕ ] n + x ≡ m - → isHLevelConnectedFun m f - → hLevelTrunc n A ≃ hLevelTrunc n B -connectedTruncEquiv2 {A = A} {B = B} n m f (x , pf) con = - connectedTruncEquiv n f (isConnectedSubtr n x f (transport (λ i → isHLevelConnectedFun (pf (~ i)) f) con)) - - ------------- - - inrConnected : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} (n : ℕ) → (f : C → A) (g : C → B) → isHLevelConnectedFun n f → isHLevelConnectedFun n {A = B} {B = Pushout f g} inr -inrConnected {A = A} {B = B} {C = C} n f g iscon = elim.isConnectedPrecompose inr n λ P → (λ l → k P l) , λ b → refl +inrConnected {A = A} {B = B} {C = C} n f g iscon = + elim.isConnectedPrecompose inr n λ P → (λ l → k P l) , λ b → refl where module _ {ℓ : Level} (P : (Pushout f g) → HLevel ℓ n) (h : (b : B) → typ (P (inr b))) - where - Q : A → HLevel _ n - Q a = (P (inl a)) - - fun : (c : C) → typ (Q (f c)) - fun c = transport (λ i → typ (P (push c (~ i)))) (h (g c)) - - k : (d : Pushout f g) → typ (P d) - k (inl x) = elim.hasSectionPrecompose f n Q (elim.isEquivPrecompose f n Q iscon) .fst fun x - k (inr x) = h x - k (push a i) = - hcomp (λ k → λ { (i = i0) → (elim.hasSectionPrecompose f n Q (elim.isEquivPrecompose f n Q iscon) .snd fun i0 a) - ; (i = i1) → transportTransport⁻ (λ j → typ (P (push a j))) (h (g a)) k }) - (transp (λ j → typ (P (push a (i ∧ j)))) - (~ i) - (elim.hasSectionPrecompose f n Q (elim.isEquivPrecompose f n Q iscon) .snd fun i a )) - + where + Q : A → HLevel _ n + Q a = (P (inl a)) + + fun : (c : C) → typ (Q (f c)) + fun c = transport (λ i → typ (P (push c (~ i)))) (h (g c)) + + k : (d : Pushout f g) → typ (P d) + k (inl x) = equiv-proof (elim.isEquivPrecompose f n Q iscon) fun .fst .fst x + k (inr x) = h x + k (push a i) = + hcomp (λ k → λ { (i = i0) → equiv-proof (elim.isEquivPrecompose f n Q iscon) fun .fst .snd i0 a + ; (i = i1) → transportTransport⁻ (λ j → typ (P (push a j))) (h (g a)) k }) + (transp (λ j → typ (P (push a (i ∧ j)))) + (~ i) + (equiv-proof (elim.isEquivPrecompose f n Q iscon) + fun .fst .snd i a)) sphereConnected : (n : ℕ) → isHLevelConnected (suc n) (S₊ n) sphereConnected zero = ∣ north ∣ , isOfHLevelTrunc 1 ∣ north ∣ @@ -332,7 +301,6 @@ sphereConnected (suc n) = λ {tt → transport (λ i → isContr (hLevelTrunc (suc n) ((isoToPath fibIso) (~ i)))) (sphereConnected n)})))) where - mapsAgree : Path ((x : Unit) → Pushout {A = S₊ n} (λ x → tt) (λ x → tt)) (λ (x : Unit) → inr tt) inr @@ -343,372 +311,3 @@ sphereConnected (suc n) = Iso.inv (fibIso) a = a , refl Iso.leftInv (fibIso) b = refl Iso.rightInv (fibIso) b = refl - --- private --- variable --- ℓ ℓ' ℓ'' : Level --- A : Type ℓ --- B : Type ℓ' --- A' : Pointed ℓ --- B' : Pointed ℓ' --- C' : Pointed ℓ'' - --- {- --- {- Lemma 8.6.1. in the HoTT book -} - --- private --- module helpers (n k : ℕ) (f : A → B) (P : B → HLevel ℓ'' (k + n)) (l : (a : A) → fst (P (f a))) (iscon : isHLevelConnectedFun n f) where --- F : _ --- F = λ(s : (b : B) → P b .fst) → s ∘ f - --- fiberPath1 : Iso (fiber F l) (Σ[ g ∈ ((b : B) → P b .fst) ] ((a : A) → g (f a) ≡ l a)) --- Iso.fun fiberPath1 p = (p .fst) , (funExt⁻ (p .snd)) --- Iso.inv fiberPath1 p = (p .fst) , (funExt (p .snd)) --- Iso.rightInv fiberPath1 p = refl --- Iso.leftInv fiberPath1 p = refl - --- throwAboutHelper : ∀{ℓ} {A : Type ℓ} {a b c : A} → (r : b ≡ c) (p : a ≡ b) (q : a ≡ c) → (p ≡ q ∙ (sym r) ) ≡ (q ≡ p ∙ r) --- throwAboutHelper {a = a} {b = b} r p q = (λ i → rUnit p i ≡ q ∙ sym r) ∙ --- (λ i → p ∙ (λ j → r (j ∧ i)) ≡ q ∙ λ j → r (i ∨ (~ j))) ∙ --- (λ i → p ∙ r ≡ rUnit q (~ i)) ∙ symId _ _ --- where --- symId : ∀{ℓ} {A : Type ℓ} (a b : A) → (a ≡ b) ≡ (b ≡ a) --- symId a b = isoToPath symiso --- where --- symiso : Iso (a ≡ b) (b ≡ a) --- Iso.fun symiso = sym --- Iso.inv symiso = sym --- Iso.rightInv symiso a = refl --- Iso.leftInv symiso a = refl - --- throwAbout : (g h : (Σ[ g ∈ ((b : B) → P b .fst) ] ((a : A) → g (f a) ≡ l a))) --- → (r : (x : B) → (g .fst) x ≡ (h .fst) x) --- → Iso (g . snd ≡ (λ x → r (f x) ∙ h .snd x)) --- ((x : A) → r (f x) ≡ (g .snd x) ∙ sym (h .snd x)) --- Iso.fun (throwAbout (g , p) (h , q) r) id x = --- transport (λ i → throwAboutHelper (sym (q x)) (p x) (r (f x)) i) --- (funExt⁻ id x) --- Iso.inv (throwAbout (g , p) (h , q) r) id = --- funExt λ x → transport (λ i → throwAboutHelper (sym (q x)) (p x) (r (f x)) (~ i)) (id x) --- Iso.rightInv (throwAbout (g , p) (h , q) r) id = --- funExt λ x → transportTransport⁻ (throwAboutHelper (sym (q x)) (p x) (r (f x))) (id x) --- Iso.leftInv (throwAbout (g , p) (h , q) r) id j = --- funExt (λ x → (transport⁻Transport (throwAboutHelper (sym (q x)) (p x) (r (f x))) (λ i → id i x)) j) - --- helper3 : (g h : (Σ[ g ∈ ((b : B) → P b .fst) ] ((a : A) → g (f a) ≡ l a))) --- → (r : (x : B) → (g .fst) x ≡ (h .fst) x) → --- (PathP (λ i → (x : A) → r (f x) i ≡ l x) (g .snd) (h .snd)) --- ≡ ((x : A) → r (f x) ≡ ((g .snd) x) ∙ sym ((h .snd) x)) --- helper3 (g , p) (h , q) r j = --- hcomp (λ k → λ{ (j = i0) → PathP (λ i → (x : A) → r (f x) i ≡ l x) p λ x → lUnit (q x) (~ k) --- ; (j = i1) → isoToPath (throwAbout (g , p) (h , q) r) k}) --- (PathP (λ i → (x : A) → r (f x) (i ∧ ~ j) ≡ l x) --- p --- λ x → (λ i → r (f x) (~ j ∨ i)) ∙ (q x)) - --- iso2 : (g h : (Σ[ g ∈ ((b : B) → P b .fst) ] ((a : A) → g (f a) ≡ l a))) --- → Iso (Σ[ p ∈ g .fst ≡ h .fst ] PathP (λ i → (a : A) → p i (f a) ≡ l a) (g .snd) (h .snd)) --- (Σ[ r ∈ ((b : B) → g .fst b ≡ h .fst b) ] r ∘ f ≡ λ a → (g .snd a) ∙ (h .snd a) ⁻¹) --- Iso.fun (iso2 (g , p) (h , q)) (funid , pathp) = funExt⁻ funid , funExt (transport (helper3 (g , p) (h , q) λ x i → funid i x) pathp) - --- Iso.inv (iso2 (g , p) (h , q)) (r , id) = (funExt r) , transport (sym (helper3 (g , p) (h , q) r)) λ x i → id i x --- Iso.rightInv (iso2 (g , p) (h , q)) a i = (fst a) , (funExt (transportTransport⁻ (helper3 (g , p) (h , q) (a .fst)) (funExt⁻ (snd a)) i)) --- Iso.leftInv (iso2 (g , p) (h , q)) a i = fst a , transport⁻Transport (helper3 (g , p) (h , q) (λ x i → a .fst i x )) (a .snd) i - - --- isConnectedFamily : (n k : ℕ) (f : A → B) (P : B → HLevel ℓ'' (k + n)) --- → isHLevelConnectedFun n f --- → isHLevelTruncatedFun k (λ(s : (b : B) → P b .fst) → s ∘ f) --- isConnectedFamily {A = A} {B = B} n zero f P iscon = equiv-proof (elim.isEquivPrecompose f n P iscon) --- isConnectedFamily {A = A} {B = B} n (suc zero) f P iscon l = --- transport (λ i → isOfHLevel 1 (isoToPath (helpers.fiberPath1 n 1 f P l iscon) (~ i))) --- λ a b → transport (λ i → (sym (ua Σ≡) ∙ (isoToPath (helpers.iso2 n 1 f P l iscon a b))) (~ i)) --- (isConnectedFamily n 0 f --- (λ x → (a .fst x ≡ b .fst x) , isOfHLevelPath' n (P x .snd) (a .fst x) (b .fst x)) --- iscon --- (λ r → (a .snd r) ∙ (b .snd r) ⁻¹) .fst ) --- isConnectedFamily {A = A} {B = B} n (suc (suc k)) f P iscon l = --- transport (λ i → isOfHLevel (2 + k) (isoToPath (helpers.fiberPath1 n (2 + k) f P l iscon) (~ i))) --- λ a b → transport (λ i → isOfHLevel (suc k) ((sym (ua Σ≡) ∙ (isoToPath (helpers.iso2 n (2 + k) f P l iscon a b))) (~ i))) --- (isConnectedFamily n (suc k) f --- (λ x → (a .fst x ≡ b .fst x) , ((P x) .snd (a .fst x) (b .fst x))) --- iscon --- λ r → (a .snd r) ∙ (b .snd r) ⁻¹) - --- -} - - --- helper : (n : ℕ) {A B : Pointed₀} → hLevelTrunc n (fst A) → hLevelTrunc n (fst B) → hLevelTrunc n (A ⋀ B) --- helper n = Trunc.rec (isOfHLevelΠ _ (λ _ → isOfHLevelTrunc _)) λ a → Trunc.rec (isOfHLevelTrunc _) λ b → ∣ inr (a , b) ∣ - --- truncTest : (n : ℕ) {A B : Pointed₀} → (hLevelTrunc n (fst A) , ∣ pt A ∣) ⋀ (hLevelTrunc n (fst B) , ∣ pt B ∣) --- → hLevelTrunc n (A ⋀ B) --- truncTest n (inl x) = ∣ inl tt ∣ --- truncTest n {A = A} {B = B} (inr (x , x₁)) = helper n x x₁ --- truncTest n {A = A} {B = B} (push (inl x) i) = helper2 x (~ i) --- where --- helper2 : (x : hLevelTrunc n (fst A)) → helper n {A = A} {B = B} x ∣ pt B ∣ ≡ ∣ inl tt ∣ --- helper2 = Trunc.elim (λ x → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a i → ∣ push (inl a) (~ i) ∣ --- truncTest n {A = A} {B = B} (push (inr x) i) = helper2 x (~ i) --- where --- helper2 : (x : hLevelTrunc n (fst B)) → helper n {A = A} {B = B} ∣ pt A ∣ x ≡ ∣ inl tt ∣ --- helper2 = Trunc.elim (λ x → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a i → ∣ push (inr a) (~ i) ∣ --- truncTest n {A = A} {B = B} (push (push tt i₁) i) = ∣ push (push tt i₁) i ∣ - --- hLevelSmash : {A B : Pointed₀} (n : ℕ) → isOfHLevel n ((hLevelTrunc n (fst A) , ∣ pt A ∣) ⋀ (hLevelTrunc n (fst B) , ∣ pt B ∣)) --- hLevelSmash {A = A} {B = B} zero = (inl tt) , {!!} --- where --- helper2 : (y : (hLevelTrunc zero (fst A) , ∣ pt A ∣) ⋀ (hLevelTrunc zero (fst B) , ∣ pt B ∣)) → inl tt ≡ y --- helper2 (inl tt) = refl --- helper2 (inr (x , x₁)) = push (inr x₁) ∙ (λ i → inr (isOfHLevelTrunc 0 .snd ∣ snd A ∣ (~ i) , x₁)) ∙ λ i → inr (((isOfHLevelTrunc 0 .snd) x i) , x₁) --- helper2 (push (inl x) i) = {!!} --- helper2 (push (inr x) i) = {!!} --- helper2 (push (push a i₁) i) = {!!} --- hLevelSmash {A = A} {B = B} (suc n) = {!!} - --- truncTestInv : {!!} --- truncTestInv = {!!} - --- isConnectedSmashIdfun2 : {A B C : Pointed₀} --- → (f : A →∙ B) (nf nc : ℕ) --- → isHLevelConnectedFun nf (fst f) --- → isHLevelConnected (2 + nc) (fst C) --- → isHLevelConnectedFun (1 + nc + nf) (Smash-map f (idfun∙ C)) --- isConnectedSmashIdfun2 {A = A , ptA} {B = B , ptB} {C = C , ptC} (f , fpt) nf nc conf conC = λ b → {!!} , {!!} --- -- elim.isConnectedPrecompose (Smash-map (f , fpt) (idfun∙ (C , ptC))) _ λ {ℓ} P → proof.h {ℓ} P , λ b → funExt (proof.hsection {ℓ} P b) -- elim.isConnectedPrecompose (Smash-map (f , fpt) (idfun∙ (C , ptC))) _ λ {ℓ} P → (λ h → proof.h {ℓ} P h) ,push (inl x)) λ b → funExt λ x → {!!} --- where --- module proof {ℓ : Level} (P : (Smash (B , ptB) (C , ptC)) → HLevel ℓ (1 + nc + nf)) --- (d : (x : Smash (A , ptA) (C , ptC)) → P ((Smash-map (f , fpt) (idfun∙ (C , ptC)) ) x) .fst) --- where --- test : (c : C) (b : B) → (hLevelTrunc (suc (nc + nf)) (fiber (Smash-map (f , fpt) (idfun∙ (C , ptC))) (proj b c))) --- ≡ ((hLevelTrunc _ (fiber f b) , conf b .fst) ⋀ {!!}) --- test = {!!} --- -- F : (c : C) → _ --- -- F c = λ(s : (b : B) → P (proj b c) .fst) → s ∘ f - --- -- step1 : (c : C) → isHLevelTruncatedFun (1 + nc) (F c) --- -- step1 c = isOfHLevelPrecomposeConnected (1 + nc) nf ((λ b → P (proj b c) .fst , P (proj b c) .snd)) f conf - --- -- codomFun : (c : C) (a : A) → P (proj (f a) c) .fst --- -- codomFun c = d ∘ λ a → proj a c - --- -- Q : C → HLevel _ (1 + nc) --- -- Q c = fiber (F c) (codomFun c) , step1 c _ - - --- -- helper : (a : A) → transport (λ i → P (gluel (f a) (~ i)) .fst) (d (basel)) --- -- ≡ d (proj a ptC) --- -- helper a = transport (PathP≡Path (λ i → P (gluel (f a) (~ i)) .fst) (d basel) (d (proj a ptC))) --- -- (transport (λ j → PathP (λ i → P (lUnit (gluel (f a)) (~ j) (~ i)) .fst) (d basel) (d (proj a ptC))) --- -- λ i → d (gluel a (~ i))) - --- -- QptC : Q ptC .fst --- -- QptC = (λ b → transport (λ i → P (gluel b (~ i)) .fst) (d basel)) , --- -- funExt helper - - --- -- Q-constructor : (c : C) → Q c .fst --- -- Q-constructor c = Iso.inv (elim.isIsoPrecompose (λ (_ : Unit) → ptC) (1 + nc) Q (isHLevelConnectedPoint _ conC ptC)) (λ _ → QptC) c - --- -- g : (b : B) (c : C) → P (proj b c) .fst --- -- g b c = Q-constructor c .fst b - --- -- Q-constructor-β : (b : B) → Q-constructor ptC .fst b ≡ transport (λ i → P (gluel b (~ i)) .fst) (d basel) --- -- Q-constructor-β b = ((λ i → (Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) (helper3 i)) .fst b)) ∙ --- -- (λ i → transportRefl (QptC) i .fst b) - --- -- where --- -- helper3 : (isHLevelConnectedPoint _ conC ptC) ptC .fst ≡ ∣ tt , refl ∣ --- -- helper3 = (isHLevelConnectedPoint _ conC ptC) ptC .snd ∣ tt , refl ∣ - --- -- moveTransport : ∀ {ℓ ℓ'} → {A : Type ℓ} {B : A → Type ℓ'} {x y : A} (p : x ≡ y) (x1 : B x) (y1 : B y) --- -- → transport (λ i → (B (p i))) x1 ≡ y1 --- -- → transport⁻ (λ i → (B (p i))) y1 ≡ x1 --- -- moveTransport {B = B} p x1 y1 transp = ((λ i → transport⁻ (λ i → B (p i)) (transp (~ i))) ∙ (transport⁻Transport (λ i → B (p i)) x1)) - --- -- otherway : (b : B) → transport (λ i → P (gluel b i) .fst) (g b ptC) ≡ d basel --- -- otherway b = (λ i → transport (λ i → P (gluel b i) .fst) (Q-constructor-β b i)) ∙ transportTransport⁻ (λ i → P (gluel b i) .fst) (d basel) - --- -- path1 : (b : B) → PathP (λ i → P (gluel b i) .fst) (g b ptC) (d basel) --- -- path1 b i = hcomp (λ k → λ {(i = i0) → g b ptC --- -- ; (i = i1) → otherway b k}) --- -- (transp (λ j → P (gluel b (i ∧ j)) .fst) (~ i) (g b ptC)) - - --- -- where --- -- massive : (c : C) → PathP _ (g ptB c) (d baser) --- -- massive c = compPathP (λ i → g (fpt (~ i)) c) (compPathP (λ i → Q-constructor c .snd i (ptA)) (λ i → d (gluer c i))) - --- -- pathPCancel : (c : C) --- -- → ((λ z → P (proj (fpt (~ z)) c) .fst) ∙ (λ i → ((λ _ → P (proj (f ptA) c) .fst) ∙ (λ i₁ → P (((λ j₁ → proj (fpt j₁) c) ∙ gluer c) i₁) .fst)) i)) --- -- ≡ (λ i → P (gluer c i) .fst) --- -- pathPCancel c = (λ i → (λ z → P (proj (fpt (~ z)) c) .fst) ∙ lUnit ((λ i₁ → P (((λ j₁ → proj (fpt j₁) c) ∙ gluer c) i₁) .fst)) (~ i)) --- -- ∙ (λ i → (λ z → P (proj (fpt (~ z ∨ i)) c) .fst) ∙ λ i₁ → P (((λ j₁ → proj (fpt (j₁ ∨ i)) c) ∙ gluer c) i₁) .fst ) --- -- ∙ (λ i → lUnit (λ i₁ → P (lUnit (gluer c) (~ i) i₁) .fst) (~ i)) - --- -- path2 : (c : C) → PathP (λ i → P (gluer c i) .fst) (g ptB c) (d baser) --- -- path2 c = transport (λ i → PathP (λ j → pathPCancel c i j) (g ptB c) (d baser)) (massive c) - --- -- h : (x : Smash (B , ptB) (C , ptC)) → P x .fst --- -- h basel = d basel --- -- h baser = d baser --- -- h (proj x y) = g x y --- -- h (gluel b i) = path1 b i --- -- h (gluer c i) = path2 c i - --- -- hsection : (x : Smash (A , ptA) (C , ptC)) → h (Smash-map (f , fpt) (idfun∙ (C , ptC)) x) ≡ d x --- -- hsection basel = refl --- -- hsection baser = refl --- -- hsection (proj x y) i = Q-constructor y .snd i x --- -- hsection (gluel a i) j = hcomp (λ k → λ { (i = i0) → transportRefl {!!} k --- -- ; (i = i1) → transportRefl (d basel) k --- -- ; (j = i0) → {!!} --- -- ; (j = i1) → {!!}}) --- -- {!hcomp!} -- hcomp {!!} {!!} --- -- where - --- -- helper2 : Path ((P (proj (f a) ptC) .fst) ≡ (P basel .fst)) --- -- (λ i → P (Smash-map (f , fpt) (idfun∙ (C , ptC)) (gluel a i)) .fst) --- -- λ i → P (gluel (f a) i) .fst --- -- helper2 = (λ j i → P ((lUnit (gluel (f a)) (~ j)) i) .fst) --- -- helper6 : (a : A) → PathP ((λ j → PathP (λ i → P {!Q-constructor ? .snd i ?!} .fst) {!P (proj (f a) ptC) .fst!} {!!})) (λ i → h (Smash-map (f , fpt) (idfun∙ (C , ptC)) (gluel a i))) (path1 (f a)) --- -- helper6 a = ((λ j i → h ((lUnit (gluel (f a)) (~ j)) i))) - --- -- helper3 : PathP {!!} (path1 (f a)) (λ i → d (gluel a i)) --- -- helper3 = compPathP {!!} {!!} --- -- helper4 : PathP {!!} (path1 (f a)) λ i → d (gluel a i) --- -- helper4 = {!!} --- -- hsection (gluer b i) = {!!} - --- -- -- isConnectedSmashIdfun : (f : A' →∙ B') (nf nc : ℕ) --- -- -- → isHLevelConnectedFun nf (fst f) --- -- -- → isHLevelConnected (2 + nc) (fst C') --- -- -- → isHLevelConnectedFun (1 + nf + nc) (f ⋀⃗ idfun∙ C') --- -- -- isConnectedSmashIdfun {A' = (A , ptA)} {B' = (B , ptB)} {C' = (C , ptC)} (f , fpt) nf nc conf conC = {!isHLel!} --- -- -- where --- -- -- module _ (P : ((B , ptB) ⋀ (C , ptC)) → HLevel (ℓ-max (ℓ-max ℓ ℓ') ℓ'') (1 + nc + nf)) --- -- -- (d : (x : (A , ptA) ⋀ (C , ptC)) → P (((f , fpt) ⋀⃗ idfun∙ (C , ptC) ) x) .fst) --- -- -- where - --- -- -- F : (c : C) → _ --- -- -- F c = λ(s : (b : B) → P (inr (b , c)) .fst) → s ∘ f - --- -- -- step1 : (c : C) → isHLevelTruncatedFun (1 + nc) (F c) --- -- -- step1 c = isOfHLevelPrecomposeConnected (1 + nc) nf ((λ b → P (inr (b , c)) .fst , P (inr (b , c)) .snd)) f conf - --- -- -- codomFun : (c : C) (a : A) → P (inr ((f a) , c)) .fst --- -- -- codomFun c = d ∘ λ a → inr (a , c) - --- -- -- Q : C → HLevel _ (1 + nc) --- -- -- Q c = fiber (F c) (codomFun c) , step1 c _ - - --- -- -- helper : (a : A) → transport (λ i → P (push (inl (f a)) i) .fst) (d (inl tt)) --- -- -- ≡ d (inr (a , ptC)) --- -- -- helper a = transport (PathP≡Path (λ i → P (push (inl (f a)) i) .fst) (d (inl tt)) (d (inr (a , ptC)))) --- -- -- (transport (λ j → PathP (λ i → P (rUnit (push (inl (f a))) (~ j) i) .fst) (d (inl tt)) (d (inr (a , ptC)))) --- -- -- λ i → d (push (inl a) i)) - - --- -- -- QptC : Q ptC .fst --- -- -- QptC = (λ b → transport (λ i → P (push (inl b) i) .fst) (d (inl tt))) , --- -- -- funExt helper --- -- -- where - - --- -- -- Q-constructor : (c : C) → Q c .fst --- -- -- Q-constructor c = Iso.inv (elim.isIsoPrecompose (λ _ → ptC) (1 + nc) Q (isHLevelConnectedPoint _ conC ptC)) (λ ( _ : Unit) → QptC) c - --- -- -- g : (b : B) (c : C) → P (inr (b , c)) .fst --- -- -- g b c = Q-constructor c .fst b - --- -- -- Q-constructor-β : (b : B) → Q-constructor ptC .fst b ≡ transport (λ i → P (push (inl b) i) .fst) (d (inl tt)) --- -- -- Q-constructor-β b = ((λ i → (Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) (helper3 i)) .fst b)) ∙ --- -- -- (λ i → transportRefl (QptC) i .fst b) - --- -- -- where --- -- -- helper3 : (isHLevelConnectedPoint _ conC ptC) ptC .fst ≡ ∣ tt , refl ∣ --- -- -- helper3 = (isHLevelConnectedPoint _ conC ptC) ptC .snd ∣ tt , refl ∣ - --- -- -- test : (b : B) (c : C) → PathP (λ i → P (push {!!} {!(~ i)!}) .fst) (Q-constructor ptC .fst b) (Q-constructor c .fst ptB) --- -- -- test b = {!!} - --- -- -- Q-constructor-β2 : (c : C) → Q-constructor c .fst ptB ≡ transport (λ i → P (push (inr c) i) .fst) (d (inl tt)) --- -- -- Q-constructor-β2 c = ((λ i → (Trunc.rec (Q c .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) ((isHLevelConnectedPoint _ conC ptC) c .fst)) .fst ptB)) ∙ --- -- -- {!isHLevelConnectedPoint _ conC ptC c .fst!} - --- -- -- gid1 : (a : A) (c : C) → g (f a) c ≡ d (inr (a , c)) --- -- -- gid1 a c i = (Q-constructor c .snd) i a - --- -- -- gid1' : {!!} --- -- -- gid1' = {!!} - --- -- -- gid1Path : (c : C) → PathP (λ i → P (push (inr c) i) .fst) (d (inl tt)) (g ptB c) --- -- -- gid1Path c i = --- -- -- hcomp (λ k → λ {(i = i0) → {!!} ; (i = i1) → {!!}}) --- -- -- {!!} - --- -- -- compPath : (c : C) → PathP _ (d (inl tt)) (g ptB c) --- -- -- compPath c = compPathP (λ i → d (push (inr c) i)) (compPathP (sym (gid1 ptA c)) (λ i → g (fpt i) c)) - - --- -- -- compPathTransport : (c : C) → ((λ z → P ((push (inr c) ∙ --- -- -- (λ i → inr (fpt (~ i) , c))) z) .fst) ∙ --- -- -- (λ i → ((λ i₁ → P (inr (f ptA , c)) .fst) ∙ (λ i₁ → P (inr (fpt i₁ , c)) .fst)) i)) --- -- -- ≡ λ i → P (push (inr c) i) .fst --- -- -- compPathTransport c = (λ k → ((λ z → P ((push (inr c) ∙ --- -- -- (λ i → inr (fpt (~ i) , c))) z) .fst) ∙ --- -- -- ((lUnit (λ i₁ → P (inr (fpt i₁ , c)) .fst) (~ k) )))) --- -- -- ∙ (λ k → ((λ z → P ((push (inr c) ∙ --- -- -- (λ i → inr (fpt (~ i ∨ k) , c))) z) .fst) ∙ --- -- -- (λ i₁ → P (inr (fpt (i₁ ∨ k) , c)) .fst) )) --- -- -- ∙ (λ k → ((λ z → P ((push (inr c) ∙ refl) z) .fst) ∙ refl)) --- -- -- ∙ (λ k → rUnit ((λ z → P ((rUnit (push (inr c)) (~ k)) z) .fst)) (~ k)) - --- -- -- compPathTransport-ptC : compPathTransport ptC ≡ {!!} --- -- -- compPathTransport-ptC = {!!} - --- -- -- compPathTrue : (c : C) → PathP (λ i → P (push (inr c) i) .fst) (d (inl tt)) (g ptB c) --- -- -- compPathTrue c = transport (λ i → PathP (λ j → compPathTransport c i j) (d (inl tt)) (g ptB c)) (compPath c) - - --- -- -- {- --- -- -- Goal: P (push (inr x) i) .fst --- -- -- ———— Boundary —————————————————————————————————————————————— --- -- -- i = i0 ⊢ d (inl tt) --- -- -- i = i1 ⊢ g ptB x --- -- -- -} - --- -- -- gid1transp : (c : C) → PathP {!!} (g ptB c) (transport (λ i → P (push (inr c) i) .fst) (d (inl tt))) --- -- -- gid1transp c = compPathP (λ i → g (fpt (~ i)) c) (compPathP (gid1 ptA c) (compPathP (λ i → d (push (inr c) (~ i))) λ i → transp (λ j → P (push (inr c) (i ∧ j)) .fst) (~ i) (d (inl tt)))) - - --- -- -- gid2 : (b : B) → g b ptC ≡ transport (λ i → P (push (inl b) i) .fst) (d (inl tt)) -- (λ i → P (push (inl b) i) .fst) (d (inl tt)) --- -- -- gid2 b = Q-constructor-β b - - --- -- -- gid1≡gid2 : (c : C) (a : A) → PathP {!gid1transp c !} (gid1 a ptC) (gid2 (f a)) --- -- -- gid1≡gid2 c a = compPathP {!λ j → (Q-constructor ptC .snd) i a!} (compPathP {!!} {!!}) - - --- -- -- gid2Path : (b : B) → PathP (λ i → P (push (inl b) i) .fst) (d (inl tt)) (g b ptC) --- -- -- gid2Path b i = --- -- -- hcomp (λ k → λ {(i = i0) → (d (inl tt)) ; (i = i1) → gid2 b (~ k)}) --- -- -- (transp (λ j → P (push (inl b) (j ∧ i)) .fst) (~ i) (d (inl tt))) - --- -- -- {- --- -- -- Goal: P (push (inl b) i) .fst --- -- -- ———— Boundary —————————————————————————————————————————————— --- -- -- i = i0 ⊢ d (inl tt) --- -- -- i = i1 ⊢ g b ptC --- -- -- -} - --- -- -- PathP2 : PathP {!!} ((gid2Path ptB)) (compPathTrue ptC) --- -- -- PathP2 = compPathP {!gid2Path ptB!} (compPathP {!!} {!compPathTrue ptC!}) - --- -- -- h : (x : (B , ptB) ⋀ (C , ptC)) → P x .fst --- -- -- h (inl x) = d (inl tt) --- -- -- h (inr (b , c)) = g b c --- -- -- h (push (inl b) i) = gid2Path b i --- -- -- h (push (inr x) i) = compPathTrue x i --- -- -- h (push (push a i₁) i) = {!!} diff --git a/Cubical/Homotopy/Freudenthal.agda b/Cubical/Homotopy/Freudenthal.agda index d48de91e8..41dfb486c 100644 --- a/Cubical/Homotopy/Freudenthal.agda +++ b/Cubical/Homotopy/Freudenthal.agda @@ -51,7 +51,7 @@ module _ {ℓ} (n : ℕ) {A : Pointed ℓ} (connA : isHLevelConnected (suc (suc ( (∀ t → isContr (fiber (fwd p a) t)) , isProp→isOfHLevelSuc n (isPropΠ λ _ → isPropIsContr) )) - + (isHLevelConnectedPoint (suc n) connA (pt A)) .equiv-proof (λ _ → Trunc.elim diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda index c57b7e1c2..f649b8475 100644 --- a/Cubical/Homotopy/Loopspace.agda +++ b/Cubical/Homotopy/Loopspace.agda @@ -77,13 +77,15 @@ Eckmann-Hilton {A = (A , pt)} zero α β = helper2 : (refl ∙ β) ⋆ (α ∙ refl) ≡ ((β ∙ refl) ⋆ (refl ∙ α)) helper2 = (λ i → (rUnit ((lUnit β) (~ i)) i) ⋆ (lUnit (rUnit α (~ i)) i)) - unit1 : {a : A} (γ : Path (a ≡ a) (λ _ → a) λ _ → a) → lUnit refl ∙ (refl ⋆ γ) ∙ sym (lUnit refl) ≡ γ + unit1 : {a : A} (γ : Path (a ≡ a) (λ _ → a) λ _ → a) + → lUnit refl ∙ (refl ⋆ γ) ∙ sym (lUnit refl) ≡ γ unit1 {a = a} γ = lUnit refl ∙ (refl ⋆ γ) ∙ sym (lUnit refl) ≡⟨ (λ i → (λ j → lUnit refl (j ∧ ~ i)) ∙ (λ j → lUnit (γ j) (~ i)) ∙ λ j → lUnit refl (~ i ∧ ~ j)) ⟩ refl ∙ γ ∙ refl ≡⟨ (λ i → lUnit (rUnit γ (~ i)) (~ i)) ⟩ γ ∎ - unit2 : {a : A} (γ : Path (a ≡ a) (λ _ → a) λ _ → a) → rUnit refl ∙ (γ ⋆ refl) ∙ sym (rUnit refl) ≡ γ + unit2 : {a : A} (γ : Path (a ≡ a) (λ _ → a) λ _ → a) + → rUnit refl ∙ (γ ⋆ refl) ∙ sym (rUnit refl) ≡ γ unit2 γ = rUnit refl ∙ (γ ⋆ refl) ∙ sym (rUnit refl) ≡⟨ (λ i → (λ j → rUnit refl (j ∧ ~ i)) ∙ (λ j → rUnit (γ j) (~ i)) ∙ λ j → rUnit refl (~ i ∧ ~ j)) ⟩ refl ∙ γ ∙ refl ≡⟨ (λ i → lUnit (rUnit γ (~ i)) (~ i)) ⟩ γ ∎ diff --git a/Cubical/ZCohomology/MayerVietoris.agda b/Cubical/ZCohomology/MayerVietoris.agda deleted file mode 100644 index 6bd2ad5cf..000000000 --- a/Cubical/ZCohomology/MayerVietoris.agda +++ /dev/null @@ -1,212 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.ZCohomology.MayerVietoris where - -open import Cubical.ZCohomology.Base -open import Cubical.HITs.S1 -open import Cubical.ZCohomology.Properties -open import Cubical.HITs.Sn -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Function -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Transport -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Univalence -open import Cubical.Data.NatMinusTwo.Base -open import Cubical.Data.Empty -open import Cubical.Data.Sigma hiding (_×_) -open import Cubical.HITs.Susp -open import Cubical.HITs.Wedge -open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) -open import Cubical.HITs.Nullification -open import Cubical.Data.Int hiding (_+_) -open import Cubical.Data.Nat -open import Cubical.Data.Prod -open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) -open import Cubical.Homotopy.Loopspace -open import Cubical.Homotopy.Connected -open import Cubical.Homotopy.Freudenthal -open import Cubical.HITs.SmashProduct.Base -open import Cubical.Data.Group.Base - -open import Cubical.ZCohomology.cupProdPrelims - - -open import Cubical.HITs.Pushout -open import Cubical.Data.Sum.Base -open import Cubical.Data.HomotopyGroup -open import Cubical.Data.Bool hiding (_⊕_) -{- -record AbGroup {ℓ} (A : Type ℓ) : Type ℓ where - constructor abgr - field - noll : A - _⊕_ : A → A → A - associate : (a b c : A) → (a ⊕ b) ⊕ c ≡ a ⊕ (b ⊕ c) - commute : (a b : A) → a ⊕ b ≡ (b ⊕ a) - r-unit : (a : A) → a ⊕ noll ≡ a - negate : (a : A) → Σ[ a⁻ ∈ A ] (a ⊕ a⁻ ≡ noll) --} -{- -grHom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (ϕ : A → B) - → AbGroup A → AbGroup B → Type (ℓ-max ℓ ℓ') -grHom {A = A} {B = B} ϕ (abgr 0A _+A_ as-A comm-A r-A neg-A) (abgr 0B _+B_ as-B comm-B r-B neg-B) = (x y : A) → ϕ (x +A y) ≡ (ϕ x +B ϕ y) - -ImHom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (ϕ : A → B) (A' : AbGroup A) (B' : AbGroup B) - → grHom ϕ A' B' → Type (ℓ-max ℓ ℓ') -ImHom {A = A} {B = B} ϕ Agr Bgr hom = Σ[ b ∈ B ] Σ[ a ∈ A ] ϕ a ≡ b - -KerHom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (ϕ : A → B) (A' : AbGroup A) (B' : AbGroup B) - → grHom ϕ A' B' → Type (ℓ-max ℓ ℓ') -KerHom {A = A} {B = B} ϕ Agr (abgr 0B _+B_ as-B comm-B r-B neg-B) hom = Σ[ a ∈ A ] ϕ a ≡ 0B --} - - ---- - - -private - variable - ℓ ℓ' ℓ'' : Level - A : Type ℓ - B : Type ℓ' - C : Type ℓ'' - A' : Pointed ℓ - B' : Pointed ℓ' - - -coHomFun : (n : ℕ) (f : A → B) → coHom n B → coHom n A -coHomFun n f = sRec setTruncIsSet λ β → ∣ β ∘ f ∣₀ - -coHomFun∙ : (n : ℕ) (f : A' →∙ B') → coHomRed n B' → coHomRed n A' -coHomFun∙ n (f , fpt) = sRec setTruncIsSet λ { (β , βpt) → ∣ β ∘ f , (λ i → β (fpt i)) ∙ βpt ∣₀} - -coHomFun∙2 : (n : ℕ) (f : A → B' .fst) → coHomRed n B' → coHom n A -coHomFun∙2 zero f = sRec setTruncIsSet λ { (β , βpt) → ∣ β ∘ f ∣₀} -coHomFun∙2 (suc n) f = sRec setTruncIsSet λ { (β , βpt) → ∣ β ∘ f ∣₀} - -module MV {ℓ ℓ' ℓ''} ((A , a₀) : Pointed ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) where - D : Pointed (ℓ-max (ℓ-max ℓ ℓ') ℓ'') - D = Pushout f g , inl a₀ - - -- σ : typ A → typ (Ω (∙Susp (typ A))) - -- σ a = merid a ∙ merid (pt A) ⁻¹ - - i : (n : ℕ) → coHomRed n D → coHomRed n (A , a₀) × coHom n B - i zero = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) - λ { (δ , δpt) → ∣ (λ x → δ (inl x)) , δpt ∣₀ , ∣ (λ x → δ (inr x)) ∣₀ } - i (suc n) = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) - λ { (δ , δpt) → ∣ (λ x → δ (inl x)) , δpt ∣₀ , ∣ (λ x → δ (inr x)) ∣₀ } - - Δ : (n : ℕ) → coHomRed n (A , a₀) × coHom n B → coHom n C - Δ n (α , β) = coHomFun∙2 n f α +ₕ (-ₕ (coHomFun n g β)) - - -- - - d-pre1 : (n : ℕ) → (C → coHomK n) → D .fst → coHomK-ptd (suc n) .fst - d-pre1 zero γ (inl x) = 0ₖ - d-pre1 zero γ (inr x) = 0ₖ - d-pre1 zero γ (push a i) = Kn→ΩKn+1 zero (γ a) i - d-pre1 (suc n) γ (inl x) = 0ₖ - d-pre1 (suc n) γ (inr x) = 0ₖ - d-pre1 (suc n) γ (push a i) = Kn→ΩKn+1 (suc n) (γ a) i - - d-pre1∙ : (n : ℕ) → (γ : (C → coHomK n)) → d-pre1 n γ (snd D) ≡ ∣ north ∣ - d-pre1∙ zero γ = refl - d-pre1∙ (suc n) γ = refl - - d-pre2 : (n : ℕ) → (C → coHomK n) → D →∙ coHomK-ptd (suc n) - d-pre2 n γ = d-pre1 n γ , d-pre1∙ n γ - - d : (n : ℕ) → coHom n C → coHomRed (suc n) D - d n = sRec setTruncIsSet λ a → ∣ d-pre2 n a ∣₀ - - iIsHom : (n : ℕ) (x y : coHomRed n D) → i n (+ₕ∙ n x y) ≡ +prod n (i n x) (i n y) - iIsHom zero = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) - λ { (f , fpt) (g , gpt) → refl} - iIsHom (suc n) = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) - λ { (f , fpt) (g , gpt) → refl} - - prodElim : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : ∥ A ∥₀ × ∥ B ∥₀ → Type ℓ''} - → ((x : ∥ A ∥₀ × ∥ B ∥₀) → isOfHLevel 2 (C x)) - → ((a : A) (b : B) → C (∣ a ∣₀ , ∣ b ∣₀)) - → (x : ∥ A ∥₀ × ∥ B ∥₀) → C x - prodElim {A = A} {B = B} {C = C} hlevel ind (a , b) = schonf a b - where - schonf : (a : ∥ A ∥₀) (b : ∥ B ∥₀) → C (a , b) - schonf = sElim (λ x → isOfHLevelΠ 2 λ y → hlevel (_ , _)) λ a → sElim (λ x → hlevel (_ , _)) - λ b → ind a b - - ΔIsHom : (n : ℕ) (x y : coHomRed n (A , a₀) × coHom n B) → Δ n (+prod n x y) ≡ (Δ n x +ₕ Δ n y) - ΔIsHom zero = prodElim (λ x → isOfHLevelΠ 2 λ y → isOfHLevelPath 2 setTruncIsSet _ _ ) - λ {(f' , fpt) x1 → prodElim (λ x → isOfHLevelPath 2 setTruncIsSet _ _) - λ {(g' , gpt) x2 → (λ i → ∣ (λ x → (f' (f x) +ₖ g' (f x)) +ₖ -distrₖ (x1 (g x)) (x2 (g x)) i) ∣₀) ∙ - (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ g' (f x)) (-ₖ x1 (g x)) (-ₖ x2 (g x)) (~ i)) ∣₀) ∙ - (λ i → ∣ (λ x → assocₖ (f' (f x)) (g' (f x)) (-ₖ x1 (g x)) i +ₖ (-ₖ x2 (g x))) ∣₀) ∙ - (λ i → ∣ (λ x → (f' (f x) +ₖ commₖ (g' (f x)) (-ₖ x1 (g x)) i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ - (λ i → ∣ (λ x → assocₖ (f' (f x)) (-ₖ x1 (g x)) (g' (f x)) (~ i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ - (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ (-ₖ x1 (g x))) (g' (f x)) (-ₖ x2 (g x)) i) ∣₀)}} - ΔIsHom (suc n) = prodElim (λ x → isOfHLevelΠ 2 λ y → isOfHLevelPath 2 setTruncIsSet _ _ ) - λ {(f' , fpt) x1 → prodElim (λ x → isOfHLevelPath 2 setTruncIsSet _ _) - λ {(g' , gpt) x2 → (λ i → ∣ (λ x → (f' (f x) +ₖ g' (f x)) +ₖ -distrₖ (x1 (g x)) (x2 (g x)) i) ∣₀) ∙ - (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ g' (f x)) (-ₖ x1 (g x)) (-ₖ x2 (g x)) (~ i)) ∣₀) ∙ - (λ i → ∣ (λ x → assocₖ (f' (f x)) (g' (f x)) (-ₖ x1 (g x)) i +ₖ (-ₖ x2 (g x))) ∣₀) ∙ - (λ i → ∣ (λ x → (f' (f x) +ₖ commₖ (g' (f x)) (-ₖ x1 (g x)) i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ - (λ i → ∣ (λ x → assocₖ (f' (f x)) (-ₖ x1 (g x)) (g' (f x)) (~ i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ - (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ (-ₖ x1 (g x))) (g' (f x)) (-ₖ x2 (g x)) i) ∣₀)}} - - helper : (h l : C → Int) (x : D .fst) → (d-pre1 zero h x +ₖ d-pre1 zero l x) ≡ d-pre1 zero (λ x → h x +ₖ l x) x - helper h l (inl x) = lUnitₖ 0ₖ - helper h l (inr x) = lUnitₖ 0ₖ - helper h l (push a i) = hcomp {!!} (cong (λ x → x i) helper1) - where - helper1 : cong₂ (λ x y → x +ₖ y) (Kn→ΩKn+1 zero (h a)) (Kn→ΩKn+1 zero (l a)) ≡ Kn→ΩKn+1 zero (h a +ₖ l a) - helper1 = {!!} - - pp : PathP (λ i → lUnitₖ (0ₖ {n = 1}) i ≡ lUnitₖ 0ₖ i) (λ i → Kn→ΩKn+1 zero (h a) i +ₖ Kn→ΩKn+1 zero (l a) i) (Kn→ΩKn+1 zero (h a +ₖ l a)) - pp = toPathP ((λ z → transport (λ i → lUnitₖ (0ₖ {n = 1}) i ≡ lUnitₖ 0ₖ i) (lUnit (rUnit (λ i → Kn→ΩKn+1 zero (h a) i +ₖ Kn→ΩKn+1 zero (l a) i) z) z )) - ∙ (λ z → transp (λ i → lUnitₖ (0ₖ {n = 1}) (i ∨ z) ≡ lUnitₖ 0ₖ (i ∨ z)) z - (((λ i → lUnitₖ (0ₖ {n = 1}) ((~ i) ∧ z)) ∙ (λ i → Kn→ΩKn+1 zero (h a) i +ₖ Kn→ΩKn+1 zero (l a) i) - ∙ λ i → lUnitₖ (0ₖ {n = 1}) (i ∧ z)))) - ∙ (λ z → sym ((λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ 1 i ∙ Kn→ΩKn+1 1 0ₖ)) ∙ - (λ i → ΩKn+1→Kn (lUnit (Kn→ΩKn+1 1 0ₖ) (~ i))) ∙ - Iso.leftInv (Iso2-Kn-ΩKn+1 1) 0ₖ) - ∙ (lUnit (rUnit (λ i → Kn→ΩKn+1 zero (h a) i +ₖ Kn→ΩKn+1 zero (l a) i) z) z) - ∙ ((λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ 1 i ∙ Kn→ΩKn+1 1 0ₖ)) ∙ - (λ i → ΩKn+1→Kn (lUnit (Kn→ΩKn+1 1 0ₖ) (~ i))) ∙ - Iso.leftInv (Iso2-Kn-ΩKn+1 1) 0ₖ)) - ∙ (λ z → sym (lUnitₖ (0ₖ {n = 1})) - ∙ (lUnit (rUnit refl z ) z ∙ (λ i → ΩKn+1→Kn (Kn→ΩKn+1 1 (Kn→ΩKn+1 zero (h a) i) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 zero (l a) i))) ∙ lUnit (rUnit refl z) z) - ∙ lUnitₖ (0ₖ {n = 1})) - ∙ ((λ z → sym (lUnitₖ (0ₖ {n = 1})) - ∙ (({!!} ∙ {!ΩKn+1→Kn (Kn→ΩKn+1 1 (Kn→ΩKn+1 zero (h a) i) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 zero (l a) i)!} ∙ λ i → {!l!}) - ∙ (λ i → Iso.leftInv (Iso2-Kn-ΩKn+1 1) {!!} {!!}) ∙ (λ i → (Iso.leftInv (Iso2-Kn-ΩKn+1 1) 0ₖ (~ i ∧ z))) ∙ {!!} ∙ {!!}) - ∙ lUnitₖ (0ₖ {n = 1}))) - ∙ {!!} - ∙ {!!} - ∙ {!!}) -- compPathP (λ j i → {!!}) (compPathP {!!} (compPathP {!!} {!!})) - - dIsHom : (n : ℕ) (x y : coHom n C) → d n (x +ₕ y) ≡ +ₕ∙ (suc n) (d n x) (d n y) - dIsHom zero = sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) - λ h l → (λ i → ∣ (λ x → helper h l x (~ i)) , lUnit (λ j → lUnitₖ 0ₖ (j ∨ (~ i))) i ∣₀) ∙ - (λ i → ∣ (λ x → d-pre1 zero h x +ₖ d-pre1 zero l x) , refl ∙ lUnitₖ 0ₖ ∣₀) - dIsHom (suc n) = {!d-pre1 zero h x +ₖ d-pre1 zero l x!} - --- record LES {ℓ} (B : ℕ → Type ℓ) (0' : {n : ℕ} → B n) (f : {n : ℕ} → B n → B (suc n)) (_+'_ : {n : ℕ} → B n → B n → B n) : Type ℓ where --- constructor les --- field --- grp : (n : ℕ) → AbGroup (B n) 0' _+'_ --- imIsKer : (n : ℕ) → {!!} - - --- {- --- record Ring {ℓ} (A : Type ℓ) (_+∙_ :) : Type (ℓ-max ℓ ℓ') where --- constructor iso --- field --- fun : A → B --- inv1 : B → A --- rightInv : section fun inv1 --- leftInv : retract fun inv1 --- -} diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index e02d08261..982fe86e5 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -72,7 +72,7 @@ coHomRed+1Equiv zero A i = ∥ helpLemma {C = (Int , pos 0)} i ∥₀ helper : (x : A ⊎ Unit) → ((helpmap.map1') (map2 (f , snd)) x) ≡ f x helper (inl x) = refl helper (inr tt) = sym snd - + coHomRed+1Equiv (suc n) A i = ∥ coHomRed+1.helpLemma A i {C = ((coHomK (suc n)) , ∣ north ∣)} i ∥₀ ----------- @@ -140,7 +140,7 @@ lCancelₖ {n = suc (suc n)} x = (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-Ω Iso.leftInv (Iso2-Kn-ΩKn+1 (suc (suc n))) 0ₖ -assocₖ : {n : ℕ} (x y z : coHomK n) → ((x +ₖ y) +ₖ z) ≡ (x +ₖ (y +ₖ z)) +assocₖ : {n : ℕ} (x y z : coHomK n) → ((x +ₖ y) +ₖ z) ≡ (x +ₖ (y +ₖ z)) assocₖ {n = n} x y z = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 n (ΩKn+1→Kn (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y)) ∙ Kn→ΩKn+1 n z)) ∙ (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 n) (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) i ∙ Kn→ΩKn+1 n z)) ∙ (λ i → ΩKn+1→Kn (assoc (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) (Kn→ΩKn+1 n z) (~ i))) ∙ @@ -175,7 +175,7 @@ private _+ₕ_ : {n : ℕ} → coHom n A → coHom n A → coHom n A _+ₕ_ = sElim2 (λ _ _ → §) λ a b → ∣ (λ x → a x +ₖ b x) ∣₀ --ₕ : {n : ℕ} → coHom n A → coHom n A +-ₕ : {n : ℕ} → coHom n A → coHom n A -ₕ = sRec § λ a → ∣ (λ x → -ₖ a x) ∣₀ 0ₕ : {n : ℕ} → coHom n A @@ -207,80 +207,13 @@ commₕ {n = n} = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) ----- Group structure of reduced cohomology groups --- - -coHomRedId : {A : Pointed ℓ} {n : ℕ} (x y : (A →∙ (coHomK-ptd n))) - → (p : Path (fst A → fst (coHomK-ptd n)) (x .fst) (y .fst)) - → transport (λ j → p j (snd A) ≡ snd (coHomK-ptd n)) (x .snd) ≡ y .snd - → Path (coHomRed n A) ∣ x ∣₀ ∣ y ∣₀ -coHomRedId {A = A}{n = n}(x , xpt) (y , ypt) p transpid i = ∣ (p i) , hcomp (λ k → λ {(i = i0) → xpt ; (i = i1) → transpid k}) (transp (λ j → p (i ∧ j) (snd A) ≡ snd (coHomK-ptd n)) (~ i) xpt) ∣₀ - -coHomRedId2 : {A : Pointed ℓ} {n : ℕ} (x y : (A →∙ (coHomK-ptd n))) - → (p : Path (fst A → fst (coHomK-ptd n)) (x .fst) (y .fst)) - → (λ j → p (~ j) (snd A)) ∙ (x .snd) ≡ y .snd - → Path (coHomRed n A) ∣ x ∣₀ ∣ y ∣₀ -coHomRedId2 {A = A} {n = n}(x , xpt) (y , ypt) p compid = - coHomRedId {n = n} ((x , xpt)) (y , ypt) p - ((λ i → transport (λ j → p j (snd A) ≡ snd (coHomK-ptd n)) (lUnit xpt i)) - ∙ (λ i → transport (λ j → p (j ∨ i) (snd A) ≡ snd (coHomK-ptd n)) ((λ j → p ((~ j) ∧ i) (snd A)) ∙ xpt)) - ∙ transportRefl ((λ j → p (~ j) (snd A)) ∙ xpt) - ∙ compid) - - -- test : {n : ℕ} → rUnitₖ (0ₖ {n = n}) ∙ sym (lUnitₖ 0ₖ) ≡ refl - -- test {n = zero} = rCancel (rUnitₖ (0ₖ {n = zero})) - -- test {n = suc n} = {!!} - -{- - inheritsStructure : ∀ {n} → (G : isGroup (coHomK (suc n))) - → isGroup.id G ≡ coHomK-ptd (suc n) .snd - → isGroup.inv G (coHomK-ptd (suc n) .snd) ≡ (coHomK-ptd (suc n) .snd) - → isGroup (coHomRed (suc n) A) - isGroup.id (inheritsStructure {n = n} (group-struct id inv comp lunit runit asso lcancel rcancel) 0id invid) = - ∣ (λ _ → id) , 0id ∣₀ - - isGroup.inv (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = - sRec § λ { (f , pt) → ∣ (λ a → inv (f a)) , (λ i → inv (pt i)) ∙ invid ∣₀} - isGroup.comp (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = - sElim2 (λ _ _ → §) λ {(f , ptf) (g , ptg) → ∣ (λ x → comp (f x) (g x)) , - (λ i → comp (ptf i) (ptg i)) - ∙ (λ i → comp (0id (~ i)) (∣ north ∣)) ∙ lunit ∣ north ∣ ∣₀} - isGroup.lUnit (inheritsStructure {n = n} (group-struct id inv comp lunit runit asso lcancel rcancel) 0id invid) = - sElim (λ _ → isOfHLevelPath 2 § _ _) λ { (f , pt) → coHomRedId2 {n = (suc n)} _ _ (funExt (λ x → lunit (f x))) - ((λ k → lUnit (sym (lunit (f (snd A))) ∙ (λ i → comp (0id i) (pt i)) - ∙ (λ i → comp (0id (~ i)) ∣ north ∣) ∙ (lunit ∣ north ∣)) k) - ∙ (λ k → (λ i → pt (i ∧ k)) ∙ (sym (lunit (pt k)) - ∙ (λ i → comp (0id i) (pt (i ∨ k))) ∙ (λ i → comp (0id (~ i)) (∣ north ∣)) - ∙ (lunit ∣ north ∣))) - ∙ helper pt (sym (lunit ∣ north ∣)) (λ i → comp (0id i) (∣ north ∣)))} - where - helper : ∀ {ℓ} {A : Type ℓ} {x y z w : A} (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) → p ∙ q ∙ r ∙ sym r ∙ sym q ≡ p - helper p q r = (λ i → p ∙ q ∙ assoc r (sym r) (sym q) i) - ∙ (λ i → p ∙ q ∙ rCancel r i ∙ sym q) - ∙ (λ i → p ∙ q ∙ lUnit (sym q) (~ i)) - ∙ (λ i → p ∙ rCancel q i) - ∙ sym (rUnit p) - isGroup.rUnit (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = - sElim (λ _ → isOfHLevelPath 2 § _ _) λ { (f , pt) → coHomRedId2 {n = (suc n)} _ _ (funExt (λ x → runit (f x))) - ((λ k → rUnit ((sym (runit (f (snd A)))) ∙ (λ i → comp (pt i) (0id i)) - ∙ (lUnit (lUnit (lUnit (lUnit (λ i → comp (0id (~ i)) ∣ north ∣) k) k) k) k) ∙ lunit ∣ north ∣) k) - ∙ (λ k → (sym (runit {!!}) ∙ (λ i → comp (pt i) {!!}) - ∙ ({!!} ∙ {!!} ∙ {!!} ∙ (λ i → {!!}) ∙ (λ i → comp (0id (~ i)) (pt (~ k)))) ∙ lunit (pt (~ k))) ∙ λ i → pt (i ∨ (~ k))) - ∙ (λ k → {!!}) - ∙ {!!} - ∙ {!!} - ∙ {!!} - ∙ {!!})} - isGroup.assoc (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = {!!} - isGroup.lCancel (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = {!!} - isGroup.rCancel (inheritsStructure {n = n} (group-struct id inv comp lunit runit assoc lcancel rcancel) 0id invid) = {!!} - --} +---- Group structure of reduced cohomology groups (in progress - might need K to compute properly first) --- +ₕ∙ : {A : Pointed ℓ} (n : ℕ) → coHomRed n A → coHomRed n A → coHomRed n A +ₕ∙ zero = sElim2 (λ _ _ → §) λ { (a , pa) (b , pb) → ∣ (λ x → a x +ₖ b x) , (λ i → (pa i +ₖ pb i)) ∣₀ } -- ∣ (λ x → a x +ₖ b x) ∣₀ -+ₕ∙ (suc n) = sElim2 (λ _ _ → §) λ { (a , pa) (b , pb) → ∣ (λ x → a x +ₖ b x) , (λ i → pa i +ₖ pb i) ∙ lUnitₖ 0ₖ ∣₀ } ++ₕ∙ (suc n) = sElim2 (λ _ _ → §) λ { (a , pa) (b , pb) → ∣ (λ x → a x +ₖ b x) , (λ i → pa i +ₖ pb i) ∙ lUnitₖ 0ₖ ∣₀ } --ₕ∙ : {A : Pointed ℓ} (n : ℕ) → coHomRed n A → coHomRed n A +-ₕ∙ : {A : Pointed ℓ} (n : ℕ) → coHomRed n A → coHomRed n A -ₕ∙ zero = sRec § λ {(a , pt) → ∣ (λ x → -ₖ a x ) , (λ i → -ₖ (pt i)) ∣₀} -ₕ∙ (suc zero) = sRec § λ {(a , pt) → ∣ (λ x → -ₖ a x ) , (λ i → -ₖ (pt i)) ∙ (λ i → ΩKn+1→Kn (sym (Kn→ΩKn+10ₖ (suc zero) i))) ∣₀} -ₕ∙ (suc (suc n)) = sRec § λ {(a , pt) → ∣ (λ x → -ₖ a x ) , (λ i → -ₖ (pt i)) ∙ @@ -291,100 +224,3 @@ coHomRedId2 {A = A} {n = n}(x , xpt) (y , ypt) p compid = 0ₕ∙ : {A : Pointed ℓ} (n : ℕ) → coHomRed n A 0ₕ∙ zero = ∣ (λ _ → 0ₖ) , refl ∣₀ 0ₕ∙ (suc n) = ∣ (λ _ → 0ₖ) , refl ∣₀ - - -module _ {A : Type ℓ} {A' : Pointed ℓ'} (n : ℕ) where - 0prod : coHomRed n A' × coHom n A - 0prod = (0ₕ∙ n) , 0ₕ - - -prod : coHomRed n A' × coHom n A → coHomRed n A' × coHom n A - -prod (a , b) = (-ₕ∙ n a) , (-ₕ b) - - +prod : coHomRed n A' × coHom n A → coHomRed n A' × coHom n A → coHomRed n A' × coHom n A - +prod (a , b) (c , d) = (+ₕ∙ n a c) , (b +ₕ d) --- rUnitₕ∙ : {n : ℕ} (x : coHomRed n A) → Path (coHomRed n A) (_+ₕ∙_ {n = n} x (0ₕ∙ {n = n})) x --- rUnitₕ∙ {n = zero} = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) --- λ { (a , pt) → {!coHomRedId2 {n = zero} ((λ x → a x +ₖ 0ₖ) , λ i → (_+ₖ_ {n = zero} (pt i) (pos 0))) (a , pt) (λ i → funExt (λ x → rUnitₖ (a x)) i) (helper a pt)!}} --- where --- helper : (a : fst A → Int) → (pt : a (snd A) ≡ pos 0) → sym (rUnitₖ (a (snd A))) ∙ (λ i → (pt i) +ₖ (pos 0)) ≡ pt --- helper a pt = (λ i → (rUnit (sym (rUnitₖ (a (snd A)))) i) ∙ (rUnit (λ i → (pt i) +ₖ (pos 0)) i)) --- ∙ (λ j → (sym (rUnitₖ (a (snd A))) ∙ λ i → rUnitₖ (a (snd A)) (i ∨ (~ j))) ∙ (λ i → rUnitₖ (pt (i ∧ (~ j))) j) ∙ {!!}) --- ∙ {!!} --- ∙ {!!} --- ∙ {!!} --- rUnitₕ∙ {n = suc n} = {!!} - --- -- lUnitₕ : {n : ℕ} (x : coHom n A) → 0ₕ +ₕ x ≡ x --- -- lUnitₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) --- -- λ a i → ∣ funExt (λ x → lUnitₖ (a x)) i ∣₀ - --- -- rCancelₕ : {n : ℕ} (x : coHom n A) → x +ₕ (-ₕ x) ≡ 0ₕ --- -- rCancelₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) --- -- λ a i → ∣ funExt (λ x → rCancelₖ (a x)) i ∣₀ - --- -- lCancelₕ : {n : ℕ} (x : coHom n A) → (-ₕ x) +ₕ x ≡ 0ₕ --- -- lCancelₕ = sElim (λ _ → isOfHLevelPath 1 (§ _ _)) --- -- λ a i → ∣ funExt (λ x → lCancelₖ (a x)) i ∣₀ - - - --- -- assocₕ : {n : ℕ} (x y z : coHom n A) → ((x +ₕ y) +ₕ z) ≡ (x +ₕ (y +ₕ z)) --- -- assocₕ = elim3 (λ _ _ _ → isOfHLevelPath 1 (§ _ _)) --- -- λ a b c i → ∣ funExt (λ x → assocₖ (a x) (b x) (c x)) i ∣₀ - - --- -- commₕ : {n : ℕ} (x y : coHom n A) → (x +ₕ y) ≡ (y +ₕ x) --- -- commₕ {n = n} = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) --- -- λ a b i → ∣ funExt (λ x → commₖ (a x) (b x)) i ∣₀ - - - - - - --- -- ---- Groups for reduced cohom --- - --- -- 0ₖ∙ : {n : ℕ} {A : Pointed ℓ} → coHomRed n A --- -- 0ₖ∙ {n = zero} {A = A} = ∣ (λ _ → 0ₖ) , refl ∣₀ --- -- 0ₖ∙ {n = suc n} {A = A} = ∣ (λ _ → 0ₖ) , refl ∣₀ - - - --- -- coHomGroup : (n : ℕ) (A : Type ℓ) → Group ℓ --- -- Group.type (coHomGroup n A) = coHom n A --- -- Group.setStruc (coHomGroup n A) = § --- -- isGroup.id (Group.groupStruc (coHomGroup n A)) = 0ₕ --- -- isGroup.inv (Group.groupStruc (coHomGroup n A)) = -ₕ --- -- isGroup.comp (Group.groupStruc (coHomGroup n A)) = _+ₕ_ --- -- isGroup.lUnit (Group.groupStruc (coHomGroup n A)) = lUnitₕ --- -- isGroup.rUnit (Group.groupStruc (coHomGroup n A)) = rUnitₕ --- -- isGroup.assoc (Group.groupStruc (coHomGroup n A)) = assocₕ --- -- isGroup.lCancel (Group.groupStruc (coHomGroup n A)) = lCancelₕ --- -- isGroup.rCancel (Group.groupStruc (coHomGroup n A)) = rCancelₕ - --- -- coHomPtdGroup : (n : ℕ) (A : Pointed ℓ) → Group ℓ --- -- Group.type (coHomPtdGroup n (A , a)) = coHomRed n (A , a) --- -- Group.setStruc (coHomPtdGroup n (A , a)) = § --- -- isGroup.id (Group.groupStruc (coHomPtdGroup zero (A , a))) = ∣ (λ _ → 0ₖ) , refl ∣₀ --- -- isGroup.id (Group.groupStruc (coHomPtdGroup (suc n) (A , a))) = ∣ (λ _ → 0ₖ) , refl ∣₀ --- -- isGroup.inv (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} --- -- isGroup.comp (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} --- -- isGroup.lUnit (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} --- -- isGroup.rUnit (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} --- -- isGroup.assoc (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} --- -- isGroup.lCancel (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} --- -- isGroup.rCancel (Group.groupStruc (coHomPtdGroup n (A , a))) = {!!} - --- -- -- -- Cup-prouct -- - --- -- -- smashₕ : (A : Type ℓ) (n m : ℕ) → Type ℓ --- -- -- smashₕ A n m = (coHom n A , ∣ (λ a → coHom-pt n) ∣₀) ⋀ (coHom n A , ∣ (λ a → coHom-pt n) ∣₀) - --- -- -- ⌣' : (n m : ℕ) → (coHomK (suc m) , coHom-pt (suc m)) ⋀ (coHomK (suc m) , coHom-pt (suc m)) --- -- -- → coHomK ((suc n) + (suc m)) --- -- -- ⌣' n m = equiv-proof (elim.isEquivPrecompose {A = ((S₊ (suc n)) , north) ⋁ ((S₊ (suc m)) , north)} --- -- -- (λ a → {!⋀!}) --- -- -- {!!} --- -- -- {!!} --- -- -- {!!}) {!!} .fst .fst - diff --git a/Cubical/ZCohomology/cupProd.agda b/Cubical/ZCohomology/cupProd.agda deleted file mode 100644 index e69de29bb..000000000 From 040147288b58ee3d8ece56a8fc1bca829709e664 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 4 May 2020 00:13:17 +0200 Subject: [PATCH 25/89] merge main --- Cubical/HITs/Truncation/Connected/Base.agda | 48 - .../Connected/FreudenthalProof/7-5-11.agda | 94 -- .../Connected/FreudenthalProof/7-5-14.agda | 56 - .../Connected/FreudenthalProof/7-5-7.agda | 103 -- .../Connected/FreudenthalProof/8-6-1.agda | 208 --- .../Connected/FreudenthalProof/8-6-2.agda | 150 -- .../Connected/FreudenthalProof/8-6-4.agda | 1355 ----------------- .../Connected/FreudenthalProof/8-6-4Code.agda | 235 --- .../Connected/FreudenthalProof/Prelim.agda | 345 ----- .../HITs/Truncation/Connected/Properties.agda | 182 --- 10 files changed, 2776 deletions(-) delete mode 100644 Cubical/HITs/Truncation/Connected/Base.agda delete mode 100644 Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-11.agda delete mode 100644 Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-14.agda delete mode 100644 Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-7.agda delete mode 100644 Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-1.agda delete mode 100644 Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-2.agda delete mode 100644 Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4.agda delete mode 100644 Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4Code.agda delete mode 100644 Cubical/HITs/Truncation/Connected/FreudenthalProof/Prelim.agda delete mode 100644 Cubical/HITs/Truncation/Connected/Properties.agda diff --git a/Cubical/HITs/Truncation/Connected/Base.agda b/Cubical/HITs/Truncation/Connected/Base.agda deleted file mode 100644 index 5535b29a8..000000000 --- a/Cubical/HITs/Truncation/Connected/Base.agda +++ /dev/null @@ -1,48 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Truncation.Connected.Base where - -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Transport -open import Cubical.Foundations.Isomorphism -open import Cubical.Data.NatMinusTwo.Base -open import Cubical.Data.Unit -open import Cubical.HITs.Truncation.Base - -{- n-connected functions -} -is-_-Connected : ∀{ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n : ℕ₋₂) (f : A → B) → Type (ℓ-max ℓ ℓ') -is-_-Connected {A = A} {B = B} n f = ((b : B) → isContr (∥ fiber f b ∥ n )) - -{- n-connected types -} -is-_-ConnectedType : ∀ {ℓ} (n : ℕ₋₂) (A : Type ℓ) → Type ℓ -is- n -ConnectedType A = isContr (∥ A ∥ n) - -is-_-ConnectedType2 : ∀ {ℓ} (n : ℕ₋₂) (A : Type ℓ) → Type ℓ -is- n -ConnectedType2 A = is- n -Connected (λ (x : A) → tt) - -private - conTypEqHelper : ∀ {ℓ} (A : Type ℓ) (n : ℕ₋₂) (b : Unit) → A ≡ fiber (λ (x : A) → b) b - conTypEqHelper A n tt = isoToPath (iso (λ x → x , refl) (λ x → fst x) (λ b i → (refl {x = fst b} i) , ((isOfHLevelSuc 1 (isPropUnit) tt tt (snd b) refl) i)) λ a → refl) - -{- The first def implies the second -} -conTyp→conTyp2 : ∀ {ℓ} (n : ℕ₋₂) (A : Type ℓ) → is- n -ConnectedType A → is- n -ConnectedType2 A -conTyp→conTyp2 n A iscon tt = transport (λ i → isContr (∥ conTypEqHelper A n tt i ∥ n)) iscon - -{- The second def implies the first -} -conTyp2→conTyp : ∀ {ℓ} (n : ℕ₋₂) (A : Type ℓ) → is- n -ConnectedType2 A → is- n -ConnectedType A -conTyp2→conTyp n A iscon = transport (λ i → isContr (∥ conTypEqHelper A n tt (~ i) ∥ n)) (iscon tt) - -{- n-truncated types -} -is-_-Truncated : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} - (n : ℕ₋₂) → (f : A → B) → Type (ℓ-max ℓ ℓ') -is-_-Truncated {A = A} {B = B} n f = (b : B) → isOfHLevel (2+ n) (fiber f b) - -{- Induced function for induction principle of n-connected functions -} -indConFun : ∀{ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} - {k : ℕ₋₂} - (f : A → B) → - (P : (B → HLevel ℓ'' (2+ k))) → - (((b : B) → fst (P b)) → (a : A) → fst (P (f a))) -indConFun f P = λ g a → g (f a) diff --git a/Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-11.agda b/Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-11.agda deleted file mode 100644 index ec56315f4..000000000 --- a/Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-11.agda +++ /dev/null @@ -1,94 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Truncation.Connected.FreudenthalProof.7-5-11 where - -open import Cubical.HITs.Truncation.Connected.Base -open import Cubical.HITs.Truncation.Connected.FreudenthalProof.7-5-7 - -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Univalence - -open import Cubical.Data.HomotopyGroup -open import Cubical.Data.NatMinusTwo.Base -open import Cubical.Data.Unit - -open import Cubical.HITs.Truncation.Base -open import Cubical.HITs.Truncation.Properties - -private - - Cor759 : ∀ {ℓ ℓ'} {A : Type ℓ} (n : ℕ₋₂) → is- n -ConnectedType A → (B : HLevel ℓ' (2+ n)) → isEquiv (λ (b : (fst B)) → λ (a : A) → b) - Cor759 {A = A} n isCon B = isEquivLemma (conInd-i→ii (λ (x : A) → tt) n (conTyp→conTyp2 n A isCon) (λ _ → B)) - where - - isEquivLemma : isEquiv (indConFun {A = A} {B = Unit} (λ x → tt) (λ _ → B)) → isEquiv (λ (b : (fst B)) → λ (a : A) → b) - isEquivLemma record { equiv-proof = eq } = record { equiv-proof = λ y → ((eq y .fst .fst tt) , (eq y .fst .snd)) , - λ z i → ((cong (λ y → (fst y) tt) ((eq y .snd) ((λ x → fst z) , snd z))) i) , - (cong (λ y → (snd y)) ((eq y .snd) ((λ x → fst z) , snd z))) i} - - Cor759← : ∀ {ℓ} {A : Type ℓ} (n : ℕ₋₂) → (∀ {ℓ'} (B : HLevel ℓ' (2+ n)) → isEquiv (λ (b : (fst B)) → λ (a : A) → b)) → is- n -ConnectedType A - Cor759← {A = A} n BEq = conTyp2→conTyp n A (conInd-iii→i (λ x → tt) n λ {ℓ'} P → conInd-ii→iii (λ x → tt) n P (compEquiv ((λ Q → Q tt) , helper P) (_ , BEq (P tt)) .snd)) - where - helper : ∀ {ℓ'} → (P : Unit → HLevel ℓ' (2+ n)) → isEquiv {A = ((b : Unit) → fst (P b))} λ Q → Q tt -- ∀ {ℓ'} P (ℓ'' : Unit) → fst (P₁ ℓ'') - helper P = isoToIsEquiv (iso (λ Q → Q tt) invFun (λ b → refl) λ b → refl) - where - invFun : fst (P tt) → (b : Unit) → fst (P b) - invFun Q tt = Q - -Lemma7-5-11 : ∀{ℓ} {A : Type ℓ} {a : A} → (n : ℕ₋₂) → (is- (suc₋₂ n) -ConnectedType A) → is- n -Connected (λ (x : Unit) → a) -Lemma7-5-11 {A = A} {a = a} n isCon = conInd-iii→i (λ _ → a) _ λ P → ((λ g → fst (helper P (g tt)))) , - (λ g → funExt (λ x → snd (helper P (g x)))) - where - helper : ∀ {ℓ} (P : (A → HLevel ℓ (2+ n) )) → (u : fst (P a)) → Σ ((a : A) → fst (P a)) λ f → f a ≡ u - helper {ℓ} P u = ((λ x → transport (λ i → Bs3 x (~ i)) (transport (λ i → Bs3 a i) u))) , - λ j → hcomp ((λ k → λ{(j = i0) → refl {x = ((transport (λ i → Bs3 a (~ i)) - (transport (λ i → Bs3 a i) u)))} k - ; (j = i1) → transportRefl (transportRefl u k) k})) - (transport (λ i → Bs3 a (~ i ∧ ~ j)) - (transport (λ i → Bs3 a (i ∧ (~ j))) u)) - where - Bs2 : Σ (HLevel ℓ (2+ n)) λ B → P ≡ λ (a : A) → B - Bs2 = (equiv-proof (Cor759 (suc₋₂ n) isCon (HLevel ℓ (2+ n) , isOfHLevelHLevel _ ))) - P .fst .fst , - sym ((equiv-proof (Cor759 (suc₋₂ n) isCon (HLevel ℓ (2+ n) , isOfHLevelHLevel _ ))) P .fst .snd) - - B* : Type ℓ - B* = fst (fst Bs2) - - Bs3 : (a : A) → fst (P a) ≡ B* - Bs3 a = cong (λ x → fst (x a)) (snd Bs2) - - - -Lemma7-5-11← : ∀{ℓ} {A : Type ℓ} {a : A} → (n : ℕ₋₂) → is- n -Connected (λ (x : Unit) → a) → (is- (suc₋₂ n) -ConnectedType A) -Lemma7-5-11← {A = A} {a = a} n iscon = Cor759← (suc₋₂ n) lemmas.isEquivMap - where - module lemmas {ℓ' : Level} (B : HLevel ℓ' (2+ (suc₋₂ n))) where - - P : (f : A → fst B) → A → Type _ - P f a2 = (f a2 ≡ f a) - - hLevelP : (f : A → fst B) → (a : A) → isOfHLevel (2+ n) (P f a) - hLevelP f a2 = helper n (f a2) (f a) (snd B) - where - helper : (n : ℕ₋₂) → (b1 b2 : (fst B)) → isOfHLevel (2+ suc₋₂ n) (fst B) → isOfHLevel (2+ n) (b1 ≡ b2) - helper neg2 b1 b2 hlvl = (hlvl b1 b2) , λ y → isProp→isSet hlvl b1 b2 _ _ - helper (-1+ n) b1 b2 hlvl = hlvl b1 b2 - - forAllP : (f : A → fst B) → (a : A) → P f a - forAllP f = equiv-proof (conInd-i→ii (λ x → a) n iscon (λ a → P f a , hLevelP f a)) (λ a → refl) .fst .fst - where - helper : (n : ℕ₋₂) → (b1 b2 : (fst B)) → isOfHLevel (2+ suc₋₂ n) (fst B) → isOfHLevel (2+ n) (b1 ≡ b2) - helper neg2 b1 b2 hlvl = (hlvl b1 b2) , λ y → isProp→isSet hlvl b1 b2 _ _ - helper (-1+ n) b1 b2 hlvl = hlvl b1 b2 - - sect : section (λ (b : fst B) (a : A) → b) λ f → f a - sect f = funExt λ y → sym ((forAllP f) y) - - isEquivMap : isEquiv (λ (b : fst B) (a : A) → b) - isEquivMap = isoToIsEquiv (iso (λ b a → b) - (λ f → f a) - sect - λ h → refl) diff --git a/Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-14.agda b/Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-14.agda deleted file mode 100644 index 203ac5178..000000000 --- a/Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-14.agda +++ /dev/null @@ -1,56 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Truncation.Connected.FreudenthalProof.7-5-14 where - -open import Cubical.HITs.Truncation.Connected.Base -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Isomorphism -open import Cubical.Data.NatMinusTwo.Base -open import Cubical.Data.Nat hiding (elim) -open import Cubical.HITs.Nullification hiding (elim) -open import Cubical.HITs.Truncation.Base -open import Cubical.HITs.Truncation.Properties - -private - variable - ℓ ℓ' : Level - A : Type ℓ - B : Type ℓ' - - -Lemma7-5-14 : (n : ℕ₋₂) (f : A → B) → (is- n -Connected f) → ∥ A ∥ n ≃ ∥ B ∥ n -Lemma7-5-14 {A = A} {B = B} neg2 f c = isoToEquiv (iso (λ _ → fst (isOfHLevelTrunc 0)) - (λ _ → fst (isOfHLevelTrunc 0)) - (λ b → snd (isOfHLevelTrunc 0) b) - (λ b → snd (isOfHLevelTrunc 0) b)) -Lemma7-5-14 {A = A} {B = B} (-1+ n) f c = isoToEquiv (iso - (∥ f ∥-fun (-1+ n)) - (elim (λ _ → isOfHLevelTrunc (suc n)) back) - (elim (λ x → (isOfHLevelSuc (suc n) (isOfHLevelTrunc (suc n))) _ _) backSection) - (elim (λ x → (isOfHLevelSuc (suc n) (isOfHLevelTrunc (suc n))) _ _) backRetract)) - where - back : B → ∥ A ∥ (-1+ n) - back y = ∥ fst ∥-fun (-1+ n) (fst (c y)) - - backSection : (b : B) → _≡_ {A = ∥ B ∥ (-1+ n)} - (elim (λ _ → isOfHLevelTrunc (suc n)) (λ a → ∣ f a ∣) (elim {n = suc n } - {B = λ _ → ∥ A ∥ (-1+ n)} (λ _ → isOfHLevelTrunc (suc n)) back ∣ b ∣)) - ∣ b ∣ - backSection b = helper {P = λ p → ∥ f ∥-fun (-1+ n) p ≡ ∣ b ∣} - (λ x → (isOfHLevelSuc (suc n) (isOfHLevelTrunc (suc n))) - (∥ f ∥-fun (-1+ n) (∥ fst ∥-fun (-1+ n) x)) ∣ b ∣) - (λ a p → cong (λ x → ∣ x ∣) p) - (fst (c b)) - - where - helper : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : A → Type ℓ'} {P : ∥ A ∥ (-1+ n) → Type ℓ''} → - ((x : ∥ Σ A B ∥ (-2+ ℕ₋₂.n (-1+ n))) → - isOfHLevel (2+ (-2+ ℕ₋₂.n (-1+ n))) (P (∥ fst ∥-fun (-1+ n) x))) → - ((a : A) (b : B a) → P ∣ a ∣) → - (p : ∥ Σ A B ∥ (-1+ n)) → P (∥ fst ∥-fun (-1+ n) p) - helper hlev pf = elim hlev λ pair → pf (fst pair) (snd pair) - - backRetract : (a : A) → ∥ fst ∥-fun (-1+ n) (fst (c (f a))) ≡ ∣ a ∣ - backRetract a = cong (∥ fst ∥-fun (-1+ n)) - (snd (c (f a)) ∣ a , refl ∣) diff --git a/Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-7.agda b/Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-7.agda deleted file mode 100644 index 8f621cc5d..000000000 --- a/Cubical/HITs/Truncation/Connected/FreudenthalProof/7-5-7.agda +++ /dev/null @@ -1,103 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Truncation.Connected.FreudenthalProof.7-5-7 where - -open import Cubical.HITs.Truncation.Connected.Base -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Equiv.Properties -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Prelude -open import Cubical.Data.Nat hiding (elim) -open import Cubical.Foundations.Isomorphism -open import Cubical.Data.NatMinusTwo.Base -open import Cubical.HITs.Nullification hiding (elim) -open import Cubical.HITs.Truncation.Base -open import Cubical.HITs.Truncation.Properties - -private - variable - ℓ ℓ' : Level - A : Type ℓ - B : Type ℓ' - - -conInd-i→ii : (f : A → B) (n : ℕ₋₂) → - is- n -Connected f → - (P : B → HLevel ℓ (2+ n)) → - isEquiv (indConFun f P) -conInd-i→ii {A = A} {B = B} f n isCon P = isoToEquiv (compIso (compIso (compIso firstIso secondIso) thirdIso) fourthIso ) .snd - where - fib₀ : (b : B) → ∥ fiber f b ∥ n - fib₀ b = isCon b .fst - - fibContr : (b : B) (y x : ∥ fiber f b ∥ n) → y ≡ x - fibContr b x y = sym (isCon b .snd x) ∙ ((isCon b .snd) y) - - firstIso : Iso ((b : B) → (fst (P b))) (((b : B) → ((∥ fiber f b ∥ n) → fst (P b)))) - firstIso = iso (λ f b x₁ → f b) - (λ f b → f b (fib₀ b)) - (λ b → funExt λ x → funExt λ y i → b x ((fibContr x (isCon x .fst) y) i)) - λ a → refl - - secondIso : Iso (((b : B) → ((∥ fiber f b ∥ n) → fst (P b)))) ((b : B) → (fiber f b) → fst (P b)) - secondIso = iso (λ g b → univTrunc.fun n {B = (P b)} (g b)) - (λ g b → univTrunc.inv n {B = (P b)} (g b)) - (λ b → funExt λ x → univTrunc.sect n {B = P x} (b x)) - λ b → funExt λ x → univTrunc.retr n (b x) - - thirdIso : Iso ((b : B) → (fiber f b) → fst (P b)) ((b : B) (a : A) (p : (f a) ≡ b) → fst (P b)) - thirdIso = iso (λ g → (λ x b x₁ → g x (b , x₁))) - (λ g b fib → g b (fib .fst) (fib .snd)) - (λ x → refl) - λ x → refl - - fourthIso : Iso ((b : B) (a : A) (p : (f a) ≡ b) → fst (P b)) ((a : A) → fst (P (f a))) - fourthIso = iso (λ g a → g (f a) a refl) - (λ g b a id → transport (λ i → fst (P (id i))) (g a)) - (λ x → funExt (λ a → transportRefl (x a))) - λ b → funExt λ x → funExt λ a → funExt (J (λ x p → transport (λ i → fst (P (p i))) (b (f a) a (λ _ → f a)) ≡ b x a p) - (transportRefl (b (f a) a refl))) - - -conInd-ii→iii : ∀ {ℓ} (f : A → B) (n : ℕ₋₂) → - (P : B → HLevel ℓ (2+ n)) → - isEquiv (indConFun f P) → - hasSection (indConFun f P) -conInd-ii→iii f n P record { equiv-proof = eqpf } = (λ g → (eqpf g) .fst .fst) , (λ b → (eqpf b) .fst .snd) - -conInd-i→iii : ∀ {ℓ} (f : A → B) (n : ℕ₋₂) → - is- n -Connected f → - (P : B → HLevel ℓ (2+ n)) → - hasSection (indConFun f P) -conInd-i→iii f n isCon P = conInd-ii→iii f n P (conInd-i→ii f n isCon P ) - -conInd-iii→i : (f : A → B) (n : ℕ₋₂) → - (∀ {ℓ} (P : B → HLevel ℓ (2+ n)) → - hasSection (indConFun f P)) → - is- n -Connected f -conInd-iii→i {A = A} {B = B} f n P→hasSection = λ b → (c n P→hasSection b) , (λ y → sym (fun n P→hasSection b y)) - where - P : (n : ℕ₋₂) → (∀ {ℓ} (P : B → HLevel ℓ (2+ n)) → hasSection (indConFun f P)) → B → Type _ - P n s b = ∥ fiber f b ∥ n - - c : (n : ℕ₋₂) → (∀ {ℓ} (P : B → HLevel ℓ (2+ n)) → hasSection (indConFun f P)) → (b : B) → ∥ fiber f b ∥ n - c n s = (s λ b → ( ∥ fiber f b ∥ n , isOfHLevelTrunc _)) .fst λ a → ∣ a , refl ∣ - - fun : (n : ℕ₋₂) → (s : (∀ {ℓ} (P : B → HLevel ℓ (2+ n)) → hasSection (indConFun f P))) → (b : B) (w : (∥ fiber f b ∥ n) ) → w ≡ c n s b - fun neg2 s b w = isOfHLevelSuc (2+ neg2) (isOfHLevelTrunc _) w (c neg2 s b) - fun (-1+ n) s b = elim (λ x → (isOfHLevelSuc (2+ (-1+ n)) (isOfHLevelTrunc {A = (fiber f b)} (suc n))) x (c (-1+ n) s b) ) (λ a → witness b (fst a) (snd a)) - where - eqtyp : ((a : A) → ∣ (a , refl {x = f a}) ∣ ≡ c (-1+ n) s (f a)) ≡ ((b : B) (a : A) (p : f (a) ≡ b) → ∣ (a , p) ∣ ≡ c (-1+ n) s b) - eqtyp = isoToPath (iso - (λ g b a → J (λ b p → ∣ a , p ∣ ≡ c (-1+ n) s b) (g a)) - (λ g a → g (f a) a refl) - (λ h → funExt λ b → - funExt λ a → - funExt (J (λ b x → (J (λ b₂ p → ∣ a , p ∣ ≡ c (-1+ n) s b₂) (h (f a) a (λ _ → f a)) x ≡ h b a x)) - (JRefl (λ b₂ p → ∣ a , p ∣ ≡ c (-1+ n) s b₂) (h (f a) a (λ _ → f a))))) - λ h → funExt λ a → JRefl (λ b₁ p → ∣ a , p ∣ ≡ c (-1+ n) s b₁) (h a)) - - c* : ((a : A) → ∣ (a , refl {x = f a}) ∣ ≡ c (-1+ n) s (f a)) - c* = λ a → sym (cong (λ x → x a) ((s λ b → ( ∥ fiber f b ∥ (-1+ n) , isOfHLevelTrunc _)) .snd λ a → ∣ a , refl ∣)) - - witness : ((b : B) (a : A) (p : f (a) ≡ b) → ∣ (a , p) ∣ ≡ c (-1+ n) s b) - witness = transport (λ i → eqtyp i) c* diff --git a/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-1.agda b/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-1.agda deleted file mode 100644 index 7e4b9f933..000000000 --- a/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-1.agda +++ /dev/null @@ -1,208 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Truncation.Connected.FreudenthalProof.8-6-1 where - - -open import Cubical.HITs.Truncation.Connected.Base -open import Cubical.HITs.Truncation.Connected.FreudenthalProof.7-5-7 -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.Transport -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Data.NatMinusTwo.Base -open import Cubical.Data.Empty -open import Cubical.Data.Sigma -open import Cubical.HITs.Nullification hiding (elim) -open import Cubical.Data.Int hiding (_+_ ; _-_ ; +-comm ) -open import Cubical.Data.Nat -open import Cubical.HITs.Truncation.Base -open import Cubical.HITs.Truncation.Properties - -open import Cubical.HITs.Pushout -open import Cubical.Data.HomotopyGroup -private - variable - ℓ ℓ' : Level - A : Type ℓ - B : Type ℓ' - -{- This file contains lemma 8.6.1. from the HoTT book -} - -private -{- The following two lemmas transform the expression as in the HoTT book -} - Lemma861-fibId : ∀{ℓ} (n : ℕ₋₂) (k : ℕ) (f : A → B) → - (is- n -Connected f) → - (P : B → HLevel ℓ (((suc k) + (2+ n)))) - (l : (a : A) → fst (P (f a))) (a b : Σ ((b : B) → fst (P b)) λ g → (a : A) → g (f a) ≡ l a ) → - (pf : ((x : B) → isOfHLevel (((k) + (2+ n))) (fst a x ≡ fst b x))) → -- silly, but needed to be able to state the lemma - (a ≡ b) ≡ (fiber (indConFun {k = -2+ (k + (2+ n)) } f (λ x → ((fst a) x ≡ (fst b) x ) , pf x )) - (λ x → ((snd a) x) ∙ (sym ((snd b) x)))) - Lemma861-fibId {A = A} {B = B} n k f isCon P l (g , p) (h , q) pf = ua (compEquiv (helper1 l (g , p) (h , q)) (helper2 l (g , p) (h , q))) ∙ - (λ j → Σ ((x : B) → g x ≡ h x) (λ r → helper3 l (g , p) (h , q) r j)) ∙ - λ j → Σ ((x : B) → g x ≡ h x) (λ r → helper4 l (g , p) (h , q) r j) - where - helper1 : (l : (a : A) → fst (P (f a))) (a b : Σ ((b : B) → fst (P b)) λ g → (a : A) → g (f a) ≡ l a ) → (a ≡ b) ≃ Σ (fst a ≡ fst b) λ r → PathP (λ i → (x : A) → (r i) (f x) ≡ l x) (snd a) (snd b) - helper1 l a b = invEquiv Σ≡ - - helper2 : (l : (a : A) → fst (P (f a))) (a b : Σ ((b : B) → fst (P b)) λ g → (a : A) → g (f a) ≡ l a ) → (Σ (fst a ≡ fst b) λ r → PathP (λ i → (x : A) → (r i) (f x) ≡ l x) (snd a) (snd b)) ≃ Σ ((x : B) → (fst a) x ≡ (fst b) x) λ r → PathP (λ i → (x : A) → r (f x) i ≡ l x) (snd a) (snd b) - helper2 l (g , p) (h , q) = isoToEquiv (iso lrFun rlFun (λ b → refl) λ b → refl) where - lrFun : (Σ (g ≡ h) λ r → PathP (λ i → (x : A) → (r i) (f x) ≡ l x) p q) → Σ ((x : B) → g x ≡ h x) λ r → PathP (λ i → (x : A) → r (f x) i ≡ l x) p q - lrFun (a , b) = funExt⁻ a , b - rlFun : (Σ ((x : B) → g x ≡ h x) λ r → PathP (λ i → (x : A) → r (f x) i ≡ l x) p q) → (Σ (g ≡ h) λ r → PathP (λ i → (x : A) → (r i) (f x) ≡ l x) p q) - rlFun (a , b) = (funExt a) , b - - helper3 : (l : (a : A) → fst (P (f a))) (a b : Σ ((b : B) → fst (P b)) λ g → (a : A) → g (f a) ≡ l a ) → (r : (x : B) → (fst a) x ≡ (fst b) x) → (PathP (λ i → (x : A) → r (f x) i ≡ l x) (snd a) (snd b)) ≡ ((x : A) → r (f x) ≡ ((snd a) x) ∙ sym ((snd b) x) ) - helper3 l (g , p) (h , q) r j = hcomp (λ k → λ{ (j = i0) → (PathP (λ i → (x : A) → r (f x) i ≡ l x) p λ x → (sym (lUnit (q x)) k )) - ; (j = i1) → throwAbout k}) - ((PathP (λ i → (x : A) → r (f x) (i ∧ ~ j) ≡ l x) p λ x → (λ i → r (f x) (~ j ∨ i)) ∙ (q x))) - where - throwAbout : (p ≡ λ x → r (f x) ∙ q x) ≡ ((x : A) → (((r (f x))) ≡ (p x) ∙ (sym (q x)))) - throwAbout = isoToPath (iso ((λ g x → transport (λ i → throwAboutHelper2 (sym (q x)) (p x) (r (f x)) i ) (funExt⁻ g x) )) - (λ g → funExt λ x → transport (λ i → throwAboutHelper2 (sym (q x)) (p x) (r (f x)) (~ i) ) (g x)) - (λ b → funExt (λ x → transportTransport⁻ (throwAboutHelper2 (sym (q x)) (p x) (r (f x))) (b x))) - λ b → λ j → funExt (λ x → (transport⁻Transport (throwAboutHelper2 (sym (q x)) (p x) (r (f x))) (λ i → b i x)) j)) - where - - symId : ∀{ℓ} {A : Type ℓ} (a b : A) → (a ≡ b) ≡ (b ≡ a) - symId a b = isoToPath (iso sym sym (λ b → refl) λ b → refl) - - throwAboutHelper2 : ∀{ℓ} {A : Type ℓ} {a b c : A} → (r : b ≡ c) (p : a ≡ b) (q : a ≡ c) → (p ≡ q ∙ (sym r) ) ≡ (q ≡ p ∙ r) - throwAboutHelper2 {a = a} {b = b} = J (λ c r → (p : a ≡ b) (q : a ≡ c) → - (p ≡ q ∙ (sym r) ) ≡ (q ≡ p ∙ r)) - λ p q → (λ i → p ≡ rUnit q (~ i)) ∙ - symId p q ∙ λ i → q ≡ (rUnit p i) - - helper4 : (l : (a : A) → fst (P (f a))) (a b : Σ ((b : B) → fst (P b)) λ g → (a : A) → g (f a) ≡ l a ) → - (r : (x : B) → (fst a) x ≡ (fst b) x) → - ((x : A) → r (f x) ≡ ((snd a) x) ∙ sym ((snd b) x) ) - ≡ ((λ (x : A) → r (f x)) ≡ (λ (x : A) → ((snd a) x) ∙ sym ((snd b) x))) - helper4 l (h , q) (g , p) r = isoToPath (iso funExt funExt⁻ (λ b → refl) λ b → refl) - - - Lemma861-fibId2 : ∀{ℓ} (n : ℕ₋₂) (k : ℕ) (f : A → B) → - (is- n -Connected f) → - (P : B → HLevel ℓ ((suc k) + (2+ n))) → - (l : (a : A) → fst (P (f a))) → - (fiber (indConFun f P) l) - ≡ Σ ((b : B) → fst (P b)) λ g → ((a : A) → g (f a) ≡ l a) - Lemma861-fibId2 n k f isCon P l = isoToPath (iso (λ p → fst p , funExt⁻ (snd p)) (λ p → fst p , funExt (snd p)) (λ b → refl) λ b → refl) - -{- - Lemma861 : ∀{ℓ} (n : ℕ₋₂) (k : ℕ) (f : A → B) → - (is- n -Connected f) → - (P : B → HLevel ℓ (k + (2+ n))) → - is- (-2+ k) -Truncated (indConFun {k = -2+ (k + (2+ n))} f P) - Lemma861 {B = B} {ℓ = ℓ} n zero f iscon P = equiv-proof (conInd-i→ii f n iscon P) - where - helper : (n : ℕ₋₂) → -2+ (ℕ₋₂.n n + zero) ≡ n - helper neg2 = refl - helper neg1 = refl - helper (ℕ→ℕ₋₂ n) i = ℕ→ℕ₋₂ (+-zero n i) - {- Because of the different HLevels, we hve to repeat essentially the same proof three times... -} - Lemma861 {A = A} {B = B} {ℓ = ℓ} neg2 (suc zero) f iscon P l = transport (λ i → isOfHLevel (suc (zero)) - (Lemma861-fibId2 neg2 (zero) f iscon P l (~ i))) - λ a b → transport (λ i → isOfHLevel (zero) - (Lemma861-fibId neg2 (zero) f iscon P l a b - (λ x → ((snd (P x)) (fst a x) (fst b x)) , - (λ y → isOfHLevelSuc 1 (snd (P x)) - (fst a x) (fst b x) _ _)) (~ i))) - ((Lemma861 neg2 (zero) f iscon - ((λ x → ((fst a) x ≡ (fst b) x ) , - ((snd (P x)) (fst a x) (fst b x) , - (λ y → isOfHLevelSuc 1 (snd (P x)) - (fst a x) (fst b x) _ _)) ))) - λ z → (snd a z) ∙ sym (snd b z)) .fst - Lemma861 {A = A} {B = B} {ℓ = ℓ} (-1+ n) (suc zero) f iscon P l = transport (λ i → isOfHLevel (suc (zero)) - (Lemma861-fibId2 (-1+ n) (zero) f iscon P l (~ i))) - λ a b → transport (λ i → isOfHLevel (zero) - (Lemma861-fibId (-1+ n) (zero) f iscon P l a b - (λ x → snd (P x) (fst a x) (fst b x)) (~ i))) - ((Lemma861 (-1+ n) (zero) f iscon - ((λ x → ((fst a) x ≡ (fst b) x ) , - (snd (P x) (fst a x) (fst b x)) ))) - λ z → (snd a z) ∙ sym (snd b z)) .fst - - - Lemma861 {A = A} {B = B} {ℓ = ℓ} n (suc (suc k)) f iscon P l = transport (λ i → isOfHLevel (suc (suc k)) - (Lemma861-fibId2 n (suc k) f iscon P l (~ i))) - λ a b → transport (λ i → isOfHLevel (suc k) - (Lemma861-fibId n (suc k) f iscon P l a b - (λ x → (snd (P x)) (fst a x) (fst b x)) (~ i))) - ((Lemma861 n (suc k) f iscon - ((λ x → ((fst a) x ≡ (fst b) x ) , - (snd (P x)) (fst a x) (fst b x) ))) - λ z → ((snd a z) ∙ sym (snd b z))) --} - - Lemma861 : ∀{ℓ} (n : ℕ₋₂) (k : ℕ) (f : A → B) → - (is- n -Connected f) → - (P : B → HLevel ℓ (k + (2+ n))) → - is- (-2+ k) -Truncated (indConFun {k = -2+ (k + (2+ n))} f P) - Lemma861 {B = B} {ℓ = ℓ} neg2 zero f iscon P = equiv-proof (conInd-i→ii f neg2 iscon P) - Lemma861 {B = B} {ℓ = ℓ} neg2 (suc zero) f iscon P l = transport (λ i → isOfHLevel (suc (zero)) - (Lemma861-fibId2 neg2 (zero) f iscon P l (~ i))) - λ a b → transport (λ i → isOfHLevel (zero) - (Lemma861-fibId neg2 (zero) f iscon P l a b - (λ x → ((snd (P x)) (fst a x) (fst b x)) , - (λ y → isOfHLevelSuc 1 (snd (P x)) - (fst a x) (fst b x) _ _)) (~ i))) - ((Lemma861 neg2 (zero) f iscon - ((λ x → ((fst a) x ≡ (fst b) x ) , - ((snd (P x)) (fst a x) (fst b x) , - (λ y → isOfHLevelSuc 1 (snd (P x)) - (fst a x) (fst b x) _ _)) ))) - λ z → (snd a z) ∙ sym (snd b z)) .fst - Lemma861 {B = B} {ℓ = ℓ} neg2 (suc (suc k)) f iscon P l = transport (λ i → isOfHLevel (suc (suc k)) - (Lemma861-fibId2 neg2 (suc k) f iscon P l (~ i))) - λ a b → transport (λ i → isOfHLevel (suc k) - (Lemma861-fibId neg2 (suc k) f iscon P l a b - (λ x → (snd (P x)) (fst a x) (fst b x)) (~ i))) - ((Lemma861 neg2 (suc k) f iscon - ((λ x → ((fst a) x ≡ (fst b) x ) , - (snd (P x)) (fst a x) (fst b x) ))) - λ z → ((snd a z) ∙ sym (snd b z))) - Lemma861 {B = B} {ℓ = ℓ} (-1+ n) zero f iscon P = equiv-proof (conInd-i→ii f (-1+ n) iscon P) - Lemma861 {B = B} {ℓ = ℓ} (-1+ n) (suc zero) f iscon P l = transport (λ i → isOfHLevel (suc (zero)) - (Lemma861-fibId2 (-1+ n) (zero) f iscon P l (~ i))) - λ a b → transport (λ i → isOfHLevel (zero) - (Lemma861-fibId (-1+ n) (zero) f iscon P l a b - (λ x → snd (P x) (fst a x) (fst b x)) (~ i))) - ((Lemma861 (-1+ n) (zero) f iscon - ((λ x → ((fst a) x ≡ (fst b) x ) , - (snd (P x) (fst a x) (fst b x)) ))) - λ z → (snd a z) ∙ sym (snd b z)) .fst - Lemma861 {B = B} {ℓ = ℓ} (-1+ n) (suc (suc k)) f iscon P l = transport (λ i → isOfHLevel (suc (suc k)) - (Lemma861-fibId2 (-1+ n) (suc k) f iscon P l (~ i))) - λ a b → transport (λ i → isOfHLevel (suc k) - (Lemma861-fibId (-1+ n) (suc k) f iscon P l a b - (λ x → (snd (P x)) (fst a x) (fst b x)) (~ i))) - ((Lemma861 (-1+ n) (suc k) f iscon - ((λ x → ((fst a) x ≡ (fst b) x ) , - (snd (P x)) (fst a x) (fst b x) ))) - λ z → ((snd a z) ∙ sym (snd b z))) - -{- The following more general versions are more useful -} -Lemma861Gen : ∀{ℓ} (n : ℕ₋₂) (k : ℕ) (expr : ℕ₋₂ → ℕ → ℕ) → - ((expr n k) ≡ (k + (2+ n))) → - (f : A → B) → - (is- n -Connected f) → - (P : B → HLevel ℓ (expr n k)) → - is- (-2+ k) -Truncated (indConFun {k = -2+ (expr n k)} f P) -Lemma861Gen {A = A} {B = B} {ℓ = ℓ} n k expr path f iscon = transport (λ i → helper i) (Lemma861 n k f iscon) - where - helper : ((P : B → HLevel ℓ (k + (2+ n))) → is- -2+ k -Truncated (indConFun f P)) ≡ ((P : B → HLevel ℓ (expr n k)) → is- -2+ k -Truncated (indConFun f P)) - helper i = ((P : B → HLevel ℓ (path (~ i))) → is- -2+ k -Truncated (indConFun f P)) - -Lemma861GenNats : ∀{ℓ} (n k : ℕ) (expr : ℕ → ℕ → ℕ) → - ((expr n k) ≡ k + (2+ ℕ→ℕ₋₂ n) ) → - (f : A → B) → - (is- (ℕ→ℕ₋₂ n) -Connected f) → - (P : B → HLevel ℓ (expr n k)) → - is- (-2+ k) -Truncated (indConFun {k = -2+ (expr n k)} f P) -Lemma861GenNats {A = A} {B = B} {ℓ = ℓ} n k expr path f iscon = transport (λ i → helper (~ i)) (Lemma861 (ℕ→ℕ₋₂ n) k f iscon) - where - helper : ((P : B → HLevel ℓ (expr n k)) → is- -2+ k -Truncated (indConFun f P)) ≡ ((P : B → HLevel ℓ (k + (2+ ℕ→ℕ₋₂ n))) → is- -2+ k -Truncated (indConFun f P)) - helper i = ((P : B → HLevel ℓ (path i)) → is- -2+ k -Truncated (indConFun f P)) diff --git a/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-2.agda b/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-2.agda deleted file mode 100644 index 93b57d7c3..000000000 --- a/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-2.agda +++ /dev/null @@ -1,150 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Truncation.Connected.FreudenthalProof.8-6-2 where - -open import Cubical.Foundations.Prelude - -open import Cubical.HITs.Truncation.Connected.Base -open import Cubical.HITs.Truncation.Connected.FreudenthalProof.Prelim -open import Cubical.HITs.Truncation.Connected.FreudenthalProof.7-5-7 -open import Cubical.HITs.Truncation.Connected.FreudenthalProof.7-5-11 -open import Cubical.HITs.Truncation.Connected.FreudenthalProof.8-6-1 -open import Cubical.HITs.Truncation.Base -open import Cubical.HITs.Truncation.Properties - -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.GroupoidLaws - -open import Cubical.Data.NatMinusTwo.Base -open import Cubical.Data.Sigma -open import Cubical.Data.Prod.Base -open import Cubical.Data.Nat - - -private - variable - ℓ ℓ' : Level - A : Type ℓ - B : Type ℓ' - - -private - transpHelper : {a b c : A} (r : b ≡ c) (q : a ≡ b) (p : a ≡ c) → - PathP (λ i → a ≡ r i) q p → - q ∙ r ≡ p - transpHelper {a = a} {b = b} = J (λ c r → (q : a ≡ b) (p : a ≡ c) → PathP (λ i → a ≡ r i) q p → q ∙ r ≡ p) λ p q path → sym (rUnit p) ∙ path - - pt-map : (A : Pointed ℓ) → Unit → typ A - pt-map A x = pt A - - conOfpt-map : (A : Pointed ℓ) (n : ℕ) → (is- (ℕ→ℕ₋₂ n) -ConnectedType (typ A)) → is- (pred₋₂ (ℕ→ℕ₋₂ n)) -Connected (pt-map A) - conOfpt-map A n conA = Lemma7-5-11 (pred₋₂ (ℕ→ℕ₋₂ n)) (transport (λ i → (is- pmId n (~ i) -ConnectedType (typ A))) conA) - - module prelims (A : Pointed ℓ) (B : Pointed ℓ') (n m : ℕ) - (conA : is- (ℕ→ℕ₋₂ n) -ConnectedType (typ A)) - (conB : is- (ℕ→ℕ₋₂ m) -ConnectedType (typ B)) - (P : (typ A) → (typ B) → HLevel (ℓ-max ℓ ℓ') (2+ (ℕ→ℕ₋₂ (n + m)))) - (f : ((a : (typ A)) → fst (P a (pt B)))) - (g : ((b : (typ B)) → fst (P (pt A) b))) (p : f (pt A) ≡ g (pt B)) - where - Q : (typ A) → Type (ℓ-max ℓ ℓ') - Q a = Σ ((b : (typ B)) → fst (P a b )) λ k → f a ≡ k (pt B) - - f-map : (a : typ A) → ((b : typ B) → fst (P a b)) → fst (P a (pt B)) - f-map a g = indConFun {A = Unit} {B = typ B} {k = (ℕ→ℕ₋₂ (n + m))} (pt-map B) (P a) g tt - - - - QisFib : (a : typ A) → Q a ≡ fiber (indConFun {A = Unit} {B = typ B} {k = (ℕ→ℕ₋₂ (n + m))} (pt-map B) (P a)) λ tt → (f a) - QisFib a = helper ∙ - isoToPath (iso (λ x → (fst x) , (funExt (λ y → snd x))) (λ x → (fst x) , - (cong (λ x → x tt) (snd x))) - (λ x → refl) - λ x → refl) - where - helper : Q a ≡ fiber (f-map a) (f a) - helper = isoToPath (iso (λ x → (x .fst) , (sym (x .snd))) (λ b → (fst b) , (sym (snd b))) (λ a → refl) λ b → refl) - -Lemma8-6-2 : (A : Pointed ℓ) (B : Pointed ℓ') (n m : ℕ) → - (is- (ℕ→ℕ₋₂ n) -ConnectedType (typ A)) → - (is- (ℕ→ℕ₋₂ m) -ConnectedType (typ B)) → - (P : (typ A) → (typ B) → HLevel (ℓ-max ℓ ℓ') (2+ (ℕ→ℕ₋₂ (n + m)))) → - (f : ((a : (typ A)) → fst (P a (pt B)))) → - (g : ((b : (typ B)) → fst (P (pt A) b))) → - (p : f (pt A) ≡ g (pt B)) → - Σ ((a : typ A) → (b : typ B) → fst (P a b)) - λ h → Σ (((a : typ A) → h a (pt B) ≡ f a) × ((b : typ B) → h (pt A) b ≡ g b)) - λ pair → p ≡ sym ((proj₁ pair) (pt A) ) ∙ (proj₂ pair) (pt B) - -Lemma8-6-2 {ℓ = ℓ} {ℓ' = ℓ'} A B n zero conA conB P f g p = - (λ a b → ((typesQ (pt A) .fst) a) .fst b) , - ((((λ a → sym ((typesQ (pt A)) .fst a .snd)))) , - (λ b → cong (λ x → (x .fst) b) (((typesQ) (pt A)) .snd))) , - (sym (transpHelper _ _ _ (cong (λ x → x .snd) (((typesQ) (pt A)) .snd)))) - where - helper : (a : typ A) → isOfHLevel (suc n) (fiber (indConFun {A = Unit} {B = typ B} {k = (ℕ→ℕ₋₂ (n + zero))} (pt-map B) (P a)) λ tt → (f a)) - helper a = Lemma861Gen neg1 (suc n) (λ x y → (2+ ℕ→ℕ₋₂ (predℕ y + zero))) (natHelper n) (pt-map B) (conOfpt-map B zero conB) (P a) (λ tt → f a) - where - natHelper : (n : ℕ) → (2+ ℕ→ℕ₋₂ (n + zero)) ≡ suc (n + (2+ neg1)) - natHelper zero = refl - natHelper (suc n) = cong (λ x → suc x) (natHelper n) - - conOfQ : (a : typ A) → isOfHLevel (2+ ((pred₋₂ (ℕ→ℕ₋₂ n)))) (prelims.Q A B n zero conA conB P f g p a) - conOfQ a = transport (λ i → isOfHLevel (2+ pred₋₂ (ℕ→ℕ₋₂ n)) (prelims.QisFib A B n zero conA conB P f g p a (~ i))) - (transport (λ i → isOfHLevel (natHelper n (~ i)) (prelims.QisFib A B n zero conA conB P f g p a i1)) - (helper a)) - where - natHelper : (n : ℕ) → (2+ pred₋₂ (ℕ→ℕ₋₂ n)) ≡ (suc n) - natHelper zero = refl - natHelper (suc n) = refl - - typesQ : (a : typ A) → Σ ((a : (typ A)) → (prelims.Q A B n zero conA conB P f g p a)) λ ℓ → ℓ (pt A) ≡ (g , p) - typesQ a = (fst (conInd-i→iii (pt-map A) - (pred₋₂ (ℕ→ℕ₋₂ n)) - (conOfpt-map A n conA) - (λ a → (prelims.Q A B n (zero) conA conB P f g p a , conOfQ a ))) - (λ x → (g , p))) , - cong (λ f → f tt) - (snd (conInd-i→iii (pt-map A) - (pred₋₂ (ℕ→ℕ₋₂ n)) - (conOfpt-map A n conA) - (λ a → (prelims.Q A B n (zero) conA conB P f g p a , conOfQ a ))) - (λ x → (g , p))) - -Lemma8-6-2 {ℓ = ℓ} {ℓ' = ℓ'} A B n (suc m) conA conB P f g p = - (λ a b → ((typesQ (pt A) .fst) a) .fst b) , - ((((λ a → sym ((typesQ (pt A)) .fst a .snd)))) , - (λ b → cong (λ x → (x .fst) b) (((typesQ) (pt A)) .snd))) , - (sym (transpHelper _ _ _ (cong (λ x → x .snd) (((typesQ) (pt A)) .snd)))) - where - helper : (a : typ A) → isOfHLevel (suc n) (fiber (indConFun {A = Unit} {B = typ B} {k = (ℕ→ℕ₋₂ (n + (suc m)))} (pt-map B) (P a)) λ tt → (f a)) - helper a = Lemma861GenNats m (suc n) (λ x y → 2+ ℕ→ℕ₋₂ (n + suc m)) (natHelper n m) (pt-map B) (conOfpt-map B (suc m) conB) (P a) λ tt → (f a) - where - natHelper : (n m : ℕ) → (2+ ℕ→ℕ₋₂ (n + suc m)) ≡ suc (n + (2+ ℕ→ℕ₋₂ m)) - natHelper zero m = refl - natHelper (suc n) m = cong (λ x → suc x) (natHelper n m) - - conOfQ : (a : typ A) → isOfHLevel (2+ ((pred₋₂ (ℕ→ℕ₋₂ n)))) (prelims.Q A B n (suc m) conA conB P f g p a) - conOfQ a = transport (λ i → isOfHLevel (natHelper n i) (prelims.Q A B n (suc m) conA conB P f g p a)) - (transport (λ i → isOfHLevel (suc n) - (prelims.QisFib A B n (suc m) conA conB P f g p a (~ i))) (helper a)) - where - natHelper : (n : ℕ) → (suc n) ≡ (2+ ((pred₋₂ (ℕ→ℕ₋₂ n)))) - natHelper zero = refl - natHelper (suc n) = refl - - typesQ : (a : typ A) → Σ ((a : (typ A)) → (prelims.Q A B n (suc m) conA conB P f g p a)) λ ℓ → ℓ (pt A) ≡ (g , p) - typesQ a = (fst (conInd-i→iii (pt-map A) - (pred₋₂ (ℕ→ℕ₋₂ n)) - (conOfpt-map A n conA) - (λ a → (prelims.Q A B n (suc m) conA conB P f g p a , conOfQ a ))) - (λ x → (g , p))) , - cong (λ f → f tt) - (snd (conInd-i→iii (pt-map A) - (pred₋₂ (ℕ→ℕ₋₂ n)) - (conOfpt-map A n conA) - (λ a → (prelims.Q A B n (suc m) conA conB P f g p a , conOfQ a ))) - (λ x → (g , p))) diff --git a/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4.agda b/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4.agda deleted file mode 100644 index 54ceac0f1..000000000 --- a/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4.agda +++ /dev/null @@ -1,1355 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Truncation.Connected.FreudenthalProof.8-6-4 where - - -open import Cubical.HITs.Truncation.Connected.Base -open import Cubical.HITs.Truncation.Connected.FreudenthalProof.Prelim -open import Cubical.HITs.Truncation.Connected.FreudenthalProof.8-6-4Code -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Function -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Transport -open import Cubical.Data.NatMinusTwo.Base -open import Cubical.Data.Sigma -open import Cubical.Data.Prod -open import Cubical.HITs.Susp -open import Cubical.HITs.Nullification hiding (elim) -open import Cubical.Data.Nat hiding (elim) -open import Cubical.HITs.Truncation.Base -open import Cubical.HITs.Truncation.Properties -open import Cubical.Data.HomotopyGroup - -private - variable - ℓ ℓ' : Level - A : Type ℓ - B : Type ℓ' - -{- this file contains a proof that σ is 2n-connected -} - ---- defining pair⁼ as in the HoTT book. Using J to be able to have easy access to the refl case ---- -pair⁼ : ∀ {ℓ'} {B : A → Type ℓ'} {x y : Σ A (λ a → B a)} (p : fst x ≡ fst y) → - transport (λ i → B (p i)) (snd x) ≡ (snd y) → - x ≡ y -pair⁼ {B = B}{x = x1 , x2} {y = y1 , y2} p = J (λ y1 p → ((y2 : B y1) → (transport (λ i → B (p i)) x2 ≡ y2) → (x1 , x2) ≡ (y1 , y2))) - (λ y2 p2 i → x1 , ((sym (transportRefl x2)) ∙ p2) i) p y2 - -pair⁼Refl : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} {a : A} {x y : B a} (p : transport (λ i → B a) x ≡ y) → - pair⁼ {B = B} {x = (a , x)} {y = (a , y)} refl p - ≡ - λ i → a , ((sym (transportRefl x)) ∙ p) i -pair⁼Refl {A = A} {B = B} {a = a} {x = x} {y = y} p = cong (λ f → f y p) - (JRefl (λ y1 p → ((y2 : B y1) → (transport (λ i → B (p i)) x ≡ y2) → - _≡_ {A = Σ A (λ a → B a)} (a , x) (y1 , y2))) - (λ y2 p2 i → a , ((sym (transportRefl x)) ∙ p2) i)) - - -{- We will also need the following idenitity for sym (pair⁼ p q). -} -pair⁼Sym : ∀ {ℓ'} {B : A → Type ℓ'} {x y : Σ A (λ a → B a)} - (p : fst x ≡ fst y) - (q : transport (λ i → B (p i)) (snd x) ≡ (snd y)) → - sym (pair⁼ {x = x} {y = y} p q) - ≡ pair⁼ (sym p) (transportLemma {B = B} p (snd x) (snd y) q ) -pair⁼Sym {ℓ} {A = A} {B = B} {x = x} {y = y} p = J {ℓ} {A} {fst x} (λ fsty p → (sndy : B (fsty)) (q : transport (λ i → B (p i)) (snd x) ≡ (sndy)) → - sym (pair⁼ p q) - ≡ pair⁼ (sym p) (transportLemma {B = B} p (snd x) (sndy) q )) - (λ sndy → J (λ sndy q → sym (pair⁼ (λ _ → fst x) q) - ≡ pair⁼ refl (transportLemma {B = B} refl (snd x) sndy q)) - ((λ j → (sym (pair⁼Refl refl j) )) ∙ - (λ k i → fst x , ((rUnit (sym (transportRefl (snd x))) (~ k)) (~ i))) ∙ - (λ k i → fst x , helper (~ k) i) ∙ - sym (pair⁼Refl (transportLemma {B = B} (λ _ → fst x) (snd x) - (transport (λ i → B (fst x)) (snd x)) refl)))) - p (snd y) - where - helper : (sym (transportRefl (transport (λ i₁ → B (fst x)) (snd x)))) ∙ - (transportLemma {B = B} (λ _ → fst x) (snd x) - (transport (λ i₂ → B (fst x)) (snd x)) - (λ _ → transport (λ i₂ → B (fst x)) (snd x))) - ≡ (transportRefl (snd x)) - helper = (λ i → sym (transportRefl (transport (λ i₁ → B (fst x)) (snd x))) ∙ - (transportLemmaRefl {B = B} (snd x) (snd x) i) (snd x) - (transport (λ i₂ → B (fst x)) (snd x)) - (λ _ → transport (λ i₂ → B (fst x)) (snd x))) ∙ - (λ i → sym (transportRefl (transport (λ i₁ → B (fst x)) (snd x))) ∙ - transportRefl (transport (λ i₂ → B (fst x)) (snd x)) ∙ - (λ _ → transport (λ i₂ → B (fst x)) (snd x)) ∙ - transportRefl (snd x)) ∙ - (λ i → sym (transportRefl (transport (λ i₁ → B (fst x)) (snd x))) ∙ - transportRefl (transport (λ i₂ → B (fst x)) (snd x)) ∙ - lUnit (transportRefl (snd x)) (~ i)) ∙ - assoc (sym (transportRefl (transport (λ i₁ → B (fst x)) (snd x)))) - (transportRefl (transport (λ i₁ → B (fst x)) (snd x))) - (transportRefl (snd x)) ∙ - (λ i → (lCancel (transportRefl (transport (λ i₁ → B (fst x)) (snd x))) i) ∙ - transportRefl (snd x)) ∙ - sym (lUnit (transportRefl (snd x))) - -{- We need something slightly stronger than functoriality of transport. - Since we will be transporting over dependent pairs, we also need to split the second component -} -functTransp2 : ∀ {ℓ ℓ'} {A : Type ℓ} {a1 b : A} {C : Σ A (λ b → a1 ≡ b) → Type ℓ' } (p : a1 ≡ b) (q : b ≡ a1) → - transport (λ i → C (pair⁼ (p ∙ q) (pairLemma2 (p ∙ q)) i) ) - ≡ λ x → transport (λ i → C (pair⁼ q ((pairLemma p q)) i)) - (transport (λ i → C (pair⁼ p (pairLemma2 p) i)) - x) -functTransp2 {ℓ} {ℓ'} {A = A} {a1 = a1} {b = b} {C = C} = - J (λ b p → (q : b ≡ a1) → - transport (λ i → C (pair⁼ {B = λ a → a1 ≡ a} {x = (a1 , refl {x = a1})} - (p ∙ q) (pairLemma2 (p ∙ q)) i) ) - ≡ λ x → transport (λ i → C (pair⁼ q ((pairLemma p q)) i)) - (transport (λ i → C (pair⁼ p (pairLemma2 p) i)) - x)) - λ q → (λ j → subst C (hej a1 q (~ j))) ∙ - (λ j → subst C (pair⁼ (λ _ → a1) (pairLemma2 (λ _ → a1)) ∙ - pair⁼ q (pairLemmaRefl q (~ j)))) ∙ - funExt λ x → (substComposite-∙ C (pair⁼ refl (pairLemma2 refl)) (pair⁼ q (pairLemma refl q)) x) - where - hej : (b : A) (q : a1 ≡ b) → (pair⁼ {A = A} {B = λ a → a1 ≡ a} - {x = (a1 , λ _ → a1)} {y = (a1 , λ _ → a1)} - (λ _ → a1) (pairLemma2 {a = a1} {b = a1} (λ _ → a1)) ∙ - pair⁼ q ((pairLemma2 q) ∙ lUnit q)) - ≡ pair⁼ (refl ∙ q) (pairLemma2 (refl {x = a1} ∙ q)) - hej b = J (λ b q → pair⁼ (λ _ → a1) (pairLemma2 (λ _ → a1)) ∙ - pair⁼ q (pairLemma2 q ∙ lUnit q) - ≡ pair⁼ (refl ∙ q) (pairLemma2 (refl ∙ q))) - ((λ i → (helper2 i) ∙ pair⁼ refl (pairLemma2 refl ∙ lUnit refl)) ∙ - sym (lUnit (pair⁼ (λ _ → a1) (pairLemma2 (λ _ → a1) ∙ lUnit (λ _ → a1)))) ∙ - desired≡) - where - helper2 : (pair⁼ {A = A} {B = λ a → a1 ≡ a} - {x = (a1 , λ _ → a1)} {y = (a1 , λ _ → a1)} - (λ _ → a1) (pairLemma2 {a = a1} {b = a1} (λ _ → a1))) - ≡ refl - helper2 = (λ i → (JRefl (λ y1 p → ((y2 : a1 ≡ y1) → (transport (λ i → a1 ≡ (p i)) refl ≡ y2) → - _≡_ {A = Σ A (λ a → a1 ≡ a)} (a1 , refl) (y1 , y2))) - (λ y2 p2 i → refl {x = a1} i , ((sym (transportRefl refl)) ∙ p2) i) i) - (λ _ → a1) - (pairLemma2 {a = a1} {b = a1} (λ _ → a1))) ∙ - (λ j → λ i → a1 , ((sym (transportRefl refl)) ∙ - JRefl (λ b p → transport (λ i → a1 ≡ p i) refl ≡ p) - (transportRefl refl) j ) i) ∙ - λ j i → a1 , (lCancel (transportRefl refl) j) i - - PathP1 : PathP (λ j → _≡_ {A = Σ A (λ a → a1 ≡ a)} - (a1 , (λ _ → a1)) - (a1 , lUnit (λ _ → a1) (~ j))) - (pair⁼ (λ _ → a1) (pairLemma2 (λ _ → a1) ∙ lUnit (λ _ → a1))) - refl - PathP1 j = hcomp (λ k → λ{(j = i0) → pair⁼ (λ _ → a1) (pairLemma2 (λ _ → a1) ∙ - lUnit (λ _ → a1)) - ; (j = i1) → ((λ i → pair⁼ (λ _ → a1) (rUnit (pairLemma2 (λ _ → a1)) (~ i)) ) ∙ - helper2) k}) - (pair⁼ refl ((pairLemma2 (λ _ → a1)) ∙ - (λ i → lUnit refl (i ∧ (~ j))))) - - PathP2 : PathP (λ j → _≡_ {A = Σ A (λ a → a1 ≡ a)} - (a1 , (λ _ → a1)) - (a1 , lUnit (λ _ → a1) j)) - refl - (pair⁼ ((λ _ → a1) ∙ refl) (pairLemma2 ((λ _ → a1) ∙ refl))) - PathP2 j = hcomp (λ k → λ{(j = i0) → helper2 k - ; (j = i1) → (pair⁼ ((λ _ → a1) ∙ refl) (pairLemma2 ((λ _ → a1) ∙ refl)))}) - (pair⁼ (lUnit (λ _ → a1) j) (pairLemma2 (lUnit (λ _ → a1) j))) - - pathsCancel : (λ j → _≡_ {A = Σ A (λ a → a1 ≡ a)} (a1 , (λ _ → a1)) (a1 , lUnit (λ _ → a1) (~ j))) ∙ - ((λ j → _≡_ {A = Σ A (λ a → a1 ≡ a)} (a1 , (λ _ → a1)) (a1 , lUnit (λ _ → a1) j))) - ≡ refl - pathsCancel = lCancel (λ j → _≡_ {A = Σ A (λ a → a1 ≡ a)} (a1 , (λ _ → a1)) (a1 , lUnit (λ _ → a1) j)) - - desired≡ : pair⁼ (λ _ → a1) (λ i i₁ → (pairLemma2 (λ _ → a1) ∙ lUnit (λ _ → a1)) i i₁) - ≡ pair⁼ ((λ _ → a1) ∙ refl) (pairLemma2 ((λ _ → a1) ∙ refl)) - desired≡ = transport (λ i → PathP (λ j → pathsCancel i j) - (pair⁼ (λ _ → a1) (pairLemma2 (λ _ → a1) ∙ lUnit (λ _ → a1))) - (pair⁼ ((λ _ → a1) ∙ refl) (pairLemma2 ((λ _ → a1) ∙ refl)))) - (compPathP PathP1 PathP2) - ------------------------------------------------------------------------------------------------------------------------- - - - - ----------------------- Lemma 8.6.10. from the HoTT book + an addition variant of it that will be useful ----------------- -Lemma8610fun : ∀{ℓ ℓ' ℓ''} {A : Type ℓ} {a1 a2 : A} {B : A → Type ℓ'} - (C : (a : A) → (B a → Type ℓ'')) - (p : a1 ≡ a2) (b : B a2) → - C a1 (subst B (sym p) b) → C a2 b -Lemma8610fun {ℓ} {ℓ'} {ℓ''} {A = A} {a1 = a1 } {a2 = a2} {B = B} C p b = transport (λ i → idHelper i ) - where - idHelper : C a1 (subst B (sym p) b) ≡ C a2 b - idHelper = sym (cong (λ x → x b) (Lemma294 {A = B} {B = λ _ → Type ℓ''} p (C a1))) ∙ - funExt⁻ (fromPathP λ j → C (p j)) b - -Lemma8610 : ∀{ℓ ℓ' ℓ''} {A : Type ℓ} {a1 a2 : A} {B : A → Type ℓ'} - (C : (a : A) → (B a → Type ℓ'')) - (p : a1 ≡ a2) (b : B a2) → - transport ((λ i → uncurry C (pair⁼ p (transpRCancel (λ i → B (p i)) b) i) )) - ≡ Lemma8610fun C p b -Lemma8610 {ℓ} {ℓ'} {ℓ''} {A = A} {a1 = a1} {B = B} C = J (λ a2 p → (b : B a2) → - transport ((λ i → uncurry C (pair⁼ p (transpRCancel (λ i → B (p i)) b) i) )) - ≡ Lemma8610fun C p b ) - λ b j → transport (helper b (~ j)) - where - helper : (b : B a1) → sym (cong (λ x → x b) (Lemma294 {A = B} {B = λ _ → Type ℓ''} refl (C a1))) ∙ - funExt⁻ (fromPathP λ j → C a1) b - ≡ (λ i → uncurry C (pair⁼ refl (transpRCancel (λ i₁ → B (refl i₁)) b) i)) - helper b = (λ i → sym (cong (λ x → x b) (Lemma294Refl {A = B} {B = λ _ → Type ℓ''} (C a1) i)) ∙ - cong (λ x → C a1 x) (transportRefl b)) ∙ - sym (lUnit (cong (λ x → C a1 x) (transportRefl b))) ∙ - (λ j i → uncurry C (a1 , lUnit (transportRefl b) j i)) ∙ - (λ j i → uncurry C (a1 , ((lCancel (transportRefl (transport refl b)) (~ j)) ∙ - (transportRefl b)) i)) ∙ - (λ j i → uncurry C (a1 , (assoc (sym (transportRefl ((transport⁻ (λ i₁ → B a1) b)))) - (transportRefl (transport refl b)) - (transportRefl b) (~ j)) i)) ∙ - (λ j i → uncurry C (a1 , ((sym (transportRefl ((transport⁻ (λ i₁ → B a1) b))) ∙ - (transpRCancelRefl b (~ j)))) i)) ∙ - (λ j i → uncurry C (pair⁼Refl (transpRCancel (λ i₁ → B a1) b) (~ j) i)) - - -{-variant for when we have have pair⁼ p refl -} -Lemma8610Reflfun : ∀{ℓ ℓ' ℓ''} {A : Type ℓ} {a1 a2 : A} {B : A → Type ℓ'} - (C : (a : A) → (B a → Type ℓ'')) - (p : a1 ≡ a2) (b : B a1) → - C a1 b ≡ C a2 (subst B p b) -Lemma8610Reflfun {ℓ'' = ℓ''} {a1 = a1} {a2 = a2} {B = B} C p b = (λ i → C a1 (transpRCancel (λ i → (B (p (~ i)))) b (~ i))) ∙ - funExt⁻ (fromPathP λ j → C (p j)) (transport (λ i → B (p i)) b) - -Lemma8610Refl : ∀{ℓ ℓ' ℓ''} {A : Type ℓ} {a1 a2 : A} {B : A → Type ℓ'} - (C : (a : A) → (B a → Type ℓ'')) - (p : a1 ≡ a2) (b : B a1) → - transport ((λ i → uncurry C (pair⁼ {x = a1 , b} p refl i) )) - ≡ transport (Lemma8610Reflfun {B = B} C p b) - -Lemma8610Refl {A = A} {a1 = a1} {B = B} C = J (λ a2 p → (b : B a1) → - transport ((λ i → uncurry C (pair⁼ {x = a1 , b} p refl i) )) - ≡ transport (Lemma8610Reflfun {B = B} C p b)) - λ b → (λ k → transport ((λ i → uncurry C (pair⁼Refl (λ _ → transport (λ i → B a1) b) k i) ))) ∙ - (λ k → transport (λ i → uncurry C (a1 , (rUnit (sym (transportRefl b)) (~ k) i)))) ∙ - (λ k → transport (rUnit (cong (λ x → C a1 x) (sym (transportRefl b))) k)) ∙ - (λ k → transport ((cong (λ x → C a1 x) (sym (transportRefl b))) ∙ - lCancel ((funExt⁻ (fromPathP λ j → C a1) (transport (λ i → B a1) b))) (~ k))) ∙ - (λ k → transport (assoc (cong (λ x → C a1 x) (sym (transportRefl b))) - (cong (λ x → C a1 x) (sym (transportRefl {A = B a1} (transport (λ _ → B a1) b)))) - (funExt⁻ (fromPathP λ j → C a1) (transport (λ i → B a1) b)) k)) ∙ - (λ k → transport (congComp2 (λ x → C a1 x) - (sym (transportRefl b)) - (sym (transportRefl {A = B a1} (transport (λ _ → B a1) b))) k ∙ - funExt⁻ (fromPathP λ j → C a1) (transport (λ i → B a1) b))) ∙ - (λ k → transport (cong (λ x → C a1 x) - (symDistr {x = transport refl (transport (λ _ → B a1) b)} - {y = transport (λ _ → B a1) b} {z = b} - (transportRefl {A = B a1} (transport (λ _ → B a1) b)) - (transportRefl b) (~ k)) ∙ - funExt⁻ (fromPathP λ j → C a1) (transport (λ i → B a1) b)) ) ∙ - (λ k → transport ((λ i → C a1 (transpRCancelRefl {A = B a1} b (~ k) (~ i))) ∙ - funExt⁻ (fromPathP λ j → C a1) (transport (λ i → B a1) b))) - - ----------------------------------------------------------------------------------------------------- - - --------------- Center of contraction for fibers of σ ----------------- - -c : (n : ℕ) (a : A) (iscon : is- (ℕ→ℕ₋₂ n) -ConnectedType A) (y : (Susp A)) (p : north ≡ y) → CODE a n iscon y p -c n a iscon y p = transport (λ i → (uncurry (CODE a n iscon) (pair⁼ p (pairLemma2 p) i))) ∣ a , (rCancel (merid a)) ∣ - -cId : (n : ℕ) (a x1 : A) (iscon : is- (ℕ→ℕ₋₂ n) -ConnectedType A) → - c n a iscon north (σ a x1) - ≡ transport (λ i → uncurry (CODE a n iscon) (pair⁼ (λ i₁ → merid a (~ i₁)) (pairLemma (merid x1) (λ i₁ → merid a (~ i₁))) i)) - (transport (λ i → uncurry (CODE a n iscon) (pair⁼ (merid x1) (pairLemma2 (merid x1)) i)) - ∣ a , rCancel (merid a) ∣) -cId n a x1 iscon = cong (λ x → x ∣ a , (rCancel (merid a)) ∣) (functTransp2 {C = uncurry (CODE a n iscon)} (merid x1) (sym (merid a))) - - ----------- path algebra -------------- -{- The goal now is to show that the center of contraction is equal to ∣ x1 , refl ∣. The following completely technical lemmas will be combined to show this. The general idea is to break the center up using functTransp2, slightly altering the forms of the inner and outer transport, apply Lemma8610 and then show that everything cancels -} - - -{- We begin by transforming the inner transport, i.e. -transport (λ i → ...) ∣ a , rCancel (merid a) ∣ to -transport (λ i → ...) ∣ a , rCancel (merid a) ∙ ... ∣ -} - -codeTranspId1 : (n : ℕ) (a x1 : A) (iscon : is- (ℕ→ℕ₋₂ n) -ConnectedType A) → - (transport (λ i → uncurry (CODE a n iscon) (pair⁼ (merid x1) (transpRCancel (λ i₁ → north ≡ merid x1 i₁) (merid x1)) i))) - ∣ a , rCancel (merid a) ∙ sym (rCancel (merid x1)) ∙ sym (pairLemma (merid x1) (sym (merid x1))) ∣ - ≡ transport (λ i → (uncurry (CODE a n iscon) (pair⁼ (merid x1) (pairLemma2 (merid x1)) i))) ∣ a , rCancel (merid a) ∣ -codeTranspId1 {ℓ} {A = A} n a x1 iscon = helper (south) (merid x1) - where - helper : (y : Susp A) (p : north ≡ y) → - (transport (λ i → uncurry (CODE a n iscon) (pair⁼ p (transpRCancel (λ i₁ → north ≡ p i₁) p) i))) - ∣ a , rCancel (merid a) ∙ sym (rCancel p) ∙ sym (pairLemma p (sym p)) ∣ - ≡ transport (λ i → (uncurry (CODE a n iscon) (pair⁼ p (pairLemma2 p) i))) - ∣ a , rCancel (merid a) ∣ - helper y = J (λ y p → (transport (λ i → uncurry (CODE a n iscon) (pair⁼ p (transpRCancel (λ i₁ → north ≡ p i₁) p) i))) - ∣ a , rCancel (merid a) ∙ sym (rCancel p) ∙ sym (pairLemma p (sym p)) ∣ - ≡ transport (λ i → (uncurry (CODE a n iscon) (pair⁼ p (pairLemma2 p) i))) ∣ a , rCancel (merid a) ∣ ) - ((transport (λ i → uncurry (CODE a n iscon) (pair⁼ refl (transpRCancel (λ i₁ → north ≡ refl i₁) refl) i))) - ∣ a , rCancel (merid a) ∙ (sym (rCancel refl)) ∙ (sym (pairLemma refl refl)) ∣ - - ≡⟨ ((λ j → transport (λ i → uncurry (CODE a n iscon) (pair⁼ refl (transpRCancelRefl refl j) i)) - ∣ a , rCancel (merid a) ∙ (sym (rCancel refl)) ∙ (sym (pairLemmaReflRefl j)) ∣)) ⟩ - - (transport (λ i → uncurry (CODE a n iscon) (pair⁼ refl (transportRefl (transport refl refl) ∙ transportRefl refl) i))) - ∣ a , rCancel (merid a) ∙ (sym (rCancel refl)) ∙ sym (transportRefl refl ∙ lUnit refl) ∣ - - ≡⟨ helper2 ⟩ - - transport (λ i → uncurry (CODE a n iscon) - (north , (sym (transportRefl ((transport (λ z → _≡_ {A = Susp A} north north) (λ _ → north)))) ∙ - transportRefl (transport (λ z → _≡_ {A = Susp A} north north) (λ _ → north)) ∙ - transportRefl (λ _ → north)) i)) - ∣ a , rCancel (merid a) ∙ (sym (rCancel refl)) ∙ sym (transportRefl refl ∙ lUnit refl) ∣ - - ≡⟨ (λ j → transport (λ i → uncurry (CODE a n iscon) - (north , assoc ((sym (transportRefl (transport (λ z → _≡_ {A = Susp A} north north) - (λ _ → north))))) - (transportRefl (transport (λ z → _≡_ {A = Susp A} north north) (λ _ → north))) - (transportRefl (λ _ → north)) j i)) - ∣ a , rCancel (merid a) ∙ (sym (rCancel refl)) ∙ sym (transportRefl refl ∙ lUnit refl) ∣) ⟩ - - transport ((λ i → uncurry (CODE a n iscon) - (north , ((sym (transportRefl (transport (λ z → _≡_ {A = Susp A} north north) (λ _ → north))) ∙ - transportRefl (transport (λ z → _≡_ {A = Susp A} north north) (λ _ → north))) ∙ - transportRefl (λ _ → north)) i))) - ∣ a , rCancel (merid a) ∙ (sym (rCancel refl)) ∙ sym (transportRefl refl ∙ lUnit refl) ∣ - - ≡⟨ ((λ j → transport (λ i → uncurry (CODE a n iscon) - (north , (lCancel (transportRefl (transport (λ z → _≡_ {A = Susp A} north north) - (λ _ → north))) j ∙ - (transportRefl (λ _ → north))) i)) - ∣ a , rCancel (merid a) ∙ (sym (rCancel refl)) ∙ sym (transportRefl refl ∙ lUnit refl) ∣)) ⟩ - - transport (λ i → uncurry (CODE a n iscon) (north , (refl ∙ (transportRefl (λ _ → north))) i)) - ∣ a , rCancel (merid a) ∙ (sym (rCancel refl)) ∙ sym (transportRefl refl ∙ lUnit refl) ∣ - - ≡⟨ ((λ j → transport (λ i → uncurry (CODE a n iscon) - (north , (lUnit (transportRefl (λ _ → north)) (~ j)) i)) - ∣ a , rCancel (merid a) ∙ (sym (rCancel refl)) ∙ sym (transportRefl refl ∙ lUnit refl) ∣) ) ⟩ - - transport (λ i → uncurry (CODE a n iscon) - (north , ((transportRefl (λ _ → north))) i)) - ∣ a , rCancel (merid a) ∙ (sym (rCancel refl)) ∙ sym (transportRefl refl ∙ lUnit refl) ∣ - - ≡⟨ ((λ j → transport (λ i → uncurry (CODE a n iscon) (north , ((transportRefl (λ _ → north))) i)) - ∣ a , cancelHelper j ∣ )) ⟩ - - transport (λ i → uncurry (CODE a n iscon) (north , ((transportRefl (λ _ → north))) i)) - ∣ a , rCancel (merid a) ∙ (sym (transportRefl refl)) ∣ - - ≡⟨ pathPtest2 ⟩ - - transport (λ i → uncurry (CODE a n iscon) (north , λ _ → north)) ∣ a , rCancel (merid a)∣ - - ≡⟨ sym backAgain ⟩ - - (transport (λ i → uncurry (CODE a n iscon) (pair⁼ refl (pairLemma2 refl) i)) - ∣ a , rCancel (merid a) ∣) ∎) - where - backAgain : transport (λ i → uncurry (CODE a n iscon) (pair⁼ refl (pairLemma2 refl) i)) - ∣ a , rCancel (merid a) ∣ - ≡ (transport (λ i → uncurry (CODE a n iscon) (north , λ _ → north)) - ∣ a , rCancel (merid a)∣) - backAgain = (λ j → transport (λ i → uncurry (CODE a n iscon) (pair⁼ refl (pairLemma2Refl j) i)) ∣ a , rCancel (merid a) ∣) ∙ - (λ j → transport (λ i → uncurry (CODE a n iscon) (pair⁼Refl (transportRefl refl) j i)) ∣ a , rCancel (merid a) ∣) - ∙ λ j → transport (λ i → uncurry (CODE a n iscon) (north , lCancel (transportRefl refl) j i)) ∣ a , rCancel (merid a) ∣ - - cancelHelper : rCancel (merid a) ∙ (λ i → rCancel (λ _ → north) (~ i)) ∙ sym (transportRefl refl ∙ lUnit refl) ≡ rCancel (merid a) ∙ sym (transportRefl refl) - cancelHelper = (λ j → rCancel (merid a) ∙ (λ i → rCancel (λ _ → north) (~ i)) ∙ symDistr (transportRefl refl) (lUnit refl) j) ∙ - (λ j → rCancel (merid a) ∙ assoc (sym (rCancel (λ _ → north))) (sym (lUnit refl)) (sym (transportRefl refl)) j) ∙ - (λ j → rCancel (merid a) ∙ lCancel (sym (lUnit refl)) j ∙ sym (transportRefl refl)) ∙ - λ j → rCancel (merid a) ∙ lUnit (sym (transportRefl refl)) (~ j) - - helper2 : transport (λ i → uncurry (CODE a n iscon) (pair⁼ (λ _ → north) (transportRefl (transport (λ _ → _≡_ {A = Susp A} north north) (λ _ → north)) ∙ transportRefl (λ _ → north)) i)) - ∣ a , rCancel (merid a) ∙ (λ i → rCancel (λ _ → north) (~ i)) ∙ (λ i → pairLemmaReflRefl i1 (~ i)) ∣ - ≡ transport (λ i → uncurry (CODE a n iscon) - (north , (sym (transportRefl ((transport (λ z → _≡_ {A = Susp A} north north) (λ _ → north)))) ∙ - transportRefl (transport (λ z → _≡_ {A = Susp A} north north) (λ _ → north)) ∙ - transportRefl (λ _ → north)) i)) - ∣ a , rCancel (merid a) ∙ (λ i → rCancel (λ _ → north) (~ i)) ∙ (λ i → pairLemmaReflRefl i1 (~ i)) ∣ - helper2 j = transport (λ i → uncurry (CODE a n iscon) - ((pair⁼Refl ((transportRefl (transport (λ _ → _≡_ {A = Susp A} north north) (λ _ → north)) ∙ - transportRefl (λ _ → north))) j) i)) - ∣ a , rCancel (merid a) ∙ (λ i → rCancel (λ _ → north) (~ i)) ∙ (λ i → pairLemmaReflRefl i1 (~ i)) ∣ - - pathPtest2 : (transport (λ i → uncurry (CODE a n iscon) - (north , (transportRefl λ _ → north) i)) - ∣ a , rCancel (merid a) ∙ (sym (transportRefl refl)) ∣) - ≡ (transport (λ i → uncurry (CODE a n iscon) (north , λ _ → north)) - ∣ a , rCancel (merid a)∣) - pathPtest2 = (transport (λ i → uncurry (CODE a n iscon) - (north , (transportRefl (λ _ → north)) i)) - ∣ a , rCancel (merid a) ∙ (sym (transportRefl (λ _ → north))) ∣) - ≡⟨ (λ j → (transport (λ i → uncurry (CODE a n iscon) - (north , (transportRefl {A = north ≡ north} (λ _ → north) (i ∨ j)))) - ∣ a , rCancel (merid a) ∙ ((λ z → transportRefl {A = north ≡ north} ((λ i → north)) ((~ z) ∨ j))) ∣)) ⟩ - (transport (λ i → uncurry {C = λ a b → Type ℓ} (CODE a n iscon) (north , λ _ → north)) - ∣ a , rCancel (merid a) ∙ (λ _ _ → north) ∣) - ≡⟨ (λ j → (transport (λ i → uncurry (CODE a n iscon) (north , λ _ → north)) - ∣ a , ((rUnit (rCancel (merid a))) (~ j)) ∣)) ⟩ - transport (λ i → uncurry (CODE a n iscon) (north , (λ _ → north))) - ∣ a , rCancel (merid a)∣ ∎ - - -{- The inner transport will have the form transport (λ i → sym (cong (λ x → x p) ... ) after applying Lemma8610. We transform this into something more useful -} -codeTranspId2 : ∀{ℓ} {X : Type ℓ} {a b : X} (q p : a ≡ b) - (A : (a ≡ a) → Type ℓ) (B : (a ≡ b) → Type ℓ) → - (f : (q₁ : a ≡ b) → A (q₁ ∙ sym q) ≡ B q₁) → - (sym (cong (λ x → x p) (Lemma294 {A = λ x → a ≡ x} {B = λ _ → Type ℓ} q A))) ∙ funExt⁻ (fromPathP (toPathP {A = λ i → a ≡ q i → Type ℓ} {x = A} {y = B} - (Lemma296-fun {X = X} {A = λ y → a ≡ y} {B = λ _ → Type ℓ} - q A - B - (helperFun {X = X} q {A = A} - {B = B} - f)))) p - ≡ (transportRefl (A (subst (λ x → a ≡ x) (sym q) p)) ∙ cong (λ x → A x) (pairLemma p (sym q))) ∙ f p -codeTranspId2 {ℓ} {X = X} {a = a} = J (λ b q → (p : a ≡ b) (A : (a ≡ a) → Type ℓ) (B : (a ≡ b) → Type ℓ) (f : (q₁ : a ≡ b) → A (q₁ ∙ sym q) ≡ B q₁) → - (sym (cong (λ x → x p) (Lemma294 {A = λ x → a ≡ x} {B = λ _ → Type ℓ} q A))) ∙ - funExt⁻ (fromPathP (toPathP {A = λ i → a ≡ q i → Type ℓ} {x = A} {y = B} - (Lemma296-fun {X = X} {A = λ y → a ≡ y} - {B = λ _ → Type ℓ} q A B - (helperFun {X = X} q {A = A} {B = B} f)))) p - ≡ (transportRefl (A (subst (λ x → a ≡ x) (sym q) p)) ∙ - cong (λ x → A x) (pairLemma p (sym q))) ∙ f p) - λ p A B f → - (λ k → ((λ i → (Lemma294Refl {A = λ x → a ≡ x} {B = λ _ → Type ℓ} A k) (~ i) p)) ∙ - (λ i → fromPathP (toPathP {A = λ i → a ≡ a → Type ℓ} {x = A} {y = B} - (Lemma296-fun {X = X} {A = λ y → a ≡ y} {B = λ _ → Type ℓ} (refl {x = a}) - A B (helperFun {X = X} (refl {x = a}) {A = A} {B = B} f))) i p)) ∙ - (λ k → lUnit ((λ i → fromPathP (toPathP {A = λ i → a ≡ a → Type ℓ} {x = A} {y = B} - (Lemma296-fun {X = X} {A = λ y → a ≡ y} {B = λ _ → Type ℓ} (refl {x = a}) - A B (helperFun {X = X} (refl {x = a}) {A = A} {B = B} f))) i p)) (~ k)) ∙ - (λ k i → (toPathCancel {A = λ i → a ≡ a → Type ℓ} {x = A} (Lemma296-fun {X = X} {A = λ y → a ≡ y} - {B = λ _ → Type ℓ} (refl {x = a}) - A B (helperFun {X = X} (refl {x = a}) {A = A} {B = B} f)) k) i p) ∙ - (λ k i → ((cong (λ x → x (helperFun {X = X} (λ _ → a) f)) (Lemma296-funRefl {A = λ y → a ≡ y} - {B = λ _ → Type ℓ} A B)) k) i p ) ∙ - (λ k i → (transportRefl {A = a ≡ a → Type ℓ} A ∙ - funExt λ z → sym (transportRefl {A = Type ℓ} (A z)) ∙ - cong (λ x → x f z) (helperFunRefl {X = X} {A = A} {B = B} ) k ∙ - λ i → (B (transportRefl {A = a ≡ a} z i))) i p) ∙ - sym (congComp2 (λ x → x p) (transportRefl A) - (funExt λ z → sym (transportRefl {A = Type ℓ} (A z)) ∙ - (transportRefl (A z) ∙ cong (λ x → A x) (rUnit z) ∙ - f z ∙ - cong (λ x → B x) (sym (transportRefl z))) ∙ - λ i → (B (transportRefl z i)))) ∙ - invCanceller (cong (λ x → x p) (transportRefl A)) - (sym (transportRefl (A p))) - (λ i → A (rUnit p i)) - (f p) - (λ i → B (transportRefl p i)) ∙ - assoc (λ i → transportRefl A i p) - (λ i → A (rUnit p i)) - (f p) ∙ - (λ k → ((λ i → transportRefl A i p) ∙ (λ i → A (rUnit p i))) ∙ f p) ∙ - (λ k → (transpLemma2 {A = A} p k ∙ (cong (λ x → A x) (rUnit p))) ∙ f p ) ∙ - (λ k → ((assoc (transportRefl (A (subst (_≡_ a) (λ i → a) p))) - (cong (λ x → A x) (transportRefl p)) - (cong (λ x → A x) (rUnit p))) (~ k)) ∙ - f p) ∙ - (λ k → ((transportRefl (A (subst (_≡_ a) (λ i → a) p)) ∙ congComp2 (λ x → A x) (transportRefl p) (rUnit p) k) ∙ f p)) ∙ - (λ k → (transportRefl (A (subst (_≡_ a) (λ i → a) p)) ∙ - (λ i → A ((transportRefl p ∙ rUnit p) i))) ∙ - f p) ∙ - (λ k → (transportRefl (A (subst (_≡_ a) (λ i → a) p)) ∙ (λ i → A (pairLemma*Refl p (~ k) i))) ∙ f p) ∙ - λ k → (transportRefl (A (subst (_≡_ a) (λ i → a) p)) ∙ (λ i → A (pairLemmaId p (λ i₁ → a ) (~ k) i))) ∙ f p - - where - transpLemma2 : ∀ {ℓ ℓ'} {X : Type ℓ} {x : X} {A : x ≡ x → Type ℓ'} (p : x ≡ x) → - (λ i → transportRefl A i p) - ≡ (transportRefl (A (transport (λ i → x ≡ x) p)) ∙ (λ i → A ((transportRefl p) i))) - transpLemma2 {ℓ' = ℓ'}{x = x} {A = A} p j = hcomp (λ k → λ{(j = i0) → (sym (lUnit (λ i → transportRefl (A ((transportRefl p) i)) i))) k - ; (j = i1) → (transportRefl (A (transport (λ i → x ≡ x) p)) ∙ - (λ i → A ((transportRefl p) i)))}) - ((λ i → transportRefl (A (transport (λ i → x ≡ x) p )) (i ∧ j)) ∙ - (λ i → transportRefl (A ((transportRefl p) i)) (i ∨ j))) - - - invCanceller : ∀ {ℓ} {A : Type ℓ} {a b c d e f : A} (p : a ≡ b) (q : b ≡ c) (r : b ≡ d) (s : d ≡ e) (t : f ≡ e) → - p ∙ q ∙ (sym q ∙ r ∙ s ∙ sym t) ∙ t ≡ p ∙ r ∙ s - invCanceller {a = a} {b = b} {c = c} {d = d} {e = e} {f = f} = - J (λ b p → (q : b ≡ c) (r : b ≡ d) (s : d ≡ e) (t : f ≡ e) → - p ∙ q ∙ (sym q ∙ r ∙ s ∙ sym t) ∙ t ≡ p ∙ r ∙ s) - (J (λ c q → (r : a ≡ d) (s : d ≡ e) (t : f ≡ e) → - refl ∙ q ∙ ((λ i → q (~ i)) ∙ r ∙ s ∙ (λ i → t (~ i))) ∙ t ≡ refl ∙ r ∙ s) - (J (λ d r → (s : d ≡ e) (t : f ≡ e) → - (λ _ → a) ∙ (λ _ → a) ∙ ((λ _ → a) ∙ r ∙ s ∙ (λ i → t (~ i))) ∙ t ≡ (λ _ → a) ∙ r ∙ s) - (J (λ e s → (t : f ≡ e) → - (λ _ → a) ∙ (λ _ → a) ∙ ((λ _ → a) ∙ (λ _ → a) ∙ s ∙ (λ i → t (~ i))) ∙ t ≡ (λ _ → a) ∙ (λ _ → a) ∙ s) - λ t → sym (lUnit ((λ _ → a) ∙ ((λ _ → a) ∙ (λ _ → a) ∙ refl ∙ (λ i → t (~ i))) ∙ t)) ∙ - sym (lUnit (((λ _ → a) ∙ (λ _ → a) ∙ refl ∙ (λ i → t (~ i))) ∙ t)) ∙ - (λ k → (lUnit (lUnit (lUnit (sym t) (~ k)) (~ k)) (~ k)) ∙ t) ∙ - lCancel t ∙ - λ k → lUnit (lUnit refl k) k))) - - - -{- We will be able to reduce the inner transport to (RlFun a x1 n iscon (merid x1)) ...... - This can now be shown to give ∣ x1 , refl ∣. Note that this is refl of (merid x1), and not refl of (merid x1) ∙ (merid a)⁻ -} -codeTranspId3 : (n : ℕ) (a x1 : A) (iscon : is- (ℕ→ℕ₋₂ n) -ConnectedType A) → - (RlFun a x1 n iscon (merid x1)) (transport (λ i → cong (λ p₁ → ∥ fiber (λ y₁ → σ a y₁) p₁ ∥ ℕ→ℕ₋₂ (n + n)) (pairLemma (merid x1) (sym (merid x1))) i) - (transport (λ i → (transportRefl (∥ fiber (λ y₁ → σ a y₁) - (subst (_≡_ north) (sym (merid x1)) (merid x1)) ∥ ℕ→ℕ₋₂ (n + n)))i) - ∣ a , rCancel (merid a) ∙ sym (rCancel (merid x1)) ∙ sym (pairLemma (merid x1) (sym (merid x1))) ∣)) - ≡ ∣ x1 , refl ∣ -codeTranspId3 n a x1 iscon = (RlFun a x1 n iscon (merid x1)) (outer (inner guy)) - ≡⟨ (λ j → (RlFun a x1 n iscon (merid x1)) (outer (innerTransp j))) ⟩ - (RlFun a x1 n iscon (merid x1)) (outer guy) - ≡⟨ (λ j → (RlFun a x1 n iscon (merid x1)) (outerTransp j)) ⟩ - (RlFun a x1 n iscon (merid x1)) guy2 ≡⟨ refl ⟩ - sufMap n x1 a a iscon (merid x1) (rCancel (merid a) ∙ sym (rCancel (merid x1))) - ≡⟨ cong (λ x → x (rCancel (merid a) ∙ sym (rCancel (merid x1)))) - (sufMapId n a x1 iscon) ⟩ - ∣ x1 , switcher (merid a) (merid x1) (merid x1) (rCancel (merid a) ∙ sym (rCancel (merid x1))) ∣ - ≡⟨ (λ j → ∣ x1 , switcherLemma (merid a) (merid x1) j ∣) ⟩ - ∣ x1 , refl ∣ ∎ - where - switcherLemma : ∀ {ℓ} {A : Type ℓ} {a b c : A} (p : a ≡ b) (q : a ≡ c) → switcher p q q (rCancel p ∙ (sym (rCancel q))) ≡ refl - switcherLemma {A = A} {a = a} {c = c} = J (λ b p → (q : a ≡ c) → switcher p q q (rCancel p ∙ (sym (rCancel q))) ≡ refl) - (J (λ c q → switcher refl q q (rCancel refl ∙ (sym (rCancel q))) ≡ refl) - ((λ j → switcher refl refl refl (rCancel (rCancel refl) j )) ∙ - cong (λ x → x refl) (switcherRefl) ∙ - (λ j → lUnit refl ∙ - cong (λ x → x ∙ refl) - (lUnit refl ∙ (lUnit (sym (lUnit (sym refl))) (~ j))) ∙ - lCancel refl) ∙ - (λ j → lUnit refl ∙ - cong (λ x → x ∙ refl) - (rCancel (lUnit refl) j ) ∙ - lCancel refl) ∙ - (λ j → lUnit refl ∙ - lUnit (lCancel refl) (~ j)) ∙ - (λ j → rCancel (lUnit refl) j))) - guy : transport (λ _ → Type _) (∥ fiber (λ y₁ → σ a y₁) (subst (_≡_ north) (λ i → merid x1 (~ i)) (merid x1)) ∥ ℕ→ℕ₋₂ (n + n)) - guy = ∣ a , rCancel (merid a) ∙ sym (rCancel (merid x1)) ∙ sym (pairLemma (merid x1) (sym (merid x1))) ∣ - - guy2 : ∥ fiber (λ y → σ a y) (merid x1 ∙ sym (merid x1)) ∥ ℕ→ℕ₋₂ (n + n) - guy2 = ∣ a , rCancel (merid a) ∙ sym (rCancel (merid x1)) ∣ - - inner : transport (λ _ → Type _) (∥ fiber (λ y₁ → σ a y₁) (subst (_≡_ north) (λ i → merid x1 (~ i)) (merid x1)) ∥ ℕ→ℕ₋₂ (n + n)) → - ∥ fiber (λ y₁ → σ a y₁) (subst (_≡_ north) (λ i → merid x1 (~ i)) (merid x1)) ∥ ℕ→ℕ₋₂ (n + n) - inner = transport (λ i → (transportRefl (∥ fiber (λ y₁ → σ a y₁) (subst (_≡_ north) (sym (merid x1)) (merid x1)) ∥ ℕ→ℕ₋₂ (n + n)))i) - - outer : transport (λ _ → Type _) (∥ fiber (λ y₁ → σ a y₁) (subst (_≡_ north) (λ i → merid x1 (~ i)) (merid x1)) ∥ ℕ→ℕ₋₂ (n + n)) → - ∥ fiber (λ y → σ a y) (merid x1 ∙ sym (merid x1)) ∥ ℕ→ℕ₋₂ (n + n) - outer = transport (λ i → cong (λ p₁ → ∥ fiber (λ y₁ → σ a y₁) p₁ ∥ ℕ→ℕ₋₂ (n + n)) (pairLemma (merid x1) (sym (merid x1))) i) - - innerTransp : inner guy ≡ guy - innerTransp = - (λ j → (transport (λ i → (transportRefl (∥ fiber (λ y₁ → σ a y₁) (subst (_≡_ north) (sym (merid x1)) (merid x1)) ∥ ℕ→ℕ₋₂ (n + n))) (i ∨ j)) guy )) - ∙ transportRefl guy - outerTransp : outer guy ≡ guy2 - outerTransp = (λ j → transport (λ i → cong (λ p₁ → ∥ fiber (λ y₁ → σ a y₁) p₁ ∥ ℕ→ℕ₋₂ (n + n)) (pairLemma (merid x1) (sym (merid x1))) i) - ∣ a , rUnit (rCancel (merid a) ∙ sym (rCancel (merid x1)) ∙ sym (pairLemma (merid x1) (sym (merid x1)))) j ∣ ) ∙ - (λ j → transport (λ i → cong (λ p₁ → ∥ fiber (λ y₁ → σ a y₁) p₁ ∥ ℕ→ℕ₋₂ (n + n)) (pairLemma (merid x1) (sym (merid x1))) (i ∨ j)) - ∣ a , (rCancel (merid a) ∙ - sym (rCancel (merid x1)) ∙ - sym (pairLemma (merid x1) (sym (merid x1)))) ∙ - (λ i → (pairLemma (merid x1) (sym (merid x1))) (i ∧ j)) ∣) ∙ - (λ j → transportRefl (∣ a , (rCancel (merid a) ∙ - sym (rCancel (merid x1)) ∙ - sym (pairLemma (merid x1) (sym (merid x1)))) ∙ - (pairLemma (merid x1) (sym (merid x1))) ∣) j) ∙ - (λ j → ∣ a , assoc (rCancel (merid a)) - (sym (rCancel (merid x1))) - (sym (pairLemma (merid x1) (sym (merid x1)))) j - ∙ (pairLemma (merid x1) (sym (merid x1))) ∣ ) ∙ - (λ j → ∣ a , assoc ((rCancel (merid a)) ∙ (sym (rCancel (merid x1)))) - (sym (pairLemma (merid x1) (sym (merid x1)))) - (pairLemma (merid x1) (sym (merid x1))) (~ j) ∣) ∙ - (λ j → ∣ a , ((rCancel (merid a)) ∙ (sym (rCancel (merid x1)))) ∙ (lCancel (pairLemma (merid x1) (sym (merid x1))) j) ∣) ∙ - λ j → ∣ a , rUnit ((rCancel (merid a)) ∙ (sym (rCancel (merid x1)))) (~ j) ∣ - - - - -{- We now need to show that the outer transport applied to ∣ x1 , refl ∣ gives ∣ x1 , refl ∣. We transform the outer transport-} -codeTranspId4 : (n : ℕ) (a x1 : A) (iscon : is- (ℕ→ℕ₋₂ n) -ConnectedType A) → - transport (λ i → uncurry (CODE a n iscon) (pair⁼ (sym (merid a)) (pairLemma (merid x1) (sym (merid a)) ) i)) - ∣ x1 , refl ∣ - ≡ transport (λ i → uncurry (CODE a n iscon) (pair⁼ (λ i₁ → merid a (~ i₁)) (transpRCancel (λ i → north ≡ (merid a (~ i))) (merid x1 ∙ (sym (merid a)))) i)) - ∣ x1 , rUnit (merid x1) ∙ - sym (cong (λ x → merid x1 ∙ x) (lCancel (merid a))) ∙ - assocJ (merid x1) (sym (merid a)) (merid a) ∙ - sym (pairLemma {a1 = north} (merid x1 ∙ (sym (merid a))) (merid a)) ∣ -codeTranspId4 {ℓ} {A = A} n a x1 iscon = sym (helper north (sym (merid a))) - where - helper : (y : Susp A) → (q : south ≡ y) → transport (λ i → uncurry (CODE a n iscon) (pair⁼ q (transpRCancel (λ i → north ≡ q i) ((merid x1) ∙ q)) i)) - ∣ x1 , rUnit (merid x1) ∙ - sym (cong (λ x → merid x1 ∙ x) (lCancel (sym q))) ∙ - assocJ (merid x1) q (sym q) ∙ sym ((pairLemma {a1 = north} (merid x1 ∙ q) (sym q))) ∣ - ≡ transport (λ i → uncurry (CODE a n iscon) (pair⁼ q (pairLemma (merid x1) q ) i)) ∣ x1 , refl ∣ - helper y = J (λ y q → transport (λ i → uncurry (CODE a n iscon) (pair⁼ q (transpRCancel (λ i → north ≡ q i) ((merid x1) ∙ q)) i)) - ∣ x1 , rUnit (merid x1) ∙ - sym (cong (λ x → merid x1 ∙ x) (lCancel (sym q))) ∙ - assocJ (merid x1) q (sym q) ∙ sym ((pairLemma {a1 = north} (merid x1 ∙ q) (sym q))) ∣ - ≡ transport (λ i → uncurry (CODE a n iscon) (pair⁼ q (pairLemma (merid x1) q ) i)) ∣ x1 , refl ∣) - (transport (λ i → uncurry (CODE a n iscon) (pair⁼ refl (transpRCancel (λ i₁ → north ≡ refl i₁) (merid x1 ∙ refl)) i)) ∣ x1 , originalPath ∣ - ≡⟨ (λ j → (transport (λ i → uncurry (CODE a n iscon) (pair⁼Refl (transpRCancel (λ i₁ → north ≡ south) (merid x1 ∙ refl)) j i))) ∣ x1 , originalPath ∣) ⟩ - (transport (λ i → uncurry (CODE a n iscon) - (south , ((sym (transportRefl (transport (λ i → _≡_ {A = Susp A} north south) (merid x1 ∙ (λ _ → south))))) ∙ - (transpRCancel {A = north ≡ south} {B = north ≡ south} (λ _ → _≡_ {A = Susp A} north south) - (merid x1 ∙ (refl {x = south})))) i) )) - ∣ x1 , originalPath ∣ - ≡⟨ (λ j → (transport (λ i → uncurry (CODE a n iscon) - (south , transportCanceller (merid x1 ∙ (λ _ → south)) j i) )) - ∣ x1 , originalPath ∣) ⟩ - (transport (λ i → uncurry (CODE a n iscon) (south , transportRefl (merid x1 ∙ (λ _ → south)) i) )) - ∣ x1 , originalPath ∣ - ≡⟨ (λ j → (transport (λ i → uncurry (CODE a n iscon) (south , transportRefl (merid x1 ∙ (λ _ → south)) i) )) - ∣ x1 , rUnit originalPath j ∣) ⟩ - (transport (λ i → uncurry (CODE a n iscon) (south , transportRefl (merid x1 ∙ (λ _ → south)) i) )) - ∣ x1 , originalPath ∙ refl ∣ - ≡⟨ (λ j → (transport (λ i → uncurry (CODE a n iscon) (south , (transportRefl {A = _≡_ {A = Susp A} north south} (merid x1 ∙ (λ _ → south)) (i ∨ j) ) ) )) - ∣ x1 , originalPath ∙ (λ k → transportRefl {A = _≡_ {A = Susp A} north south} (merid x1 ∙ (λ _ → south)) (k ∧ j)) ∣) ⟩ - (transport (λ i → uncurry (CODE a n iscon) (south , (merid x1 ∙ (λ _ → south)) ) )) - ∣ x1 , originalPath ∙ transportRefl {A = _≡_ {A = Susp A} north south} (merid x1 ∙ (λ _ → south)) ∣ - ≡⟨ (λ j → transport (λ i → uncurry (CODE a n iscon) (south , merid x1 ∙ (λ _ → south))) - ∣ x1 , pathId (merid x1) j ∣) ⟩ - transport (λ i → uncurry (CODE a n iscon) (south , (merid x1 ∙ (λ _ → south)))) ∣ x1 , (λ _ → merid x1) ∙ rUnit (merid x1) ∣ - ≡⟨ (λ j → transport (λ i → uncurry (CODE a n iscon) (south , rUnit (merid x1) (i ∨ (~ j) ))) ∣ x1 , (λ _ → merid x1) ∙ ( λ i → rUnit (merid x1) (i ∧ (~ j))) ∣) ⟩ - transport (λ i → uncurry (CODE a n iscon) (south , rUnit (merid x1) i)) ∣ x1 , (λ _ → merid x1) ∙ refl ∣ - ≡⟨ (λ j → transport (λ i → uncurry (CODE a n iscon) (south , rUnit (merid x1) i)) ∣ x1 , rUnit (λ _ → merid x1) (~ j) ∣) ⟩ - transport (λ i → uncurry (CODE a n iscon) (south , rUnit (merid x1) i)) ∣ x1 , (λ _ → merid x1) ∣ - ≡⟨ (λ j → transport (λ i → uncurry (CODE a n iscon) (south , transportCanceller2 (merid x1) (~ j) i)) ∣ x1 , (λ _ → merid x1) ∣) ⟩ - transport (λ i → uncurry (CODE a n iscon) (south , (sym (transportRefl (merid x1)) ∙ (pairLemma (merid x1) refl)) i)) ∣ x1 , (λ _ → merid x1) ∣ - ≡⟨ (λ j → transport (λ i → uncurry (CODE a n iscon) (pair⁼Refl (pairLemma (merid x1) (λ _ → south)) (~ j) i )) ∣ x1 , (λ _ → merid x1) ∣) ⟩ - transport (λ i → uncurry (CODE a n iscon) (pair⁼ refl (pairLemma (merid x1) refl) i)) ∣ x1 , (λ _ → merid x1) ∣ ∎) - - where - originalPath : merid x1 ≡ transport (λ i₁ → _≡_ {A = Susp A} north south) (merid x1 ∙ (λ _ → south)) - originalPath = rUnit (merid x1) ∙ (λ i → merid x1 ∙ lCancel (refl {x = south}) (~ i)) ∙ assocJ (merid x1) refl refl ∙ - (λ i → pairLemma (merid x1 ∙ refl) refl (~ i)) - - pathId : ∀ {A : Type ℓ} {x y : A} (p : x ≡ y) → (rUnit p ∙ - (λ i → p ∙ lCancel (refl) (~ i)) ∙ - assocJ p refl refl ∙ - (λ i → pairLemma (p ∙ refl) refl (~ i))) ∙ - transportRefl (p ∙ refl) - ≡ - refl ∙ rUnit p - pathId {x = x} = J (λ y p → (rUnit p ∙ - (λ i → p ∙ lCancel (refl) (~ i)) ∙ - assocJ p refl refl ∙ - (λ i → pairLemma (p ∙ refl) refl (~ i))) ∙ - transportRefl (p ∙ refl) - ≡ - refl ∙ rUnit p) - ((λ j → (rUnit refl ∙ - (λ i → refl ∙ lCancel (refl) (~ i)) ∙ - assocJRefl j ∙ - sym (pairLemmaId (refl ∙ refl) refl j)) ∙ - transportRefl (refl ∙ refl)) ∙ - (λ j → (rUnit refl ∙ - (λ i → refl ∙ lCancel (refl) (~ i)) ∙ - ((λ i → refl ∙ rCancel refl i) ∙ rUnit (refl ∙ refl)) ∙ - sym (pairLemma*Refl (refl ∙ refl) j)) ∙ - transportRefl (refl ∙ refl)) ∙ - (λ j → (rUnit refl ∙ - (λ i → refl ∙ lCancel (refl) (~ i)) ∙ - ((λ i → refl ∙ rCancel refl i) ∙ rUnit (refl ∙ refl)) ∙ - symDistr (transportRefl (refl ∙ refl)) (rUnit (refl ∙ refl)) j ) ∙ - transportRefl (refl ∙ refl)) ∙ - invKiller (rUnit refl) (λ i → refl ∙ lCancel (refl) (~ i)) (rUnit (refl ∙ refl)) (sym (transportRefl (refl ∙ refl))) ∙ - lUnit (rUnit refl)) - where - invKiller : ∀ {ℓ} {A : Type ℓ} {a b c d e : A} (p : a ≡ b) (q : b ≡ c) (r : b ≡ d) (s : b ≡ e) → - (p ∙ q ∙ (sym q ∙ r) ∙ sym r ∙ s) ∙ (sym s) ≡ p - invKiller {a = a} {b = b} {c = c} {d = d} {e = e} p = J (λ c q → (r : b ≡ d) (s : b ≡ e) → - (p ∙ q ∙ (sym q ∙ r) ∙ sym r ∙ s) ∙ (sym s) ≡ p) - (J (λ d r → (s : b ≡ e) → (p ∙ refl ∙ (refl ∙ r) ∙ sym r ∙ s) ∙ (sym s) ≡ p) - (J (λ e s → (p ∙ refl ∙ (refl ∙ refl) ∙ refl ∙ s) ∙ (sym s) ≡ p) - ((λ i → rUnit (p ∙ (λ _ → b) ∙ (rUnit refl (~ i)) ∙ refl ∙ refl) (~ i)) ∙ - (λ i → p ∙ (lUnit (lUnit (lUnit refl (~ i)) (~ i)) (~ i) )) ∙ - sym (rUnit p)))) - - - transportCanceller : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) → - sym (transportRefl (transport (λ i → x ≡ y) p )) ∙ transpRCancel (λ _ → x ≡ y) p ≡ transportRefl p - transportCanceller {x = x} {y = y} p = (λ j → sym (transportRefl (transport (λ i → x ≡ y) p)) ∙ (transpRCancelRefl p j)) ∙ - assoc (sym (transportRefl (transport (λ i → x ≡ y) p))) - ((transportRefl (transport (λ i → x ≡ y) p))) - (transportRefl p) ∙ - (λ j → lCancel (transportRefl (transport (λ i → x ≡ y) p)) j ∙ transportRefl p) ∙ - sym (lUnit (transportRefl p)) - - transportCanceller2 : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) → - (sym (transportRefl p) ∙ (pairLemma p refl)) ≡ rUnit p - transportCanceller2 {x = x} = J (λ y p → (sym (transportRefl p) ∙ (pairLemma p refl)) ≡ rUnit p) - ((λ j → sym (transportRefl refl) ∙ pairLemmaRefl refl j) ∙ - (λ j → sym (transportRefl refl) ∙ pairLemma2Refl j ∙ lUnit refl) ∙ - assoc (sym (transportRefl refl)) (transportRefl refl) (lUnit refl) ∙ - (λ j → lCancel (transportRefl refl) j ∙ lUnit refl) ∙ - sym (lUnit (lUnit refl)) ) - - - - -{- We will use transportLemma to move the transport to the other side. Applying Lemma8610Refl, we will get a transport into a path like the one below. codeTranspId5 transforms this expression into something more useful -} -codeTranspId5 : ∀ {ℓ ℓ'} {X : Type ℓ} {x y : X} {A : x ≡ x → Type ℓ'} {B : x ≡ y → Type ℓ'} (ma mx1 : x ≡ y) → - (r : ((mx1 : x ≡ y) → A (mx1 ∙ (sym ma)) ≡ B mx1)) → - (λ i → A (transpRCancel (λ i₁ → x ≡ ma (~ i₁)) (mx1 ∙ sym ma) (~ i))) ∙ - funExt⁻ (Lemma296-fun {X = X} {A = λ y → x ≡ y} ma A B (helperFun {X = X} ma - {A = A} - {B = B} - r)) - (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma)) - ≡ - r mx1 ∙ - cong B (sym (pairLemma (mx1 ∙ sym ma) ma ∙ sym (assocJ mx1 (sym ma) ma) ∙ (λ i → mx1 ∙ lCancel ma i) ∙ sym (rUnit mx1))) -codeTranspId5 {ℓ' = ℓ'} {X = X} {x = x} {A = A} {B = B} ma mx1 r = - J (λ y ma → (A : x ≡ x → Type ℓ') (B : x ≡ y → Type ℓ') (mx1 : x ≡ y) → - (r : ((mx1 : x ≡ y) → A (mx1 ∙ (sym ma)) ≡ B mx1)) → - (λ i → A (transpRCancel (λ i₁ → x ≡ ma (~ i₁)) (mx1 ∙ sym ma) (~ i))) ∙ - funExt⁻ (Lemma296-fun {X = X} {A = λ y → x ≡ y} ma A B (helperFun {X = X} ma - {A = A} - {B = B} - r)) - (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma)) - ≡ - r mx1 ∙ - cong B (sym (pairLemma (mx1 ∙ sym ma) ma ∙ sym (assocJ mx1 (sym ma) ma) ∙ (λ i → mx1 ∙ lCancel ma i) ∙ sym (rUnit mx1)))) - (λ A B mx1 r → (λ k → (λ i → A (transpRCancel (λ i₁ → x ≡ x) (mx1 ∙ (λ _ → x)) (~ i))) ∙ - ((codeTranspId5' {X = X} {x = x} {A = A} {B = B} (λ _ → x) mx1 r) ∙ - (codeTranspId5'' {X = X} {x = x} {A = A} {B = B} (λ _ → x) mx1 r)) k) ∙ - (λ k → ((λ i → A (transpRCancelRefl (mx1 ∙ (λ _ → x)) k (~ i)))) ∙ - (cong (λ f → f (transport (λ i → x ≡ x) (mx1 ∙ (λ _ → x)))) (Lemma294Refl {A = λ y → x ≡ y} {B = λ _ → Type ℓ'} A k) ∙ - transportRefl (A (transport (λ i → x ≡ x) (transport (λ i → x ≡ x) (mx1 ∙ (λ _ → x))))) ∙ - cong A (transpRCancelRefl (mx1 ∙ (λ _ → x)) k)) ∙ - r mx1 ∙ - cong B (sym (pairLemma (mx1 ∙ (λ _ → x)) (λ _ → x) ∙ sym (assocJ mx1 (λ _ → x) (λ _ → x)) ∙ - (λ i → mx1 ∙ lCancel (λ _ → x) i) ∙ sym (rUnit mx1)))) ∙ - (λ k → (λ i → A ((transportRefl (transport refl (mx1 ∙ (λ _ → x))) ∙ transportRefl (mx1 ∙ (λ _ → x))) (~ i))) ∙ - lUnit (transportRefl (A (transport (λ i → x ≡ x) (transport (λ i → x ≡ x) (mx1 ∙ (λ _ → x))))) ∙ - cong A (transportRefl (transport refl (mx1 ∙ (λ _ → x))) ∙ transportRefl (mx1 ∙ (λ _ → x)))) (~ k) ∙ - r mx1 ∙ - cong B (sym (pairLemma (mx1 ∙ (λ _ → x)) (λ _ → x) ∙ sym (assocJ mx1 (λ _ → x) (λ _ → x)) ∙ - (λ i → mx1 ∙ lCancel (λ _ → x) i) ∙ sym (rUnit mx1)))) ∙ - (λ k → (λ i → A ((transportRefl (transport refl (mx1 ∙ (λ _ → x))) ∙ transportRefl (mx1 ∙ (λ _ → x))) (~ i))) ∙ - (lUnit (cong A (transportRefl (transport refl (mx1 ∙ (λ _ → x))) ∙ transportRefl (mx1 ∙ (λ _ → x)))) (~ k)) ∙ - r mx1 ∙ - cong B (sym (pairLemma (mx1 ∙ (λ _ → x)) (λ _ → x) ∙ sym (assocJ mx1 (λ _ → x) (λ _ → x)) ∙ - (λ i → mx1 ∙ lCancel (λ _ → x) i) ∙ sym (rUnit mx1)))) ∙ - (assoc (sym (cong A (transportRefl (transport refl (mx1 ∙ (λ _ → x))) ∙ transportRefl (mx1 ∙ (λ _ → x))))) - (cong A (transportRefl (transport refl (mx1 ∙ (λ _ → x))) ∙ transportRefl (mx1 ∙ (λ _ → x)))) - (r mx1 ∙ - cong B (sym (pairLemma (mx1 ∙ (λ _ → x)) (λ _ → x) ∙ sym (assocJ mx1 (λ _ → x) (λ _ → x)) ∙ - (λ i → mx1 ∙ lCancel (λ _ → x) i) ∙ sym (rUnit mx1))))) ∙ - (λ k → lCancel (cong A (transportRefl (transport refl (mx1 ∙ (λ _ → x))) ∙ transportRefl (mx1 ∙ (λ _ → x)))) k ∙ - r mx1 ∙ - cong B (sym (pairLemma (mx1 ∙ (λ _ → x)) (λ _ → x) ∙ sym (assocJ mx1 (λ _ → x) (λ _ → x)) ∙ - (λ i → mx1 ∙ lCancel (λ _ → x) i) ∙ sym (rUnit mx1)))) ∙ - sym (lUnit (r mx1 ∙ - cong B (sym (pairLemma (mx1 ∙ (λ _ → x)) (λ _ → x) ∙ sym (assocJ mx1 (λ _ → x) (λ _ → x)) ∙ - (λ i → mx1 ∙ lCancel (λ _ → x) i) ∙ sym (rUnit mx1)))))) - ma A B mx1 r - - where - codeTranspId5' : ∀ {ℓ ℓ'} {X : Type ℓ} {x y : X} {A : x ≡ x → Type ℓ'} {B : x ≡ y → Type ℓ'} (ma mx1 : x ≡ y) → - (r : ((mx1 : x ≡ y) → A (mx1 ∙ (sym ma)) ≡ B mx1)) → - funExt⁻ (Lemma296-fun {X = X} {A = λ y → x ≡ y} - ma A - B - (helperFun {X = X} ma - {A = A} - {B = B} - r)) - (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma)) - ≡ - cong (λ f → f (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma))) (Lemma294 {A = λ y → x ≡ y} {B = λ _ → Type ℓ'} ma A) ∙ - transportRefl (A (transport (λ i → x ≡ ma (~ i)) (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma)))) ∙ - cong A (pairLemma (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma)) (sym ma)) ∙ - r (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma)) - codeTranspId5' {ℓ' = ℓ'} {X = X} {x = x} {y = y} {A = A} {B = B} ma mx1 = - - J (λ y ma → {A : x ≡ x → Type ℓ'} {B : x ≡ y → Type ℓ'} (mx1 : x ≡ y) → - (r : ((mx1 : x ≡ y) → A (mx1 ∙ (sym ma)) ≡ B mx1)) → - funExt⁻ (Lemma296-fun {X = X} {A = λ y → x ≡ y} - ma A - B - (helperFun {X = X} ma - {A = A} - {B = B} - r)) - (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma)) - ≡ - cong (λ f → f (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma))) (Lemma294 {A = λ y → x ≡ y} {B = λ _ → Type ℓ'} ma A) ∙ - transportRefl (A (transport (λ i → x ≡ ma (~ i)) (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma)))) ∙ - cong A (pairLemma (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma)) (sym ma)) ∙ - r (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma))) - (λ {A} {B} mx1 r → (λ k → funExt⁻ (Lemma296-funRefl {X = X} {A = λ y → x ≡ y} A B k - (helperFunRefl {X = X} {A = A} {B = B} k r)) - (transport (λ i → x ≡ x) (mx1 ∙ (λ _ → x)))) ∙ - (λ k → funExt⁻ (transportRefl A ∙ - funExt λ z → sym (transportRefl (A z)) ∙ - (transportRefl (A z) ∙ - cong A (rUnit z) ∙ - r z ∙ - cong B (sym (transportRefl z))) ∙ - λ i → (B (transportRefl z i))) - (transport (λ i → x ≡ x) (mx1 ∙ (λ _ → x)))) ∙ - (λ k → funExt⁻ (transportRefl A ∙ - funExt λ z → littleCanceller (sym (transportRefl (A z))) - (cong A (rUnit z)) - (r z) - (cong B (sym (transportRefl z))) k) - (transport (λ i → x ≡ x) (mx1 ∙ (λ _ → x)))) ∙ - (λ k → funExt⁻ (transportRefl A) (transport (λ i → x ≡ x) (mx1 ∙ (λ _ → x))) ∙ - cong A (rUnit (transport (λ i → x ≡ x) (mx1 ∙ (λ _ → x)))) ∙ r (transport (λ i → x ≡ x) (mx1 ∙ (λ _ → x)))) ∙ - (assoc (funExt⁻ (transportRefl A) (transport (λ i → x ≡ x) (mx1 ∙ (λ _ → x)))) - ( cong A (rUnit (transport (λ i → x ≡ x) (mx1 ∙ (λ _ → x))))) - (r (transport (λ i → x ≡ x) (mx1 ∙ (λ _ → x))))) ∙ - (λ k → lemma1 {X = X} {A = A} mx1 k ∙ - r (transport (λ i → x ≡ x) (mx1 ∙ (λ _ → x)))) ∙ - (sym (assoc (transportRefl - (A (transport (λ i → x ≡ x) - (transport (λ i → x ≡ x) (mx1 ∙ (λ i → x)))))) - ((λ i → A (pairLemma (transport (λ i₁ → x ≡ x) (mx1 ∙ (λ i₁ → x))) (λ i₁ → x) i))) - (r (transport (λ i → x ≡ x) (mx1 ∙ (λ i → x)))))) ∙ - (lUnit (transportRefl - (A (transport (λ i → x ≡ x) - (transport (λ i → x ≡ x) (mx1 ∙ (λ i → x))))) ∙ - (λ i → A (pairLemma (transport (λ i₁ → x ≡ x) (mx1 ∙ (λ i₁ → x))) (λ i₁ → x) i)) ∙ - r (transport (λ i → x ≡ x) (mx1 ∙ (λ i → x))))) ∙ - λ k → (λ i → Lemma294Refl {A = λ y → x ≡ y} {B = λ _ → Type ℓ'} A (~ k) i - (transport (λ i₁ → x ≡ x) (mx1 ∙ (λ i₁ → x)))) ∙ - transportRefl (A (transport (λ i → x ≡ x) - (transport (λ i → x ≡ x) (mx1 ∙ (λ i → x))))) ∙ - (λ i → A (pairLemma (transport (λ i₁ → x ≡ x) (mx1 ∙ (λ i₁ → x))) (λ i₁ → x) i)) ∙ - r (transport (λ i → x ≡ x) (mx1 ∙ (λ i → x)))) - - ma {A} {B} mx1 - where - lemma1 : ∀ {ℓ ℓ'} {X : Type ℓ} {x : X} {A : x ≡ x → Type ℓ'} (mx1 : x ≡ x) → - funExt⁻ (transportRefl A) (transport (λ i → x ≡ x) (mx1 ∙ (λ _ → x))) ∙ - cong A (rUnit (transport (λ i → x ≡ x) (mx1 ∙ (λ _ → x)))) ≡ - transportRefl (A (transport (λ i → x ≡ x) (transport (λ i → x ≡ x) (mx1 ∙ (λ _ → x))))) ∙ - (λ i → A (pairLemma (transport (λ i₁ → x ≡ x) (mx1 ∙ (λ _ → x))) (λ _ → x) i)) - lemma1 {x = x} {A = A} mx1 = - sym ((λ k → transportRefl (A (transport (λ i → x ≡ x) (transport (λ i → x ≡ x) (mx1 ∙ (λ _ → x))))) ∙ - (λ i → A (pairLemmaId (transport (λ i₁ → x ≡ x) (mx1 ∙ (λ _ → x))) (λ _ → x) k i))) ∙ - (λ k → transportRefl (A (transport (λ i → x ≡ x) (transport (λ i → x ≡ x) (mx1 ∙ (λ _ → x))))) ∙ - (λ i → A (pairLemma*Refl (transport (λ i₁ → x ≡ x) (mx1 ∙ (λ _ → x))) k i))) ∙ - (λ k → transportRefl (A (transport (λ i → x ≡ x) (transport (λ i → x ≡ x) (mx1 ∙ (λ _ → x))))) ∙ - λ i → A ((transportRefl (transport (λ i₁ → x ≡ x) (mx1 ∙ (λ _ → x))) ∙ - rUnit (transport (λ i₁ → x ≡ x) (mx1 ∙ (λ _ → x)))) i)) ∙ - (λ k → transportRefl (A (transport (λ i → x ≡ x) (transport (λ i → x ≡ x) (mx1 ∙ (λ _ → x))))) ∙ - congComp2 A (transportRefl (transport (λ i₁ → x ≡ x) (mx1 ∙ (λ _ → x)))) - (rUnit (transport (λ i₁ → x ≡ x) (mx1 ∙ (λ _ → x)))) (~ k)) ∙ - (λ k → (λ i → transportRefl (A (transport (λ i → x ≡ x) (transport (λ i → x ≡ x) (mx1 ∙ (λ _ → x))))) (i ∨ k)) ∙ - cong A (transportRefl (transport (λ i₁ → x ≡ x) (mx1 ∙ (λ _ → x)))) ∙ - cong A (rUnit (transport (λ i₁ → x ≡ x) (mx1 ∙ (λ _ → x))))) ∙ - (sym (lUnit (cong A (transportRefl (transport (λ i₁ → x ≡ x) (mx1 ∙ (λ _ → x)))) ∙ - cong A (rUnit (transport (λ i₁ → x ≡ x) (mx1 ∙ (λ _ → x)))))))) - - littleCanceller : ∀ {ℓ} {A : Type ℓ} {a b c d e : A} → (p : a ≡ b) (q : a ≡ c) (r : c ≡ d) (s : d ≡ e) → - p ∙ (sym p ∙ q ∙ r ∙ s) ∙ sym s ≡ q ∙ r - littleCanceller {a = a} {b = b} {c = c} p q r s = - p ∙ (sym p ∙ q ∙ r ∙ s) ∙ sym s ≡⟨ assoc p (sym p ∙ q ∙ r ∙ s) (sym s) ⟩ - (p ∙ sym p ∙ q ∙ r ∙ s) ∙ sym s ≡⟨ (λ j → assoc p (sym p) (q ∙ r ∙ s) j ∙ sym s) ⟩ - ((p ∙ sym p) ∙ q ∙ r ∙ s) ∙ sym s ≡⟨ (λ j → (rCancel p j ∙ q ∙ r ∙ s) ∙ sym s) ⟩ - (refl ∙ q ∙ r ∙ s) ∙ sym s ≡⟨ (λ j → lUnit (q ∙ r ∙ s) (~ j) ∙ sym s) ⟩ - (q ∙ r ∙ s) ∙ sym s ≡⟨ (λ j → assoc q r s j ∙ sym s) ⟩ - ((q ∙ r) ∙ s) ∙ sym s ≡⟨ sym (assoc (q ∙ r) s (sym s)) ⟩ - (q ∙ r) ∙ s ∙ sym s ≡⟨ (λ j → (q ∙ r) ∙ (rCancel s j)) ⟩ - (q ∙ r) ∙ refl ≡⟨ sym (rUnit (q ∙ r)) ⟩ - q ∙ r ∎ - - codeTranspId5'' : ∀ {ℓ ℓ'} {X : Type ℓ} {x y : X} {A : x ≡ x → Type ℓ'} {B : x ≡ y → Type ℓ'} (ma mx1 : x ≡ y) → - (r : ((mx1 : x ≡ y) → A (mx1 ∙ (sym ma)) ≡ B mx1)) → - cong (λ f → f (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma))) (Lemma294 {A = λ y → x ≡ y} {B = λ _ → Type ℓ'} ma A) ∙ - transportRefl (A (transport (λ i → x ≡ ma (~ i)) (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma)))) ∙ - cong A (pairLemma (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma)) (sym ma)) ∙ - r (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma)) - ≡ - (cong (λ f → f (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma))) (Lemma294 {A = λ y → x ≡ y} {B = λ _ → Type ℓ'} ma A) ∙ - transportRefl (A (transport (λ i → x ≡ ma (~ i)) (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma)))) ∙ - cong A (transpRCancel (λ i → x ≡ ma (~ i)) (mx1 ∙ sym ma))) ∙ - r mx1 ∙ - cong B (sym (pairLemma (mx1 ∙ sym ma) ma ∙ sym (assocJ mx1 (sym ma) ma) ∙ (λ i → mx1 ∙ lCancel ma i) ∙ sym (rUnit mx1))) - codeTranspId5'' {ℓ' = ℓ'} {X = X} {x = x} {y = y} {A = A} {B = B} ma mx1 = - J (λ y ma → ( mx1 : x ≡ y) {A : x ≡ x → Type ℓ'} {B : x ≡ y → Type ℓ'} - (r : ((mx1 : x ≡ y) → A (mx1 ∙ (sym ma)) ≡ B mx1)) → - cong (λ f → f (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma))) (Lemma294 {A = λ y → x ≡ y} {B = λ _ → Type ℓ'} ma A) ∙ - transportRefl (A (transport (λ i → x ≡ ma (~ i)) (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma)))) ∙ - cong A (pairLemma (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma)) (sym ma)) ∙ - r (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma)) - ≡ - (cong (λ f → f (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma))) (Lemma294 {A = λ y → x ≡ y} {B = λ _ → Type ℓ'} ma A) ∙ - transportRefl (A (transport (λ i → x ≡ ma (~ i)) (transport (λ i → x ≡ ma i) (mx1 ∙ sym ma)))) ∙ - cong A (transpRCancel (λ i → x ≡ ma (~ i)) (mx1 ∙ sym ma))) ∙ - r mx1 ∙ - cong B (sym (pairLemma (mx1 ∙ sym ma) ma ∙ sym (assocJ mx1 (sym ma) ma) ∙ (λ i → mx1 ∙ lCancel ma i) ∙ sym (rUnit mx1)))) - (λ p {A} {B} r → (λ k → cong (λ f → f (transport (λ i → x ≡ x) (p ∙ (λ _ → x)))) (Lemma294Refl {A = λ y → x ≡ y} {B = λ _ → Type ℓ'} A k) ∙ - transportRefl (A (transport (λ i → x ≡ x) (transport (λ i → x ≡ x) (p ∙ (λ _ → x))))) ∙ - cong A (pairLemma (transport (λ i → x ≡ x) (p ∙ (λ _ → x))) (λ _ → x)) ∙ - r (transport (λ i → x ≡ x) (p ∙ (λ _ → x)))) ∙ - (sym (lUnit (transportRefl (A (transport (λ i → x ≡ x) (transport (λ i → x ≡ x) (p ∙ (λ _ → x))))) ∙ - cong A (pairLemma (transport (λ i → x ≡ x) (p ∙ (λ _ → x))) (λ _ → x)) ∙ - r (transport (λ i → x ≡ x) (p ∙ (λ _ → x)))))) ∙ - (sym (lUnit (cong A (pairLemma (transport (λ i → x ≡ x) (p ∙ (λ _ → x))) (λ _ → x)) ∙ - r (transport (λ i → x ≡ x) (p ∙ (λ _ → x)))))) ∙ - (λ k → cong A (pairLemmaId (transport (λ i → x ≡ x) (p ∙ (λ _ → x))) (λ _ → x) k) ∙ - r (transport (λ i → x ≡ x) (p ∙ (λ _ → x)))) ∙ - (λ k → cong A (pairLemma*Refl (transport (λ i → x ≡ x) (p ∙ (λ _ → x))) k) ∙ - r (transport (λ i → x ≡ x) (p ∙ (λ _ → x)))) ∙ - assocCancel1 (cong A (transportRefl ((transport (λ i → x ≡ x) (p ∙ (λ _ → x)))) ∙ - rUnit (transport (λ i → x ≡ x) (p ∙ (λ _ → x))))) - (r (transport (λ i → x ≡ x) (p ∙ (λ _ → x)))) ∙ - (λ k → cong A (transportRefl ((transport (λ i → x ≡ x) (p ∙ (λ _ → x)))) ∙ - rUnit (transport (λ i → x ≡ x) (p ∙ (λ _ → x)))) ∙ - cong A (λ i → (transportRefl (p ∙ (λ _ → x)) (k ∧ i)) ∙ refl) ∙ refl ∙ - r (transportRefl (p ∙ (λ _ → x)) k) ∙ - refl ∙ cong B λ i → (transportRefl (p ∙ (λ _ → x)) (k ∧ (~ i)))) ∙ - (λ k → cong A (transportRefl ((transport (λ i → x ≡ x) (p ∙ (λ _ → x)))) ∙ - rUnit (transport (λ i → x ≡ x) (p ∙ (λ _ → x)))) ∙ - cong A (λ i → (transportRefl (p ∙ (λ _ → x)) i) ∙ refl) ∙ cong A (λ i → rUnit p (~ k ∨ (~ i)) ∙ refl) ∙ - r (rUnit p (~ k)) ∙ - cong B (λ i → rUnit p (~ k ∨ i)) ∙ cong B λ i → ((transportRefl (p ∙ (λ _ → x)) (~ i)))) ∙ - (λ k → congComp2 A (transportRefl ((transport (λ i → x ≡ x) (p ∙ (λ _ → x))))) (rUnit (transport (λ i → x ≡ x) (p ∙ (λ _ → x)))) - (~ k) ∙ - cong A (λ i → (transportRefl (p ∙ (λ _ → x)) i) ∙ refl) ∙ - lUnit (cong A (λ i → rUnit p (~ i) ∙ refl)) k ∙ - r p ∙ - cong B (λ i → rUnit p i) ∙ cong B λ i → (transportRefl (p ∙ (λ _ → x)) (~ i))) ∙ - reflAdder (cong A (transportRefl ((transport (λ i → x ≡ x) (p ∙ (λ _ → x)))))) - (cong A ((rUnit (transport (λ i → x ≡ x) (p ∙ (λ _ → x)))))) - (cong A (λ i → (transportRefl (p ∙ (λ _ → x)) i) ∙ refl)) - (cong A (λ i → rUnit p (~ i) ∙ refl)) - (r p) - (cong B (λ i → rUnit p i) ∙ cong B λ i → (transportRefl (p ∙ (λ _ → x)) (~ i))) ∙ - (λ k → (cong A (transportRefl ((transport (λ i → x ≡ x) (p ∙ (λ _ → x))))) ∙ - helper {A = A} {B = B} p k ) ∙ - r p ∙ - cong B (λ i → rUnit p i) ∙ cong B λ i → (transportRefl (p ∙ (λ _ → x)) (~ i))) ∙ - (λ k → congComp2 A (transportRefl ((transport (λ i → x ≡ x) (p ∙ (λ _ → x))))) (transportRefl (p ∙ (λ _ → x))) k ∙ - r p ∙ - (congComp2 B (rUnit p) (λ i → (transportRefl (p ∙ (λ _ → x)) (~ i))) k)) ∙ - (λ k → cong A (transportRefl (transport refl (p ∙ (λ _ → x))) ∙ transportRefl (p ∙ (λ _ → x))) ∙ - r p ∙ - cong B (symDistr (transportRefl (p ∙ sym (λ _ → x))) (sym (rUnit p)) (~ k))) ∙ - (λ k → cong A (transportRefl (transport refl (p ∙ (λ _ → x))) ∙ transportRefl (p ∙ (λ _ → x))) ∙ - r p ∙ - cong B (sym (pairLemmaCancel p (~ k)))) ∙ - (λ k → cong A (transpRCancelRefl (p ∙ (λ _ → x)) (~ k)) ∙ - r p ∙ - cong B (sym (pairLemma (p ∙ sym (λ _ → x)) (λ _ → x) ∙ - sym (assocJ p (λ _ → x) (λ _ → x)) ∙ - (λ i → p ∙ lCancel (λ _ → x) i) ∙ sym (rUnit p)))) ∙ - (λ k → lUnit (lUnit (cong A (transpRCancel (λ i → x ≡ x) (p ∙ (λ _ → x)))) k) k ∙ - r p ∙ - cong B (sym (pairLemma (p ∙ (λ _ → x)) (λ _ → x) ∙ - sym (assocJ p (λ _ → x) (λ _ → x)) ∙ - (λ i → p ∙ lCancel (λ _ → x) i) ∙ sym (rUnit p)))) ∙ - λ k → (cong (λ f → f (transport (λ i → x ≡ x) (p ∙ (λ _ → x)))) (Lemma294Refl {A = λ y → x ≡ y} {B = λ _ → Type ℓ'} A (~ k)) ∙ - transportRefl (A (transport (λ i → x ≡ x) (transport (λ i → x ≡ x) (p ∙ (λ _ → x))))) ∙ - cong A (transpRCancel (λ i → x ≡ x) (p ∙ (λ _ → x)))) ∙ - r p ∙ - cong B (sym (pairLemma (p ∙ (λ _ → x)) (λ _ → x) ∙ - sym (assocJ p (λ _ → x) (λ _ → x)) ∙ - (λ i → p ∙ lCancel (λ _ → x) i) ∙ sym (rUnit p)))) - ma mx1 {A} {B} - - where - - - - reflAdder : ∀ {ℓ} {A : Type ℓ} {a b c d e f g : A} (p : a ≡ b) (q : b ≡ c) (r : c ≡ d) (s : d ≡ e) (t : e ≡ f) (u : f ≡ g) → - (p ∙ q) ∙ r ∙ (refl ∙ s) ∙ t ∙ u ≡ (p ∙ q ∙ r ∙ refl ∙ s) ∙ t ∙ u - reflAdder p q r s t u = (p ∙ q) ∙ r ∙ (refl ∙ s) ∙ t ∙ u ≡⟨ (λ k → (p ∙ q) ∙ r ∙ (lUnit s (~ k)) ∙ t ∙ u) ⟩ - (p ∙ q) ∙ r ∙ s ∙ t ∙ u ≡⟨ assoc (p ∙ q) r (s ∙ t ∙ u) ⟩ - ((p ∙ q) ∙ r) ∙ s ∙ t ∙ u ≡⟨ assoc ((p ∙ q) ∙ r) s (t ∙ u) ⟩ - (((p ∙ q) ∙ r) ∙ s) ∙ t ∙ u ≡⟨ (λ k → assoc (p ∙ q) r s (~ k) ∙ t ∙ u) ⟩ - ((p ∙ q) ∙ r ∙ s) ∙ t ∙ u ≡⟨ (λ k → assoc p q (r ∙ s) (~ k) ∙ t ∙ u) ⟩ - (p ∙ q ∙ r ∙ s) ∙ t ∙ u ≡⟨ (λ k → (p ∙ q ∙ r ∙ lUnit s k) ∙ t ∙ u) ⟩ - (p ∙ q ∙ r ∙ refl ∙ s) ∙ t ∙ u ∎ - - assocCancel1 : ∀ {ℓ} {A : Type ℓ} {a b c : A} (p : a ≡ b) (q : b ≡ c) → - p ∙ q ≡ p ∙ refl ∙ refl ∙ q ∙ refl ∙ refl - assocCancel1 p q = p ∙ q ≡⟨ (λ k → p ∙ lUnit (lUnit (rUnit (rUnit q k) k) k) k) ⟩ - p ∙ refl ∙ refl ∙ (q ∙ refl) ∙ refl ≡⟨ (λ k → p ∙ refl ∙ refl ∙ assoc q refl refl (~ k)) ⟩ - p ∙ refl ∙ refl ∙ q ∙ refl ∙ refl ∎ - - helper : ∀ {ℓ ℓ'} {X : Type ℓ} {x y : X} {A B : x ≡ y → Type ℓ'} (p : x ≡ y) → - cong A (rUnit (transport (λ i → x ≡ y) (p ∙ (λ _ → y)))) ∙ - cong A (λ i → (transportRefl (p ∙ (λ _ → y)) i) ∙ refl) ∙ - refl ∙ - cong A (λ i → rUnit p (~ i) ∙ refl) - ≡ - cong A (transportRefl (p ∙ (λ _ → y))) - helper {ℓ' = ℓ'} {x = x} {y = y} {A = A} {B = B} p = - J (λ y p → (A B : x ≡ y → Type ℓ') → - cong A (rUnit (transport (λ i → x ≡ y) (p ∙ (λ _ → y)))) ∙ - (cong A (λ i → (transportRefl (p ∙ (λ _ → y)) i) ∙ refl)) ∙ - refl ∙ - (cong A (λ i → rUnit p (~ i) ∙ refl)) - ≡ - cong A (transportRefl (p ∙ (λ _ → y)))) - (λ A B → ((λ k → (λ i → A (rUnit (transport (λ i₁ → x ≡ x) (refl ∙ (λ _ → x))) (i ∧ (~ k)))) ∙ - (λ i → A (rUnit (transportRefl (refl ∙ (λ _ → x)) i) (~ k))) ∙ - (λ i → A (rUnit (rUnit ((λ _ → x)) ((~ k) ∨ (~ i))) ((~ k) ∨ i))) ∙ - λ i → A ((rUnit (λ _ → x) ((~ i) ∧ (~ k))) ∙ refl)) ) ∙ - (λ k → lUnit ((λ i → A (transportRefl (refl ∙ (λ _ → x)) i)) ∙ - (rUnit (λ i → A (rUnit (rUnit ((λ _ → x)) (~ i)) i )) (~ k))) (~ k)) ∙ - (λ k → cong A (transportRefl ((λ _ → x) ∙ (λ _ → x))) ∙ - λ i → A (inv-rUnit x k i)) ∙ - sym (rUnit (cong A (transportRefl ((λ _ → x) ∙ (λ _ → x)))))) - p A B - - pairLemmaCancel : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) → - pairLemma (p ∙ sym (λ _ → y)) (λ _ → y) ∙ sym (assocJ p (λ _ → y) (λ _ → y)) ∙ (λ i → p ∙ lCancel (λ _ → y) i) ∙ sym (rUnit p) - ≡ - transportRefl (p ∙ sym (λ _ → y)) ∙ sym (rUnit p) - pairLemmaCancel {x = x} = - J (λ y p → pairLemma (p ∙ sym (λ _ → y)) (λ _ → y) ∙ sym (assocJ p (λ _ → y) (λ _ → y)) ∙ (λ i → p ∙ lCancel (λ _ → y) i) ∙ sym (rUnit p) - ≡ - transportRefl (p ∙ sym (λ _ → y)) ∙ sym (rUnit p)) - ((λ k → pairLemmaId (refl ∙ (λ i → x)) (λ _ → x) k ∙ - (λ i → assocJ refl (λ _ → x) (λ _ → x) (~ i)) ∙ - (λ i → refl ∙ lCancel (λ _ → x) i) ∙ (λ i → rUnit refl (~ i))) ∙ - (λ k → pairLemma*Refl (refl ∙ (λ i → x)) k ∙ - (λ i → assocJRefl k (~ i)) ∙ - (λ i → refl ∙ lCancel (λ _ → x) i) ∙ (λ i → rUnit refl (~ i))) ∙ - (λ k → (transportRefl (refl ∙ (λ _ → x)) ∙ rUnit (refl ∙ (λ _ → x))) ∙ - (symDistr (λ i → refl ∙ rCancel refl i) (rUnit (refl ∙ refl)) k ) ∙ - (λ i → refl ∙ lCancel (λ _ → x) i) ∙ (λ i → rUnit refl (~ i))) ∙ - littleCanceller2 (transportRefl (refl ∙ (λ _ → x))) - (rUnit (refl ∙ (λ _ → x))) - (λ i → refl ∙ rCancel refl (~ i)) - (λ i → rUnit refl (~ i))) - where - littleCanceller2 : ∀ {ℓ} {A : Type ℓ} {a b c d e : A} → (p : a ≡ b) (q : b ≡ c) (r : b ≡ d) (s : b ≡ e) → - (p ∙ q) ∙ (sym q ∙ r) ∙ sym r ∙ s ≡ p ∙ s - littleCanceller2 p q r s = (p ∙ q) ∙ ((sym q) ∙ r) ∙ sym r ∙ s ≡⟨ (λ k → (p ∙ q) ∙ assoc (sym q ∙ r) (sym r) s k) ⟩ - (p ∙ q) ∙ ((sym q ∙ r) ∙ sym r) ∙ s ≡⟨ (λ k → (p ∙ q) ∙ assoc (sym q) r (sym r) (~ k) ∙ s) ⟩ - (p ∙ q) ∙ (sym q ∙ r ∙ sym r) ∙ s ≡⟨ (λ k → (p ∙ q) ∙ (sym q ∙ rCancel r k) ∙ s) ⟩ - (p ∙ q) ∙ (sym q ∙ refl) ∙ s ≡⟨ (λ k → (p ∙ q) ∙ rUnit (sym q) (~ k) ∙ s) ⟩ - (p ∙ q) ∙ sym q ∙ s ≡⟨ assoc (p ∙ q) (sym q) s ⟩ - ((p ∙ q) ∙ sym q) ∙ s ≡⟨ (λ k → assoc p q (sym q) (~ k) ∙ s) ⟩ - (p ∙ q ∙ sym q) ∙ s ≡⟨ (λ k → (p ∙ rCancel q k) ∙ s) ⟩ - (p ∙ refl) ∙ s ≡⟨ (λ k → rUnit p (~ k) ∙ s) ⟩ - p ∙ s ∎ - -{- We will need RlFun (merid x1) to send ∣ x1 , (λ _ → σ a x1) ∣ to ∣ x1 , refl ∣ -} -funPart : (n : ℕ) (a x1 : A) (iscon : is- (ℕ→ℕ₋₂ n) -ConnectedType A) → RlFun a a n iscon (merid x1) ∣ x1 , (λ _ → σ a x1) ∣ ≡ ∣ x1 , refl ∣ -funPart n a x1 iscon = (λ k → sufMap n a a x1 iscon (merid x1) refl) ∙ - sufMapId2 n a x1 iscon ∙ - (λ k → ∣ x1 , cancellerLemma (sym (merid a)) (merid x1) k ∣) - where - cancellerLemma : ∀ {ℓ} {A : Type ℓ} {a b c : A} (r : b ≡ c) (p : a ≡ b) → canceller r p p refl ≡ refl - cancellerLemma {a = a} {b = b} {c = c} = J (λ c r → (p : a ≡ b) → canceller r p p refl ≡ refl) - (J (λ b p → canceller refl p p (λ _ → p ∙ refl) ≡ (λ _ → p)) - ((λ k → cancellerRefl k refl) ∙ - (λ k → rUnit refl ∙ (lUnit (sym (rUnit refl)) (~ k))) ∙ - rCancel (rUnit refl))) - - - - -codeTranspId6 : (n : ℕ) (a x1 : A) (iscon : is- (ℕ→ℕ₋₂ n) -ConnectedType A) → - transport (λ i → uncurry (CODE a n iscon) (pair⁼ (λ i₁ → merid a (~ i₁)) - (transpRCancel (λ i → north ≡ (merid a (~ i))) (merid x1 ∙ (sym (merid a)))) i)) - ∣ x1 , rUnit (merid x1) ∙ sym (cong (λ x → merid x1 ∙ x) (lCancel (merid a))) ∙ - assocJ (merid x1) (sym (merid a)) (merid a) ∙ - sym (pairLemma {a1 = north} (merid x1 ∙ (sym (merid a))) (merid a)) ∣ - ≡ ∣ x1 , refl ∣ -codeTranspId6 {ℓ}{A = A} n a x1 iscon = transportLemma {B = uncurry (CODE a n iscon)} - (sym (pair⁼ (λ i₁ → merid a (~ i₁)) - (transpRCancel (λ i → north ≡ (merid a (~ i))) (merid x1 ∙ (sym (merid a)))))) - (∣ x1 , refl ∣) - (∣ x1 , rUnit (merid x1) ∙ sym (cong (λ x → merid x1 ∙ x) (lCancel (merid a))) ∙ - assocJ (merid x1) (sym (merid a)) (merid a) ∙ - sym (pairLemma {a1 = north} (merid x1 ∙ (sym (merid a))) (merid a)) ∣) - ((λ k → transport (λ i → uncurry (CODE a n iscon) - (pair⁼Sym (λ i₁ → merid a (~ i₁)) - (transpRCancel (λ i₁ → north ≡ merid a (~ i₁)) - (merid x1 ∙ (λ i₁ → merid a (~ i₁)))) k i)) - ∣ x1 , (λ _ → σ a x1) ∣) - ∙ (λ k → transport (λ i → uncurry (CODE a n iscon) - (pair⁼ (merid a) (transportLemma {B = λ y → north ≡ y} - (sym (merid a)) - (transport (λ i₁ → north ≡ merid a i₁) - (merid x1 ∙ (sym (merid a)))) - (merid x1 ∙ (sym (merid a))) - ((transpRCancel (λ i₁ → north ≡ merid a (~ i₁)) - (merid x1 ∙ - (λ i₁ → merid a (~ i₁)))))) i) ) - ∣ x1 , (λ _ → σ a x1) ∣) ∙ - (λ k → transport (λ i → uncurry (CODE a n iscon) - (pair⁼ (merid a) (helper south (merid a) (merid x1) k) i) ) - ∣ x1 , (λ _ → σ a x1) ∣) ∙ - cong (λ x → x ∣ x1 , (λ _ → σ a x1) ∣) (Lemma8610Refl (CODE a n iscon) (merid a) (σ a x1)) ∙ - (λ k → transport ((λ i → ∥ fiber (λ y → σ a y) - (transpRCancel - (λ i₁ → north ≡ merid a (~ i₁)) (σ a x1) (~ i)) ∥ ℕ→ℕ₋₂ (n + n)) ∙ - funExt⁻ (toPathCancel {A = λ i → north ≡ merid a i → Type ℓ} - {x = λ p → (∥ fiber (λ y → σ a y) p ∥ (ℕ→ℕ₋₂ (n + n)))} - {y = λ q → ∥ fiber merid q ∥ (ℕ→ℕ₋₂ (n + n))} - (Lemma296-fun {X = Susp A} {A = λ y → north ≡ y} {B = λ _ → Type ℓ} - (merid a) (λ p → ∥ fiber (λ y → σ a y) p ∥ ℕ→ℕ₋₂ (n + n)) - (λ q → ∥ fiber merid q ∥ ℕ→ℕ₋₂ (n + n)) - (helperFun {X = Susp A} (merid a) - {A = λ q → ∥ fiber (λ y → σ a y) q ∥ ℕ→ℕ₋₂ (n + n)} - {B = λ q → ∥ fiber merid q ∥ ℕ→ℕ₋₂ (n + n)} - (λ q → ua' (RlFun a a n iscon q , RlFunEquiv a a n iscon q)))) k ) - (transport (λ i → north ≡ merid a i) (σ a x1))) - ∣ x1 , (λ _ → σ a x1) ∣) ∙ - (λ k → transport ((((λ i → ∥ fiber (λ y → σ a y) - (transpRCancel - (λ i₁ → north ≡ merid a (~ i₁)) (σ a x1) (~ i)) ∥ ℕ→ℕ₋₂ (n + n)))) ∙ - funExt⁻ (Lemma296-fun {X = Susp A} {A = λ y → north ≡ y} {B = λ _ → Type ℓ} - (merid a) (λ p → ∥ fiber (λ y → σ a y) p ∥ ℕ→ℕ₋₂ (n + n)) - (λ q → ∥ fiber merid q ∥ ℕ→ℕ₋₂ (n + n)) - (helperFun {X = Susp A} (merid a) - {A = λ q → ∥ fiber (λ y → σ a y) q ∥ ℕ→ℕ₋₂ (n + n)} - {B = λ q → ∥ fiber merid q ∥ ℕ→ℕ₋₂ (n + n)} - (λ q → ua' (RlFun a a n iscon q , RlFunEquiv a a n iscon q)))) - (transport (λ i → north ≡ merid a i) (σ a x1))) - ∣ x1 , ((λ _ → σ a x1)) ∣) ∙ - (λ k → transport (codeTranspId5 {A = λ p → ∥ fiber (λ y → σ a y) p ∥ ℕ→ℕ₋₂ (n + n)} - {B = λ q → ∥ fiber merid q ∥ ℕ→ℕ₋₂ (n + n)} - (merid a) (merid x1) - (λ q → ua' (RlFun a a n iscon q , RlFunEquiv a a n iscon q)) k) - ∣ x1 , (((λ _ → σ a x1))) ∣) ∙ - (λ k → transpFunct {B = λ x → x} - (ua' (RlFun a a n iscon (merid x1) , RlFunEquiv a a n iscon (merid x1))) - (cong (λ q → ∥ fiber merid q ∥ ℕ→ℕ₋₂ (n + n)) (sym (pairLemma (mx1 ∙ sym ma) ma ∙ - sym (assocJ mx1 (sym ma) ma) ∙ (λ i → mx1 ∙ lCancel ma i) ∙ sym (rUnit mx1)))) - ∣ x1 , (((λ _ → σ a x1))) ∣ (~ k)) ∙ - (λ k → transport (cong (λ q → ∥ fiber merid q ∥ ℕ→ℕ₋₂ (n + n)) (sym (pairLemma (mx1 ∙ sym ma) ma ∙ - sym (assocJ mx1 (sym ma) ma) ∙ (λ i → mx1 ∙ lCancel ma i) ∙ sym (rUnit mx1)))) - (ua'Cancel ((RlFun a a n iscon (merid x1) , RlFunEquiv a a n iscon (merid x1))) k - ∣ x1 , (((λ _ → σ a x1))) ∣ )) ∙ - (λ k → transport (cong (λ q → ∥ fiber merid q ∥ ℕ→ℕ₋₂ (n + n)) (sym (pairLemma (mx1 ∙ sym ma) ma ∙ - sym (assocJ mx1 (sym ma) ma) ∙ (λ i → mx1 ∙ lCancel ma i) ∙ sym (rUnit mx1)))) - (funPart n a x1 iscon k)) ∙ - finalStep) - - where - ma : _≡_ {A = Susp A} north south - ma = merid a - mx1 : _≡_ {A = Susp A} north south - mx1 = merid x1 - - guyid : (transport (λ i → north ≡ merid a i) (σ a x1)) ≡ merid x1 - guyid = pairLemma {a1 = north} (σ a x1) (merid a) ∙ sym (assoc (merid x1) (sym (merid a)) (merid a)) ∙ - (λ i → merid x1 ∙ (lCancel (merid a)) i) ∙ sym (rUnit (merid x1)) - - - - helper : (y : Susp A) (ma mx1 : north ≡ y) → - (transportLemma {B = λ y → north ≡ y} (sym ma) - (transport (λ i₁ → north ≡ ma i₁) (mx1 ∙ (sym ma))) (mx1 ∙ (sym ma)) - (transpRCancel (λ i₁ → north ≡ ma (~ i₁)) (mx1 ∙ (sym ma)))) - ≡ refl - helper y = J (λ y ma → (mx1 : north ≡ y) → (transportLemma {B = λ y → north ≡ y} (sym ma) - (transport (λ i₁ → north ≡ ma i₁) (mx1 ∙ (sym ma))) (mx1 ∙ (sym ma)) - (transpRCancel (λ i₁ → north ≡ ma (~ i₁)) (mx1 ∙ (sym ma)))) - ≡ refl) - λ p → (λ i → transportLemmaRefl {a = north} {B = λ y → _≡_ {A = Susp A} north y} - (λ _ → north) (λ _ → north) i - (transport (λ i₁ → _≡_ {A = Susp A} north north) (p ∙ (λ _ → north))) - (p ∙ (λ _ → north)) - (transpRCancel {A = _≡_ {A = Susp A} north north} - {B = _≡_ {A = Susp A} north north} - (λ i₁ → _≡_ {A = Susp A} north north) - (p ∙ (λ _ → north)))) ∙ - (λ k → transportRefl (p ∙ (λ _ → north)) ∙ - sym (transpRCancel {A = _≡_ {A = Susp A} north north} - {B = _≡_ {A = Susp A} north north} - (λ i₁ → _≡_ {A = Susp A} north north) - (p ∙ (λ _ → north))) ∙ - transportRefl (transport (λ i₁ → _≡_ {A = Susp A} north north) (p ∙ (λ _ → north)))) ∙ - (λ k → transportRefl (p ∙ (λ _ → north)) ∙ - sym (transpRCancelRefl {A = north ≡ north} (p ∙ (λ _ → north)) k) ∙ - transportRefl (transport (λ i₁ → _≡_ {A = Susp A} north north) (p ∙ (λ _ → north)))) ∙ - (λ k → transportRefl (p ∙ (λ _ → north)) ∙ - symDistr (transportRefl (transport refl (p ∙ (λ _ → north)))) (transportRefl (p ∙ (λ _ → north))) k ∙ - transportRefl (transport (λ i₁ → _≡_ {A = Susp A} north north) (p ∙ (λ _ → north)))) ∙ - (λ k → transportRefl (p ∙ (λ _ → north)) ∙ - assoc (sym (transportRefl (p ∙ (λ _ → north)))) - (sym (transportRefl (transport refl (p ∙ (λ _ → north))))) - (transportRefl (transport (λ i₁ → _≡_ {A = Susp A} north north) (p ∙ (λ _ → north)))) (~ k) ) ∙ - (λ k → transportRefl (p ∙ (λ _ → north)) ∙ - sym (transportRefl (p ∙ (λ _ → north))) ∙ - lCancel (transportRefl (transport (λ i₁ → _≡_ {A = Susp A} north north) (p ∙ (λ _ → north)))) k) ∙ - (λ k → transportRefl (p ∙ (λ _ → north)) ∙ (rUnit (sym (transportRefl (p ∙ (λ _ → north)))) (~ k))) ∙ - (rCancel (transportRefl (p ∙ (λ _ → north))) ) - - finalStep : transport (cong (λ q → ∥ fiber merid q ∥ ℕ→ℕ₋₂ (n + n)) (sym (pairLemma ((merid x1) ∙ sym (merid a)) (merid a) ∙ - sym (assocJ (merid x1) (sym (merid a)) (merid a)) ∙ - (λ i → (merid x1) ∙ lCancel (merid a) i) ∙ sym (rUnit (merid x1))))) - ∣ x1 , refl ∣ - ≡ ∣ x1 , rUnit (merid x1) ∙ sym (cong (λ x → merid x1 ∙ x) (lCancel (merid a))) ∙ - assocJ (merid x1) (sym (merid a)) (merid a) ∙ - sym (pairLemma {a1 = north} (merid x1 ∙ (sym (merid a))) (merid a)) ∣ - finalStep = (λ k → transport (cong (λ q → ∥ fiber merid q ∥ ℕ→ℕ₋₂ (n + n)) - (λ i → (sym (pairLemma ((merid x1) ∙ sym (merid a)) (merid a) ∙ - sym (assocJ (merid x1) (sym (merid a)) (merid a)) ∙ - (λ i → (merid x1) ∙ lCancel (merid a) i) ∙ - sym (rUnit (merid x1)))) (i ∨ k) )) - ∣ x1 , (λ i → (sym (pairLemma ((merid x1) ∙ sym (merid a)) (merid a) ∙ - sym (assocJ (merid x1) (sym (merid a)) (merid a)) ∙ - (λ i → (merid x1) ∙ lCancel (merid a) i) ∙ - sym (rUnit (merid x1)))) (i ∧ k)) ∣) ∙ - (transportRefl ∣ x1 , sym (pairLemma ((merid x1) ∙ sym (merid a)) (merid a) ∙ - sym (assocJ (merid x1) (sym (merid a)) (merid a)) ∙ - (λ i → (merid x1) ∙ lCancel (merid a) i) ∙ - sym (rUnit (merid x1))) ∣) ∙ - (λ k → ∣ x1 , symDestroyer (pairLemma ((merid x1) ∙ sym (merid a)) (merid a)) - (sym (assocJ (merid x1) (sym (merid a)) (merid a))) - (λ i → (merid x1) ∙ lCancel (merid a) i) - (sym (rUnit (merid x1))) k ∣) - where - symDestroyer : ∀ {ℓ} {A : Type ℓ} {a b c d e : A} (p : a ≡ b) (q : b ≡ c) (r : c ≡ d) (s : d ≡ e) → - sym (p ∙ q ∙ r ∙ s) ≡ sym s ∙ sym r ∙ sym q ∙ sym p - symDestroyer p q r s = sym (p ∙ q ∙ r ∙ s) ≡⟨ symDistr p (q ∙ r ∙ s) ⟩ - (sym (q ∙ r ∙ s)) ∙ sym p ≡⟨ (λ k → symDistr q (r ∙ s) k ∙ sym p) ⟩ - (sym (r ∙ s) ∙ sym q) ∙ sym p ≡⟨ (λ k → (symDistr r s k ∙ sym q) ∙ sym p) ⟩ - ((sym s ∙ sym r) ∙ sym q) ∙ sym p ≡⟨ sym (assoc (sym s ∙ sym r) (sym q) (sym p)) ⟩ - (sym s ∙ sym r) ∙ sym q ∙ sym p ≡⟨ sym (assoc (sym s) (sym r) (sym q ∙ sym p) ) ⟩ - sym s ∙ sym r ∙ sym q ∙ sym p ∎ - - -{- Finally, the we are done -} -reflCase : (n : ℕ) (a x1 : A) (iscon : is- (ℕ→ℕ₋₂ n) -ConnectedType A) → c n a iscon north (σ a x1) ≡ ∣ x1 , refl ∣ -reflCase {ℓ} {A} n a x1 iscon = transport (λ i → (uncurry (CODE a n iscon) (pair⁼ (σ a x1) (pairLemma2 (σ a x1)) i))) - ∣ a , (rCancel (merid a)) ∣ - ≡⟨ cId n a x1 iscon ⟩ - transport (λ i → uncurry (CODE a n iscon) (pair⁼ (λ i₁ → merid a (~ i₁)) - (pairLemma (merid x1) (λ i₁ → merid a (~ i₁))) i)) - (transport (λ i → uncurry (CODE a n iscon) (pair⁼ (merid x1) (pairLemma2 (merid x1)) i)) - ∣ a , rCancel (merid a) ∣) - ≡⟨ (λ k → transport (λ i → uncurry (CODE a n iscon) (pair⁼ (λ i₁ → merid a (~ i₁)) - (pairLemma (merid x1) (λ i₁ → merid a (~ i₁))) i)) - (codeTranspId1 n a x1 iscon (~ k))) ⟩ - transport (λ i → uncurry (CODE a n iscon) (pair⁼ (λ i₁ → merid a (~ i₁)) - (pairLemma (merid x1) (λ i₁ → merid a (~ i₁))) i)) - ((transport (λ i → uncurry (CODE a n iscon) - (pair⁼ (merid x1) (transpRCancel (λ i₁ → north ≡ merid x1 i₁) (merid x1)) i))) - ∣ a , rCancel (merid a) ∙ sym (rCancel (merid x1)) ∙ sym (pairLemma (merid x1) (sym (merid x1))) ∣) - ≡⟨ (λ k → transport (λ i → uncurry (CODE a n iscon) (pair⁼ (λ i₁ → merid a (~ i₁)) - (pairLemma (merid x1) (λ i₁ → merid a (~ i₁))) i)) - ((Lemma8610 (CODE a n iscon) (merid x1) (merid x1) k) - ∣ a , rCancel (merid a) ∙ sym (rCancel (merid x1)) ∙ sym (pairLemma (merid x1) (sym (merid x1))) ∣)) ⟩ - transport (λ i → uncurry (CODE a n iscon) (pair⁼ (λ i₁ → merid a (~ i₁)) - (pairLemma (merid x1) (λ i₁ → merid a (~ i₁))) i)) - (transport ((sym (cong (λ x → x (merid x1)) (Lemma294 {A = λ x → north ≡ x} - {B = λ _ → Type ℓ} - (merid x1) - ((CODE a n iscon) north)))) ∙ - funExt⁻ (fromPathP (toPathP {A = λ i → north ≡ merid x1 i → Type ℓ} - {x = λ p → (∥ fiber (λ y → σ a y) p ∥ (ℕ→ℕ₋₂ (n + n)))} - {y = λ q → ∥ fiber merid q ∥ (ℕ→ℕ₋₂ (n + n))} - (Lemma296-fun {X = Susp A} {A = λ y → north ≡ y} {B = λ _ → Type ℓ} - (merid x1) - (λ p → (∥ fiber (λ y → σ a y) p ∥ (ℕ→ℕ₋₂ (n + n)))) - (λ q → ∥ fiber merid q ∥ (ℕ→ℕ₋₂ (n + n))) - (helperFun {X = Susp A} (merid x1) - {A = λ q → ∥ fiber (λ y → σ a y) q ∥ ℕ→ℕ₋₂ (n + n)} - {B = λ q → ∥ fiber merid q ∥ (ℕ→ℕ₋₂ (n + n))} - λ q → ua' (RlFun a x1 n iscon q , RlFunEquiv a x1 n iscon q))))) (merid x1)) - ∣ a , rCancel (merid a) ∙ sym (rCancel (merid x1)) ∙ sym (pairLemma (merid x1) (sym (merid x1))) ∣) - ≡⟨ (λ k → transport (λ i → uncurry (CODE a n iscon) (pair⁼ (λ i₁ → merid a (~ i₁)) - (pairLemma (merid x1) (λ i₁ → merid a (~ i₁))) i)) - (transport (codeTranspId2 (merid x1) (merid x1) - (λ q → ∥ fiber (λ y → σ a y) q ∥ ℕ→ℕ₋₂ (n + n)) - (λ q → ∥ fiber merid q ∥ (ℕ→ℕ₋₂ (n + n))) - (λ q → ua' (RlFun a x1 n iscon q , RlFunEquiv a x1 n iscon q)) k) - ∣ a , rCancel (merid a) ∙ sym (rCancel (merid x1)) ∙ sym (pairLemma (merid x1) (sym (merid x1))) ∣)) ⟩ - transport (λ i → uncurry (CODE a n iscon) (pair⁼ (λ i₁ → merid a (~ i₁)) - (pairLemma (merid x1) (λ i₁ → merid a (~ i₁))) i)) - (transport ((transportRefl (∥ fiber (λ y → σ a y) - (subst (_≡_ north) (sym (merid x1)) (merid x1)) ∥ ℕ→ℕ₋₂ (n + n)) ∙ - cong (λ q → ∥ fiber (λ y → σ a y) q ∥ ℕ→ℕ₋₂ (n + n)) - (pairLemma (merid x1) (sym (merid x1)))) ∙ - ua' (RlFun a x1 n iscon (merid x1) , RlFunEquiv a x1 n iscon (merid x1))) - ∣ a , rCancel (merid a) ∙ sym (rCancel (merid x1)) ∙ sym (pairLemma (merid x1) (sym (merid x1))) ∣) - ≡⟨ (λ k → transport (λ i → uncurry (CODE a n iscon) (pair⁼ (λ i₁ → merid a (~ i₁)) - (pairLemma (merid x1) (λ i₁ → merid a (~ i₁))) i)) - (transpFunct {B = λ x → x} - (transportRefl (∥ fiber (λ y → σ a y) - (subst (_≡_ north) (sym (merid x1)) (merid x1)) ∥ ℕ→ℕ₋₂ (n + n)) ∙ - cong (λ q → ∥ fiber (λ y → σ a y) q ∥ ℕ→ℕ₋₂ (n + n)) - (pairLemma (merid x1) (sym (merid x1)))) - (ua' (RlFun a x1 n iscon (merid x1) , RlFunEquiv a x1 n iscon (merid x1))) - ∣ a , rCancel (merid a) ∙ sym (rCancel (merid x1)) ∙ - sym (pairLemma (merid x1) (sym (merid x1))) ∣ (~ k))) ⟩ - transport (λ i → uncurry (CODE a n iscon) (pair⁼ (λ i₁ → merid a (~ i₁)) - (pairLemma (merid x1) (λ i₁ → merid a (~ i₁))) i)) - (transport (ua' (RlFun a x1 n iscon (merid x1) , RlFunEquiv a x1 n iscon (merid x1))) - (transport (λ i → (transportRefl - (∥ fiber (λ y → σ a y) - (subst (_≡_ north) (λ i₁ → merid x1 (~ i₁)) (merid x1)) ∥ ℕ→ℕ₋₂ (n + n)) ∙ - (λ i₁ → ∥ fiber (λ y → σ a y) - (pairLemma (merid x1) (λ i₂ → merid x1 (~ i₂)) i₁) ∥ ℕ→ℕ₋₂ (n + n))) i) - ∣ a , rCancel (merid a) ∙ sym (rCancel (merid x1)) ∙ sym (pairLemma (merid x1) (sym (merid x1))) ∣)) - ≡⟨ (λ k → transport (λ i → uncurry (CODE a n iscon) (pair⁼ (λ i₁ → merid a (~ i₁)) - (pairLemma (merid x1) (λ i₁ → merid a (~ i₁))) i)) - ((ua'Cancel (RlFun a x1 n iscon (merid x1) , RlFunEquiv a x1 n iscon (merid x1)) k) - (transport (λ i → (transportRefl - (∥ fiber (λ y → σ a y) - (subst (_≡_ north) (λ i₁ → merid x1 (~ i₁)) (merid x1)) ∥ ℕ→ℕ₋₂ (n + n)) ∙ - (λ i₁ → ∥ fiber (λ y → σ a y) - (pairLemma (merid x1) (λ i₂ → merid x1 (~ i₂)) i₁) ∥ ℕ→ℕ₋₂ (n + n))) i) - ∣ a , rCancel (merid a) ∙ sym (rCancel (merid x1)) ∙ sym (pairLemma (merid x1) (sym (merid x1))) ∣))) ⟩ - transport (λ i → uncurry (CODE a n iscon) (pair⁼ (λ i₁ → merid a (~ i₁)) - (pairLemma (merid x1) (λ i₁ → merid a (~ i₁))) i)) - (RlFun a x1 n iscon (merid x1) - (transport (λ i → (transportRefl - (∥ fiber (λ y → σ a y) - (subst (_≡_ north) (λ i₁ → merid x1 (~ i₁)) (merid x1)) ∥ ℕ→ℕ₋₂ (n + n)) ∙ - (λ i₁ → ∥ fiber (λ y → σ a y) - (pairLemma (merid x1) (λ i₂ → merid x1 (~ i₂)) i₁) ∥ ℕ→ℕ₋₂ (n + n))) i) - ∣ a , rCancel (merid a) ∙ sym (rCancel (merid x1)) ∙ sym (pairLemma (merid x1) (sym (merid x1))) ∣)) - ≡⟨ (λ k → transport (λ i → uncurry (CODE a n iscon) (pair⁼ (λ i₁ → merid a (~ i₁)) - (pairLemma (merid x1) (λ i₁ → merid a (~ i₁))) i)) - (RlFun a x1 n iscon (merid x1) - (transpFunct {B = λ x → x} - (transportRefl - (∥ fiber (λ y → σ a y) - (subst (_≡_ north) (λ i₁ → merid x1 (~ i₁)) (merid x1)) ∥ ℕ→ℕ₋₂ (n + n))) - (λ i₁ → ∥ fiber (λ y → σ a y) - (pairLemma (merid x1) (λ i₂ → merid x1 (~ i₂)) i₁) ∥ ℕ→ℕ₋₂ (n + n)) - ∣ a , rCancel (merid a) ∙ sym (rCancel (merid x1)) ∙ - sym (pairLemma (merid x1) (sym (merid x1))) ∣ (~ k)))) ⟩ - transport (λ i → uncurry (CODE a n iscon) (pair⁼ (λ i₁ → merid a (~ i₁)) - (pairLemma (merid x1) (λ i₁ → merid a (~ i₁))) i)) - (RlFun a x1 n iscon (merid x1) - (transport (λ i₁ → ∥ fiber (λ y → σ a y) - (pairLemma (merid x1) (λ i₂ → merid x1 (~ i₂)) i₁) ∥ ℕ→ℕ₋₂ (n + n)) - (transport (transportRefl - (∥ fiber (λ y → σ a y) - (subst (_≡_ north) (λ i₁ → merid x1 (~ i₁)) (merid x1)) ∥ ℕ→ℕ₋₂ (n + n))) - ∣ a , rCancel (merid a) ∙ sym (rCancel (merid x1)) ∙ - sym (pairLemma (merid x1) (sym (merid x1))) ∣))) - ≡⟨ (λ k → transport (λ i → uncurry (CODE a n iscon) (pair⁼ (λ i₁ → merid a (~ i₁)) - (pairLemma (merid x1) (λ i₁ → merid a (~ i₁))) i)) - (codeTranspId3 n a x1 iscon k)) ⟩ - transport (λ i → uncurry (CODE a n iscon) (pair⁼ (λ i₁ → merid a (~ i₁)) - (pairLemma (merid x1) (λ i₁ → merid a (~ i₁))) i)) - ∣ x1 , refl ∣ - ≡⟨ codeTranspId4 n a x1 iscon ⟩ - transport (λ i → uncurry (CODE a n iscon) - (pair⁼ (λ i₁ → merid a (~ i₁)) - (transpRCancel (λ i₁ → north ≡ merid a (~ i₁)) (merid x1 ∙ sym (merid a))) i)) - ∣ x1 , rUnit (merid x1) ∙ sym (cong (_∙_ (merid x1)) (lCancel (merid a))) ∙ - assocJ (merid x1) (sym (merid a)) (merid a) ∙ sym (pairLemma (merid x1 ∙ sym (merid a)) (merid a)) ∣ - ≡⟨ codeTranspId6 n a x1 iscon ⟩ - ∣ x1 , refl ∣ ∎ - - -Thm8-6-4 : (n : ℕ) (a : A) - (iscon : is- (ℕ→ℕ₋₂ n) -ConnectedType A) {y : Susp A} → - (p : north ≡ y) → isContr (CODE a n iscon y p) -Thm8-6-4 {ℓ} {A} n a iscon {y = y} = J {ℓ} {Susp A} {north} (λ y p → (isContr (CODE a n iscon y p))) (c n a iscon north refl , (λ y → sym (mainId refl y))) - where - mainId : (p : north ≡ north) (d : CODE a n iscon north p) → - d ≡ c n a iscon north p - mainId p = elim (λ x → isOfHLevelSuc (suc (n + n)) (isOfHLevelTrunc {A = fiber (λ y → σ a y) p} (2 + (n + n)) x (c n a iscon north p))) - base - where - base : (x : fiber (λ y → σ a y) p) → ∣ x ∣ ≡ c n a iscon north p - base (x1 , r) = J (λ p r → ∣ x1 , r ∣ ≡ c n a iscon north p) (sym (reflCase n a x1 iscon)) r diff --git a/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4Code.agda b/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4Code.agda deleted file mode 100644 index 1afa8f10d..000000000 --- a/Cubical/HITs/Truncation/Connected/FreudenthalProof/8-6-4Code.agda +++ /dev/null @@ -1,235 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Truncation.Connected.FreudenthalProof.8-6-4Code where - -open import Cubical.HITs.Truncation.Connected.Base -open import Cubical.HITs.Truncation.Connected.FreudenthalProof.Prelim -open import Cubical.HITs.Truncation.Connected.FreudenthalProof.7-5-7 -open import Cubical.HITs.Truncation.Connected.FreudenthalProof.7-5-11 -open import Cubical.HITs.Truncation.Connected.FreudenthalProof.8-6-2 -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Univalence -open import Cubical.Data.NatMinusTwo.Base -open import Cubical.Data.Prod -open import Cubical.HITs.Susp -open import Cubical.HITs.Nullification hiding (elim) -open import Cubical.Data.Nat hiding (elim) -open import Cubical.HITs.Truncation.Base -open import Cubical.HITs.Truncation.Properties - -open import Cubical.Data.HomotopyGroup -private - variable - ℓ ℓ' : Level - A : Type ℓ - B : Type ℓ' - -σ : A → A → typ (Ω ((Susp A) , north)) -σ a x = merid x ∙ sym (merid a) - -{- To show that σ is 2n-connected we define CODE (see proof in the HoTT book). -} - - -{- We first need to show that (a variant of) the canceller function defined in FreudenthalProof.Prelim is an equivalence -} - -isEquivCancel : ∀ {ℓ} {A : Type ℓ}{a : A} (n : ℕ) (q : north ≡ south) → - isEquiv {A = ∥ fiber (λ y → σ a y) - (q ∙ sym (merid a)) ∥ ℕ→ℕ₋₂ (n + n)} - {B = ∥ fiber merid q ∥ ℕ→ℕ₋₂ (n + n)} - (elim (λ _ → isOfHLevelTrunc (2 + (n + n) )) - λ b → ∣ (fst b) , canceller (sym (merid a)) (merid (fst b)) q (snd b) ∣) -isEquivCancel {a = a} n q = isoToEquiv (iso - ((elim (λ _ → isOfHLevelTrunc (2 + (n + n))) λ b → ∣ (fst b) , canceller (sym (merid a)) (merid (fst b)) q (snd b) ∣)) - (elim (λ _ → isOfHLevelTrunc (2 + (n + n))) λ s → ∣ (fst s) , cancellerInv (λ i → (merid a) (~ i)) (merid (fst s)) q (snd s) ∣) - (λ b → elim {B = λ b → ((elim (λ _ → isOfHLevelTrunc (2 + (n + n))) - λ b → ∣ (fst b) , canceller (sym (merid a)) (merid (fst b)) q (snd b) ∣)) - ((elim (λ _ → isOfHLevelTrunc (2 + (n + n))) - λ s → ∣ (fst s) , cancellerInv (sym (merid a)) (merid (fst s)) q (snd s) ∣) b) - ≡ b} - (λ x → isOfHLevelSuc _ (isOfHLevelTrunc (2 + (n + n)) _ x)) - (λ b i → ∣ fst b , cancellerSection (sym (merid a)) (merid (fst b)) q (snd b) i ∣) b) - (λ b → elim {B = λ b → ((elim (λ _ → isOfHLevelTrunc (2 + (n + n))) - λ b → ∣ (fst b) , cancellerInv (sym (merid a)) (merid (fst b)) q (snd b) ∣)) - ((elim (λ _ → isOfHLevelTrunc (2 + (n + n))) - λ s → ∣ (fst s) , canceller (sym (merid a)) (merid (fst s)) q (snd s) ∣) b) - ≡ b} - (λ x → isOfHLevelSuc (suc (n + n)) (isOfHLevelTrunc (2 + (n + n)) _ x)) - (λ b i → ∣ fst b , cancellerRetract (sym (merid a)) (merid (fst b)) q (snd b) i ∣) b)) .snd - -{- CODE will be defined by means of univalence applied to an equivalence -∥ fiber (λ y → σ a y) (q ∙ sym (merid c)) ∥ (ℕ→ℕ₋₂ (n + n)) ≃ ∥ fiber merid q ∥ (ℕ→ℕ₋₂ (n + n)). -To define this, it suffices to define the following map -} -sufMap : (n : ℕ) (c a x₂ : A) - (iscon : is- (ℕ→ℕ₋₂ n) -ConnectedType A) - (q : north ≡ south) → - (merid x₂ ∙ sym (merid a)) - ≡ (q ∙ sym (merid c)) → ∥ Σ A (λ x → merid x ≡ q) ∥ (ℕ→ℕ₋₂ (n + n)) -sufMap {A = A} n c a x₂ iscon q = Lemma8-6-2 (A , a) (A , a) n n iscon iscon - (λ x₂ c → ((merid x₂ ∙ sym (merid a)) ≡ (q ∙ sym (merid c)) → ∥ Σ A (λ x → merid x ≡ q) ∥ (ℕ→ℕ₋₂ (n + n))) , - isOfHLevelΠ (2+ ((ℕ→ℕ₋₂ (n + n)))) λ _ → isOfHLevelTrunc (2 + (n + n))) - (λ x r → ∣ x , canceller (sym (merid a)) (merid x) q r ∣) - (λ x r → ∣ x , switcher (merid a) q (merid x) r ∣) - (funExt (λ x i → ∣ (refl i) , ((switcherCancellerIdGeneral (merid a) q - (canceller (sym (merid a)) - (merid a) q x ) x i) ) ∣)) .fst x₂ c - -{- We define the function -} -RlFun : (a c : A) (n : ℕ) → (iscon : is- (ℕ→ℕ₋₂ n) -ConnectedType A) → (q : north ≡ south) → - (∥ fiber (λ y → σ a y) (q ∙ sym (merid c)) ∥ (ℕ→ℕ₋₂ (n + n))) → - ∥ fiber merid q ∥ (ℕ→ℕ₋₂ (n + n)) -RlFun a c n iscon q = elim (λ x → isOfHLevelTrunc ((2 + (n + n)))) - λ r → sufMap n c a (fst r) iscon q (snd r) - -------------- - -{- To show RlFun is an equivalence, we need the following identity -} -RlFunId : (a : A) (n : ℕ) - (iscon : is- (ℕ→ℕ₋₂ n) -ConnectedType A) - (q : north ≡ south) - (b : fiber (λ y → σ a y) (q ∙ sym (merid a))) → - (RlFun a a n iscon q ∣ b ∣) ≡ ∣ (fst b) , canceller (sym (merid a)) (merid (fst b)) q (snd b)∣ -RlFunId {A = A} a n iscon q b = cong (λ x → x (snd b)) - (proj₁ (((Lemma8-6-2 (A , a) (A , a) n n iscon iscon - (λ x₂ c → (((merid x₂ ∙ sym (merid a)) - ≡ (q ∙ sym (merid c)) → ∥ Σ A (λ x → merid x ≡ q) ∥ (ℕ→ℕ₋₂ (n + n)) ) , - isOfHLevelΠ (2+ ((ℕ→ℕ₋₂ (n + n)))) λ _ → isOfHLevelTrunc (2 + (n + n)))) - (λ x r → ∣ x , canceller (sym (merid a)) (merid x) q r ∣) - (λ b s → ∣ b , switcher (merid a) q (merid b) s ∣) - (funExt (λ x → λ j → ∣ (refl j) , (switcherCancellerIdGeneral (merid a) q - (canceller (sym (merid a)) - (merid a) q x ) x j) ∣))) - .snd) .fst) - (fst b)) - -{- We show that RlFun is an equivalence for q : north ≡ south -} -baseCase : (a : A) (n : ℕ) - (iscon : is- (ℕ→ℕ₋₂ n) -ConnectedType A) - (q : north ≡ south) → - isEquiv (RlFun a a n iscon q) -baseCase {A = A} a n iscon q = transport (λ i → isEquiv (helper (~ i))) - (isEquivCancel n q ) - where - helper : RlFun a a n iscon q - ≡ elim (λ _ → isOfHLevelTrunc (2 + (n + n))) - λ b → ∣ (fst b) , canceller (sym (merid a)) (merid (fst b)) q (snd b) ∣ - helper = funExt (elim {B = λ y → RlFun a a n iscon q y - ≡ (elim (λ _ → isOfHLevelTrunc (2 + (n + n))) - λ b → ∣ (fst b) , canceller (sym (merid a)) (merid (fst b)) q (snd b) ∣) y } - ((λ z y p → (isOfHLevelSuc (suc (n + n)) ((isOfHLevelTrunc {A = (fiber merid q)} (2 + (n + n))) - (RlFun a a n iscon q z) - ((elim (λ _ → isOfHLevelTrunc (2 + (n + n))) - λ b → ∣ (fst b) , canceller (sym (merid a)) (merid (fst b)) - q (snd b) ∣) z) )) y p)) - (RlFunId a n iscon q)) - -{- By induction for connected functions, the previous lemma is enough to show that RlFun is an equivalence -} -RlFunEquiv : (a c : A) (n : ℕ) - (iscon : is- (ℕ→ℕ₋₂ n) -ConnectedType A) - (q : north ≡ south) → - isEquiv (RlFun a c n iscon q) -RlFunEquiv {A = A} a c n iscon q = fst (conInd-i→iii {A = Unit} {B = A} (λ x → a) (-1+ n) - (Lemma7-5-11 _ iscon) - λ c → (isEquiv (RlFun a c n iscon q)) , - transport (λ i → isOfHLevel (natHelper n i) (isEquiv (RlFun a c n iscon q))) - (isOfHLevelPlus {A = isEquiv (RlFun a c n iscon q)} n - (isPropIsEquiv (RlFun a c n iscon q)))) - (λ _ → baseCase a n iscon q) - c - where - natHelper : (n : ℕ) → n + 1 ≡ (2+ (-1+ n)) - natHelper zero = refl - natHelper (suc n) = cong (λ x → suc x) (natHelper n) - -{- We will also need the following function to get things on the right form -} - -helperFun : ∀ {ℓ ℓ'} {X : Type ℓ} {x y : X} (p : x ≡ y) {A : x ≡ x → Type ℓ'} {B : x ≡ y → Type ℓ'} → - ((q : x ≡ y) → A (q ∙ (sym p)) ≡ B q) → - (q : x ≡ x) → - transport (λ i₁ → Type ℓ') (A q) - ≡ B (transport (λ i → x ≡ p i) q ) -helperFun {ℓ' = ℓ'} {x = x} = J (λ y p → {A : x ≡ x → Type ℓ'} {B : x ≡ y → Type ℓ'} → - ((q : x ≡ y) → A (q ∙ (sym p)) ≡ B q) → - (q : x ≡ x) → - transport (λ i₁ → Type ℓ') (A q) - ≡ B (transport (λ i → x ≡ p i) q ) ) - λ {A} {B} k q → transportRefl (A q) ∙ - cong (λ x → A x) (rUnit q) ∙ - k q ∙ - cong (λ x → B x) (sym (transportRefl q)) - -helperFunRefl : ∀ {ℓ ℓ'} {X : Type ℓ} {x : X} - {A : x ≡ x → Type ℓ'} - {B : x ≡ x → Type ℓ'} → - helperFun {X = X} (refl {x = x}) {A = A} {B = B} - ≡ λ k q → transportRefl (A q) ∙ - cong (λ x → A x) (rUnit q) ∙ - k q ∙ - cong (λ x → B x) (sym (transportRefl q)) -helperFunRefl {ℓ' = ℓ'} {x = x} {A = A} {B = B} = cong (λ x → x {A} {B}) - (JRefl (λ y p → {A : x ≡ x → Type ℓ'} {B : x ≡ y → Type ℓ'} → - ((q : x ≡ y) → A (q ∙ (sym p)) ≡ B q) → - (q : x ≡ x) → - transport (λ i₁ → Type ℓ') (A q) - ≡ B (transport (λ i → x ≡ p i) q ) ) - λ {A} {B} k q → transportRefl (A q) ∙ - cong (λ x → A x) (rUnit q) ∙ - k q ∙ - cong (λ x → B x) (sym (transportRefl q))) - -{- For now things do not compute properly, so we use an abstract version of univalence-} - -abstract - ua' : ∀ {A B : Type ℓ} → A ≃ B → A ≡ B - ua' = ua - - ua'Cancel : ∀ {A B : Type ℓ} → (X : A ≃ B) → transport (λ i → ua' X i ) ≡ (fst X) - ua'Cancel X = funExt λ x → uaβ X x - -{- We finally define CODE -} -CODE : ∀{ℓ} {A : Type ℓ} (a : A) (n : ℕ) → is- (ℕ→ℕ₋₂ n) -ConnectedType A → (y : Susp A) → (north ≡ y → Type ℓ) -CODE {A = A} a n iscon north p = (∥ fiber (λ y → σ a y) p ∥ (ℕ→ℕ₋₂ (n + n))) -CODE {ℓ} {A = A} a n iscon south q = ∥ fiber merid q ∥ (ℕ→ℕ₋₂ (n + n)) -CODE {ℓ} {A = A} a n iscon (merid c i) = toPathP {A = λ i → north ≡ merid c i → Type ℓ} {x = λ p → (∥ fiber (λ y → σ a y) p ∥ (ℕ→ℕ₋₂ (n + n)))} - {y = λ q → ∥ fiber merid q ∥ (ℕ→ℕ₋₂ (n + n))} - (Lemma296-fun {X = Susp A} {A = λ y → north ≡ y} {B = λ _ → Type ℓ} - (merid c) (λ p → ∥ fiber (λ y → σ a y) p ∥ ℕ→ℕ₋₂ (n + n)) - (λ q → ∥ fiber merid q ∥ ℕ→ℕ₋₂ (n + n)) - (helperFun {X = Susp A} (merid c) {A = λ q → ∥ fiber (λ y → σ a y) q ∥ ℕ→ℕ₋₂ (n + n)} - {B = λ q → ∥ fiber merid q ∥ ℕ→ℕ₋₂ (n + n)} - λ q → (ua' (RlFun a c n iscon q , RlFunEquiv a c n iscon q)))) i -------------------------- - -{- We will need the following identites to prove that σ is 2n-connected -} -sufMapId : (n : ℕ) (a x1 : A) - (iscon : is- (ℕ→ℕ₋₂ n) -ConnectedType A) → - sufMap n x1 a a iscon (merid x1) - ≡ λ r → ∣ x1 , switcher (merid a) (merid x1) (merid x1) r ∣ -sufMapId {A = A} n a x1 iscon = (proj₂ (Lemma8-6-2 (A , a) (A , a) n n iscon iscon - (λ x₂ c → ((merid x₂ ∙ sym (merid a)) - ≡ ((merid x1) ∙ sym (merid c)) → - ∥ Σ A (λ x → merid x ≡ (merid x1)) ∥ (ℕ→ℕ₋₂ (n + n))) , - isOfHLevelΠ (2+ ((ℕ→ℕ₋₂ (n + n)))) λ _ → isOfHLevelTrunc (2 + (n + n))) - (λ x r → ∣ x , canceller (sym (merid a)) (merid x) (merid x1) r ∣) - (λ x r → ∣ x , switcher (merid a) (merid x1) (merid x) r ∣) - (funExt (λ x i → ∣ (refl i) , ((switcherCancellerIdGeneral (merid a) (merid x1) - (canceller (sym (merid a)) (merid a) - (merid x1) x ) x i) ) ∣)) .snd .fst) x1) - -sufMapId2 : (n : ℕ) (a x1 : A) - (iscon : is- (ℕ→ℕ₋₂ n) -ConnectedType A) → - sufMap n a a x1 iscon (merid x1) refl ≡ ∣ x1 , canceller (sym (merid a)) (merid x1) (merid x1) refl ∣ -sufMapId2 {A = A} n a x1 iscon i = (proj₁ (Lemma8-6-2 (A , a) (A , a) n n iscon iscon - (λ x₂ c → ((merid x₂ ∙ sym (merid a)) - ≡ ((merid x1) ∙ sym (merid c)) → ∥ Σ A (λ x → merid x ≡ (merid x1)) ∥ (ℕ→ℕ₋₂ (n + n))) , - isOfHLevelΠ (2+ ((ℕ→ℕ₋₂ (n + n)))) λ _ → isOfHLevelTrunc (2 + (n + n))) - (λ x r → ∣ x , canceller (sym (merid a)) (merid x) (merid x1) r ∣) - (λ x r → ∣ x , switcher (merid a) (merid x1) (merid x) r ∣) - (funExt (λ x i → ∣ (refl i) , ((switcherCancellerIdGeneral (merid a) (merid x1) - (canceller (sym (merid a)) (merid a) (merid x1) x) - x i) ) ∣)) - .snd .fst) x1) i - refl diff --git a/Cubical/HITs/Truncation/Connected/FreudenthalProof/Prelim.agda b/Cubical/HITs/Truncation/Connected/FreudenthalProof/Prelim.agda deleted file mode 100644 index 1715b8669..000000000 --- a/Cubical/HITs/Truncation/Connected/FreudenthalProof/Prelim.agda +++ /dev/null @@ -1,345 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Truncation.Connected.FreudenthalProof.Prelim where - -open import Cubical.HITs.Truncation.Connected.Base -open import Cubical.Foundations.Path -open import Cubical.Foundations.Transport -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Data.NatMinusTwo.Base -open import Cubical.Data.Nat -open import Cubical.Data.Unit -open import Cubical.HITs.Truncation.Base -open import Cubical.HITs.Truncation.Properties - -open import Cubical.Data.Sum.Base -open import Cubical.Data.HomotopyGroup -open import Cubical.Data.Bool -private - variable - ℓ ℓ' : Level - A : Type ℓ - B : Type ℓ' - --- some cong stuff -- -congComp2 : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {a b c : A} (f : A → B) (p : a ≡ b) (q : b ≡ c) → - (cong f p ∙ cong f q) ≡ cong f (p ∙ q) -congComp2 {A = A}{a = a} {b = b} {c = c} f p q = J (λ b p → (c : A) (q : b ≡ c) → (cong f p ∙ cong f q) ≡ cong f (p ∙ q)) - (λ c → J (λ c q → (λ i → f a) ∙ (λ i → f (q i)) ≡ (λ i → f ((refl ∙ q) i))) - ((λ j → rUnit (refl {x = f a}) (~ j)) ∙ - λ j i → f ((lUnit (refl {x = a}) j) i))) - p c q - -congComp3 : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {a b c d : A} (f : A → B) (p : a ≡ b) (q : b ≡ c) (r : c ≡ d) → - (cong f p ∙ cong f q) ∙ cong f r ≡ cong f (p ∙ q ∙ r) -congComp3 {A = A}{a = a} {b = b} {c = c} {d = d} f p q r = - J (λ b p → (c d : A) (q : b ≡ c) (r : c ≡ d) → (cong f p ∙ cong f q) ∙ cong f r ≡ cong f (p ∙ q ∙ r)) - (λ c d → J (λ c q → (r : c ≡ d) → ((λ i → f (refl i)) ∙ (λ i → f (q i))) ∙ (λ i → f (r i)) ≡ (λ i → f ((refl ∙ q ∙ r) i))) - (J (λ d r → ((λ i → f a) ∙ (λ i → f a)) ∙ (λ i → f (r i)) ≡ (λ i → f (((λ _ → a) ∙ refl ∙ r) i))) - ((λ j → rUnit (rUnit (λ i → f a) (~ j) ) (~ j) ) ∙ - λ j i → f ( (lUnit (lUnit (λ _ → a) j) j) i) ))) - p c d q r - ---- ℕ₋₂ stuff --- - -{- -_-_ : ℕ → ℕ → ℕ -zero - m = zero -suc n - zero = suc n -suc n - suc m = n - m --} -pmId : (n : ℕ) → suc₋₂ (pred₋₂ (ℕ→ℕ₋₂ n)) ≡ (ℕ→ℕ₋₂ n) -pmId zero = refl -pmId (suc n) = refl - -{- -minusId : (n : ℕ) → n - n ≡ 0 -minusId zero = refl -minusId (suc n) = minusId n - --assocHelp : {a b : ℕ} → ((a + b) - b) ≡ a --assocHelp {zero} {zero} = refl --assocHelp {zero} {suc b} = -assocHelp {zero} {b} --assocHelp {suc a} {zero} i = suc (+-zero a i) --assocHelp {suc a} {suc b} = (λ i → ((+-suc a b i) - b)) ∙ (λ i → ((suc (+-comm a b i)) - b)) ∙ - (λ i → (+-suc b a (~ i)) - b) ∙ ((λ i → ((+-comm b (suc a) i) - b))) ∙ - -assocHelp {suc a} {b} --} -------------------- -{- Functions used in definitions of Code -} -canceller : ∀ {ℓ} {A : Type ℓ} {a b c : A} → (r : b ≡ c) (p q : a ≡ b) → p ∙ r ≡ q ∙ r → p ≡ q -canceller {ℓ} {A} {a} {b} {c} = J {ℓ} {A} {b} (λ c r → ((p q : a ≡ b) → p ∙ r ≡ q ∙ r → p ≡ q)) λ p q id → (rUnit p) ∙ id ∙ sym (rUnit q) - -cancellerReflCase : ∀ {ℓ} {A : Type ℓ} {a b : A} → (p q : a ≡ b) → canceller refl p q ≡ λ id → (rUnit p) ∙ id ∙ sym (rUnit q) -cancellerReflCase {a = a} {b = b} p q = cong (λ x → x p q) (JRefl (λ c r → ((p q : a ≡ b) → p ∙ r ≡ q ∙ r → p ≡ q)) λ p q id → (rUnit p) ∙ id ∙ sym (rUnit q)) - -cancellerInv : ∀ {ℓ} {A : Type ℓ} {a b c : A} → (r : b ≡ c) (p q : a ≡ b) → (p ≡ q) → p ∙ r ≡ q ∙ r -cancellerInv {a = a} {b = b} = J (λ c r → (p q : a ≡ b) → (p ≡ q) → p ∙ r ≡ q ∙ r) λ p q id → sym (rUnit p) ∙ id ∙ rUnit q - -cancellerInvReflCase : ∀ {ℓ} {A : Type ℓ} {a b : A} → (p q : a ≡ b) → cancellerInv refl p q ≡ λ id → sym (rUnit p) ∙ id ∙ rUnit q -cancellerInvReflCase {a = a} {b = b} p q = cong (λ x → x p q) (JRefl (λ c r → (p q : a ≡ b) → (p ≡ q) → p ∙ r ≡ q ∙ r) λ p q id → sym (rUnit p) ∙ id ∙ rUnit q) - -cancellerSection : ∀ {ℓ} {A : Type ℓ} {a b c : A} → (r : b ≡ c) (p q : a ≡ b) → section (canceller r p q) (cancellerInv r p q) -cancellerSection {a = a} {b = b} {c = c} = J (λ c r → (p q : a ≡ b) → section (canceller r p q) (cancellerInv r p q) ) λ p q → transport (λ i → section (cancellerReflCase p q (~ i)) (cancellerInvReflCase p q (~ i))) (λ b → assoc (rUnit p) ((λ i → rUnit p (~ i)) ∙ b ∙ rUnit q) (λ i → rUnit q (~ i)) ∙ - (λ i → ((assoc (rUnit p) (sym (rUnit p)) (b ∙ rUnit q)) i) ∙ (λ i → rUnit q (~ i))) ∙ - (λ i → ((rCancel (rUnit p) i) ∙ (b ∙ rUnit q)) ∙ (sym (rUnit q))) ∙ - (λ i → lUnit (b ∙ rUnit q) (~ i) ∙ (sym (rUnit q))) ∙ - (sym (assoc b (rUnit q) (sym (rUnit q)))) ∙ - (λ i → b ∙ (rCancel (rUnit q) i)) ∙ - sym (rUnit b)) - -cancellerRetract : ∀ {ℓ} {A : Type ℓ} {a b c : A} → (r : b ≡ c) (p q : a ≡ b) → retract (canceller r p q) (cancellerInv r p q) -cancellerRetract {a = a} {b = b} {c = c} = J (λ c r → (p q : a ≡ b) → retract (canceller r p q) (cancellerInv r p q) ) λ p q → transport (λ i → retract (cancellerReflCase p q (~ i)) (cancellerInvReflCase p q (~ i))) λ b → assoc (sym (rUnit p)) ((λ i → rUnit p (i)) ∙ b ∙ (sym (rUnit q))) (rUnit q) ∙ - (λ i → ((assoc (sym (rUnit p)) (rUnit p) (b ∙ sym (rUnit q))) i) ∙ (rUnit q)) ∙ - (λ i → ((lCancel (rUnit p) i) ∙ (b ∙ sym (rUnit q))) ∙ ((rUnit q))) ∙ - (λ i → (lUnit (b ∙ sym (rUnit q)) (~ i)) ∙ (rUnit q)) ∙ - (sym (assoc b (sym (rUnit q)) (rUnit q))) ∙ - (λ i → b ∙ (rCancel (sym (rUnit q)) i)) ∙ - sym (rUnit b) - - -cancellerIsEquiv : ∀ {ℓ} {A : Type ℓ} {a b c : A} → (r : b ≡ c) (p q : a ≡ b) → isEquiv (canceller r p q) -cancellerIsEquiv {ℓ} {A} {a} {b} {c} = J {ℓ} {A} {b} (λ c r → ((p q : a ≡ b) → isEquiv (canceller r p q))) λ p q → transport (λ i → isEquiv (cancellerHelp p q (~ i))) (helper p q) - where - helper : (p q : a ≡ b) → isEquiv (λ id → ((rUnit p) ∙ id ∙ sym (rUnit q))) - helper p q = isoToEquiv (iso ((λ id → ((rUnit p) ∙ id ∙ sym (rUnit q)))) (λ id → sym (rUnit p) ∙ id ∙ rUnit q) - ((λ b → assoc (rUnit p) ((λ i → rUnit p (~ i)) ∙ b ∙ rUnit q) (λ i → rUnit q (~ i)) ∙ - (λ i → ((assoc (rUnit p) (sym (rUnit p)) (b ∙ rUnit q)) i) ∙ (λ i → rUnit q (~ i))) ∙ - (λ i → ((rCancel (rUnit p) i) ∙ (b ∙ rUnit q)) ∙ (sym (rUnit q))) ∙ - (λ i → lUnit (b ∙ rUnit q) (~ i) ∙ (sym (rUnit q))) ∙ - (sym (assoc b (rUnit q) (sym (rUnit q)))) ∙ - (λ i → b ∙ (rCancel (rUnit q) i)) ∙ - sym (rUnit b))) - λ b → assoc (sym (rUnit p)) ((λ i → rUnit p (i)) ∙ b ∙ (sym (rUnit q))) (rUnit q) ∙ - (λ i → ((assoc (sym (rUnit p)) (rUnit p) (b ∙ sym (rUnit q))) i) ∙ (rUnit q)) ∙ - (λ i → ((lCancel (rUnit p) i) ∙ (b ∙ sym (rUnit q))) ∙ ((rUnit q))) ∙ - (λ i → (lUnit (b ∙ sym (rUnit q)) (~ i)) ∙ (rUnit q)) ∙ - (sym (assoc b (sym (rUnit q)) (rUnit q))) ∙ - (λ i → b ∙ (rCancel (sym (rUnit q)) i)) ∙ - sym (rUnit b)) .snd - - cancellerHelp : (p q : a ≡ b) → canceller refl p q ≡ λ id → (rUnit p) ∙ id ∙ sym (rUnit q) - cancellerHelp p q = cong (λ x → x p q) (JRefl (λ c r → ((p q : a ≡ b) → p ∙ r ≡ q ∙ r → p ≡ q)) λ p q id → (rUnit p) ∙ id ∙ sym (rUnit q)) - -switcher : ∀ {ℓ} {A : Type ℓ} {a b c : A} → (p : a ≡ b) (q r : a ≡ c) → p ∙ (sym p) ≡ q ∙ (sym r) → r ≡ q -switcher {ℓ} {A} {a} {b} {c} = J {ℓ} {A} {a} (λ b p → ((q r : a ≡ c) → p ∙ (sym p) ≡ q ∙ (sym r) → r ≡ q)) - (J {ℓ} {A} {a} (λ c q → (r : a ≡ c) → refl ∙ (sym refl) ≡ q ∙ (sym r) → r ≡ q) λ r id → (lUnit r) ∙ cong (λ x → x ∙ r) ((rUnit (refl) ∙ id ∙ (sym (lUnit (sym r))))) ∙ (lCancel r)) - -- ((λ j → (λ i → (lUnit (sym r) j) (~ i) ))) ∙ cong (λ x → sym x) (sym id) ∙ ((λ j → (λ i → (sym (rUnit (refl {x = a})) j) (~ i) )))) -cancellerRefl : ∀ {ℓ} {A : Type ℓ} {a : A} → canceller {a = a} {b = a} {c = a} refl refl refl ≡ λ id → rUnit refl ∙ id ∙ sym (rUnit refl) -cancellerRefl {ℓ} {A} {a} = cong (λ x → x refl refl) (JRefl (λ c r → (p q : a ≡ a) → p ∙ r ≡ q ∙ r → p ≡ q) λ p q id → rUnit p ∙ id ∙ sym (rUnit q)) - -switcherRefl : ∀ {ℓ} {A : Type ℓ} {a : A} → switcher (refl {x = a}) refl refl ≡ λ id → ((lUnit (refl {x = a})) ∙ cong (λ x → x ∙ refl) ((rUnit (refl {x = a}) ∙ id ∙ (sym (lUnit (sym refl))))) ∙ (lCancel refl)) -switcherRefl {ℓ} {A} {a} = cong (λ x → x refl refl) (JRefl {ℓ} {A} {a} ((λ b p → ((q r : a ≡ a) → p ∙ (sym p) ≡ q ∙ (sym r) → r ≡ q))) (J {ℓ} {A} {a} (λ c q → (r : a ≡ c) → refl ∙ (sym refl) ≡ q ∙ (sym r) → r ≡ q) λ r id → (lUnit r) ∙ cong (λ x → x ∙ r) ((rUnit (refl) ∙ id ∙ (sym (lUnit (sym r))))) ∙ (lCancel r))) ∙ - cong (λ x → x refl) (JRefl {ℓ} {A} {a} (λ c q → (r : a ≡ c) → (λ _ → a) ∙ (λ i → a) ≡ q ∙ (λ i → r (~ i)) → r ≡ q) (λ r id → lUnit r ∙ (λ i → (rUnit (λ _ → a) ∙ id ∙ (λ i₁ → lUnit (λ i₂ → r (~ i₂)) (~ i₁))) i ∙ r) ∙ lCancel r)) - - -switcherCancellerId : ∀ {ℓ} {A : Type ℓ} {a : A} (id : (refl {x = a}) ∙ refl ≡ refl ∙ refl) → canceller (refl {x = a}) refl refl id ≡ switcher refl refl refl id -switcherCancellerId id = (λ i → (cancellerRefl i) id) ∙ sym (finalReflPath id) ∙ λ i → (switcherRefl (~ i)) id - where - finalReflPath : {a : A} → (id : refl ∙ refl ≡ refl ∙ refl ) → ((lUnit (refl {x = a})) ∙ cong (λ x → x ∙ refl) ((rUnit (refl {x = a}) ∙ id ∙ (sym (lUnit (sym refl))))) ∙ (lCancel refl)) ≡ (rUnit (refl {x = a}) ∙ id ∙ (sym (lUnit (sym refl)))) - finalReflPath {a = a} id j = hcomp (λ k → λ{(j = i0) → refl {x = (lUnit (refl {x = a})) - ∙ cong (λ x → x ∙ refl) ((rUnit (refl {x = a}) ∙ id - ∙ (sym (lUnit (sym refl))))) ∙ (lCancel refl)} k - ; (j = i1) → ((sym (lUnit ((cong (λ x → x) (rUnit (refl {x = a}) ∙ id ∙ (sym (lUnit (sym refl))))) ∙ refl))) - ∙ (sym (rUnit ((cong (λ x → x) (rUnit (refl {x = a}) ∙ id ∙ (sym (lUnit (sym refl))))))))) k}) - ((λ i → (lUnit (refl {x = a})) (i ∧ ~ j) ) ∙ (cong (λ x → (sym (rUnit x)) j) (rUnit (refl {x = a}) - ∙ id ∙ (sym (lUnit (sym refl))))) ∙ (λ i → (lCancel refl) (i ∨ j) )) - - -switcherCancellerIdGeneral : ∀ {ℓ} {A : Type ℓ} {a b : A} → (p q : a ≡ b) → (p ≡ q) → (P : p ∙ (sym p) ≡ q ∙ (sym p)) → canceller (sym p) p q P ≡ switcher p q p P -switcherCancellerIdGeneral {ℓ} {A} {a} {b} p q = J {ℓ} {a ≡ b} {p} (λ q _ → (P : p ∙ (sym p) ≡ q ∙ (sym p)) → canceller (sym p) p q P ≡ switcher p q p P) (J {ℓ} {A} {a} (λ b p → (P : p ∙ (sym p) ≡ p ∙ (sym p)) → canceller (sym p) p p P ≡ switcher p p p P) (λ P → switcherCancellerId P) p) - -symTypeId : {a b : A} → (a ≡ b) ≡ (b ≡ a) -symTypeId {A = A} {a = a} {b = b} = isoToPath (iso sym sym (λ x → refl) λ x → refl) - ------------------------ -{- Lemmas used in definitions of Code/Freudenthal proof -} -assocJ : ∀ {ℓ} {A : Type ℓ} {x y z w : A} → (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) → - p ∙ q ∙ r ≡ (p ∙ q) ∙ r -assocJ {x = x} {y = y} {z = z} {w = w} = J (λ y p → (q : y ≡ z) (r : z ≡ w) → p ∙ q ∙ r ≡ (p ∙ q) ∙ r) - (J (λ z q → (r : z ≡ w) → refl ∙ q ∙ r ≡ (refl ∙ q) ∙ r) - (J (λ w r → (λ _ → x) ∙ refl ∙ r ≡ ((λ _ → x) ∙ refl) ∙ r) - (((λ i → refl ∙ rCancel refl i)) ∙ rUnit (refl ∙ refl)))) - -assocJRefl : ∀ {ℓ} {A : Type ℓ} {x : A} → assocJ (λ _ → x) (λ _ → x) (λ _ → x) ≡ (λ i → refl ∙ rCancel refl i) ∙ (rUnit (refl ∙ refl)) -assocJRefl {x = x} = (cong (λ x → x refl refl) (JRefl (λ y p → (q : y ≡ x) (r : x ≡ x) → p ∙ q ∙ r ≡ (p ∙ q) ∙ r) - (J (λ z q → (r : z ≡ x) → refl ∙ q ∙ r ≡ (refl ∙ q) ∙ r) - (J (λ w r → (λ _ → x) ∙ refl ∙ r ≡ ((λ _ → x) ∙ refl) ∙ r) - ((λ i → refl ∙ rCancel refl i) ∙ rUnit (refl ∙ refl)))))) ∙ - cong (λ x → x refl) (JRefl (λ z q → (r : z ≡ x) → refl ∙ q ∙ r ≡ (refl ∙ q) ∙ r) - (J (λ w r → (λ _ → x) ∙ refl ∙ r ≡ ((λ _ → x) ∙ refl) ∙ r) - ((λ i → refl ∙ rCancel refl i) ∙ rUnit (refl ∙ refl)))) ∙ - JRefl (λ w r → (λ _ → x) ∙ refl ∙ r ≡ ((λ _ → x) ∙ refl) ∙ r) - ((λ i → refl ∙ rCancel refl i) ∙ rUnit (refl ∙ refl)) - - - -transpFunct : ∀ {ℓ ℓ'} {A : Type ℓ} {x y z : A} {B : A → Type ℓ'} (p : x ≡ y) (q : y ≡ z) (u : B x) → - transport (λ i → B (q i)) (transport (λ i → B (p i)) u) ≡ transport (λ i → B ((p ∙ q) i)) u -transpFunct {A = A} {x = x} {y = y} {z = z} {B = B} p = - J (λ y p → (z : A) (q : y ≡ z) (u : B x) → - transport (λ i → B (q i)) (transport (λ i → B (p i)) u) ≡ transport (λ i → B ((p ∙ q) i)) u) - (λ z → J (λ z q → (u : B x) → - transport (λ i → B (q i)) (transport (λ i → B (refl {x = x} i)) u) ≡ transport (λ i → B ((refl ∙ q) i)) u) - (λ u → transportRefl ((transport (λ i → B x) u)) ∙ - λ j → transport (λ i → B ((lUnit ((λ _ → x)) j) i)) u) {z}) - p z - -transpFunctRefl : ∀ {ℓ ℓ'} {A : Type ℓ} {x : A} {B : A → Type ℓ'} (u : B x) → - transpFunct {A = A} {B = B} (λ _ → x) (λ _ → x) u ≡ - (transportRefl ((transport (λ i → B x) u)) ∙ - λ j → transport (λ i → B ((lUnit ((λ _ → x)) j) i)) u) -transpFunctRefl {A = A} {x = x} {B = B} u = - (λ i → JRefl (λ y p → (z : A) (q : y ≡ z) (u : B x) → - transport (λ i → B (q i)) (transport (λ i → B (p i)) u) ≡ transport (λ i → B ((p ∙ q) i)) u) - (λ z p u → J (λ z q → (u : B x) → - transport (λ i → B (q i)) (transport (λ i → B (refl {x = x} i)) u) ≡ transport (λ i → B ((refl ∙ q) i)) u) - (λ u → transportRefl ((transport (λ i → B x) u)) ∙ - λ j → transport (λ i → B ((lUnit ((λ _ → x)) j) i)) u) p u ) i x refl u) - ∙ cong (λ x → x u) (JRefl (λ z q → (u : B x) → - transport (λ i → B (q i)) (transport (λ i → B (refl {x = x} i)) u) ≡ transport (λ i → B ((refl ∙ q) i)) u) - (λ u → transportRefl ((transport (λ i → B x) u)) ∙ - λ j → transport (λ i → B ((lUnit ((λ _ → x)) j) i)) u)) - -cancelReflMid : ∀ {ℓ} {A : Type ℓ} {a b : A} (p : a ≡ b) (q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q -cancelReflMid {ℓ = ℓ}{A = A} {a = a} {b = b} p q = J {ℓ} {A} {a} {ℓ} (λ b p → ((q : b ≡ a) → p ∙ refl ∙ q ≡ p ∙ q)) (λ r → sym (lUnit (refl ∙ r ))) {b} p q - -{- (2.9.4) in the HoTT Book -} -Lemma294 : ∀ {ℓ ℓ' ℓ''} {X : Type ℓ} {A : X → Type ℓ'} {B : X → Type ℓ''} {x1 x2 : X} (p : x1 ≡ x2) (f : A (x1) → B (x1)) → transport (λ i → A (p i) → B (p i)) f ≡ λ x → transport (λ i → B (p i)) (f (transport (λ i → A (p (~ i))) x)) -Lemma294 {A = A} {B = B} {x1 = x1} = J (λ x2 p → (f : A (x1) → B (x1)) → transport (λ i → A (p i) → B (p i)) f ≡ λ x → transport (λ i → B (p i)) (f (transport (λ i → A (p (~ i))) x)) ) λ f → refl - -Lemma294Refl : ∀ {ℓ ℓ' ℓ''} {X : Type ℓ} {A : X → Type ℓ'} {B : X → Type ℓ''} {x1 : X} (f : A (x1) → B (x1)) → Lemma294 {A = A} {B = B} (λ _ → x1) f ≡ refl -Lemma294Refl {A = A} {B = B} {x1 = x1} f = cong (λ g → g f) (JRefl (λ x2 p → (f : A (x1) → B (x1)) → transport (λ i → A (p i) → B (p i)) f ≡ λ x → transport (λ i → B (p i)) (f (transport (λ i → A (p (~ i))) x)) ) λ f → refl) - -{- The (inverse) function from Lemma 2.9.6 in the HoTT book -} -abstract - Lemma296-fun : ∀ {ℓ ℓ' ℓ''} {X : Type ℓ} {x y : X} {A : X → Type ℓ'} {B : X → Type ℓ''} - (p : x ≡ y) - (f : (A x) → (B x)) - (g : (A y) → (B y)) → - ((a : A x) → transport (λ i → B (p i)) (f a) ≡ g (transport (λ i → A (p i)) a)) → - (transport (λ i → (A (p i)) → (B (p i))) f ≡ g) - Lemma296-fun {x = x} {A = A} {B = B} p f g = - J (λ y p → (f : (A x) → (B x)) (g : (A y) → (B y)) → - ((a : A x) → transport (λ i → B (p i)) (f a) ≡ g (transport (λ i → A (p i)) a)) → - (transport (λ i → (A (p i)) → (B (p i))) f ≡ g)) - (λ f g h → transportRefl f ∙ funExt λ z → sym (transportRefl (f z)) ∙ - (h z) ∙ - λ i → (g (transportRefl z i))) - p f g - Lemma296-funRefl : ∀ {ℓ ℓ' ℓ''} {X : Type ℓ} {x : X} {A : X → Type ℓ'} {B : X → Type ℓ''} - (f : (A x) → (B x)) - (g : (A x) → (B x)) → - Lemma296-fun {A = A} {B = B} (refl {x = x}) f g - ≡ λ h → transportRefl f ∙ funExt λ z → sym (transportRefl (f z)) ∙ - (h z) ∙ - λ i → (g (transportRefl z i)) - Lemma296-funRefl {x = x} {A = A} {B = B} f g = - cong (λ h → h f g) - (JRefl (λ y p → (f : (A x) → (B x)) (g : (A y) → (B y)) → - ((a : A x) → transport (λ i → B (p i)) (f a) ≡ g (transport (λ i → A (p i)) a)) → - (transport (λ i → (A (p i)) → (B (p i))) f ≡ g)) - (λ f g h → transportRefl f ∙ funExt λ z → sym (transportRefl (f z)) ∙ - (h z) ∙ - λ i → (g (transportRefl z i)))) - - - -Lemma296 : ∀{ℓ ℓ' ℓ''} {X : Type ℓ} {x y : X} {A : X → Type ℓ'} {B : X → Type ℓ''} - (p : x ≡ y) - (f : (A x) → (B x)) - (g : (A y) → (B y)) → - (transport (λ i → (A (p i)) → (B (p i))) f ≡ g) ≡ ((a : A x) → transport (λ i → B (p i)) (f a) ≡ g (transport (λ i → A (p i)) a)) -Lemma296 {ℓ = ℓ} {X = X} {x = x} {y = y} {A = A} {B = B} = J {ℓ} {X} {x} - (λ y p → (f : (A x) → (B x)) (g : (A y) → (B y)) → - (transport (λ i → (A (p i)) → (B (p i))) f ≡ g) ≡ - ((a : A x) → transport (λ i → B (p i)) (f a) ≡ g (transport (λ i → A (p i)) a))) - (λ f g → transport (λ i → reflTransport f g (~ i)) (reflCase f g)) - where - reflTransport : (f g : A x → B x) → ((transport (λ i → A (refl {x = x} i) → B (refl {x = x} i)) f ≡ g) ≡ ((a : A x) → transport (λ i → B (refl {x = x} i)) (f a) ≡ g (transport (λ i → A (refl {x = x} i)) a))) ≡ ((f ≡ g) ≡ ((a : A x) → f a ≡ g a)) - reflTransport f g j = (transportRefl f j ≡ g) ≡ ((a : A x) → transportRefl (f a) j ≡ g (transportRefl a j)) - - reflCase : (f g : A x → B x) → ((f ≡ g) ≡ ((a : A x) → f a ≡ g a)) - reflCase f g = isoToPath (iso (λ p → funExt⁻ p) (λ p → funExt p) (λ x → refl) λ x → refl) - - - - --------------------- Some useful lemmas. Using J to prove them makes life a lot easier later ------------------------------ -symDistr : {x y z : A} (p : x ≡ y) (q : y ≡ z) → sym (p ∙ q) ≡ (sym q) ∙ (sym p) -symDistr {x = x} {z = z} = J (λ y p → (q : y ≡ z) → sym (p ∙ q) ≡ (sym q) ∙ (sym p)) - (J (λ z q → sym (refl ∙ q) ≡ (sym q) ∙ (sym refl) ) - refl) - -toPathCancel : ∀ {ℓ} {A : I → Type ℓ} {x : A i0} {y : A i1} → (z : transp A i0 x ≡ y) → fromPathP (toPathP {A = A} {x = x} {y = y} z) ≡ z -toPathCancel {A = A} {x = x} {y = y} z = cong (λ x → fst x) (equiv-proof (toPathP-isEquiv A {x = x} {y = y}) (toPathP z) .snd (z , refl)) - - -transportLemma : {a b : A} {B : (z : A) → Type ℓ'} (p : a ≡ b) (x : B a) (y : B b) → transport (λ i → B (p i)) x ≡ y → transport (λ i → B (p (~ i))) y ≡ x -transportLemma {a = a} {B = B} x y = J (λ b p → (x : B a) (y : B b) → transport (λ i → B (p i)) x ≡ y → transport (λ i → B (p (~ i))) y ≡ x) - (λ x y id → transportRefl y ∙ (sym id) ∙ transportRefl x) - x y -transportLemmaRefl : {a : A} {B : (z : A) → Type ℓ'} → (x y : B a) → transportLemma {B = B} (λ _ → a) ≡ λ x y id → transportRefl y ∙ (sym id) ∙ transportRefl x -transportLemmaRefl {ℓ} {A = A} {a = a} {B = B} x y = JRefl {ℓ} {A} {a} (λ b p → (x y : B a) → transport (λ i → B a) x ≡ y → transport (λ i → B a) y ≡ x) - (λ x y id → transportRefl y ∙ (sym id) ∙ transportRefl x) - -transpRCancel : (p : A ≡ B) (b : B) → transport p (transport⁻ p b) ≡ b -transpRCancel {A = A} = J (λ B p → (b : B) → transport p (transport⁻ p b) ≡ b) λ b → transportRefl (transport refl b) ∙ transportRefl b - -transpRCancelRefl : (a : A) → transpRCancel refl a ≡ transportRefl (transport refl a) ∙ transportRefl a -transpRCancelRefl {A = A} a = cong (λ x → x a) (JRefl (λ A p → (a : A) → transport p (transport⁻ p a) ≡ a) λ b → transportRefl (transport refl b) ∙ transportRefl b) - -pairLemma2 : ∀ {ℓ} {A : Type ℓ} {a b : A} (p : a ≡ b) → transport (λ i → a ≡ p i) refl ≡ p -pairLemma2 {ℓ} {A} {a} = J (λ b p → transport (λ i → a ≡ p i) refl ≡ p) (transportRefl refl) - -pairLemma2Refl : ∀ {ℓ} {A : Type ℓ} {a : A} → pairLemma2 (refl {x = a}) ≡ (transportRefl refl) -pairLemma2Refl {ℓ} {A} {a} = JRefl (λ b p → transport (λ i → a ≡ p i) refl ≡ p) (transportRefl refl) - -pairLemma : ∀ {ℓ} {A : Type ℓ} {a1 b c : A} (p : a1 ≡ b) (q : b ≡ c) → transport (λ i₁ → a1 ≡ q i₁) p ≡ p ∙ q -pairLemma {A = A} {a1 = a1} {b = b} {c = c} p q = J (λ b p → (c : A) → (q : b ≡ c) → transport (λ i₁ → a1 ≡ q i₁) p ≡ p ∙ q ) - (λ c q → pairLemma2 q ∙ lUnit q) - p c q - -pairLemmaRefl : ∀ {ℓ} {A : Type ℓ} {a1 : A} (q : a1 ≡ a1) → pairLemma (λ _ → a1) q ≡ pairLemma2 q ∙ lUnit q -pairLemmaRefl {A = A} {a1 = a1} q = cong (λ x → x a1 q) (JRefl (λ b p → (c : A) → (q : b ≡ c) → transport (λ i₁ → a1 ≡ q i₁) p ≡ p ∙ q ) (λ c q → pairLemma2 q ∙ lUnit q)) - - -pairLemma* : ∀ {ℓ} {A : Type ℓ} {a1 b c : A} (q : b ≡ c) (p : a1 ≡ b) → transport (λ i₁ → a1 ≡ q i₁) p ≡ p ∙ q -pairLemma* {a1 = a1} {b = b} = J (λ c q → (p : a1 ≡ b) → transport (λ i₁ → a1 ≡ q i₁) p ≡ p ∙ q) - λ p → transportRefl p ∙ rUnit p - -pairLemma*Refl : ∀ {ℓ} {A : Type ℓ} {a1 b : A} (p : a1 ≡ b) → pairLemma* refl p ≡ transportRefl p ∙ rUnit p -pairLemma*Refl {a1 = a1} {b = b} p = cong (λ x → x p) (JRefl (λ c q → (p : a1 ≡ b) → transport (λ i₁ → a1 ≡ q i₁) p ≡ p ∙ q) - λ p → transportRefl p ∙ rUnit p) - -pairLemmaId : ∀ {ℓ} {A : Type ℓ} {a1 b c : A} (p : a1 ≡ b) (q : b ≡ c) → pairLemma p q ≡ pairLemma* q p -pairLemmaId {ℓ} {A} {a1} {b} {c} = J (λ b p → (q : b ≡ c) → pairLemma p q ≡ pairLemma* q p) - (J (λ c q → pairLemma refl q ≡ pairLemma* q refl) - (pairLemmaRefl refl ∙ sym (pairLemma*Refl refl ∙ λ k → (pairLemma2Refl (~ k)) ∙ rUnit refl))) - -pairLemmaReflRefl : {a1 : A} → pairLemma (λ _ → a1) (λ _ → a1) ≡ (transportRefl refl) ∙ lUnit refl -pairLemmaReflRefl = pairLemmaRefl refl ∙ λ i → pairLemma2Refl i ∙ lUnit refl - -substComposite-∙ : ∀ {ℓ ℓ′} {A : Type ℓ} → (B : A → Type ℓ′) - → {x y z : A} (p : x ≡ y) (q : y ≡ z) (u : B x) - → subst B (p ∙ q) u ≡ subst B q (subst B p u) -substComposite-∙ = substComposite - -inv-rUnit : ∀ {ℓ} {A : Type ℓ} (x : A) → (λ i → rUnit (rUnit (λ _ → x) (~ i)) i ) ≡ refl -inv-rUnit x = transport (λ i → PathP (λ j → (lCancel (λ k → (λ _ → x) ∙ (λ _ → x) ≡ rUnit (λ _ → x) k) i) j) - (λ i → rUnit (rUnit (λ _ → x) (~ i)) i ) - refl) - (helper x) - where - helper : ∀ {ℓ} {A : Type ℓ} (x : A) → - PathP (λ i → ((λ k → (λ _ → x) ∙ (λ _ → x) ≡ rUnit (λ _ → x) (~ k)) ∙ - λ k → (λ _ → x) ∙ (λ _ → x) ≡ rUnit (λ _ → x) k) i) - (λ i → rUnit (rUnit (λ _ → x) (~ i)) i ) - refl - helper x = compPathP (λ k i → rUnit (rUnit (λ _ → x) (~ i)) (i ∧ (~ k))) - λ k i → rUnit (λ _ → x) (~ i ∨ k) diff --git a/Cubical/HITs/Truncation/Connected/Properties.agda b/Cubical/HITs/Truncation/Connected/Properties.agda deleted file mode 100644 index 8b7d16f09..000000000 --- a/Cubical/HITs/Truncation/Connected/Properties.agda +++ /dev/null @@ -1,182 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.HITs.Truncation.Connected.Properties where - -open import Cubical.HITs.Truncation.Connected.Base -open import Cubical.HITs.Truncation.Connected.FreudenthalProof.7-5-7 -open import Cubical.HITs.Truncation.Connected.FreudenthalProof.7-5-11 -open import Cubical.HITs.Truncation.Connected.FreudenthalProof.7-5-14 -open import Cubical.HITs.Truncation.Connected.FreudenthalProof.8-6-1 -open import Cubical.HITs.Truncation.Connected.FreudenthalProof.8-6-2 -open import Cubical.HITs.Truncation.Connected.FreudenthalProof.8-6-4 -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Equiv.Properties -open import Cubical.Foundations.Function -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Transport -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.Isomorphism -open import Cubical.Data.NatMinusTwo.Base -open import Cubical.Data.Prod.Base -open import Cubical.HITs.Susp -open import Cubical.Data.Nat -open import Cubical.HITs.Truncation.Base -open import Cubical.HITs.Truncation.Properties renaming (elim to trElim) -open import Cubical.HITs.Nullification -open import Cubical.HITs.Pushout.Base -open import Cubical.HITs.Sn.Base -open import Cubical.Data.Unit.Properties - -open import Cubical.Data.HomotopyGroup - -private - variable - ℓ ℓ' : Level - A : Type ℓ - B : Type ℓ' - -isConnectedSubtr : (n m : ℕ) (f : A → B) → is- (-2+ (n + m)) -Connected f → is- (-2+ n) -Connected f -isConnectedSubtr n m f iscon b = transport (λ i → isContr (ua (truncOfTruncEq {A = fiber f b} n m) (~ i) )) - (∣ iscon b .fst ∣ , - trElim (λ x → isOfHLevelPath n (isOfHLevelTrunc n) _ _) - λ a → cong ∣_∣ (iscon b .snd a)) --- - -{- -The "induction principle" for n-connected functions from the HoTT book (see Lemma 7.5.7): - -For n ≥ -2, f : A → B, P : B → Type, TFAE: -1. f in n-connected -2. For every B → n-Type, the map (indConFun f p) is an equivalence -3. For every B → n-Type, the map (indConFun f p) has a section --} - -module elim (f : A → B) (n : ℕ₋₂) where - 1→2 : ∀ {ℓ} → is- n -Connected f → - (P : B → HLevel ℓ (2+ n)) → - isEquiv (indConFun f P) - 1→2 = conInd-i→ii f n - - 2→3 : ∀ {ℓ} (P : B → HLevel ℓ (2+ n)) → - isEquiv (indConFun f P) → - hasSection (indConFun f P) - 2→3 = conInd-ii→iii f n - - 3→1 : (∀ {ℓ} (P : B → HLevel ℓ (2+ n)) → - hasSection (indConFun f P)) → - is- n -Connected f - 3→1 = conInd-iii→i f n - -{- a function from Unit to an (n + 1)-connected type is n-connected -} -trivFunCon : ∀{ℓ} {A : Type ℓ} {a : A} → (n : ℕ₋₂) → (is- (suc₋₂ n) -ConnectedType A) → is- n -Connected (λ (x : Unit) → a) -trivFunCon = Lemma7-5-11 - -trivFunCon← : ∀{ℓ} {A : Type ℓ} {a : A} → (n : ℕ₋₂) → is- n -Connected (λ (x : Unit) → a) → (is- (suc₋₂ n) -ConnectedType A) -trivFunCon← = Lemma7-5-11← - -{- n-connected functions induce equivalences between n-truncations -} -conEquiv : (n : ℕ₋₂) (f : A → B) → (is- n -Connected f) → ∥ A ∥ n ≃ ∥ B ∥ n -conEquiv = Lemma7-5-14 - -conEquiv2 : (n m : ℕ) (f : A → B) → (is- (-2+ (n + m)) -Connected f) → ∥ A ∥ (-2+ n) ≃ ∥ B ∥ (-2+ n) -conEquiv2 n m f iscon = conEquiv (-2+ n) f (isConnectedSubtr n m f iscon) - -conEquiv3 : (n m : ℕ) (f : A → B) → Σ[ x ∈ ℕ ] n + x ≡ m → (is- (-2+ m) -Connected f) → ∥ A ∥ (-2+ n) ≃ ∥ B ∥ (-2+ n) -conEquiv3 n m f (x , pf) iscon = conEquiv (-2+ n) f (isConnectedSubtr n x f (transport (λ i → is- (-2+ pf (~ i)) -Connected f) iscon)) -- - -{- Wedge connectivity lemma (Lemma 8.6.2 in the HoTT book) -} -WedgeConn : (A : Pointed ℓ) (B : Pointed ℓ') (n m : ℕ) → - (is- (ℕ→ℕ₋₂ n) -ConnectedType (typ A)) → - (is- (ℕ→ℕ₋₂ m) -ConnectedType (typ B)) → - (P : (typ A) → (typ B) → HLevel (ℓ-max ℓ ℓ') (2+ (ℕ→ℕ₋₂ (n + m)))) → - (f : ((a : (typ A)) → fst (P a (pt B)))) → - (g : ((b : (typ B)) → fst (P (pt A) b))) → - (p : f (pt A) ≡ g (pt B)) → - Σ ((a : typ A) → (b : typ B) → fst (P a b)) - λ h → Σ (((a : typ A) → h a (pt B) ≡ f a) × ((b : typ B) → h (pt A) b ≡ g b)) - λ pair → p ≡ sym ((proj₁ pair) (pt A) ) ∙ (proj₂ pair) (pt B) -WedgeConn = Lemma8-6-2 - ---------- Freudenthal ------- - -FthalFun : A → A → typ (Ω ((Susp A) , north)) -FthalFun a x = merid x ∙ sym (merid a) - -abstract - FthalFun-2nConnected : (n : ℕ) (a : A) - (iscon : is- (ℕ→ℕ₋₂ n) -ConnectedType A) → - is- ℕ→ℕ₋₂ (n + n) -Connected (FthalFun a) - FthalFun-2nConnected n a iscon = Thm8-6-4 n a iscon - - Freudenthal : (n : ℕ) (A : Pointed ℓ) → - is- (ℕ→ℕ₋₂ n) -ConnectedType (typ A) → - ∥ typ A ∥ (ℕ→ℕ₋₂ (n + n)) ≃ ∥ typ (Ω (Susp (typ A) , north)) ∥ ((ℕ→ℕ₋₂ (n + n))) - Freudenthal n A iscon = conEquiv _ (FthalFun (pt A)) (FthalFun-2nConnected n (pt A) iscon) - - ------------------------------- - - -inrConnected : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} (n : ℕ₋₂) - (f : C → A) (g : C → B) → - is- n -Connected f → - is-_-Connected {A = B} {B = Pushout f g} n inr -inrConnected {A = A} {B = B} {C = C} n f g iscon = elim.3→1 inr n λ P → (λ k → helpLemmas.k P k) , λ b → refl - where - module helpLemmas {ℓ : Level} (P : (Pushout f g) → HLevel ℓ (2+ n)) - (h : (b : B) → typ (P (inr b))) - where - Q : A → HLevel _ (2+ n) - Q a = (P (inl a)) - - fun : (c : C) → typ (Q (f c)) - fun c = transport (λ i → typ (P (push c (~ i)))) (h (g c)) - - k : (d : Pushout f g) → typ (P d) - k (inl x) = elim.2→3 f n Q (elim.1→2 f n iscon Q) .fst fun x - k (inr x) = h x - k (push a i) = hcomp (λ k → λ{(i = i0) → elim.2→3 f n Q - (elim.1→2 f n iscon Q) .fst - fun (f a) ; - (i = i1) → transportTransport⁻ (λ j → typ (P (push a j))) (h (g a)) k}) - (transp (λ j → typ (P (push a (i ∧ j)))) - (~ i) - (elim.2→3 f n Q - (elim.1→2 f n iscon Q) .snd fun i a)) - - -sphereConnected : (n : ℕ) → is- (-1+ n) -ConnectedType (S₊ n) -sphereConnected zero = ∣ north ∣ , (isOfHLevelTrunc 1 ∣ north ∣) -sphereConnected (suc n) = transport (λ i → is- ℕ→ℕ₋₂ n -ConnectedType (Susp≡Push {A = S₊ n} (~ i))) - (trivFunCon← {A = Pushout {A = S₊ n} (λ x → tt) λ x → tt} {a = inr tt } _ - (transport (λ i → is- (-1+ n) -Connected (mapsAgree (~ i))) - (inrConnected _ (λ x → tt) (λ x → tt) λ{tt → transport (λ i → isContr (∥ fibUnit (~ i) ∥ (-1+ n))) (sphereConnected n)}))) - where - mapsAgree : Path ((x : Unit) → Pushout {A = S₊ n} (λ x → tt) (λ x → tt)) (λ (x : Unit) → inr tt) inr - mapsAgree = funExt λ {tt → refl} - - fibUnit : fiber (λ (x : S₊ n) → tt) tt ≡ S₊ n - fibUnit = isoToPath (iso (λ b → fst b) (λ a → a , refl) (λ b → refl) λ b i → (fst b) , isProp→isSet isPropUnit tt tt refl (snd b) i) - - Susp≡Push : Susp A ≡ Pushout {A = A} (λ a → tt) λ a → tt - Susp≡Push {A = A} = isoToPath (iso fun inverse sect retr) - where - fun : Susp A → Pushout {A = A} (λ a → tt) (λ a → tt) - fun north = inl tt - fun south = inr tt - fun (merid a i) = push a i - inverse : Pushout {A = A} (λ a → tt) (λ a → tt) → Susp A - inverse (inl tt) = north - inverse (inr tt) = south - inverse (push a i) = merid a i - - sect : section fun inverse - sect (inl tt) = refl - sect (inr tt) = refl - sect (push a i) = refl - - retr : retract fun inverse - retr north = refl - retr south = refl - retr (merid a i) = refl From d4305ca8c4508960e96c93725a489c3d074bf50f Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 4 May 2020 01:02:35 +0200 Subject: [PATCH 26/89] minor fixes --- Cubical/Foundations/Equiv/Properties.agda | 1 + Cubical/Foundations/Isomorphism.agda | 3 +- Cubical/Foundations/Path.agda | 1 - Cubical/Foundations/Transport.agda | 2 - Cubical/HITs/Pushout/Base.agda | 2 - Cubical/HITs/SetTruncation/Properties.agda | 7 -- Cubical/HITs/Truncation/Properties.agda | 76 ++++++++-------------- Cubical/Homotopy/Connected.agda | 9 ++- Cubical/ZCohomology/cupProdPrelims.agda | 6 +- 9 files changed, 37 insertions(+), 70 deletions(-) diff --git a/Cubical/Foundations/Equiv/Properties.agda b/Cubical/Foundations/Equiv/Properties.agda index 3e61d76d2..5c2e609de 100644 --- a/Cubical/Foundations/Equiv/Properties.agda +++ b/Cubical/Foundations/Equiv/Properties.agda @@ -57,6 +57,7 @@ postCompEquiv : ∀ {ℓ ℓ′} {A B : Type ℓ} {C : Type ℓ′} (e : A ≃ B → (C → A) ≃ (C → B) postCompEquiv e = (λ φ x → fst e (φ x)) , isEquivPostComp e + hasSection : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → (A → B) → Type _ hasSection {A = A} {B} f = Σ[ g ∈ (B → A) ] section f g diff --git a/Cubical/Foundations/Isomorphism.agda b/Cubical/Foundations/Isomorphism.agda index 5b107079d..61d600f4b 100644 --- a/Cubical/Foundations/Isomorphism.agda +++ b/Cubical/Foundations/Isomorphism.agda @@ -111,7 +111,8 @@ compIso (iso fun inv rightInv leftInv) (iso fun₁ inv₁ rightInv₁ leftInv₁ compIso (iso fun inv rightInv leftInv) (iso fun₁ inv₁ rightInv₁ leftInv₁) .Iso.leftInv a = cong inv (leftInv₁ (fun a) ) ∙ leftInv a -composesToId→Iso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (G : Iso A B) → (g : B → A) → (Iso.fun G) ∘ g ≡ idfun B → Iso B A +composesToId→Iso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (G : Iso A B) (g : B → A) + → (Iso.fun G) ∘ g ≡ idfun B → Iso B A Iso.fun (composesToId→Iso (iso fun inv rightInv leftInv) g path) = g Iso.inv (composesToId→Iso (iso fun inv rightInv leftInv) g path) = fun Iso.rightInv (composesToId→Iso (iso fun inv rightInv leftInv) g path) b = diff --git a/Cubical/Foundations/Path.agda b/Cubical/Foundations/Path.agda index f9f530ab4..592c58492 100644 --- a/Cubical/Foundations/Path.agda +++ b/Cubical/Foundations/Path.agda @@ -6,7 +6,6 @@ open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Transport -open import Cubical.Foundations.Univalence private variable diff --git a/Cubical/Foundations/Transport.agda b/Cubical/Foundations/Transport.agda index 5e04fab04..a8d5ace02 100644 --- a/Cubical/Foundations/Transport.agda +++ b/Cubical/Foundations/Transport.agda @@ -79,8 +79,6 @@ substComposite B p q Bx = sym (substRefl {B = B} _) ∙ helper where helper : subst B refl (subst B (p ∙ q) Bx) ≡ subst B q (subst B p Bx) helper i = subst B (λ k → compSq (~ i ∧ ~ k) (~ i ∨ k)) (subst B (λ k → compSq (~ i ∨ ~ k) (~ i ∧ k)) Bx) - - -- substitution commutes with morphisms in slices substCommSlice : ∀ {ℓ ℓ′} {A : Type ℓ} → (B C : A → Type ℓ′) diff --git a/Cubical/HITs/Pushout/Base.agda b/Cubical/HITs/Pushout/Base.agda index 9c297e49a..38ab84273 100644 --- a/Cubical/HITs/Pushout/Base.agda +++ b/Cubical/HITs/Pushout/Base.agda @@ -4,8 +4,6 @@ module Cubical.HITs.Pushout.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Pointed -open import Cubical.Data.Prod open import Cubical.Data.Unit diff --git a/Cubical/HITs/SetTruncation/Properties.agda b/Cubical/HITs/SetTruncation/Properties.agda index 3196eca14..f700075fa 100644 --- a/Cubical/HITs/SetTruncation/Properties.agda +++ b/Cubical/HITs/SetTruncation/Properties.agda @@ -92,10 +92,3 @@ setSigmaIso {A = A} {B = B} = iso fun funinv sect retr retr : retract fun funinv retr = elim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ { (a , p) → refl } - - -setSigmaIso2 : ∀ {ℓ} {B : A → Type ℓ} (f g : (x : A) → B x) (a : A) (p1 p2 : f a ≡ g a) - → ((x : A) → isSet (B x)) - → Path ∥ Σ[ x ∈ A ] f x ≡ g x ∥₀ ∣ (a , p1) ∣₀ ∣ (a , p2) ∣₀ -setSigmaIso2 f g a p1 p2 set i = ∣ a , set a (f a) (g a) p1 p2 i ∣₀ - diff --git a/Cubical/HITs/Truncation/Properties.agda b/Cubical/HITs/Truncation/Properties.agda index 3c2112a22..4aee594d4 100644 --- a/Cubical/HITs/Truncation/Properties.agda +++ b/Cubical/HITs/Truncation/Properties.agda @@ -149,10 +149,6 @@ elim3 {n = n} hB g = elim2 (λ _ _ → isOfHLevelΠ n (hB _ _)) (λ a b → elim (λ _ → hB _ _ _) (λ c → g a b c)) - -∥_∥-fun : ∀ {ℓ'} {B : Type ℓ'} (f : A → B) (n : ℕ₋₂) → ∥ A ∥ n → ∥ B ∥ n -∥ f ∥-fun n = elim (λ _ → isOfHLevelTrunc (2+ n)) λ a → ∣ f a ∣ - HLevelTruncModality : ∀ {ℓ} (n : ℕ) → Modality ℓ isModal (HLevelTruncModality n) = isOfHLevel n isModalIsProp (HLevelTruncModality n) = isPropIsOfHLevel n @@ -171,23 +167,12 @@ truncIdempotent n hA = ua (invEquiv (truncIdempotent≃ n hA)) -- universal property -module univTrunc where - fun : ∀ {ℓ} (n : ℕ₋₂) {B : HLevel ℓ (2+ n)} → (∥ A ∥ n → (fst B)) → (A → (fst B)) - fun n {B , lev} = λ g a → g ∣ a ∣ - - inv : ∀ {ℓ} (n : ℕ₋₂) {B : HLevel ℓ (2+ n)} → (A → (fst B)) → (∥ A ∥ n → (fst B)) - inv n {B , lev} = elim (λ _ → lev) - - sect : ∀ {ℓ} (n : ℕ₋₂) {B : HLevel ℓ (2+ n)} → section {A = (∥ A ∥ n → (fst B))} {B = (A → (fst B))} (fun n {B}) (inv n {B}) - sect n {B , lev} b = refl - - retr : ∀ {ℓ} (n : ℕ₋₂) {B : HLevel ℓ (2+ n)} → retract {A = (∥ A ∥ n → (fst B))} {B = (A → (fst B))} (fun n {B}) (inv n {B}) - retr neg2 {B , lev} b = funExt λ x → sym ((snd lev) (elim (λ _ → lev) (λ a → b ∣ a ∣) x)) ∙ (snd lev) (b x) - retr (-2+ suc n) {B , lev} b = funExt (elim (λ x → (isOfHLevelSuc (2+ (-2+ (suc n)) ) lev) ((elim (λ _ → lev) (λ a → b ∣ a ∣) x)) (b x)) λ a → refl) - - univTrunc : ∀ {ℓ} (n : ℕ₋₂) {B : HLevel ℓ (2+ n)} → (∥ A ∥ n → (fst B)) ≃ (A → (fst B)) - univTrunc n {B} = isoToEquiv (iso (fun n {B}) (inv n {B}) (sect n {B}) (retr n {B})) - +univTrunc : ∀ {ℓ} (n : ℕ) {B : HLevel ℓ n} → Iso (hLevelTrunc n A → B .fst) (A → B .fst) +Iso.fun (univTrunc n {B , lev}) g a = g ∣ a ∣ +Iso.inv (univTrunc n {B , lev}) = elim λ _ → lev +Iso.rightInv (univTrunc n {B , lev}) b = refl +Iso.leftInv (univTrunc n {B , lev}) b = funExt (elim (λ x → isOfHLevelPath _ lev _ _) + λ a → refl) -- functorial action @@ -269,12 +254,11 @@ module ΩTrunc where P {n = n} x y = elim2 (λ _ _ → isOfHLevelHLevel2 (2+ n)) (λ a b → ∥ a ≡ b ∥ n , isOfHLevelTrunc (2+ n)) x y .fst - {- We will need P to be of hLevel n + 3 -} - abstract - hLevelP : {n : ℕ₋₂} (a b : ∥ B ∥ (suc₋₂ n)) → isOfHLevel (2+ (suc₋₂ n)) (P a b) - hLevelP {n = n} = - elim2 (λ x y → isProp→isOfHLevelSuc (2+ n) (isPropIsOfHLevel (2+ suc₋₂ n))) - (λ a b → isOfHLevelSuc (2+ n) (isOfHLevelTrunc (2+ n))) + {- We will need P to be of hLevel n + 3 -} + hLevelP : {n : ℕ₋₂} (a b : ∥ B ∥ (suc₋₂ n)) → isOfHLevel (2+ (suc₋₂ n)) (P a b) + hLevelP {n = n} = + elim2 (λ x y → isProp→isOfHLevelSuc (2+ n) (isPropIsOfHLevel (2+ suc₋₂ n))) + (λ a b → isOfHLevelSuc (2+ n) (isOfHLevelTrunc (2+ n))) {- decode function from P x y to x ≡ y -} decode-fun : {n : ℕ₋₂} (x y : ∥ B ∥ (suc₋₂ n)) → P x y → x ≡ y @@ -358,28 +342,20 @@ PathIdTrunc n = isoToPath (ΩTrunc.IsoFinal n _ _) PathΩ : {a : A} (n : ℕ₋₂) → (Path (∥ A ∥ (suc₋₂ n)) ∣ a ∣ ∣ a ∣) ≡ (∥ a ≡ a ∥ n) PathΩ n = PathIdTrunc n -truncEquivΩ : {a : A} (n : ℕ₋₂) → (∥ a ≡ a ∥ n) ≃ (_≡_ {A = ∥ A ∥ (suc₋₂ n)} ∣ a ∣ ∣ a ∣) -truncEquivΩ {a = a} n = isoToEquiv (ΩTrunc.IsoFinal2 ∣ a ∣ ∣ a ∣) - -------------------------- -truncOfTruncEq : (n m : ℕ) → ∥ A ∥ (-2+ n) ≃ ∥ ∥ A ∥ (-2+ (n + m)) ∥ (-2+ n) -truncOfTruncEq {A = A} n m = isoToEquiv (iso fun funInv sect retr) - where - fun : ∥ A ∥ (-2+ n) → ∥ ∥ A ∥ (-2+ (n + m)) ∥ (-2+ n) - fun = elim (λ _ → isOfHLevelTrunc n) λ a → ∣ ∣ a ∣ ∣ - funInv : ∥ ∥ A ∥ (-2+ (n + m)) ∥ (-2+ n) → ∥ A ∥ (-2+ n) - funInv = elim (λ _ → isOfHLevelTrunc n) - (elim (λ _ → transport (λ i → isOfHLevel (+-comm n m (~ i)) (∥ A ∥ (-2+ n))) - (isOfHLevelPlus m (isOfHLevelTrunc n ))) - λ a → ∣ a ∣) - - sect : section fun funInv - sect = elim (λ x → isOfHLevelPath n (isOfHLevelTrunc n) _ _ ) - (elim (λ x → isOfHLevelPath (n + m) (transport (λ i → isOfHLevel (+-comm n m (~ i)) - (∥ ∥ A ∥ (-2+ (n + m)) ∥ (-2+ n))) - (isOfHLevelPlus m (isOfHLevelTrunc n))) _ _ ) - λ a → refl) - - retr : retract fun funInv - retr = elim (λ x → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → refl + +truncOfTruncIso : (n m : ℕ) → Iso (hLevelTrunc n A) (hLevelTrunc n (hLevelTrunc (m + n) A)) +Iso.fun (truncOfTruncIso n m) = elim (λ _ → isOfHLevelTrunc n) λ a → ∣ ∣ a ∣ ∣ +Iso.inv (truncOfTruncIso {A = A} n m) = + elim (λ _ → isOfHLevelTrunc n) + (elim (λ _ → (isOfHLevelPlus m (isOfHLevelTrunc n ))) + λ a → ∣ a ∣) +Iso.rightInv (truncOfTruncIso {A = A} n m) = + elim (λ x → isOfHLevelPath n (isOfHLevelTrunc n) _ _ ) + (elim (λ x → isOfHLevelPath (m + n) (isOfHLevelPlus m (isOfHLevelTrunc n)) _ _ ) + λ a → refl) +Iso.leftInv (truncOfTruncIso n m) = elim (λ x → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → refl + +truncOfTruncEq : (n m : ℕ) → (hLevelTrunc n A) ≃ (hLevelTrunc n (hLevelTrunc (m + n) A)) +truncOfTruncEq n m = isoToEquiv (truncOfTruncIso n m) diff --git a/Cubical/Homotopy/Connected.agda b/Cubical/Homotopy/Connected.agda index bd06687a9..f32459db5 100644 --- a/Cubical/Homotopy/Connected.agda +++ b/Cubical/Homotopy/Connected.agda @@ -35,7 +35,7 @@ isHLevelTruncatedFun : ∀ {ℓ ℓ'} (n : ℕ) {A : Type ℓ} {B : Type ℓ'} ( isHLevelTruncatedFun n f = ∀ b → isOfHLevel n (fiber f b) isConnectedSubtr : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n m : ℕ) (f : A → B) - → isHLevelConnectedFun (n + m) f + → isHLevelConnectedFun (m + n) f → isHLevelConnectedFun n f isConnectedSubtr n m f iscon b = transport (λ i → isContr (ua (truncOfTruncEq {A = fiber f b} n m) (~ i) )) @@ -238,9 +238,8 @@ connectedTruncIso {A = A} {B = B} (suc n) f con = g → P (map fst p) helper P hlev pf = Trunc.elim hlev λ pair → pf (fst pair) (snd pair) - backRetract : (a : A) → ∥ fst ∥-fun (-1+ n) (fst (con (f a))) ≡ ∣ a ∣ - backRetract a = cong (∥ fst ∥-fun (-1+ n)) - (snd (con (f a)) ∣ a , refl ∣) + backRetract : (a : A) → map fst (con (f a) .fst) ≡ ∣ a ∣ + backRetract a = cong (map fst) (con (f a) .snd ∣ a , refl ∣) g : Iso (hLevelTrunc (suc n) A) (hLevelTrunc (suc n) B) Iso.fun g = map f @@ -251,7 +250,7 @@ connectedTruncIso {A = A} {B = B} (suc n) f con = g backSection connectedTruncIso2 : ∀ {ℓ} {A B : Type ℓ} (n m : ℕ) (f : A → B) - → Σ[ x ∈ ℕ ] n + x ≡ m + → Σ[ x ∈ ℕ ] x + n ≡ m → isHLevelConnectedFun m f → Iso (hLevelTrunc n A) (hLevelTrunc n B) connectedTruncIso2 {A = A} {B = B} n m f (x , pf) con = diff --git a/Cubical/ZCohomology/cupProdPrelims.agda b/Cubical/ZCohomology/cupProdPrelims.agda index ca01e9584..7c913df4e 100644 --- a/Cubical/ZCohomology/cupProdPrelims.agda +++ b/Cubical/ZCohomology/cupProdPrelims.agda @@ -254,10 +254,12 @@ Iso.leftInv (decodeIso n x) b = cong (ΩTrunc.encode-fun ∣ x ∣ ∣ x ∣) (f Iso-Kn-ΩKn+1 : (n : ℕ) → Iso (coHomK n) (typ (Ω (coHomK-ptd (suc n)))) Iso-Kn-ΩKn+1 zero = compIso isolooper (congIso (truncIdempotent≃ _ isOfHLevelS1)) Iso-Kn-ΩKn+1 (suc zero) = compIso Iso∣ϕ∣ (decodeIso _ north) -Iso-Kn-ΩKn+1 (suc (suc n)) = compIso (connectedTruncIso2 (4 + n) _ (ϕ north) (n , (λ i → suc (suc (suc (+-suc n n (~ i))))) ∙ - (λ i → suc (suc (+-suc n (suc n) ((~ i)))))) +Iso-Kn-ΩKn+1 (suc (suc n)) = compIso (connectedTruncIso2 (4 + n) _ (ϕ north) (n , helper) (isConnectedσ (suc n) (sphereConnected _))) (decodeIso _ north) + where + helper : n + (4 + n) ≡ 2 + (n + (2 + n)) + helper = +-suc n (3 + n) ∙ (λ i → suc (+-suc n (2 + n) i)) mapId2 : (n : ℕ) → Kn→ΩKn+1 n ≡ Iso.fun (Iso-Kn-ΩKn+1 n) mapId2 zero = refl From 6ce924f70fc714e4304c5cf5b103765f14832a3e Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 4 May 2020 01:07:21 +0200 Subject: [PATCH 27/89] delete abgroup --- Cubical/Data/Group/Base.agda | 7 ------- 1 file changed, 7 deletions(-) diff --git a/Cubical/Data/Group/Base.agda b/Cubical/Data/Group/Base.agda index 9ad745dba..6b26fdbef 100644 --- a/Cubical/Data/Group/Base.agda +++ b/Cubical/Data/Group/Base.agda @@ -7,7 +7,6 @@ open import Cubical.Foundations.Prelude hiding ( comp ) import Cubical.Foundations.Isomorphism as I import Cubical.Foundations.Equiv as E import Cubical.Foundations.Equiv.HalfAdjoint as HAE -open import Cubical.Data.Prod.Base record isGroup {ℓ} (A : Type ℓ) : Type ℓ where constructor group-struct @@ -28,12 +27,6 @@ record Group ℓ : Type (ℓ-suc ℓ) where setStruc : isSet type groupStruc : isGroup type -isAbelianGroup : ∀ {ℓ} → Type (ℓ-suc ℓ) -isAbelianGroup {ℓ} = Σ[ G ∈ Group ℓ ] (∀ a b → isGroup.comp (Group.groupStruc G) a b ≡ isGroup.comp (Group.groupStruc G) b a) - --- (group type setStruc (group-struct id inv comp lUnit rUnit assoc lCancel rCancel)) = - -- (a b : type) → comp a b ≡ comp b a - isMorph : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (f : (Group.type G → Group.type H)) → Type (ℓ-max ℓ ℓ') isMorph (group G Gset (group-struct _ _ _⊙_ _ _ _ _ _)) (group H Hset (group-struct _ _ _∘_ _ _ _ _ _)) f From eaaa4409789fd9dda65f453f8c8f4d7647271c4d Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 4 May 2020 01:16:32 +0200 Subject: [PATCH 28/89] rename --- .../ZCohomology/{cupProdPrelims.agda => KcompPrelims.agda} | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) rename Cubical/ZCohomology/{cupProdPrelims.agda => KcompPrelims.agda} (99%) diff --git a/Cubical/ZCohomology/cupProdPrelims.agda b/Cubical/ZCohomology/KcompPrelims.agda similarity index 99% rename from Cubical/ZCohomology/cupProdPrelims.agda rename to Cubical/ZCohomology/KcompPrelims.agda index 7c913df4e..4a7189e62 100644 --- a/Cubical/ZCohomology/cupProdPrelims.agda +++ b/Cubical/ZCohomology/KcompPrelims.agda @@ -1,5 +1,5 @@ {-# OPTIONS --cubical --safe #-} -module Cubical.ZCohomology.cupProdPrelims where +module Cubical.ZCohomology.KcompPrelims where open import Cubical.ZCohomology.Base open import Cubical.HITs.S1 @@ -204,7 +204,7 @@ private --------------------------------- -- We cheat when n = 1 and use J to prove the following lemmma. There is an obvious dependent path between ϕ base and ϕ north. Since the first one is an iso, so is the other. - -- + -- So far this hasn't been an issue. pointFunIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ} {C : (A : Type ℓ) (a1 : A) → Type ℓ'} (p : A ≡ B) (a : A) (b : B) → From 816f6d833b6af5a3b4bcbb41e39b9266c2ebd76d Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 4 May 2020 01:19:39 +0200 Subject: [PATCH 29/89] renaming --- Cubical/ZCohomology/Properties.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index 982fe86e5..c96c5d372 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -35,7 +35,7 @@ open import Cubical.Data.Sum.Base open import Cubical.Data.HomotopyGroup open import Cubical.Data.Bool hiding (_⊕_) -open import Cubical.ZCohomology.cupProdPrelims +open import Cubical.ZCohomology.KcompPrelims private variable From 6d89eab98a012d58656ef8f0a2c4cbaf2e519324 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 4 May 2020 01:42:02 +0200 Subject: [PATCH 30/89] whitespace... --- Cubical/HITs/S1/Properties.agda | 2 -- Cubical/HITs/Truncation/Properties.agda | 4 ++-- 2 files changed, 2 insertions(+), 4 deletions(-) diff --git a/Cubical/HITs/S1/Properties.agda b/Cubical/HITs/S1/Properties.agda index 10bcedf2f..1806d1a65 100644 --- a/Cubical/HITs/S1/Properties.agda +++ b/Cubical/HITs/S1/Properties.agda @@ -32,5 +32,3 @@ toPropElim {B = B} isprop b base = b toPropElim {B = B} isprop b (loop i) = hcomp (λ k → λ {(i = i0) → b ; (i = i1) → isprop base (subst B (loop) b) b k}) (transp (λ j → B (loop (i ∧ j))) (~ i) b) - - diff --git a/Cubical/HITs/Truncation/Properties.agda b/Cubical/HITs/Truncation/Properties.agda index 4aee594d4..c91511fe6 100644 --- a/Cubical/HITs/Truncation/Properties.agda +++ b/Cubical/HITs/Truncation/Properties.agda @@ -244,7 +244,7 @@ groupoidTrunc≡Trunc1 = ua groupoidTrunc≃Trunc1 abstract isOfHLevelHLevel2 : ∀ n → isOfHLevel (suc n) (HLevel ℓ n) - isOfHLevelHLevel2 n = isOfHLevelHLevel n + isOfHLevelHLevel2 n = isOfHLevelHLevel n {- Proofs of Theorem 7.3.12. and Corollary 7.3.13. in the HoTT book -} @@ -347,7 +347,7 @@ PathΩ n = PathIdTrunc n truncOfTruncIso : (n m : ℕ) → Iso (hLevelTrunc n A) (hLevelTrunc n (hLevelTrunc (m + n) A)) Iso.fun (truncOfTruncIso n m) = elim (λ _ → isOfHLevelTrunc n) λ a → ∣ ∣ a ∣ ∣ -Iso.inv (truncOfTruncIso {A = A} n m) = +Iso.inv (truncOfTruncIso {A = A} n m) = elim (λ _ → isOfHLevelTrunc n) (elim (λ _ → (isOfHLevelPlus m (isOfHLevelTrunc n ))) λ a → ∣ a ∣) From 457cd7b13365d019b9ca6517c1d7da282471bbec Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Wed, 6 May 2020 18:28:39 +0200 Subject: [PATCH 31/89] MVunreduced --- Cubical/Data/Group/Base.agda | 44 ++++ Cubical/Foundations/GroupoidLaws.agda | 10 + Cubical/Foundations/Prelude.agda | 1 + Cubical/ZCohomology/KcompPrelims.agda | 22 ++ Cubical/ZCohomology/MayerVietoris.agda | 235 ++++++++++++++++++ Cubical/ZCohomology/MayerVietorisReduced.agda | 192 ++++++++++++++ Cubical/ZCohomology/Properties.agda | 192 ++++++++++---- Cubical/ZCohomology/altToFreudenthal.agda | 201 +++++++++++++++ 8 files changed, 850 insertions(+), 47 deletions(-) create mode 100644 Cubical/ZCohomology/MayerVietoris.agda create mode 100644 Cubical/ZCohomology/MayerVietorisReduced.agda create mode 100644 Cubical/ZCohomology/altToFreudenthal.agda diff --git a/Cubical/Data/Group/Base.agda b/Cubical/Data/Group/Base.agda index 6b26fdbef..bdb7b030d 100644 --- a/Cubical/Data/Group/Base.agda +++ b/Cubical/Data/Group/Base.agda @@ -3,6 +3,8 @@ module Cubical.Data.Group.Base where open import Cubical.Foundations.Prelude hiding ( comp ) +open import Cubical.Foundations.HLevels +open import Cubical.Data.Prod import Cubical.Foundations.Isomorphism as I import Cubical.Foundations.Equiv as E @@ -160,3 +162,45 @@ compIso {ℓ} {F} {G} {H} ret f = h→f (f→h f) ≡⟨ cong g→f (hg (f→g f)) ⟩ g→f (f→g f) ≡⟨ gf _ ⟩ f ∎ + + +Im : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (Group.type G → Group.type H) + → Type (ℓ-max ℓ ℓ') +Im (group A setA strucA) (group B setB strucB) f = + Σ[ b ∈ B ] Σ[ a ∈ A ] f a ≡ b + +isInIm : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (Group.type G → Group.type H) + → Group.type H → Type _ +isInIm G H ϕ h = Σ[ g ∈ Group.type G ] ϕ g ≡ h + +Ker : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (f : (Group.type G → Group.type H)) + → Type (ℓ-max ℓ ℓ') +Ker (group A setA strA) (group B setB strB) f = + Σ[ a ∈ A ] f a ≡ isGroup.id strB + +isInKer : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (Group.type G → Group.type H) + → Group.type G → Type _ +isInKer G H ϕ g = ϕ g ≡ isGroup.id (Group.groupStruc H) + + + +{- direct products of groups -} +dirProd : ∀ {ℓ ℓ'} (A : Group ℓ) (B : Group ℓ') → Group (ℓ-max ℓ ℓ') +Group.type (dirProd (group A setA strucA) (group B setB strucB)) = A × B +Group.setStruc (dirProd (group A setA strucA) (group B setB strucB)) = isOfHLevelProd 2 setA setB +isGroup.id (Group.groupStruc (dirProd (group A setA strucA) (group B setB strucB))) = + (isGroup.id strucA) , (isGroup.id strucB) +isGroup.inv (Group.groupStruc (dirProd (group A setA strucA) (group B setB strucB))) = + map (isGroup.inv strucA) (isGroup.inv strucB) +isGroup.comp (Group.groupStruc (dirProd (group A setA strucA) (group B setB strucB))) (a1 , b1) (a2 , b2) = + (isGroup.comp strucA a1 a2) , isGroup.comp strucB b1 b2 +isGroup.lUnit (Group.groupStruc (dirProd (group A setA strucA) (group B setB strucB))) (a , b) i = + (isGroup.lUnit strucA a i) , (isGroup.lUnit strucB b i) +isGroup.rUnit (Group.groupStruc (dirProd (group A setA strucA) (group B setB strucB))) (a , b) i = + (isGroup.rUnit strucA a i) , (isGroup.rUnit strucB b i) +isGroup.assoc (Group.groupStruc (dirProd (group A setA strucA) (group B setB strucB))) (a1 , b1) (a2 , b2) (a3 , b3) i = + (isGroup.assoc strucA a1 a2 a3 i) , (isGroup.assoc strucB b1 b2 b3 i) +isGroup.lCancel (Group.groupStruc (dirProd (group A setA strucA) (group B setB strucB))) (a , b) i = + (isGroup.lCancel strucA a i) , (isGroup.lCancel strucB b i) +isGroup.rCancel (Group.groupStruc (dirProd (group A setA strucA) (group B setB strucB))) (a , b) i = + (isGroup.rCancel strucA a i) , (isGroup.rCancel strucB b i) diff --git a/Cubical/Foundations/GroupoidLaws.agda b/Cubical/Foundations/GroupoidLaws.agda index cc0eaa033..2f2c9a10d 100644 --- a/Cubical/Foundations/GroupoidLaws.agda +++ b/Cubical/Foundations/GroupoidLaws.agda @@ -286,6 +286,16 @@ congFunct f p q i = hcomp (λ j → λ{(i = i0) → rUnit (cong f (p ∙ q)) (~ (i = i1) → cong f (rUnit p (~ j)) ∙ cong f q}) (cong f (p ∙ (λ k → q (k ∧ (~ i)))) ∙ cong f λ k → q ((~ i) ∨ k) ) +cong₂Funct : (f : A → A → A) → + (p : x ≡ y) → + {u v : A} (q : u ≡ v) → + cong₂ f p q ≡ cong (λ x → f x u) p ∙ cong (f y) q +cong₂Funct {x = x} {y = y} f p {u = u} {v = v} q j i = + hcomp (λ k → λ { (i = i0) → f x u + ; (i = i1) → f y (q k) + ; (j = i0) → f (p i) (q (i ∧ k))}) + (f (p i) u) + symDistr : (p : x ≡ y) (q : y ≡ z) → sym (p ∙ q) ≡ sym q ∙ sym p symDistr p q i = hcomp (λ j → λ{(i = i0) → rUnit (sym (p ∙ q)) (~ j) ; (i = i1) → sym (lUnit q (~ j)) ∙ sym p}) diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index 6a3ced803..2a0850ad0 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -61,6 +61,7 @@ cong₂ : ∀ {C : (a : A) → (b : B a) → Type ℓ} → PathP (λ i → C (p i) (q i)) (f x u) (f y v) cong₂ f p q i = f (p i) (q i) + {- The most natural notion of homogenous path composition in a cubical setting is double composition: diff --git a/Cubical/ZCohomology/KcompPrelims.agda b/Cubical/ZCohomology/KcompPrelims.agda index 4a7189e62..f8d9f2219 100644 --- a/Cubical/ZCohomology/KcompPrelims.agda +++ b/Cubical/ZCohomology/KcompPrelims.agda @@ -281,3 +281,25 @@ Iso.leftInv (Iso2-Kn-ΩKn+1 n) a = linv abstract linv : Iso.inv (Iso-Kn-ΩKn+1 n) (Kn→ΩKn+1 n a) ≡ a linv = cong (Iso.inv (Iso-Kn-ΩKn+1 n)) (funExt⁻ (mapId2 n) a) ∙ Iso.leftInv (Iso-Kn-ΩKn+1 n) a + +--- even more abstract + +abstract + absInv : (n : ℕ) → typ (Ω (coHomK-ptd (2 + n))) → coHomK (1 + n) + absInv n = Iso.inv (Iso-Kn-ΩKn+1 (1 + n)) + + absSect : (n : ℕ) (a : typ (Ω (coHomK-ptd (2 + n)))) → Kn→ΩKn+1 (1 + n) (absInv n a) ≡ a + absSect n a = funExt⁻ (mapId2 (1 + n)) _ ∙ Iso.rightInv (Iso-Kn-ΩKn+1 (1 + n)) a + + absRetr : (n : ℕ) (a : coHomK (1 + n)) → absInv n (Kn→ΩKn+1 (1 + n) a) ≡ a + absRetr n a = cong (Iso.inv (Iso-Kn-ΩKn+1 (1 + n))) (funExt⁻ (mapId2 (1 + n)) a) ∙ Iso.leftInv (Iso-Kn-ΩKn+1 (1 + n)) a + + +Iso3-Kn-ΩKn+1 : (n : ℕ) → Iso (coHomK n) (typ (Ω (coHomK-ptd (suc n)))) +Iso.fun (Iso3-Kn-ΩKn+1 n) = Kn→ΩKn+1 n +Iso.inv (Iso3-Kn-ΩKn+1 zero) = Iso.inv (Iso-Kn-ΩKn+1 zero) +Iso.inv (Iso3-Kn-ΩKn+1 (suc n)) = absInv n +Iso.rightInv (Iso3-Kn-ΩKn+1 zero) a = funExt⁻ (mapId2 zero) (Iso.inv isolooper2 (Iso.inv (congIso (truncIdempotent≃ 3 isOfHLevelS1)) a)) ∙ Iso.rightInv (Iso-Kn-ΩKn+1 zero) a +Iso.rightInv (Iso3-Kn-ΩKn+1 (suc n)) = absSect n +Iso.leftInv (Iso3-Kn-ΩKn+1 zero) a = cong (Iso.inv (Iso-Kn-ΩKn+1 zero)) (funExt⁻ (mapId2 zero) a) ∙ Iso.leftInv (Iso-Kn-ΩKn+1 zero) a +Iso.leftInv (Iso3-Kn-ΩKn+1 (suc n)) = absRetr n diff --git a/Cubical/ZCohomology/MayerVietoris.agda b/Cubical/ZCohomology/MayerVietoris.agda new file mode 100644 index 000000000..08db1c8e7 --- /dev/null +++ b/Cubical/ZCohomology/MayerVietoris.agda @@ -0,0 +1,235 @@ + +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.MayerVietoris where + +open import Cubical.ZCohomology.Base +open import Cubical.HITs.S1 +open import Cubical.ZCohomology.Properties +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence +open import Cubical.Data.NatMinusTwo.Base +open import Cubical.Data.Empty +open import Cubical.Data.Sigma hiding (_×_) +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.Nullification +open import Cubical.Data.Int hiding (_+_) +open import Cubical.Data.Nat +open import Cubical.Data.Prod +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Freudenthal +open import Cubical.HITs.SmashProduct.Base +open import Cubical.Data.Group.Base + + +open import Cubical.ZCohomology.KcompPrelims + + +open import Cubical.HITs.Pushout +open import Cubical.Data.Sum.Base +open import Cubical.Data.HomotopyGroup +open import Cubical.Data.Bool hiding (_⊕_) + + +-- {- +-- record AbGroup {ℓ} (A : Type ℓ) : Type ℓ where +-- constructor abgr +-- field +-- noll : A +-- _⊕_ : A → A → A +-- associate : (a b c : A) → (a ⊕ b) ⊕ c ≡ a ⊕ (b ⊕ c) +-- commute : (a b : A) → a ⊕ b ≡ (b ⊕ a) +-- r-unit : (a : A) → a ⊕ noll ≡ a +-- negate : (a : A) → Σ[ a⁻ ∈ A ] (a ⊕ a⁻ ≡ noll) +-- -} +-- {- +-- grHom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (ϕ : A → B) +-- → AbGroup A → AbGroup B → Type (ℓ-max ℓ ℓ') +-- grHom {A = A} {B = B} ϕ (abgr 0A _+A_ as-A comm-A r-A neg-A) (abgr 0B _+B_ as-B comm-B r-B neg-B) = (x y : A) → ϕ (x +A y) ≡ (ϕ x +B ϕ y) +-- ImHom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (ϕ : A → B) (A' : AbGroup A) (B' : AbGroup B) +-- → grHom ϕ A' B' → Type (ℓ-max ℓ ℓ') +-- ImHom {A = A} {B = B} ϕ Agr Bgr hom = Σ[ b ∈ B ] Σ[ a ∈ A ] ϕ a ≡ b +-- KerHom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (ϕ : A → B) (A' : AbGroup A) (B' : AbGroup B) +-- → grHom ϕ A' B' → Type (ℓ-max ℓ ℓ') +-- KerHom {A = A} {B = B} ϕ Agr (abgr 0B _+B_ as-B comm-B r-B neg-B) hom = Σ[ a ∈ A ] ϕ a ≡ 0B +-- -} + + +-- --- + + +private + variable + ℓ ℓ' ℓ'' : Level + A : Type ℓ + B : Type ℓ' + C : Type ℓ'' + A' : Pointed ℓ + B' : Pointed ℓ' + + +-- coHomFun : (n : ℕ) (f : A → B) → coHom n B → coHom n A +-- coHomFun n f = sRec setTruncIsSet λ β → ∣ β ∘ f ∣₀ + +-- coHomFun∙ : (n : ℕ) (f : A' →∙ B') → coHomRed n B' → coHomRed n A' +-- coHomFun∙ n (f , fpt) = sRec setTruncIsSet λ { (β , βpt) → ∣ β ∘ f , (λ i → β (fpt i)) ∙ βpt ∣₀} + +-- coHomFun∙2 : (n : ℕ) (f : A → B' .fst) → coHomRed n B' → coHom n A +-- coHomFun∙2 zero f = sRec setTruncIsSet λ { (β , βpt) → ∣ β ∘ f ∣₀} +-- coHomFun∙2 (suc n) f = sRec setTruncIsSet λ { (β , βpt) → ∣ β ∘ f ∣₀} + +-- module MV {ℓ ℓ' ℓ''} ((A , a₀) : Pointed ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) where +-- D : Pointed (ℓ-max (ℓ-max ℓ ℓ') ℓ'') +-- D = Pushout f g , inl a₀ + +-- -- σ : typ A → typ (Ω (∙Susp (typ A))) +-- -- σ a = merid a ∙ merid (pt A) ⁻¹ + +-- i : (n : ℕ) → coHomRed n D → coHomRed n (A , a₀) × coHom n B +-- i zero = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) +-- λ { (δ , δpt) → ∣ (λ x → δ (inl x)) , δpt ∣₀ , ∣ (λ x → δ (inr x)) ∣₀ } +-- i (suc n) = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) +-- λ { (δ , δpt) → ∣ (λ x → δ (inl x)) , δpt ∣₀ , ∣ (λ x → δ (inr x)) ∣₀ } + +-- Δ : (n : ℕ) → coHomRed n (A , a₀) × coHom n B → coHom n C +-- Δ n (α , β) = coHomFun∙2 n f α +ₕ (-ₕ (coHomFun n g β)) + +-- -- + +-- d-pre1 : (n : ℕ) → (C → coHomK n) → D .fst → coHomK-ptd (suc n) .fst +-- d-pre1 zero γ (inl x) = 0ₖ +-- d-pre1 zero γ (inr x) = 0ₖ +-- d-pre1 zero γ (push a i) = Kn→ΩKn+1 zero (γ a) i +-- d-pre1 (suc n) γ (inl x) = 0ₖ +-- d-pre1 (suc n) γ (inr x) = 0ₖ +-- d-pre1 (suc n) γ (push a i) = Kn→ΩKn+1 (suc n) (γ a) i + +-- d-pre1∙ : (n : ℕ) → (γ : (C → coHomK n)) → d-pre1 n γ (snd D) ≡ ∣ north ∣ +-- d-pre1∙ zero γ = refl +-- d-pre1∙ (suc n) γ = refl + +-- d-pre2 : (n : ℕ) → (C → coHomK n) → D →∙ coHomK-ptd (suc n) +-- d-pre2 n γ = d-pre1 n γ , d-pre1∙ n γ + +-- d : (n : ℕ) → coHom n C → coHomRed (suc n) D +-- d n = sRec setTruncIsSet λ a → ∣ d-pre2 n a ∣₀ + +-- iIsHom : (n : ℕ) (x y : coHomRed n D) → i n (+ₕ∙ n x y) ≡ +prod n (i n x) (i n y) +-- iIsHom zero = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- λ { (f , fpt) (g , gpt) → refl} +-- iIsHom (suc n) = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- λ { (f , fpt) (g , gpt) → refl} + +-- prodElim : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : ∥ A ∥₀ × ∥ B ∥₀ → Type ℓ''} +-- → ((x : ∥ A ∥₀ × ∥ B ∥₀) → isOfHLevel 2 (C x)) +-- → ((a : A) (b : B) → C (∣ a ∣₀ , ∣ b ∣₀)) +-- → (x : ∥ A ∥₀ × ∥ B ∥₀) → C x +-- prodElim {A = A} {B = B} {C = C} hlevel ind (a , b) = schonf a b +-- where +-- schonf : (a : ∥ A ∥₀) (b : ∥ B ∥₀) → C (a , b) +-- schonf = sElim (λ x → isOfHLevelΠ 2 λ y → hlevel (_ , _)) λ a → sElim (λ x → hlevel (_ , _)) +-- λ b → ind a b + +-- ΔIsHom : (n : ℕ) (x y : coHomRed n (A , a₀) × coHom n B) → Δ n (+prod n x y) ≡ (Δ n x +ₕ Δ n y) +-- ΔIsHom zero = prodElim (λ x → isOfHLevelΠ 2 λ y → isOfHLevelPath 2 setTruncIsSet _ _ ) +-- λ {(f' , fpt) x1 → prodElim (λ x → isOfHLevelPath 2 setTruncIsSet _ _) +-- λ {(g' , gpt) x2 → (λ i → ∣ (λ x → (f' (f x) +ₖ g' (f x)) +ₖ -distrₖ (x1 (g x)) (x2 (g x)) i) ∣₀) ∙ +-- (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ g' (f x)) (-ₖ x1 (g x)) (-ₖ x2 (g x)) (~ i)) ∣₀) ∙ +-- (λ i → ∣ (λ x → assocₖ (f' (f x)) (g' (f x)) (-ₖ x1 (g x)) i +ₖ (-ₖ x2 (g x))) ∣₀) ∙ +-- (λ i → ∣ (λ x → (f' (f x) +ₖ commₖ (g' (f x)) (-ₖ x1 (g x)) i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ +-- (λ i → ∣ (λ x → assocₖ (f' (f x)) (-ₖ x1 (g x)) (g' (f x)) (~ i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ +-- (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ (-ₖ x1 (g x))) (g' (f x)) (-ₖ x2 (g x)) i) ∣₀)}} +-- ΔIsHom (suc n) = prodElim (λ x → isOfHLevelΠ 2 λ y → isOfHLevelPath 2 setTruncIsSet _ _ ) +-- λ {(f' , fpt) x1 → prodElim (λ x → isOfHLevelPath 2 setTruncIsSet _ _) +-- λ {(g' , gpt) x2 → (λ i → ∣ (λ x → (f' (f x) +ₖ g' (f x)) +ₖ -distrₖ (x1 (g x)) (x2 (g x)) i) ∣₀) ∙ +-- (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ g' (f x)) (-ₖ x1 (g x)) (-ₖ x2 (g x)) (~ i)) ∣₀) ∙ +-- (λ i → ∣ (λ x → assocₖ (f' (f x)) (g' (f x)) (-ₖ x1 (g x)) i +ₖ (-ₖ x2 (g x))) ∣₀) ∙ +-- (λ i → ∣ (λ x → (f' (f x) +ₖ commₖ (g' (f x)) (-ₖ x1 (g x)) i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ +-- (λ i → ∣ (λ x → assocₖ (f' (f x)) (-ₖ x1 (g x)) (g' (f x)) (~ i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ +-- (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ (-ₖ x1 (g x))) (g' (f x)) (-ₖ x2 (g x)) i) ∣₀)}} + +-- niceHelper : (n : ℕ) (h l : C → coHomK n) (x : D .fst) → (d-pre1 n h x +ₖ d-pre1 n l x) ≡ d-pre1 n (λ x → h x +ₖ l x) x +-- niceHelper zero h l (inl x) = {!!} +-- niceHelper zero h l (inr x) = {!!} +-- niceHelper zero h l (push a i₁) = {!!} +-- niceHelper (suc n) h l (inl x) = {!!} -- lUnitₖ 0ₖ +-- niceHelper (suc n) h l (inr x) = {!!} -- lUnitₖ 0ₖ +-- niceHelper (suc n) h l (push a i) j = +-- hcomp (λ k → λ {(i = i0) → {!lUnitₖ 0ₖ j!} +-- ; (i = i1) → {!!} +-- ; (j = i0) → {!!} +-- ; (j = i1) → {!!}}) +-- {!!} +-- where +-- side : I → I → coHomK n +-- side i j = +-- hcomp {!!} +-- {!!} + +-- helper : (h l : C → Int) (x : D .fst) → (d-pre1 zero h x +ₖ d-pre1 zero l x) ≡ d-pre1 zero (λ x → h x +ₖ l x) x +-- helper h l (inl x) = lUnitₖ 0ₖ +-- helper h l (inr x) = lUnitₖ 0ₖ +-- helper h l (push a i) = {!!} +-- where +-- helper1 : Kn→ΩKn+1 zero (h a +ₖ l a) ≡ cong₂ (λ x y → x +ₖ y) (Kn→ΩKn+1 zero (h a)) (Kn→ΩKn+1 zero (l a)) +-- helper1 = (λ i → Kn→ΩKn+1 zero (h a +ₖ l a)) ∙ +ₖ→∙ zero (h a) (l a) ∙ +-- {!!} ∙ +-- sym (cong₂Funct (λ x y → x +ₖ y) (Kn→ΩKn+1 zero (h a)) (Kn→ΩKn+1 zero (l a))) ∙ +-- (λ i j → Kn→ΩKn+1 zero (h a) j +ₖ Kn→ΩKn+1 zero (l a) j) ∙ +-- λ i j → ΩKn+1→Kn (Kn→ΩKn+1 1 (Kn→ΩKn+1 zero (h a) j) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 zero (l a) j)) +-- helper2 : PathP (λ i → rUnitₖ 0ₖ i ≡ rUnitₖ 0ₖ i ) (cong (λ x → x +ₖ ∣ north ∣) (Kn→ΩKn+1 zero (h a)) ∙ cong (λ y → ∣ north ∣ +ₖ y) (Kn→ΩKn+1 zero (l a))) +-- (Kn→ΩKn+1 zero (h a) ∙ Kn→ΩKn+1 zero (l a)) +-- helper2 j i = hcomp (λ k → λ { (i = i0) → rUnitₖ ∣ north ∣ j +-- ; (i = i1) → {!!}}) +-- (rUnitₖ (Kn→ΩKn+1 zero (h a) i) j) -- (cong (λ x → {!rUnitₖ {n = 1} x i!}) (Kn→ΩKn+1 zero (h a)) ∙ cong (λ y → lUnitₖ {n = 1} y i) (Kn→ΩKn+1 zero (l a))) + +-- helper3 : PathP (λ i → rUnitₖ 0ₖ i ≡ rUnitₖ 0ₖ i) (cong (λ x → x +ₖ ∣ north ∣) (Kn→ΩKn+1 zero (h a))) (Kn→ΩKn+1 zero (h a)) +-- helper3 i = cong (λ x → rUnitₖ x i) (Kn→ΩKn+1 zero (h a)) + +-- heler4 : PathP (λ i → lUnitₖ 0ₖ i ≡ lUnitₖ 0ₖ i) (cong (λ x → ∣ north ∣ +ₖ x) (Kn→ΩKn+1 zero (l a))) (Kn→ΩKn+1 zero (l a)) +-- heler4 i = cong (λ x → lUnitₖ x i) (Kn→ΩKn+1 zero (l a)) + +-- ttest : {!!} +-- ttest = congFunct +-- -- test2 : (n : ℕ) (a : C) (h l : C → coHomK n) → cong₂ (λ x y → x +ₖ y) (Kn→ΩKn+1 n (h a)) (Kn→ΩKn+1 n (l a)) ≡ {!!} ∙ Kn→ΩKn+1 n (h a +ₖ l a) ∙ {!!} +-- -- test2 = {!!} + +-- -- pp : PathP (λ i → lUnitₖ (0ₖ {n = 1}) i ≡ lUnitₖ 0ₖ i) (λ i → Kn→ΩKn+1 zero (h a) i +ₖ Kn→ΩKn+1 zero (l a) i) (Kn→ΩKn+1 zero (h a +ₖ l a)) +-- -- pp = toPathP ((λ z → transport (λ i → lUnitₖ (0ₖ {n = 1}) i ≡ lUnitₖ 0ₖ i) (lUnit (rUnit (λ i → Kn→ΩKn+1 zero (h a) i +ₖ Kn→ΩKn+1 zero (l a) i) z) z )) +-- -- ∙ (λ z → transp (λ i → lUnitₖ (0ₖ {n = 1}) (i ∨ z) ≡ lUnitₖ 0ₖ (i ∨ z)) z +-- -- (((λ i → lUnitₖ (0ₖ {n = 1}) ((~ i) ∧ z)) ∙ (λ i → Kn→ΩKn+1 zero (h a) i +ₖ Kn→ΩKn+1 zero (l a) i) +-- -- ∙ λ i → lUnitₖ (0ₖ {n = 1}) (i ∧ z)))) +-- -- ∙ (λ z → sym ((λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ 1 i ∙ Kn→ΩKn+1 1 0ₖ)) ∙ +-- -- (λ i → ΩKn+1→Kn (lUnit (Kn→ΩKn+1 1 0ₖ) (~ i))) ∙ +-- -- Iso.leftInv (Iso2-Kn-ΩKn+1 1) 0ₖ) +-- -- ∙ (lUnit (rUnit (λ i → Kn→ΩKn+1 zero (h a) i +ₖ Kn→ΩKn+1 zero (l a) i) z) z) +-- -- ∙ ((λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ 1 i ∙ Kn→ΩKn+1 1 0ₖ)) ∙ +-- -- (λ i → ΩKn+1→Kn (lUnit (Kn→ΩKn+1 1 0ₖ) (~ i))) ∙ +-- -- Iso.leftInv (Iso2-Kn-ΩKn+1 1) 0ₖ)) +-- -- ∙ (λ z → sym (lUnitₖ (0ₖ {n = 1})) +-- -- ∙ (lUnit (rUnit refl z ) z ∙ (λ i → ΩKn+1→Kn (Kn→ΩKn+1 1 (Kn→ΩKn+1 zero (h a) i) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 zero (l a) i))) ∙ lUnit (rUnit refl z) z) +-- -- ∙ lUnitₖ (0ₖ {n = 1})) +-- -- ∙ ((λ z → sym (lUnitₖ (0ₖ {n = 1})) +-- -- ∙ (({!!} ∙ {!ΩKn+1→Kn (Kn→ΩKn+1 1 (Kn→ΩKn+1 zero (h a) i) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 zero (l a) i)!} ∙ λ i → {!l!}) +-- -- ∙ (λ i → Iso.leftInv (Iso2-Kn-ΩKn+1 1) {!!} {!!}) ∙ (λ i → (Iso.leftInv (Iso2-Kn-ΩKn+1 1) 0ₖ (~ i ∧ z))) ∙ {!!} ∙ {!!}) +-- -- ∙ lUnitₖ (0ₖ {n = 1}))) +-- -- ∙ {!!} +-- -- ∙ {!!} +-- -- ∙ {!!}) -- compPathP (λ j i → {!!}) (compPathP {!!} (compPathP {!!} {!!})) + +-- -- dIsHom : (n : ℕ) (x y : coHom n C) → d n (x +ₕ y) ≡ +ₕ∙ (suc n) (d n x) (d n y) +-- -- dIsHom zero = sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- -- λ h l → (λ i → ∣ (λ x → helper h l x (~ i)) , lUnit (λ j → lUnitₖ 0ₖ (j ∨ (~ i))) i ∣₀) ∙ +-- -- (λ i → ∣ (λ x → d-pre1 zero h x +ₖ d-pre1 zero l x) , refl ∙ lUnitₖ 0ₖ ∣₀) +-- -- dIsHom (suc n) = {!d-pre1 zero h x +ₖ d-pre1 zero l x!} diff --git a/Cubical/ZCohomology/MayerVietorisReduced.agda b/Cubical/ZCohomology/MayerVietorisReduced.agda new file mode 100644 index 000000000..6979f236c --- /dev/null +++ b/Cubical/ZCohomology/MayerVietorisReduced.agda @@ -0,0 +1,192 @@ + +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.MayerVietorisReduced where + +open import Cubical.ZCohomology.Base +open import Cubical.HITs.S1 +open import Cubical.ZCohomology.Properties +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence +open import Cubical.Data.NatMinusTwo.Base +open import Cubical.Data.Empty +open import Cubical.Data.Sigma hiding (_×_) +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.Nullification +open import Cubical.Data.Int hiding (_+_) +open import Cubical.Data.Nat +open import Cubical.Data.Prod +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Freudenthal +open import Cubical.HITs.SmashProduct.Base +open import Cubical.Data.Group.Base + + +open import Cubical.ZCohomology.KcompPrelims + + +open import Cubical.HITs.Pushout +open import Cubical.Data.Sum.Base +open import Cubical.Data.HomotopyGroup +open import Cubical.Data.Bool hiding (_⊕_) + + +-- {- +-- record AbGroup {ℓ} (A : Type ℓ) : Type ℓ where +-- constructor abgr +-- field +-- noll : A +-- _⊕_ : A → A → A +-- associate : (a b c : A) → (a ⊕ b) ⊕ c ≡ a ⊕ (b ⊕ c) +-- commute : (a b : A) → a ⊕ b ≡ (b ⊕ a) +-- r-unit : (a : A) → a ⊕ noll ≡ a +-- negate : (a : A) → Σ[ a⁻ ∈ A ] (a ⊕ a⁻ ≡ noll) +-- -} +-- {- +-- grHom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (ϕ : A → B) +-- → AbGroup A → AbGroup B → Type (ℓ-max ℓ ℓ') +-- grHom {A = A} {B = B} ϕ (abgr 0A _+A_ as-A comm-A r-A neg-A) (abgr 0B _+B_ as-B comm-B r-B neg-B) = (x y : A) → ϕ (x +A y) ≡ (ϕ x +B ϕ y) +-- ImHom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (ϕ : A → B) (A' : AbGroup A) (B' : AbGroup B) +-- → grHom ϕ A' B' → Type (ℓ-max ℓ ℓ') +-- ImHom {A = A} {B = B} ϕ Agr Bgr hom = Σ[ b ∈ B ] Σ[ a ∈ A ] ϕ a ≡ b +-- KerHom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (ϕ : A → B) (A' : AbGroup A) (B' : AbGroup B) +-- → grHom ϕ A' B' → Type (ℓ-max ℓ ℓ') +-- KerHom {A = A} {B = B} ϕ Agr (abgr 0B _+B_ as-B comm-B r-B neg-B) hom = Σ[ a ∈ A ] ϕ a ≡ 0B +-- -} + + +-- --- + + +private + variable + ℓ ℓ' ℓ'' : Level + A : Type ℓ + B : Type ℓ' + C : Type ℓ'' + A' : Pointed ℓ + B' : Pointed ℓ' + + +coHomFun : (n : ℕ) (f : A → B) → coHom n B → coHom n A +coHomFun n f = sRec setTruncIsSet λ β → ∣ β ∘ f ∣₀ + +module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) where + D : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + D = Pushout f g + + -- σ : typ A → typ (Ω (∙Susp (typ A))) + -- σ a = merid a ∙ merid (pt A) ⁻¹ + + i : (n : ℕ) → coHom n D → coHom n A × coHom n B + i zero = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) + λ { δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀} + i (suc n) = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) + λ { δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀} + + Δ : (n : ℕ) → coHom n A × coHom n B → coHom n C + Δ n (α , β) = coHomFun n f α +ₕ -ₕ (coHomFun n g β) -- coHomFun∙2 n f α +ₕ (-ₕ (coHomFun n g β)) + + d-pre : (n : ℕ) → (C → coHomK n) → D → coHomK (suc n) + d-pre n γ (inl x) = 0ₖ + d-pre n γ (inr x) = 0ₖ + d-pre zero γ (push a i) = Kn→ΩKn+1 zero (γ a) i + d-pre (suc n) γ (push a i) = Kn→ΩKn+1 (suc n) (γ a) i + + d : (n : ℕ) → coHom n C → coHom (suc n) D + d n = sRec setTruncIsSet λ a → ∣ d-pre n a ∣₀ + + iIsHom : (n : ℕ) → isMorph (coHomGr n D) (×coHomGr n A B) (i n) + iIsHom zero = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ f g → refl + iIsHom (suc n) = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ f g → refl + + + prodElim : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : ∥ A ∥₀ × ∥ B ∥₀ → Type ℓ''} + → ((x : ∥ A ∥₀ × ∥ B ∥₀) → isOfHLevel 2 (C x)) + → ((a : A) (b : B) → C (∣ a ∣₀ , ∣ b ∣₀)) + → (x : ∥ A ∥₀ × ∥ B ∥₀) → C x + prodElim {A = A} {B = B} {C = C} hlevel ind (a , b) = schonf a b + where + schonf : (a : ∥ A ∥₀) (b : ∥ B ∥₀) → C (a , b) + schonf = sElim (λ x → isOfHLevelΠ 2 λ y → hlevel (_ , _)) λ a → sElim (λ x → hlevel (_ , _)) + λ b → ind a b + + + prodElim2 : ∀ {ℓ ℓ' ℓ'' ℓ''' ℓ''''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} + {E : (∥ A ∥₀ × ∥ B ∥₀) → (∥ C ∥₀ × ∥ D ∥₀) → Type ℓ''''} + → ((x : ∥ A ∥₀ × ∥ B ∥₀) (y : ∥ C ∥₀ × ∥ D ∥₀) → isOfHLevel 2 (E x y)) + → ((a : A) (b : B) (c : C) (d : D) → E (∣ a ∣₀ , ∣ b ∣₀) (∣ c ∣₀ , ∣ d ∣₀)) + → ((x : ∥ A ∥₀ × ∥ B ∥₀) (y : ∥ C ∥₀ × ∥ D ∥₀) → (E x y)) + prodElim2 isset f = prodElim (λ _ → isOfHLevelΠ 2 λ _ → isset _ _) + λ a b → prodElim (λ _ → isset _ _) + λ c d → f a b c d + + ΔIsHom : (n : ℕ) → isMorph (×coHomGr n A B) (coHomGr n C) (Δ n) + ΔIsHom zero = prodElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _ ) + λ f' x1 g' x2 → (λ i → ∣ (λ x → (f' (f x) +ₖ g' (f x)) +ₖ -distrₖ (x1 (g x)) (x2 (g x)) i) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ g' (f x)) (-ₖ x1 (g x)) (-ₖ x2 (g x)) (~ i)) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x)) (g' (f x)) (-ₖ x1 (g x)) i +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → (f' (f x) +ₖ commₖ (g' (f x)) (-ₖ x1 (g x)) i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x)) (-ₖ x1 (g x)) (g' (f x)) (~ i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ (-ₖ x1 (g x))) (g' (f x)) (-ₖ x2 (g x)) i) ∣₀) + ΔIsHom (suc n) = prodElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _ ) + λ f' x1 g' x2 → (λ i → ∣ (λ x → (f' (f x) +ₖ g' (f x)) +ₖ -distrₖ (x1 (g x)) (x2 (g x)) i) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ g' (f x)) (-ₖ x1 (g x)) (-ₖ x2 (g x)) (~ i)) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x)) (g' (f x)) (-ₖ x1 (g x)) i +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → (f' (f x) +ₖ commₖ (g' (f x)) (-ₖ x1 (g x)) i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x)) (-ₖ x1 (g x)) (g' (f x)) (~ i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ (-ₖ x1 (g x))) (g' (f x)) (-ₖ x2 (g x)) i) ∣₀) + + dHomHelperPath : (n : ℕ) (h l : C → coHomK n) (a : C) → I → I → coHomK (suc n) + dHomHelperPath n h l a i j = + hcomp (λ k → λ { (i = i0) → lUnitₖ 0ₖ (~ j) + ; (i = i1) → lUnitₖ 0ₖ (~ j) + ; (j = i0) → +ₖ→∙ n (h a) (l a) (~ k) i + ; (j = i1) → cong₂Funct (λ x y → x +ₖ y) (Kn→ΩKn+1 n (h a)) (Kn→ΩKn+1 n (l a)) (~ k) i}) + (bottom i j) + where + bottom : I → I → coHomK (suc n) + bottom i j = hcomp (λ k → λ { (i = i0) → lUnitₖ 0ₖ (~ j) + ; (i = i1) → cong (λ x → lUnitₖ x (~ j)) (Kn→ΩKn+1 n (l a)) k}) + (anotherbottom i j) + where + anotherbottom : I → I → coHomK (suc n) + anotherbottom i j = hcomp (λ k → λ { (i = i0) → rUnitlUnit0 k (~ j) + ; (i = i1) → rUnitlUnit0 k (~ j) + ; (j = i0) → Kn→ΩKn+1 n (h a) i + ; (j = i1) → cong (λ x → x +ₖ 0ₖ) (Kn→ΩKn+1 n (h a)) i}) + (cong (λ x → rUnitₖ x (~ j)) (Kn→ΩKn+1 n (h a)) i) + + + dHomHelper : (n : ℕ) (h l : C → coHomK n) (x : D) → d-pre n (λ x → h x +ₖ l x) x ≡ d-pre n h x +ₖ d-pre n l x + dHomHelper n h l (inl x) = sym (lUnitₖ 0ₖ) + dHomHelper n h l (inr x) = sym (lUnitₖ 0ₖ) + dHomHelper zero h l (push a i) j = dHomHelperPath zero h l a i j + dHomHelper (suc n) h l (push a i) j = dHomHelperPath (suc n) h l a i j + + dIsHom : (n : ℕ) → isMorph (coHomGr n C) (coHomGr (suc n) D) (d n) + dIsHom zero = sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f g i → ∣ funExt (λ x → dHomHelper zero f g x) i ∣₀ + dIsHom (suc n) = sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f g i → ∣ funExt (λ x → dHomHelper (suc n) f g x) i ∣₀ + + + -- Long Exact Sequence + + + Im-d⊂Ker-i : (n : ℕ) (x : coHom (suc n) D) → isInIm (coHomGr n C) (coHomGr (suc n) D) (d n) x → isInKer (coHomGr (suc n) D) (×coHomGr (suc n) A B) (i (suc n)) x + Im-d⊂Ker-i n = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ a (δ , p) → {!!} diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index c96c5d372..4a8c42667 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -78,10 +78,10 @@ coHomRed+1Equiv (suc n) A i = ∥ coHomRed+1.helpLemma A i {C = ((coHomK (suc n) ----------- Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) -Kn→ΩKn+1 n = Iso.fun (Iso2-Kn-ΩKn+1 n) +Kn→ΩKn+1 n = Iso.fun (Iso3-Kn-ΩKn+1 n) ΩKn+1→Kn : {n : ℕ} → typ (Ω (coHomK-ptd (suc n))) → coHomK n -ΩKn+1→Kn {n = n} = Iso.inv (Iso2-Kn-ΩKn+1 n) +ΩKn+1→Kn {n = n} = Iso.inv (Iso3-Kn-ΩKn+1 n) Kn≃ΩKn+1 : {n : ℕ} → coHomK n ≃ typ (Ω (coHomK-ptd (suc n))) Kn≃ΩKn+1 {n = n} = isoToEquiv (Iso-Kn-ΩKn+1 n) @@ -103,70 +103,96 @@ _+ₖ_ {n = n} x y = ΩKn+1→Kn (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) Kn→ΩKn+10ₖ : (n : ℕ) → Kn→ΩKn+1 n 0ₖ ≡ refl Kn→ΩKn+10ₖ zero = refl -Kn→ΩKn+10ₖ (suc n) = (λ i → cong ∣_∣ (rCancel (merid north) i)) +Kn→ΩKn+10ₖ (suc n) = (λ i → cong ∣_∣ (rCancel (merid north) i)) -- could also use refl for n = 1, but for computational reasons I don't want to expand the definition if not necessary. -rUnitₖ : {n : ℕ} (x : coHomK n) → x +ₖ 0ₖ ≡ x -rUnitₖ {n = zero} x = cong ΩKn+1→Kn (sym (rUnit (Kn→ΩKn+1 zero x))) ∙ - Iso.leftInv (Iso2-Kn-ΩKn+1 zero) x -rUnitₖ {n = suc n} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 (suc n) x ∙ Kn→ΩKn+10ₖ (suc n) i)) ∙ - (λ i → ΩKn+1→Kn (rUnit (Kn→ΩKn+1 (suc n) x) (~ i))) ∙ - Iso.leftInv (Iso2-Kn-ΩKn+1 (suc n)) x ++ₖ→∙ : (n : ℕ) (a b : coHomK n) → Kn→ΩKn+1 n (a +ₖ b) ≡ Kn→ΩKn+1 n a ∙ Kn→ΩKn+1 n b ++ₖ→∙ n a b = Iso.rightInv (Iso3-Kn-ΩKn+1 n) (Kn→ΩKn+1 n a ∙ Kn→ΩKn+1 n b) + +commₖ : {n : ℕ} (x y : coHomK n) → (x +ₖ y) ≡ (y +ₖ x) +commₖ {n = n} x y i = ΩKn+1→Kn (EH-instance (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) i) + where + EH-instance : (p q : typ (Ω ((∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) , ∣ north ∣))) → p ∙ q ≡ q ∙ p + EH-instance = transport (λ i → (p q : K-Id n (~ i)) → p ∙ q ≡ q ∙ p) + λ p q → Eckmann-Hilton 0 p q + where + K-Id : (n : ℕ) → typ (Ω (hLevelTrunc (3 + n) (S₊ (1 + n)) , ∣ north ∣)) ≡ typ ((Ω^ 2) (hLevelTrunc (4 + n) (S₊ (2 + n)) , ∣ north ∣ )) + K-Id n = (λ i → typ (Ω (isoToPath (Iso3-Kn-ΩKn+1 (suc n)) i , hcomp (λ k → λ { (i = i0) → ∣ north ∣ + ; (i = i1) → transportRefl (λ j → ∣ rCancel (merid north) k j ∣) k}) + (transp (λ j → isoToPath (Iso3-Kn-ΩKn+1 (suc n)) (i ∧ j)) (~ i) ∣ north ∣)))) lUnitₖ : {n : ℕ} (x : coHomK n) → 0ₖ +ₖ x ≡ x lUnitₖ {n = zero} x = cong ΩKn+1→Kn (sym (lUnit (Kn→ΩKn+1 zero x))) ∙ - Iso.leftInv (Iso2-Kn-ΩKn+1 zero) x + Iso.leftInv (Iso3-Kn-ΩKn+1 zero) x lUnitₖ {n = suc n} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc n) i ∙ Kn→ΩKn+1 (suc n) x)) ∙ - (λ i → ΩKn+1→Kn (lUnit (Kn→ΩKn+1 (suc n) x) (~ i))) ∙ - Iso.leftInv (Iso2-Kn-ΩKn+1 (suc n)) x + (cong ΩKn+1→Kn (sym (lUnit (Kn→ΩKn+1 (suc n) x)))) ∙ + Iso.leftInv (Iso3-Kn-ΩKn+1 (suc n)) x + +rUnitₖ : {n : ℕ} (x : coHomK n) → x +ₖ 0ₖ ≡ x +rUnitₖ {n = zero} x = cong ΩKn+1→Kn (sym (rUnit (Kn→ΩKn+1 zero x))) ∙ + Iso.leftInv (Iso3-Kn-ΩKn+1 zero) x +rUnitₖ {n = suc n} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 (suc n) x ∙ Kn→ΩKn+10ₖ (suc n) i)) ∙ + (cong ΩKn+1→Kn (sym (rUnit (Kn→ΩKn+1 (suc n) x)))) ∙ + Iso.leftInv (Iso3-Kn-ΩKn+1 (suc n)) x +-- + + +rUnitₖ' : {n : ℕ} (x : coHomK n) → x +ₖ 0ₖ ≡ x +rUnitₖ' {n = n} x = commₖ x 0ₖ ∙ lUnitₖ x + +-- rCancelₖ : {n : ℕ} (x : coHomK n) → x +ₖ (-ₖ x) ≡ 0ₖ +-- rCancelₖ {n = zero} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 zero x ∙ Iso.rightInv (Iso3-Kn-ΩKn+1 zero) (sym (Kn→ΩKn+1 zero x)) i)) ∙ +-- cong ΩKn+1→Kn (rCancel (Kn→ΩKn+1 zero x)) +-- rCancelₖ {n = suc zero} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 1 x ∙ Iso.rightInv (Iso3-Kn-ΩKn+1 1) (sym (Kn→ΩKn+1 1 x)) i)) ∙ +-- cong ΩKn+1→Kn (rCancel (Kn→ΩKn+1 1 x)) +-- rCancelₖ {n = suc (suc n)} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 (2 + n) x ∙ Iso.rightInv (Iso3-Kn-ΩKn+1 (2 + n)) (sym (Kn→ΩKn+1 (2 + n) x)) i)) ∙ +-- cong ΩKn+1→Kn (rCancel (Kn→ΩKn+1 (2 + n) x)) ∙ +-- (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc (suc n)) (~ i))) ∙ +-- Iso.leftInv (Iso3-Kn-ΩKn+1 (suc (suc n))) 0ₖ + +-- lCancelₖ : {n : ℕ} (x : coHomK n) → (-ₖ x) +ₖ x ≡ 0ₖ +-- lCancelₖ {n = zero} x = (λ i → ΩKn+1→Kn (Iso.rightInv (Iso3-Kn-ΩKn+1 zero) (sym (Kn→ΩKn+1 zero x)) i ∙ Kn→ΩKn+1 zero x)) ∙ +-- cong ΩKn+1→Kn (lCancel (Kn→ΩKn+1 zero x)) +-- lCancelₖ {n = suc zero} x = ((λ i → ΩKn+1→Kn (Iso.rightInv (Iso3-Kn-ΩKn+1 1) (sym (Kn→ΩKn+1 1 x)) i ∙ Kn→ΩKn+1 1 x))) ∙ +-- cong ΩKn+1→Kn (lCancel (Kn→ΩKn+1 1 x)) +-- lCancelₖ {n = suc (suc n)} x = (λ i → ΩKn+1→Kn (Iso.rightInv (Iso3-Kn-ΩKn+1 (2 + n)) (sym (Kn→ΩKn+1 (2 + n) x)) i ∙ Kn→ΩKn+1 (2 + n) x)) ∙ +-- cong ΩKn+1→Kn (lCancel (Kn→ΩKn+1 (2 + n) x)) ∙ +-- (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc (suc n)) (~ i))) ∙ +-- Iso.leftInv (Iso3-Kn-ΩKn+1 (suc (suc n))) 0ₖ + rCancelₖ : {n : ℕ} (x : coHomK n) → x +ₖ (-ₖ x) ≡ 0ₖ -rCancelₖ {n = zero} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 zero x ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 zero) (sym (Kn→ΩKn+1 zero x)) i)) ∙ +rCancelₖ {n = zero} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 zero x ∙ Iso.rightInv (Iso3-Kn-ΩKn+1 zero) (sym (Kn→ΩKn+1 zero x)) i)) ∙ cong ΩKn+1→Kn (rCancel (Kn→ΩKn+1 zero x)) -rCancelₖ {n = suc zero} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 1 x ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 1) (sym (Kn→ΩKn+1 1 x)) i)) ∙ - cong ΩKn+1→Kn (rCancel (Kn→ΩKn+1 1 x)) -rCancelₖ {n = suc (suc n)} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 (2 + n) x ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 (2 + n)) (sym (Kn→ΩKn+1 (2 + n) x)) i)) ∙ - cong ΩKn+1→Kn (rCancel (Kn→ΩKn+1 (2 + n) x)) ∙ - (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc (suc n)) (~ i))) ∙ - Iso.leftInv (Iso2-Kn-ΩKn+1 (suc (suc n))) 0ₖ +rCancelₖ {n = suc n} x = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 (1 + n) x ∙ Iso.rightInv (Iso3-Kn-ΩKn+1 (1 + n)) (sym (Kn→ΩKn+1 (1 + n) x)) i)) ∙ + cong ΩKn+1→Kn (rCancel (Kn→ΩKn+1 (1 + n) x)) ∙ + (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc n) (~ i))) ∙ + Iso.leftInv (Iso3-Kn-ΩKn+1 (suc n)) 0ₖ lCancelₖ : {n : ℕ} (x : coHomK n) → (-ₖ x) +ₖ x ≡ 0ₖ -lCancelₖ {n = zero} x = (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 zero) (sym (Kn→ΩKn+1 zero x)) i ∙ Kn→ΩKn+1 zero x)) ∙ +lCancelₖ {n = zero} x = (λ i → ΩKn+1→Kn (Iso.rightInv (Iso3-Kn-ΩKn+1 zero) (sym (Kn→ΩKn+1 zero x)) i ∙ Kn→ΩKn+1 zero x)) ∙ cong ΩKn+1→Kn (lCancel (Kn→ΩKn+1 zero x)) -lCancelₖ {n = suc zero} x = ((λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 1) (sym (Kn→ΩKn+1 1 x)) i ∙ Kn→ΩKn+1 1 x))) ∙ - cong ΩKn+1→Kn (lCancel (Kn→ΩKn+1 1 x)) -lCancelₖ {n = suc (suc n)} x = (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 (2 + n)) (sym (Kn→ΩKn+1 (2 + n) x)) i ∙ Kn→ΩKn+1 (2 + n) x)) ∙ - cong ΩKn+1→Kn (lCancel (Kn→ΩKn+1 (2 + n) x)) ∙ - (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc (suc n)) (~ i))) ∙ - Iso.leftInv (Iso2-Kn-ΩKn+1 (suc (suc n))) 0ₖ +lCancelₖ {n = suc n} x = (λ i → ΩKn+1→Kn (Iso.rightInv (Iso3-Kn-ΩKn+1 (1 + n)) (sym (Kn→ΩKn+1 (1 + n) x)) i ∙ Kn→ΩKn+1 (1 + n) x)) ∙ + cong ΩKn+1→Kn (lCancel (Kn→ΩKn+1 (1 + n) x)) ∙ + (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc n) (~ i))) ∙ + Iso.leftInv (Iso3-Kn-ΩKn+1 (suc n)) 0ₖ assocₖ : {n : ℕ} (x y z : coHomK n) → ((x +ₖ y) +ₖ z) ≡ (x +ₖ (y +ₖ z)) assocₖ {n = n} x y z = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 n (ΩKn+1→Kn (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y)) ∙ Kn→ΩKn+1 n z)) ∙ - (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 n) (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) i ∙ Kn→ΩKn+1 n z)) ∙ + (λ i → ΩKn+1→Kn (Iso.rightInv (Iso3-Kn-ΩKn+1 n) (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) i ∙ Kn→ΩKn+1 n z)) ∙ (λ i → ΩKn+1→Kn (assoc (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) (Kn→ΩKn+1 n z) (~ i))) ∙ - (λ i → ΩKn+1→Kn ((Kn→ΩKn+1 n x) ∙ Iso.rightInv (Iso2-Kn-ΩKn+1 n) ((Kn→ΩKn+1 n y ∙ Kn→ΩKn+1 n z)) (~ i))) + (λ i → ΩKn+1→Kn ((Kn→ΩKn+1 n x) ∙ Iso.rightInv (Iso3-Kn-ΩKn+1 n) ((Kn→ΩKn+1 n y ∙ Kn→ΩKn+1 n z)) (~ i))) -commₖ : {n : ℕ} (x y : coHomK n) → (x +ₖ y) ≡ (y +ₖ x) -commₖ {n = n} x y i = ΩKn+1→Kn (EH-instance (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) i) - where - EH-instance : (p q : typ (Ω ((∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) , ∣ north ∣))) → p ∙ q ≡ q ∙ p - EH-instance = transport (λ i → (p q : K-Id n (~ i)) → p ∙ q ≡ q ∙ p) - λ p q → Eckmann-Hilton 0 p q - where - K-Id : (n : ℕ) → typ (Ω (hLevelTrunc (3 + n) (S₊ (1 + n)) , ∣ north ∣)) ≡ typ ((Ω^ 2) (hLevelTrunc (4 + n) (S₊ (2 + n)) , ∣ north ∣ )) - K-Id n = (λ i → typ (Ω (isoToPath (Iso2-Kn-ΩKn+1 (suc n)) i , hcomp (λ k → λ { (i = i0) → ∣ north ∣ - ; (i = i1) → transportRefl (λ j → ∣ rCancel (merid north) k j ∣) k}) - (transp (λ j → isoToPath (Iso2-Kn-ΩKn+1 (suc n)) (i ∧ j)) (~ i) ∣ north ∣)))) -distrₖ : {n : ℕ} → (x y : coHomK n) → -ₖ (x +ₖ y) ≡ (-ₖ x) +ₖ (-ₖ y) -distrₖ {n = n} x y = (λ i → ΩKn+1→Kn (sym (Kn→ΩKn+1 n (ΩKn+1→Kn (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y))))) ∙ - (λ i → ΩKn+1→Kn (sym (Iso.rightInv (Iso2-Kn-ΩKn+1 n) (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) i))) ∙ + (λ i → ΩKn+1→Kn (sym (Iso.rightInv (Iso3-Kn-ΩKn+1 n) (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) i))) ∙ (λ i → ΩKn+1→Kn (symDistr (Kn→ΩKn+1 n x) (Kn→ΩKn+1 n y) i)) ∙ - (λ i → ΩKn+1→Kn (Iso.rightInv (Iso2-Kn-ΩKn+1 n) (sym (Kn→ΩKn+1 n y)) (~ i) ∙ (Iso.rightInv (Iso2-Kn-ΩKn+1 n) (sym (Kn→ΩKn+1 n x)) (~ i)))) ∙ + (λ i → ΩKn+1→Kn (Iso.rightInv (Iso3-Kn-ΩKn+1 n) (sym (Kn→ΩKn+1 n y)) (~ i) ∙ (Iso.rightInv (Iso3-Kn-ΩKn+1 n) (sym (Kn→ΩKn+1 n x)) (~ i)))) ∙ commₖ (-ₖ y) (-ₖ x) ----- Group structure of cohomology groups --- +-- Group structure of cohomology groups --- private § : isSet ∥ A ∥₀ @@ -207,20 +233,92 @@ commₕ {n = n} = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) +rUnitlUnit0 : {n : ℕ} → rUnitₖ {n = n} 0ₖ ≡ lUnitₖ 0ₖ +rUnitlUnit0 {n = zero} = refl +rUnitlUnit0 {n = suc n} = + (assoc (λ i → ΩKn+1→Kn (Kn→ΩKn+1 (suc n) 0ₖ ∙ Kn→ΩKn+10ₖ (suc n) i)) + (cong ΩKn+1→Kn (sym (rUnit (Kn→ΩKn+1 (suc n) 0ₖ)))) + (Iso.leftInv (Iso3-Kn-ΩKn+1 (suc n)) 0ₖ)) + ∙ (λ j → helper j + ∙ Iso.leftInv (Iso3-Kn-ΩKn+1 (suc n)) 0ₖ) + ∙ sym (assoc (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc n) i ∙ Kn→ΩKn+1 (suc n) 0ₖ)) (cong ΩKn+1→Kn (sym (lUnit (Kn→ΩKn+1 (suc n) 0ₖ)))) (Iso.leftInv (Iso3-Kn-ΩKn+1 (suc n)) 0ₖ)) + + where + helper : (λ i → ΩKn+1→Kn (Kn→ΩKn+1 (suc n) 0ₖ ∙ Kn→ΩKn+10ₖ (suc n) i)) + ∙ (cong ΩKn+1→Kn (sym (rUnit (Kn→ΩKn+1 (suc n) 0ₖ)))) + ≡ (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc n) i ∙ Kn→ΩKn+1 (suc n) 0ₖ)) + ∙ (cong ΩKn+1→Kn (sym (lUnit (Kn→ΩKn+1 (suc n) 0ₖ)))) + helper = (λ j → lUnit (rUnit ((λ i → ΩKn+1→Kn (Kn→ΩKn+1 (suc n) 0ₖ ∙ Kn→ΩKn+10ₖ (suc n) i))) j) j + ∙ rUnit (cong ΩKn+1→Kn (sym (rUnit (Kn→ΩKn+1 (suc n) 0ₖ)))) j) + ∙ (λ j → ((λ z → ΩKn+1→Kn (Kn→ΩKn+1 (suc n) ∣ north ∣ ∙ (λ i → ∣ rCancel (merid north) (z ∧ j) i ∣))) + ∙ (λ i → ΩKn+1→Kn (Kn→ΩKn+1 (suc n) 0ₖ ∙ Kn→ΩKn+10ₖ (suc n) (i ∨ j))) + ∙ λ z → ΩKn+1→Kn (((λ i → ∣ rCancel (merid north) (z ∧ j) i ∣)) ∙ refl)) + ∙ cong ΩKn+1→Kn (sym (rUnit (Kn→ΩKn+10ₖ (suc n) j))) + ∙ λ z → ΩKn+1→Kn (λ i → ∣ rCancel (merid north) ((~ z) ∧ j) i ∣)) + ∙ (λ j → (((λ z → ΩKn+1→Kn (Kn→ΩKn+1 (suc n) ∣ north ∣ ∙ (λ i → ∣ rCancel (merid north) z i ∣)))) + ∙ (λ i → ΩKn+1→Kn (Kn→ΩKn+1 (suc n) 0ₖ ∙ refl)) + ∙ λ z → ΩKn+1→Kn ((λ i → ∣ rCancel (merid north) z i ∣) ∙ λ i → ∣ rCancel (merid north) ((~ z) ∨ (~ j)) i ∣)) + ∙ cong ΩKn+1→Kn (sym (lUnit (Kn→ΩKn+10ₖ (suc n) (~ j)))) + ∙ λ z → ΩKn+1→Kn (λ i → ∣ rCancel (merid north) (~ z ∧ (~ j)) i ∣)) + ∙ (λ j → ((λ z → ΩKn+1→Kn (Kn→ΩKn+1 (suc n) ∣ north ∣ ∙ λ i → ∣ rCancel (merid north) (z ∧ (~ j)) i ∣)) + ∙ (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc n) (i ∧ j) ∙ Kn→ΩKn+10ₖ (suc n) (~ j))) + ∙ λ z → ΩKn+1→Kn ((λ i → ∣ rCancel (merid north) (z ∨ j) i ∣) ∙ λ i → ∣ rCancel (merid north) (~ z ∧ ~ j) i ∣)) + ∙ rUnit (cong ΩKn+1→Kn (sym (lUnit (Kn→ΩKn+1 (suc n) 0ₖ)))) (~ j)) + ∙ (λ j → lUnit (rUnit (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc n) i ∙ Kn→ΩKn+1 (suc n) 0ₖ)) (~ j)) (~ j) + ∙ cong ΩKn+1→Kn (sym (lUnit (Kn→ΩKn+1 (suc n) 0ₖ)))) + + + ---- Group structure of reduced cohomology groups (in progress - might need K to compute properly first) --- +ₕ∙ : {A : Pointed ℓ} (n : ℕ) → coHomRed n A → coHomRed n A → coHomRed n A +ₕ∙ zero = sElim2 (λ _ _ → §) λ { (a , pa) (b , pb) → ∣ (λ x → a x +ₖ b x) , (λ i → (pa i +ₖ pb i)) ∣₀ } -- ∣ (λ x → a x +ₖ b x) ∣₀ +ₕ∙ (suc n) = sElim2 (λ _ _ → §) λ { (a , pa) (b , pb) → ∣ (λ x → a x +ₖ b x) , (λ i → pa i +ₖ pb i) ∙ lUnitₖ 0ₖ ∣₀ } +-- -ₕ∙ : {A : Pointed ℓ} (n : ℕ) → coHomRed n A → coHomRed n A +-- -ₕ∙ zero = sRec § λ {(a , pt) → ∣ (λ x → -ₖ a x ) , (λ i → -ₖ (pt i)) ∣₀} +-- -ₕ∙ (suc zero) = sRec § λ {(a , pt) → ∣ (λ x → -ₖ a x ) , (λ i → -ₖ (pt i)) ∙ (λ i → ΩKn+1→Kn (sym (Kn→ΩKn+10ₖ (suc zero) i))) ∣₀} +-- -ₕ∙ (suc (suc n)) = sRec § λ {(a , pt) → ∣ (λ x → -ₖ a x ) , (λ i → -ₖ (pt i)) ∙ +-- (λ i → ΩKn+1→Kn (sym (Kn→ΩKn+10ₖ (suc (suc n)) i))) ∙ +-- (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc (suc n)) (~ i))) ∙ +-- Iso.leftInv (Iso3-Kn-ΩKn+1 (suc (suc n))) ∣ north ∣ ∣₀} + -ₕ∙ : {A : Pointed ℓ} (n : ℕ) → coHomRed n A → coHomRed n A -ₕ∙ zero = sRec § λ {(a , pt) → ∣ (λ x → -ₖ a x ) , (λ i → -ₖ (pt i)) ∣₀} --ₕ∙ (suc zero) = sRec § λ {(a , pt) → ∣ (λ x → -ₖ a x ) , (λ i → -ₖ (pt i)) ∙ (λ i → ΩKn+1→Kn (sym (Kn→ΩKn+10ₖ (suc zero) i))) ∣₀} --ₕ∙ (suc (suc n)) = sRec § λ {(a , pt) → ∣ (λ x → -ₖ a x ) , (λ i → -ₖ (pt i)) ∙ - (λ i → ΩKn+1→Kn (sym (Kn→ΩKn+10ₖ (suc (suc n)) i))) ∙ - (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc (suc n)) (~ i))) ∙ - Iso.leftInv (Iso2-Kn-ΩKn+1 (suc (suc n))) ∣ north ∣ ∣₀} +-ₕ∙ (suc n) = sRec § λ {(a , pt) → ∣ (λ x → -ₖ a x ) , (λ i → -ₖ (pt i)) ∙ + (λ i → ΩKn+1→Kn (sym (Kn→ΩKn+10ₖ (suc n) i))) ∙ + (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ (suc n) (~ i))) ∙ + Iso.leftInv (Iso3-Kn-ΩKn+1 (suc n)) ∣ north ∣ ∣₀} 0ₕ∙ : {A : Pointed ℓ} (n : ℕ) → coHomRed n A 0ₕ∙ zero = ∣ (λ _ → 0ₖ) , refl ∣₀ 0ₕ∙ (suc n) = ∣ (λ _ → 0ₖ) , refl ∣₀ + + + +module _ {A : Type ℓ} {A' : Pointed ℓ'} (n : ℕ) where + 0prod : coHomRed n A' × coHom n A + 0prod = (0ₕ∙ n) , 0ₕ + + -prod : coHomRed n A' × coHom n A → coHomRed n A' × coHom n A + -prod (a , b) = (-ₕ∙ n a) , (-ₕ b) + + +prod : coHomRed n A' × coHom n A → coHomRed n A' × coHom n A → coHomRed n A' × coHom n A + +prod (a , b) (c , d) = (+ₕ∙ n a c) , (b +ₕ d) + + + +coHomGr : ∀ {ℓ} (n : ℕ) (A : Type ℓ) → Group ℓ +Group.type (coHomGr n A) = coHom n A +Group.setStruc (coHomGr n A) = § +isGroup.id (Group.groupStruc (coHomGr n A)) = 0ₕ +isGroup.inv (Group.groupStruc (coHomGr n A)) = -ₕ +isGroup.comp (Group.groupStruc (coHomGr n A)) = _+ₕ_ +isGroup.lUnit (Group.groupStruc (coHomGr n A)) = lUnitₕ +isGroup.rUnit (Group.groupStruc (coHomGr n A)) = rUnitₕ +isGroup.assoc (Group.groupStruc (coHomGr n A)) = assocₕ +isGroup.lCancel (Group.groupStruc (coHomGr n A)) = lCancelₕ +isGroup.rCancel (Group.groupStruc (coHomGr n A)) = rCancelₕ + +×coHomGr : (n : ℕ) (A : Type ℓ) (B : Type ℓ') → Group (ℓ-max ℓ ℓ') +×coHomGr n A B = dirProd (coHomGr n A) (coHomGr n B) diff --git a/Cubical/ZCohomology/altToFreudenthal.agda b/Cubical/ZCohomology/altToFreudenthal.agda new file mode 100644 index 000000000..9a5d4e492 --- /dev/null +++ b/Cubical/ZCohomology/altToFreudenthal.agda @@ -0,0 +1,201 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.altToFreudenthal where + +open import Cubical.ZCohomology.Base +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence +open import Cubical.Data.NatMinusTwo.Base +open import Cubical.Data.Empty +open import Cubical.Data.Sigma hiding (_×_) +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.Nullification +open import Cubical.Data.Int hiding (_+_) +open import Cubical.Data.Nat +open import Cubical.Data.Prod +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Freudenthal +open import Cubical.HITs.SmashProduct.Base +open import Cubical.Data.Group.Base hiding (_≃_ ; Iso) + +private + variable + ℓ ℓ' : Level + A : Type ℓ + B : Type ℓ' + A' : Pointed ℓ + B' : Pointed ℓ' + + +test : ∀ n → hLevelTrunc (3 + n) (S₊ (1 + n)) → hLevelTrunc (3 + n) (Path (S₊ (2 + n)) north south) +test n = trRec (isOfHLevelTrunc (3 + n)) λ x → ∣ merid x ∣ + +isEquivMerid : ∀ n → isHLevelConnectedFun (3 + n) (test n) +isEquivMerid n = trElim {!!} λ a → isHLevelConnectedPoint2 _ (∣ {!!} ∣ , {!a!}) {!!} + where + ou : (b : Path (S₊ (suc (suc n))) north south) → Iso (fiber (test n) ∣ b ∣) (fiber {A = Unit} {B = hLevelTrunc (suc (suc (suc n))) (Path (S₊ (suc (suc n))) north south)} (λ {tt → ∣ b ∣}) ∣ b ∣) + Iso.fun (ou b) (p , q) = tt , refl + Iso.inv (ou b) (p , q) = ∣ {!!} ∣ , {!!} + Iso.rightInv (ou b) = {!!} + Iso.leftInv (ou b) = {!!} + where + king : (b : Path (S₊ (suc (suc n))) north south) → Iso (fiber (test n) ∣ b ∣) ( hLevelTrunc (suc (suc (suc n))) (S₊ (2 + n))) + Iso.fun (king b) (p , q) = trRec {A = (S₊ (suc n))} (isOfHLevelTrunc (3 + n)) (λ p → ∣ fun p ∣) p + where + fun : S₊ (suc n) → S₊ (suc (suc n)) + fun north = north + fun south = south + fun (merid a i) = merid (merid a i) i + Iso.inv (king b) = {!!} + Iso.rightInv (king b) = {!merid ? ?!} + Iso.leftInv (king b) = {!!} + +-- id : (n : ℕ) → typ ((Ω^ (1 + n)) A') +-- id zero = refl +-- id (suc n) = λ _ → refl + +-- indEquiv : (n : ℕ) → ((x : S₊ (suc n)) → typ (Ω ((Ω^ n) (S₊ (suc n) , x)))) → typ (Ω (Ω ((Ω^ n) (Type₀ , S₊ (suc n))))) +-- indEquiv zero f = {!!} +-- indEquiv (suc n) f = {!!} + + +-- funcTypeEq : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → Iso (typ ((Ω^ (2 + n)) (Type₀ , S₊ (1 + n)))) ((x : S₊ (1 + n)) → typ ((Ω^ (1 + n)) (S₊ (1 + n) , x))) +-- Iso.fun (funcTypeEq zero) f x = {!transport !} +-- Iso.fun (funcTypeEq (suc zero)) f x i = {!cong !} +-- Iso.fun (funcTypeEq (suc (suc n))) f x i = {!!} +-- Iso.inv (funcTypeEq n) f i = {!f north !} +-- Iso.rightInv (funcTypeEq n) = {!!} +-- Iso.leftInv (funcTypeEq n) = {!!} + +-- idToIso : ∀ {ℓ} {A : Type ℓ} (p : A ≡ A) → Iso A A +-- Iso.fun (idToIso {A = A} p) = transport p +-- Iso.inv (idToIso {A = A} p) = transport (sym p) +-- Iso.rightInv (idToIso {A = A} p) = transportTransport⁻ p +-- Iso.leftInv (idToIso {A = A} p) = transport⁻Transport p + +-- test : ∀ {ℓ} {A : Type ℓ} (p : A ≡ A) → (typ (Ω ((A ≡ A) , p))) ≃ (typ (Ω ((A ≃ A) , pathToEquiv p))) +-- test {ℓ = ℓ} {A = A} p = transport (λ i → helper2 (~ i) ≃ typ (Ω ((A ≃ A) , (pathToEquiv p)))) (compEquiv (invEquiv helper3) (pathToEquiv λ i → typ (Ω (_ , transportRefl (pathToEquiv p) i)))) +-- where +-- helper2 : typ (Ω ((A ≡ A) , p)) ≡ typ (Ω (Lift (A ≃ A) , transport (univalencePath {A = A} {B = A}) p )) +-- helper2 i = typ (Ω ((univalencePath {A = A} {B = A} i) , transp (λ j → univalencePath {A = A} {B = A} (j ∧ i)) (~ i) p)) +-- helper3 : typ (Ω ((A ≃ A) , _)) +-- ≃ typ (Ω (Lift (A ≃ A) , transport (univalencePath {A = A} {B = A}) p)) + +-- helper3 = congEquiv LiftEquiv + +-- test2 : ∀ {ℓ} {A : Type ℓ} (p : A ≡ A) → typ (Ω ((A ≃ A) , pathToEquiv p)) ≃ (typ (Ω ((A → A) , transport p))) +-- test2 {A = A} p = EquivCat +-- where +-- EquivCat : (pathToEquiv p ≡ pathToEquiv p) ≃ (pathToEquiv p .fst ≡ pathToEquiv p .fst) +-- EquivCat = compEquiv (invEquiv Σ≡) helper -- {!sym ua Σ≡!} ∙ {!!} ∙ {!!} +-- where +-- helper : Σ (fst (pathToEquiv p) ≡ fst (pathToEquiv p)) (λ p₁ → PathP (λ i → isEquiv (p₁ i)) (snd (pathToEquiv p)) (snd (pathToEquiv p))) ≃ (pathToEquiv p .fst ≡ pathToEquiv p .fst) +-- helper = helper2 λ x → toPathP (isPropIsEquiv _ _ _) , λ y → {!!} +-- where +-- helper2 : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} → ((x : A) → isContr (B x)) → Σ A B ≃ A +-- helper2 contr = isoToEquiv (iso (λ { (a , p) → a}) (λ a → a , (contr a .fst)) (λ b → refl) λ { (a , p) i → a , (contr a .snd p) i }) + +-- withJ : (p q : A ≃ B) → Iso (Path (A ≃ B) p q) (Path (A → B) (p .fst) (q .fst)) +-- Iso.fun (withJ {A = A} {B = B} p q) = cong fst +-- Iso.inv (withJ {A = A} p q) path = J (λ f r → (x : isEquiv f) → p ≡ (f , x)) (λ x i → (path i) , {!!}) refl (q .snd) +-- Iso.rightInv (withJ p q) a = {!a!} +-- Iso.leftInv (withJ p q) = {!!} + +-- together : ∀ {ℓ} (n : ℕ) {A : Type ℓ} (p : A ≡ A) → (typ ((Ω^ (1 + n)) ((A ≡ A) , p))) ≃ (typ ((Ω^ (1 + n)) ((A → A) , transport p))) +-- toghelper : ∀ {ℓ} (n : ℕ) {A : Type ℓ} (p : A ≡ A) → together n p .fst (snd (Ω ((Ω^ n) ((A ≡ A) , p)))) ≡ snd (Ω ((Ω^ n) ((A → A) , transport p))) +-- toghelper zero {A = A} p = {!!} + +-- toghelper (suc n) {A = A} p = {!refl!} + +-- together zero {A = A} p = helper2 -- (compEquiv (test p) (test2 p)) +-- where +-- helper : ∀ B p → typ (Ω ((A ≃ B) , univalence .fst p)) ≃ typ (Ω ((A → B) , transport p)) +-- helper B p = isoToEquiv (iso (cong fst) {!!} {!!} {!!}) -- isoToEquiv (iso (cong fst) {!univalence .fst p!} {!!} {!!}) +-- helper2 : typ (Ω ((A ≡ A) , p)) ≃ typ (Ω ((A → A) , transport p)) +-- helper2 = compEquiv (congEquiv univalence) {!!} +-- together (suc n) {A = A} p = isoToEquiv helper +-- where +-- helper : Iso (typ (Ω (Ω ((Ω^ n) ((A ≡ A) , p))))) (typ (Ω (Ω ((Ω^ n) ((A → A) , transport p))))) +-- Iso.fun helper a = sym (toghelper n p) ∙ cong (together n p .fst) a ∙ toghelper n p +-- Iso.inv helper a = {!? ∙ ? ∙ ?!} +-- Iso.rightInv helper = {!!} +-- Iso.leftInv helper = {!!} + +-- test6 : Iso (A ≡ A) (A ≃ A) +-- Iso.fun test6 p = transportEquiv p +-- Iso.inv test6 p = ua p +-- Iso.rightInv test6 b i = funExt (uaβ b) i , hcomp {!!} (transp (λ j → isEquiv (funExt (λ x → transportRefl (b .fst x)) (i ∧ j))) (~ i) (transp (λ i₁ → isEquiv (transp (λ j → ua b (i₁ ∧ j)) (~ i₁))) i0 (idIsEquiv A))) +-- Iso.leftInv test6 b = {! .fst!} + + +-- -- S∙ : (n : ℕ) → Pointed₀ +-- -- S∙ n = (S₊ n) , north + +-- -- ΩS : (n : ℕ) → Type₀ +-- -- ΩS n = typ (Ω' (S₊ n , north)) + +-- -- promote : (n : ℕ) → S₊ n → ΩS (1 + n) +-- -- promote n north = merid north ∙ sym (merid north) +-- -- promote n south = merid south ∙ sym (merid north) +-- -- promote n (merid a i) = merid (merid a i) ∙ sym (merid north) + +-- -- decode' : (n : ℕ) → hLevelTrunc (2 + n) (S₊ n) → hLevelTrunc (2 + n) (ΩS (1 + n)) +-- -- decode' n = trRec (isOfHLevelTrunc (2 + n)) +-- -- λ a → ∣ promote n a ∣ + + +-- -- Codes : (n : ℕ) → S₊ (suc n) → Type₀ +-- -- Codes n north = hLevelTrunc (3 + n) (S₊ (1 + n)) +-- -- Codes n south = hLevelTrunc (3 + n) (S₊ (1 + n)) +-- -- Codes n (merid a i) = {!!} + +-- -- SEquiv1 : (m : ℕ) → Iso (ℕ → S₊ (suc m)) (S₊ (suc m) → S₊ (suc m)) +-- -- Iso.fun (SEquiv1 m) f north = north +-- -- Iso.fun (SEquiv1 m) f south = south +-- -- Iso.fun (SEquiv1 m) f (merid a i) = {!!} +-- -- Iso.inv (SEquiv1 m) f k = {!!} +-- -- Iso.rightInv (SEquiv1 m) = {!!} +-- -- Iso.leftInv (SEquiv1 m) = {!!} + + + + + + + +-- -- -- SEquiv1 : (n : ℕ) → S₊ (suc n) → S₊ (suc n) → Iso (S₊ (suc n)) (S₊ (suc n)) +-- -- -- Iso.fun (SEquiv1 n north north) north = north +-- -- -- Iso.fun (SEquiv1 n north north) south = north +-- -- -- Iso.fun (SEquiv1 n north north) (merid a i) = (merid a ∙ sym (merid north)) i +-- -- -- Iso.fun (SEquiv1 n north south) north = north +-- -- -- Iso.fun (SEquiv1 n north south) south = south +-- -- -- Iso.fun (SEquiv1 n north south) (merid a i) = merid a i +-- -- -- Iso.fun (SEquiv1 n north (merid a i)) = {!!} +-- -- -- Iso.fun (SEquiv1 n south t) = {!!} +-- -- -- Iso.fun (SEquiv1 n (merid a i) t) = {!!} +-- -- -- Iso.inv (SEquiv1 n s t) = {!!} +-- -- -- Iso.rightInv (SEquiv1 n s t) = {!!} +-- -- -- Iso.leftInv (SEquiv1 n s t) = {!!} + +-- -- -- SEquiv : (n : ℕ) → S₊ (suc n) ≃ S₊ (suc n) +-- -- -- SEquiv n = isoToEquiv {!!} +-- -- -- where +-- -- -- isot : Iso (S₊ (suc n)) (S₊ (suc n)) +-- -- -- Iso.fun isot north = {!!} +-- -- -- Iso.fun isot south = {!!} +-- -- -- Iso.fun isot (merid a i) = {!!} +-- -- -- Iso.inv isot = {!!} +-- -- -- Iso.rightInv isot = {!!} +-- -- -- Iso.leftInv isot = {!!} From 5dcd3c8fca7128c5de78dcb4185a3369b28692ff Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 12 May 2020 18:47:00 +0200 Subject: [PATCH 32/89] MV-almostDone --- Cubical/Data/Group/Base.agda | 55 +++- Cubical/Foundations/Isomorphism.agda | 3 + Cubical/HITs/SetTruncation/Properties.agda | 22 +- Cubical/HITs/Sn/Properties.agda | 241 +++++++++++++++++- Cubical/HITs/Truncation/Properties.agda | 72 +++--- Cubical/Homotopy/Connected.agda | 238 +++++++++++++++++ Cubical/ZCohomology/MayerVietorisReduced.agda | 219 +++++++++++++--- Cubical/ZCohomology/Properties.agda | 25 +- 8 files changed, 789 insertions(+), 86 deletions(-) diff --git a/Cubical/Data/Group/Base.agda b/Cubical/Data/Group/Base.agda index bdb7b030d..9c6f08f0b 100644 --- a/Cubical/Data/Group/Base.agda +++ b/Cubical/Data/Group/Base.agda @@ -2,9 +2,10 @@ module Cubical.Data.Group.Base where -open import Cubical.Foundations.Prelude hiding ( comp ) +open import Cubical.Foundations.Prelude hiding (comp) open import Cubical.Foundations.HLevels open import Cubical.Data.Prod +open import Cubical.HITs.PropositionalTruncation hiding (map) import Cubical.Foundations.Isomorphism as I import Cubical.Foundations.Equiv as E @@ -37,6 +38,46 @@ isMorph (group G Gset (group-struct _ _ _⊙_ _ _ _ _ _)) morph : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → Type (ℓ-max ℓ ℓ') morph G H = Σ (Group.type G → Group.type H) (isMorph G H) +morph0→0 : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (f : (Group.type G → Group.type H)) + → isMorph G H f + → f (isGroup.id (Group.groupStruc G)) ≡ isGroup.id (Group.groupStruc H) +morph0→0 (group G strucG (group-struct idG invG compG lUnitG rUnitG assocG lCancelG rCancelG)) + (group H strucH (group-struct idH invH compH lUnitH rUnitH assocH lCancelH rCancelH)) f morph = + f idG ≡⟨ sym (rUnitH (f idG)) ⟩ + compH (f idG) idH ≡⟨ (λ i → compH (f idG) (rCancelH (f idG) (~ i))) ⟩ + compH (f idG) (compH (f idG) (invH (f idG))) ≡⟨ sym (assocH (f idG) (f idG) (invH (f idG))) ⟩ + compH (compH (f idG) (f idG)) (invH (f idG)) ≡⟨ sym (cong (λ x → compH x (invH (f idG))) (sym (cong f (lUnitG idG)) ∙ morph idG idG)) ⟩ + compH (f idG) (invH (f idG)) ≡⟨ rCancelH (f idG) ⟩ + idH ∎ + +morphMinus : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (f : (Group.type G → Group.type H)) + → isMorph G H f + → (g : Group.type G) → f (isGroup.inv (Group.groupStruc G) g) ≡ isGroup.inv (Group.groupStruc H) (f g) +morphMinus G H f morph g = + let idG = isGroup.id (Group.groupStruc G) + idH = isGroup.id (Group.groupStruc H) + invG = isGroup.inv (Group.groupStruc G) + invH = isGroup.inv (Group.groupStruc H) + lCancelG = isGroup.lCancel (Group.groupStruc G) + rCancelH = isGroup.rCancel (Group.groupStruc H) + lUnitH = isGroup.lUnit (Group.groupStruc H) + rUnitH = isGroup.rUnit (Group.groupStruc H) + assocH = isGroup.assoc (Group.groupStruc H) + compG = isGroup.comp (Group.groupStruc G) + compH = isGroup.comp (Group.groupStruc H) + helper : compH (f (invG g)) (f g) ≡ idH + helper = sym (morph (invG g) g) ∙ (λ i → f (lCancelG g i)) ∙ morph0→0 G H f morph + in f (invG g) ≡⟨ sym (rUnitH (f (invG g))) ⟩ + compH (f (invG g)) idH ≡⟨ (λ i → compH (f (invG g)) (rCancelH (f g) (~ i))) ⟩ + compH (f (invG g)) (compH (f g) (invH (f g))) ≡⟨ sym (assocH (f (invG g)) (f g) (invH (f g))) ⟩ + compH (compH (f (invG g)) (f g)) (invH (f g)) ≡⟨ cong (λ x → compH x (invH (f g))) helper ⟩ + compH idH (invH (f g)) ≡⟨ lUnitH (invH (f g)) ⟩ + invH (f g) ∎ + + + +-- sym (rUnitH (f (invG g))) ∙ (λ i → compH (f (invG g)) (rCancelH (f g) (~ i))) ∙ sym (assocH (f (invG g)) (f g) (invH (f g))) ∙ cong (λ x → compH x (invH (f g))) helper ∙ lUnitH (invH (f g)) + record Iso {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') : Type (ℓ-max ℓ ℓ') where constructor iso field @@ -164,19 +205,9 @@ compIso {ℓ} {F} {G} {H} f ∎ -Im : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (Group.type G → Group.type H) - → Type (ℓ-max ℓ ℓ') -Im (group A setA strucA) (group B setB strucB) f = - Σ[ b ∈ B ] Σ[ a ∈ A ] f a ≡ b - isInIm : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (Group.type G → Group.type H) → Group.type H → Type _ -isInIm G H ϕ h = Σ[ g ∈ Group.type G ] ϕ g ≡ h - -Ker : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (f : (Group.type G → Group.type H)) - → Type (ℓ-max ℓ ℓ') -Ker (group A setA strA) (group B setB strB) f = - Σ[ a ∈ A ] f a ≡ isGroup.id strB +isInIm G H ϕ h = ∃[ g ∈ Group.type G ] ϕ g ≡ h isInKer : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (Group.type G → Group.type H) → Group.type G → Type _ diff --git a/Cubical/Foundations/Isomorphism.agda b/Cubical/Foundations/Isomorphism.agda index 61d600f4b..c4aab5bf8 100644 --- a/Cubical/Foundations/Isomorphism.agda +++ b/Cubical/Foundations/Isomorphism.agda @@ -42,6 +42,9 @@ record Iso {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') : Type (ℓ-max ℓ ℓ') w isIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → (A → B) → Type _ isIso {A = A} {B = B} f = Σ[ g ∈ (B → A) ] Σ[ _ ∈ section f g ] retract f g +symIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → Iso A B → Iso B A +symIso (iso f g rinv linv) = iso g f linv rinv + -- Any iso is an equivalence module _ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (i : Iso A B) where open Iso i renaming ( fun to f diff --git a/Cubical/HITs/SetTruncation/Properties.agda b/Cubical/HITs/SetTruncation/Properties.agda index f700075fa..fdf9aabba 100644 --- a/Cubical/HITs/SetTruncation/Properties.agda +++ b/Cubical/HITs/SetTruncation/Properties.agda @@ -16,7 +16,8 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv open import Cubical.Foundations.HLevels open import Cubical.Foundations.Univalence -open import Cubical.Data.Sigma +open import Cubical.Data.Sigma hiding (_×_) +open import Cubical.Data.Prod private variable @@ -92,3 +93,22 @@ setSigmaIso {A = A} {B = B} = iso fun funinv sect retr retr : retract fun funinv retr = elim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ { (a , p) → refl } + +sigmaElim : ∀ {ℓ ℓ'} {B : ∥ A ∥₀ → Type ℓ} {C : Σ ∥ A ∥₀ B → Type ℓ'} + (Bset : (x : Σ ∥ A ∥₀ B) → isSet (C x)) + (g : (a : A) (b : B ∣ a ∣₀) → C (∣ a ∣₀ , b)) + (x : Σ ∥ A ∥₀ B) → C x +sigmaElim {B = B} {C = C} set g (x , y) = elim {B = λ x → (y : B x) → C (x , y)} + (λ _ → isOfHLevelΠ 2 λ _ → set (_ , _)) + g x y + +sigmaProdElim : ∀ {ℓ ℓ' ℓ''} {B : Type ℓ} {C : ∥ A ∥₀ × ∥ B ∥₀ → Type ℓ'} {D : Σ (∥ A ∥₀ × ∥ B ∥₀) C → Type ℓ''} + (Bset : (x : Σ (∥ A ∥₀ × ∥ B ∥₀) C) → isSet (D x)) + (g : (a : A) (b : B) (c : C (∣ a ∣₀ , ∣ b ∣₀)) → D ((∣ a ∣₀ , ∣ b ∣₀) , c)) + (x : Σ (∥ A ∥₀ × ∥ B ∥₀) C) → D x +sigmaProdElim {B = B} {C = C} {D = D} set g ((x , y) , c) = + elim {B = λ x → (y : ∥ B ∥₀) (c : C (x , y)) → D ((x , y) , c)} + (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelΠ 2 λ _ → set _) + (λ x → elim (λ _ → isOfHLevelΠ 2 λ _ → set _) + λ y c → g x y c) + x y c diff --git a/Cubical/HITs/Sn/Properties.agda b/Cubical/HITs/Sn/Properties.agda index a8d360420..3cf4b4e97 100644 --- a/Cubical/HITs/Sn/Properties.agda +++ b/Cubical/HITs/Sn/Properties.agda @@ -19,6 +19,7 @@ open import Cubical.HITs.Susp open import Cubical.Data.Unit open import Cubical.HITs.Join open import Cubical.HITs.Pushout +open import Cubical.HITs.Wedge open import Cubical.HITs.SmashProduct private @@ -27,8 +28,6 @@ private A : Type ℓ B : Type ℓ' - - --- Some silly lemmas on S1 --- S¹≡S1 : S₊ 1 ≡ S¹ @@ -67,7 +66,7 @@ SuspBool≃S1 = isoToEquiv iso1 Iso.leftInv iso1 (merid true i) = refl --- map between S¹ ∧ A and Susp A. +-- map between S₊ private f' : {a : A} → A → S₊ 1 → Susp A f' {a = pt} A north = north @@ -93,9 +92,6 @@ module smashS1→susp {(A , pt) : Pointed ℓ} where fun⁻ (merid a i) = (push (inr a) ∙∙ cong (λ x → proj' {A = S₊ 1 , north} {B = A , pt} x a) (merid south ∙ sym (merid north)) ∙∙ sym (push (inr a))) i - {- TODO : Prove that they cancel out -} - -{- Map used in definition of cup product. Maybe mover there later -} sphereSmashMap : (n m : ℕ) → (S₊ (suc n) , north) ⋀ (S₊ (suc m) , north) → S₊ (2 + n + m) sphereSmashMap zero m = smashS1→susp.fun sphereSmashMap (suc n) m = @@ -103,3 +99,236 @@ sphereSmashMap (suc n) m = (idfun∙ _ ⋀⃗ (sphereSmashMap n m , refl)) ∘ ⋀-associate ∘ ((smashS1→susp.fun⁻ , refl) ⋀⃗ idfun∙ _) + +private + f* : {a : A} → S¹ → A → Susp A + f* base b = north + f* {a = a} (loop i) = funExt (λ x → merid x ∙ sym (merid a)) i + + f'' : {a : A} → Susp Bool → A → Susp A + f'' north a = north + f'' south a = north + f'' {a = a} (merid p i) = funExt (λ x → merid x ∙ sym (merid a)) i + + +proj'' : {A : Pointed ℓ} {B : Pointed ℓ'} → typ A → typ B → A ⋀ B +proj'' a b = inr (a , b) + +projᵣ : {A : Pointed ℓ} {B : Pointed ℓ'} (a : typ A) → proj'' {A = A} a (pt B) ≡ inl tt +projᵣ a = sym (push (inl a)) + +projₗ : {A : Pointed ℓ} {B : Pointed ℓ'} (b : typ B) → proj'' {B = B} (pt A) b ≡ inl tt +projₗ b = sym (push (inr b)) + +projᵣₗ : {A : Pointed ℓ} {B : Pointed ℓ'} → projᵣ (pt A) ≡ projₗ (pt B) +projᵣₗ {A = A} {B = B} i j = push (push tt i) (~ j) -- cong sym (cong push (push tt)) + + +compLrFiller : {a b c d : A} (p : a ≡ b) (q : b ≡ c) (s : c ≡ d) → PathP (λ i → p i ≡ s (~ i)) (p ∙ q ∙ s) q +compLrFiller {A = A} {a = a} {b = b} {c = c} {d = d} p q s i j = filler j i1 i + where + rhs-filler : I → I → I → A + rhs-filler i j z = hfill (λ k → λ {(i = i0) → b ; + (i = i1) → s (~ z ∧ k) ; + (z = i1) → q i}) + (inS (q i)) j + filler : I → I → I → A + filler i j z = hfill (λ k → λ {(i = i0) → p z ; + (i = i1) → rhs-filler k i1 z ; + (z = i1) → q (k ∧ i)}) + (inS (p (i ∨ z))) + j + + +-- module S1-smash-Iso ((A , pt) : Pointed ℓ) where + +-- funInv : Susp A → (S¹ , base) ⋀ (A , pt) +-- funInv north = inl tt +-- funInv south = inl tt +-- funInv (merid a i) = (sym (projₗ a) ∙ +-- cong (λ x → proj'' {A = S¹ , base} {B = A , pt} x a) loop ∙ +-- projₗ a) i + + + +-- fun : (S¹ , base) ⋀ (A , pt) → (Susp A) +-- fun (inl tt) = north +-- fun (inr (x , x₁)) = f* {a = pt} x x₁ +-- fun (push (inl base) i) = north +-- fun (push (inl (loop i₁)) i) = rCancel (merid pt) (~ i) i₁ -- rCancel (merid pt) (~ i) i₁ +-- fun (push (inr x) i) = north +-- fun (push (push a i₁) i) = north + +-- funSect : (x : Susp A) → fun (funInv x) ≡ x +-- funSect north = refl +-- funSect south = merid pt +-- funSect (merid a i) = hcomp (λ k → λ {(i = i0) → ((λ j → rCancel refl j ∙ refl ∙ refl) ∙ +-- λ j → lUnit (lUnit refl (~ j)) (~ j)) k ; +-- (i = i1) → ((λ j → lUnit refl (~ j) ∙ sym (lUnit (rUnit (merid a ∙ sym (merid pt)) (~ j)) (~ j)) ∙ merid a) ∙ +-- (λ j → lUnit (symDistr (merid a) (sym (merid pt)) j ∙ merid a) (~ j)) ∙ +-- (sym (assoc (merid pt) (sym (merid a)) (merid a))) ∙ +-- (λ j → merid pt ∙ lCancel (merid a) j) ∙ +-- sym (rUnit (merid pt)) ) k}) +-- (helper2 i ∙ filler i) +-- where + +-- filler : (i : I) → (refl ∙ (merid a ∙ sym (merid pt)) ∙ refl) i ≡ merid a i +-- filler i = (λ j → (refl ∙ (merid a ∙ sym (merid pt)) ∙ refl) (i ∧ (~ j))) ∙ λ j → merid a (j ∧ i) + +-- helper2 : (i : I) → fun ((sym (projₗ a) ∙ +-- cong (λ x → proj'' {A = S¹ , base} {B = A , pt} x a) loop ∙ +-- projₗ a) i) +-- ≡ ((λ i → fun (sym (projₗ a) i)) ∙ +-- (λ i → fun (cong (λ x → proj'' {A = S¹ , base} {B = A , pt} x a) loop i)) ∙ +-- λ i → fun (projₗ a i)) i +-- helper2 i = (λ j → congFunct fun (sym (projₗ a)) (cong (λ x → proj'' {A = S¹ , base} {B = A , pt} x a) loop ∙ projₗ a) j i) ∙ +-- (λ j → (cong fun (sym (projₗ a)) ∙ congFunct fun (cong (λ x → proj'' {A = S¹ , base} {B = A , pt} x a) loop) (projₗ a) j) i) + +-- retrHelper : (x : S¹) (a : A) → funInv (f* {a = pt} x a) ≡ proj'' x a +-- retrHelper base a = sym (projₗ a) +-- retrHelper (loop i) a j = hcomp (λ k → λ { (i = i0) → push (inr a) j +-- ; (i = i1) → push (inr a) j +-- ; (j = i0) → id3 (~ k) i +-- ; (j = i1) → proj'' (loop i) a}) +-- (compLrFiller (push (inr a)) (λ i → proj'' (loop i) a) (sym (push (inr a))) j i) + +-- module _ where + +-- invCancelHelper : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) (q : y ≡ y) → (p ∙ q ∙ sym p) ∙ p ∙ sym q ≡ p +-- invCancelHelper p q = (p ∙ q ∙ sym p) ∙ p ∙ sym q ≡⟨ (λ i → assoc p q (sym p) i ∙ p ∙ sym q) ⟩ +-- ((p ∙ q) ∙ sym p) ∙ p ∙ sym q ≡⟨ assoc ((p ∙ q) ∙ sym p) p (sym q) ⟩ +-- (((p ∙ q) ∙ sym p) ∙ p) ∙ sym q ≡⟨ (λ i → assoc (p ∙ q) (sym p) p (~ i) ∙ sym q) ⟩ +-- ((p ∙ q) ∙ sym p ∙ p) ∙ sym q ≡⟨ (λ i → ((p ∙ q) ∙ lCancel p i) ∙ sym q) ⟩ +-- ((p ∙ q) ∙ refl) ∙ sym q ≡⟨ (λ i → rUnit (p ∙ q) (~ i) ∙ sym q) ⟩ +-- (p ∙ q) ∙ sym q ≡⟨ sym (assoc p q (sym q)) ⟩ +-- p ∙ q ∙ sym q ≡⟨ (λ i → p ∙ rCancel q i) ⟩ +-- p ∙ refl ≡⟨ sym (rUnit p) ⟩ +-- p ∎ + +-- filler : funInv (f* {a = pt} (loop i) a) ≡ proj'' (loop i) a +-- filler = (λ j → funInv (funExt (λ x → merid x ∙ sym (merid pt)) (i ∨ j) a)) ∙ sym (projₗ a) ∙ λ j → proj'' (loop (i ∨ (~ j))) a + +-- id1 : sym (projᵣ base) ∙ cong ((λ x → proj'' {A = S¹ , base} {B = A , pt} x pt)) loop ∙ projᵣ base ≡ refl +-- id1 = (λ i → (push (inl (loop i))) ∙ (λ j → inr (loop (i ∨ j) , pt)) ∙ sym (push (inl base))) ∙ +-- (λ i → push (inl base) ∙ lUnit (sym (push (inl base))) (~ i)) ∙ +-- rCancel (push (inl base)) + +-- id2 : cong funInv (sym (merid pt)) ≡ (λ _ → inl tt) +-- id2 = (λ i j → (sym (projₗ pt) ∙ (cong (λ x → proj'' {A = S¹ , base} {B = A , pt} x pt) loop) ∙ projₗ pt) (~ j)) ∙ +-- (λ i j → (sym ((projᵣₗ (~ i))) ∙ cong (λ x → proj'' {A = S¹ , base} {B = A , pt} x pt) loop ∙ projᵣₗ (~ i)) (~ j)) ∙ +-- (λ i j → id1 i (~ j)) + +-- id3 : (λ i → funInv (f* {a = pt} (loop i) a)) ≡ ((sym (projₗ a) ∙ cong (λ x → proj'' {A = S¹ , base} {B = A , pt} x a) loop ∙ (projₗ a))) +-- id3 = (λ j i → congFunct funInv (merid a) (sym (merid pt)) j i ) ∙ +-- (λ j i → (cong funInv (merid a) ∙ id2 j) i) ∙ +-- (λ j i → rUnit (cong funInv (merid a)) (~ j) i) + + + +-- test5 : ((sym (projₗ a) ∙ cong (λ x → proj'' {A = S¹ , base} {B = A , pt} x a) loop ∙ (projₗ a))) ≡ λ j → funInv (merid a j) +-- test5 = refl + +-- test6 : (ia ib : I) → (x : S¹) → retrHelper (loop ia) pt ≡ {!? ∙ sym (cong (projₗ) loop) ∙ ?!} +-- test6 = {!!} +-- where +-- test7 : ((sym (projₗ pt) ∙ cong (λ x → proj'' {A = S¹ , base} {B = A , pt} x pt) loop ∙ (projₗ pt))) ≡ refl +-- test7 = (λ i j → funInv (merid pt j)) ∙ λ i → {!!} +-- {- +-- test2 : (i : I) → (retrHelper (loop i) pt) ≡ (λ j → funInv (rCancel (merid pt) j i)) ∙ sym (projᵣ (loop i)) +-- test2 i = (λ z → hfill ((λ k → λ { (i = i0) → ((λ j → id3 i pt j ∙ sym (projₗ pt) ∙ λ j → proj'' (loop (~ j)) pt) ∙ invCancelHelper i pt _ _) k +-- ; (i = i1) → lUnit (rUnit (sym (projₗ pt)) (~ k)) (~ k) })) +-- (inS ((λ j → funInv (funExt (λ x → merid x ∙ sym (merid pt)) (i ∨ j) pt)) ∙ sym (projₗ pt) ∙ λ j → proj'' (loop (i ∨ (~ j))) pt)) (~ z)) ∙ +-- (λ z → (λ j → funInv (funExt (λ x → merid x ∙ sym (merid pt)) (i ∨ j) pt)) ∙ sym (projₗ pt) ∙ λ j → proj'' (loop (i ∨ (~ j))) pt) ∙ +-- (λ z → (λ j → funInv {!funExt (λ x → merid x ∙ sym (merid pt)) (i ∨ j) pt -- rCancel (merid pt) z (i ∨ j)!}) ∙ {!!} ∙ {!!}) ∙ +-- {!!} +-- -} +-- isoHelperPre : (x : S¹ × A) (v : (S¹ , base) ⋁ (A , pt)) → {!!} +-- isoHelperPre = {!!} + +-- isoHelper : (x : S¹ × A) (v : (S¹ , base) ⋁ (A , pt)) (q : inr x ≡ inl tt) (i : I) → (funInv (fun (q i)) ≡ q i) ≡ (funInv (fun ((q ∙ push v) i)) ≡ (q ∙ push v) i) +-- isoHelper x v q i j = funInv (fun (test2 j)) ≡ test2 j +-- where +-- test2 : q i ≡ (q ∙ push v) i +-- test2 = (λ j → q (i ∧ (~ j))) ∙ λ j → (q ∙ push v) (i ∧ j) + +-- idIsEquiv2 : ∀ {ℓ} {A B : Type ℓ} → A ≡ B → A ≃ B +-- idIsEquiv2 = J (λ B p → _ ≃ B) (idEquiv _) + +-- helper2 : (x : S¹ × A) → funInv (fun (inr x)) ≡ inr x +-- helper2 (x , x₁) = retrHelper x x₁ +-- funRetr2 : (x : (S¹ , base) ⋀ (A , pt)) → funInv (fun x) ≡ x +-- funRetr2 (inl tt) = {!!} +-- funRetr2 (inr x) = {!!} +-- funRetr2 (push a i) = ElimR.elimL (λ b path → funInv (fun (path (~ i))) ≡ path (~ i)) (λ b path → funInv (fun ((path) (~ i))) ≡ (path) (~ i)) +-- (helper2 (refl i)) -- (helper2 (refl i)) +-- (λ z q → idIsEquiv2 {!!})-- isoToEquiv (isoHelper _ a q (~ i))) +-- tt (sym (push a)) +-- where +-- test3 : (a₁ : (S¹ , base) ⋁ (A , pt)) → (q : inr (i∧ a) ≡ inl tt) → Iso ((funInv (fun (q (~ i)))) ≡ q (~ i)) (funInv (fun ((q ∙ push a₁) (~ i))) ≡ (q ∙ push a₁) (~ i)) +-- test3 (inl x) q = {!!} +-- test3 (inr x) q = {!!} +-- test3 (push a i) q = {!!} + +-- test2 : (x : S¹ × A) (v : (S¹ , base) ⋁ (A , pt)) (q : Path ((S¹ , base) ⋀ (A , pt)) (inr x) (inl tt)) → q (~ i) ≡ (q ∙ push v) (~ i) +-- test2 x v q = (λ j → q ((~ i) ∧ (~ j))) ∙ λ j → (q ∙ push v) ((~ i) ∧ j) + +-- funRetr : (x : (S¹ , base) ⋀ (A , pt)) → funInv (fun x) ≡ x +-- funRetr (inl tt) = refl +-- funRetr (inr (x , a)) = retrHelper x a +-- funRetr (push (inl base) i) j = hcomp (λ k → λ { (i = i0) → inl tt +-- ; (i = i1) → sym (projᵣₗ k) j +-- ; (j = i0) → inl tt +-- ; (j = i1) → projᵣ base (~ i)}) +-- (projᵣ base ((~ j) ∨ (~ i))) +-- funRetr (push (inl (loop z)) i) j = {!(projᵣ base (~ j ∨ ~ i))!} +-- {- hcomp (λ k → λ {(i₁ = i0) → ((λ l → hcomp (λ r → λ {(i = i0) → (rCancel (λ _ → inl tt)) r +-- ; (i = i1) → (lUnit ((λ j → push (inl base) (i ∧ j))) (~ l)) }) +-- (lUnit ((λ j → push (inl base) (i ∧ j))) (~ i)))) k +-- ; (i₁ = i1) → ((λ l → hcomp (λ r → λ {(i = i0) → (rCancel (λ _ → inl tt)) r +-- ; (i = i1) → (lUnit ((λ j → push (inl base) (i ∧ j))) (~ l)) }) +-- (lUnit ((λ j → push (inl base) (i ∧ j))) (~ i)))) k +-- ; (i = i0) → ({!!} ∙ {!!} ∙ {!push inl !}) k -- rCancel refl k +-- ; (i = i1) → {!!}}) +-- ((hcomp (λ k → λ {(i = i0) → (rCancel (λ _ → inl tt)) k +-- ; (i = i1) → _ }) ((λ j → funInv (rCancel (merid pt) (~ i ∨ j) i₁)) ∙ +-- λ j → push (inl (loop i₁)) (i ∧ j)))) -} +-- where + +-- bottom : funInv (rCancel (merid pt) (~ i) z) ≡ push (inl (loop z)) i +-- bottom = (λ j → funInv(rCancel (merid pt) (~ i) (z ∨ j))) ∙ (λ j → push (inl (loop z)) (i ∧ j)) + +-- i=i1 : (λ j → funInv(rCancel (merid pt) i0 (z ∨ j))) ∙ (λ j → push (inl (loop z)) j) ≡ retrHelper (loop z) pt +-- i=i1 = {!-- rCancel (merid pt) i1!} + +-- frontCubeFiller : (i j z : I) → (S¹ , base) ⋀ (A , pt) -- bottom is i j , z is height +-- frontCubeFiller i j z = hfill (λ k → λ { (i = i0) → inl tt +-- ; (i = i1) → sym (projᵣₗ k) j +-- ; (j = i0) → inl tt +-- ; (j = i1) → projᵣ base (~ i)}) (inS (push (inl base) (i ∧ j))) z + +-- bottomCubeFiller : (i j z : I) → (S¹ , base) ⋀ (A , pt) +-- bottomCubeFiller i j z = funInv (rCancel (merid pt) (~ i) (z ∧ j)) + + +-- theHComp : (i j z : I) → (S¹ , base) ⋀ (A , pt) +-- theHComp i j z = {!!} + + + + +-- funRetr (push (inr x) i) = λ j → (projₗ x ((~ i) ∨ (~ j))) +-- funRetr (push (push tt i₁) i) = {! + +-- i₁ = i0 ⊢ hcomp +-- (λ { k (i = i0) → λ _ → inl tt +-- ; k (i = i1) → λ i₂ → projᵣₗ k (~ i₂) +-- }) +-- (λ i₂ → projᵣ base (~ i₂ ∨ ~ i)) +-- i₁ = i1 ⊢ hcomp +-- (λ { k (i = i0) → λ _ → inl tt +-- ; k (i = i1) → λ i₂ → projᵣₗ k (~ i₂) +-- }) +-- (λ i₂ → projᵣ base (~ i₂ ∨ ~ i)) +-- i = i0 ⊢ λ _ → inl tt + +-- i = i1 ⊢ !} diff --git a/Cubical/HITs/Truncation/Properties.agda b/Cubical/HITs/Truncation/Properties.agda index c91511fe6..ef2058b69 100644 --- a/Cubical/HITs/Truncation/Properties.agda +++ b/Cubical/HITs/Truncation/Properties.agda @@ -7,6 +7,7 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.HLevels open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.Equiv.PathSplit open isPathSplitEquiv open import Cubical.Modalities.Modality @@ -192,50 +193,56 @@ mapId {n = n} = -- equivalences to prop/set/groupoid truncations + +-- -1 -- +propTruncTrunc-1Iso : Iso ∥ A ∥₋₁ (∥ A ∥ -1) +Iso.fun propTruncTrunc-1Iso = PropTrunc.elim (λ _ → isOfHLevelTrunc 1) ∣_∣ +Iso.inv propTruncTrunc-1Iso = elim (λ _ → squash₋₁) ∣_∣₋₁ +Iso.rightInv propTruncTrunc-1Iso = elim (λ _ → isOfHLevelPath 1 (isOfHLevelTrunc 1) _ _) (λ _ → refl) +Iso.leftInv propTruncTrunc-1Iso = PropTrunc.elim (λ _ → isOfHLevelPath 1 squash₋₁ _ _) (λ _ → refl) + propTrunc≃Trunc-1 : ∥ A ∥₋₁ ≃ ∥ A ∥ -1 -propTrunc≃Trunc-1 = - isoToEquiv - (iso - (PropTrunc.elim (λ _ → isOfHLevelTrunc 1) ∣_∣) - (elim (λ _ → squash₋₁) ∣_∣₋₁) - (elim (λ _ → isOfHLevelPath 1 (isOfHLevelTrunc 1) _ _) (λ _ → refl)) - (PropTrunc.elim (λ _ → isOfHLevelPath 1 squash₋₁ _ _) (λ _ → refl))) +propTrunc≃Trunc-1 = isoToEquiv propTruncTrunc-1Iso propTrunc≡Trunc-1 : ∥ A ∥₋₁ ≡ ∥ A ∥ -1 propTrunc≡Trunc-1 = ua propTrunc≃Trunc-1 +-- 0 -- +setTruncTrunc0Iso : Iso ∥ A ∥₀ (∥ A ∥ 0) +Iso.fun setTruncTrunc0Iso = SetTrunc.elim (λ _ → isOfHLevelTrunc 2) ∣_∣ +Iso.inv setTruncTrunc0Iso = elim (λ _ → squash₀) ∣_∣₀ +Iso.rightInv setTruncTrunc0Iso = elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) (λ _ → refl) +Iso.leftInv setTruncTrunc0Iso = SetTrunc.elim (λ _ → isOfHLevelPath 2 squash₀ _ _) (λ _ → refl) + setTrunc≃Trunc0 : ∥ A ∥₀ ≃ ∥ A ∥ 0 -setTrunc≃Trunc0 = - isoToEquiv - (iso - (SetTrunc.elim (λ _ → isOfHLevelTrunc 2) ∣_∣) - (elim (λ _ → squash₀) ∣_∣₀) - (elim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) (λ _ → refl)) - (SetTrunc.elim (λ _ → isOfHLevelPath 2 squash₀ _ _) (λ _ → refl))) +setTrunc≃Trunc0 = isoToEquiv setTruncTrunc0Iso propTrunc≡Trunc0 : ∥ A ∥₀ ≡ ∥ A ∥ -0 propTrunc≡Trunc0 = ua setTrunc≃Trunc0 +-- 1 -- +groupoidTrunc≃Trunc1Iso : Iso ∥ A ∥₁ (∥ A ∥ 1) +Iso.fun groupoidTrunc≃Trunc1Iso = GpdTrunc.elim (λ _ → isOfHLevelTrunc 3) ∣_∣ +Iso.inv groupoidTrunc≃Trunc1Iso = elim (λ _ → squash₁) ∣_∣₁ +Iso.rightInv groupoidTrunc≃Trunc1Iso = elim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) (λ _ → refl) +Iso.leftInv groupoidTrunc≃Trunc1Iso = GpdTrunc.elim (λ _ → isOfHLevelPath 3 squash₁ _ _) (λ _ → refl) + groupoidTrunc≃Trunc1 : ∥ A ∥₁ ≃ ∥ A ∥ 1 -groupoidTrunc≃Trunc1 = - isoToEquiv - (iso - (GpdTrunc.elim (λ _ → isOfHLevelTrunc 3) ∣_∣) - (elim (λ _ → squash₁) ∣_∣₁) - (elim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) (λ _ → refl)) - (GpdTrunc.elim (λ _ → isOfHLevelPath 3 squash₁ _ _) (λ _ → refl))) +groupoidTrunc≃Trunc1 = isoToEquiv groupoidTrunc≃Trunc1Iso groupoidTrunc≡Trunc1 : ∥ A ∥₁ ≡ ∥ A ∥ 1 groupoidTrunc≡Trunc1 = ua groupoidTrunc≃Trunc1 + +-- 2 -- +2GroupoidTrunc≃Trunc2Iso : Iso ∥ A ∥₂ (∥ A ∥ 2) +Iso.fun 2GroupoidTrunc≃Trunc2Iso = 2GpdTrunc.elim (λ _ → isOfHLevelTrunc 4) ∣_∣ +Iso.inv 2GroupoidTrunc≃Trunc2Iso = elim (λ _ → squash₂) ∣_∣₂ +Iso.rightInv 2GroupoidTrunc≃Trunc2Iso = elim (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) (λ _ → refl) +Iso.leftInv 2GroupoidTrunc≃Trunc2Iso = 2GpdTrunc.elim (λ _ → isOfHLevelPath 4 squash₂ _ _) (λ _ → refl) + 2GroupoidTrunc≃Trunc2 : ∥ A ∥₂ ≃ ∥ A ∥ 2 -2GroupoidTrunc≃Trunc2 = - isoToEquiv - (iso - (2GpdTrunc.elim (λ _ → isOfHLevelTrunc 4) ∣_∣) - (elim (λ _ → squash₂) ∣_∣₂) - (elim (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) (λ _ → refl)) - (2GpdTrunc.elim (λ _ → isOfHLevelPath 4 squash₂ _ _) (λ _ → refl))) +2GroupoidTrunc≃Trunc2 = isoToEquiv 2GroupoidTrunc≃Trunc2Iso 2GroupoidTrunc≡Trunc2 : ∥ A ∥₂ ≡ ∥ A ∥ 2 2GroupoidTrunc≡Trunc2 = ua 2GroupoidTrunc≃Trunc2 @@ -342,7 +349,14 @@ PathIdTrunc n = isoToPath (ΩTrunc.IsoFinal n _ _) PathΩ : {a : A} (n : ℕ₋₂) → (Path (∥ A ∥ (suc₋₂ n)) ∣ a ∣ ∣ a ∣) ≡ (∥ a ≡ a ∥ n) PathΩ n = PathIdTrunc n --------------------------- +{- Special cases using old defs of truncations -} +PathIdTrunc₀Iso : {a b : A} → Iso (∣ a ∣₀ ≡ ∣ b ∣₀) ∥ a ≡ b ∥₋₁ +PathIdTrunc₀Iso = compIso (congIso setTrunc≃Trunc0) + (compIso (ΩTrunc.IsoFinal _ ∣ _ ∣ ∣ _ ∣) + (symIso propTruncTrunc-1Iso)) + + +------------------------- truncOfTruncIso : (n m : ℕ) → Iso (hLevelTrunc n A) (hLevelTrunc n (hLevelTrunc (m + n) A)) diff --git a/Cubical/Homotopy/Connected.agda b/Cubical/Homotopy/Connected.agda index f32459db5..8251b83e4 100644 --- a/Cubical/Homotopy/Connected.agda +++ b/Cubical/Homotopy/Connected.agda @@ -310,3 +310,241 @@ sphereConnected (suc n) = Iso.inv (fibIso) a = a , refl Iso.leftInv (fibIso) b = refl Iso.rightInv (fibIso) b = refl + + +private + variable + ℓ ℓ' ℓ'' : Level + A : Type ℓ + B : Type ℓ' + A' : Pointed ℓ + B' : Pointed ℓ' + C' : Pointed ℓ'' + + + +-- isConnectedSmashIdfun2 : {A B C : Pointed₀} +-- → (f : A →∙ B) (nf nc : ℕ) +-- → isHLevelConnectedFun nf (fst f) +-- → isHLevelConnected (2 + nc) (fst C) +-- → isHLevelConnectedFun (1 + nc + nf) (Smash-map f (idfun∙ C)) +-- isConnectedSmashIdfun2 {A = A , ptA} {B = B , ptB} {C = C , ptC} (f , fpt) nf nc conf conC = λ b → {!!} , {!!} +-- where +-- module proof {ℓ : Level} (P : (Smash (B , ptB) (C , ptC)) → HLevel ℓ (1 + nc + nf)) +-- (d : (x : Smash (A , ptA) (C , ptC)) → P ((Smash-map (f , fpt) (idfun∙ (C , ptC)) ) x) .fst) +-- where +-- F : (c : C) → _ +-- F c = λ(s : (b : B) → P (proj b c) .fst) → s ∘ f + +-- step1 : (c : C) → isHLevelTruncatedFun (1 + nc) (F c) +-- step1 c = isOfHLevelPrecomposeConnected (1 + nc) nf ((λ b → P (proj b c) .fst , P (proj b c) .snd)) f conf + +-- codomFun : (c : C) (a : A) → P (proj (f a) c) .fst +-- codomFun c = d ∘ λ a → proj a c + +-- Q : C → HLevel _ (1 + nc) +-- Q c = fiber (F c) (codomFun c) , step1 c _ + + +-- helper : (a : A) → transport (λ i → P (gluel (f a) (~ i)) .fst) (d (basel)) +-- ≡ d (proj a ptC) +-- helper a = transport (PathP≡Path (λ i → P (gluel (f a) (~ i)) .fst) (d basel) (d (proj a ptC))) +-- (transport (λ j → PathP (λ i → P (lUnit (gluel (f a)) (~ j) (~ i)) .fst) (d basel) (d (proj a ptC))) +-- λ i → d (gluel a (~ i))) + +-- QptC : Q ptC .fst +-- QptC = (λ b → transport (λ i → P (gluel b (~ i)) .fst) (d basel)) , +-- funExt helper + + +-- Q-constructor : (c : C) → Q c .fst +-- Q-constructor c = Iso.inv (elim.isIsoPrecompose (λ (_ : Unit) → ptC) (1 + nc) Q (isHLevelConnectedPoint _ conC ptC)) (λ _ → QptC) c + +-- g : (b : B) (c : C) → P (proj b c) .fst +-- g b c = Q-constructor c .fst b + +-- Q-constructor-β : (b : B) → Q-constructor ptC .fst b ≡ transport (λ i → P (gluel b (~ i)) .fst) (d basel) +-- Q-constructor-β b = ((λ i → (Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) (helper3 i)) .fst b)) ∙ +-- (λ i → transportRefl (QptC) i .fst b) + +-- where +-- helper3 : (isHLevelConnectedPoint _ conC ptC) ptC .fst ≡ ∣ tt , refl ∣ +-- helper3 = (isHLevelConnectedPoint _ conC ptC) ptC .snd ∣ tt , refl ∣ + +-- moveTransport : ∀ {ℓ ℓ'} → {A : Type ℓ} {B : A → Type ℓ'} {x y : A} (p : x ≡ y) (x1 : B x) (y1 : B y) +-- → transport (λ i → (B (p i))) x1 ≡ y1 +-- → transport⁻ (λ i → (B (p i))) y1 ≡ x1 +-- moveTransport {B = B} p x1 y1 transp = ((λ i → transport⁻ (λ i → B (p i)) (transp (~ i))) ∙ (transport⁻Transport (λ i → B (p i)) x1)) + +-- otherway : (b : B) → transport (λ i → P (gluel b i) .fst) (g b ptC) ≡ d basel +-- otherway b = (λ i → transport (λ i → P (gluel b i) .fst) (Q-constructor-β b i)) ∙ transportTransport⁻ (λ i → P (gluel b i) .fst) (d basel) + +-- path1 : (b : B) → PathP (λ i → P (gluel b i) .fst) (g b ptC) (d basel) +-- path1 b i = hcomp (λ k → λ {(i = i0) → g b ptC +-- ; (i = i1) → otherway b k}) +-- (transp (λ j → P (gluel b (i ∧ j)) .fst) (~ i) (g b ptC)) + + +-- where +-- massive : (c : C) → PathP _ (g ptB c) (d baser) +-- massive c = compPathP (λ i → g (fpt (~ i)) c) (compPathP (λ i → Q-constructor c .snd i (ptA)) (λ i → d (gluer c i))) + +-- pathPCancel : (c : C) +-- → ((λ z → P (proj (fpt (~ z)) c) .fst) ∙ (λ i → ((λ _ → P (proj (f ptA) c) .fst) ∙ (λ i₁ → P (((λ j₁ → proj (fpt j₁) c) ∙ gluer c) i₁) .fst)) i)) +-- ≡ (λ i → P (gluer c i) .fst) +-- pathPCancel c = (λ i → (λ z → P (proj (fpt (~ z)) c) .fst) ∙ lUnit ((λ i₁ → P (((λ j₁ → proj (fpt j₁) c) ∙ gluer c) i₁) .fst)) (~ i)) +-- ∙ (λ i → (λ z → P (proj (fpt (~ z ∨ i)) c) .fst) ∙ λ i₁ → P (((λ j₁ → proj (fpt (j₁ ∨ i)) c) ∙ gluer c) i₁) .fst ) +-- ∙ (λ i → lUnit (λ i₁ → P (lUnit (gluer c) (~ i) i₁) .fst) (~ i)) + +-- path2 : (c : C) → PathP (λ i → P (gluer c i) .fst) (g ptB c) (d baser) +-- path2 c = transport (λ i → PathP (λ j → pathPCancel c i j) (g ptB c) (d baser)) (massive c) + +-- h : (x : Smash (B , ptB) (C , ptC)) → P x .fst +-- h basel = d basel +-- h baser = d baser +-- h (proj x y) = g x y +-- h (gluel b i) = {!!} -- path1 b i +-- h (gluer c i) = {!!} -- path2 c i + +-- hsection : (x : Smash (A , ptA) (C , ptC)) → h (Smash-map (f , fpt) (idfun∙ (C , ptC)) x) ≡ d x +-- hsection basel = refl +-- hsection baser = refl +-- hsection (proj x y) i = Q-constructor y .snd i x +-- hsection (gluel a i) j = hcomp (λ k → λ { (i = i0) → transportRefl {!!} k +-- ; (i = i1) → transportRefl (d basel) k +-- ; (j = i0) → {!!} +-- ; (j = i1) → {!!}}) +-- {!hcomp!} -- hcomp {!!} {!!} +-- where + +-- helper2 : Path ((P (proj (f a) ptC) .fst) ≡ (P basel .fst)) +-- (λ i → P (Smash-map (f , fpt) (idfun∙ (C , ptC)) (gluel a i)) .fst) +-- λ i → P (gluel (f a) i) .fst +-- helper2 = (λ j i → P ((lUnit (gluel (f a)) (~ j)) i) .fst) +-- helper6 : (a : A) → PathP ((λ j → PathP (λ i → P {!Q-constructor ? .snd i ?!} .fst) {!P (proj (f a) ptC) .fst!} {!!})) (λ i → h (Smash-map (f , fpt) (idfun∙ (C , ptC)) (gluel a i))) (path1 (f a)) +-- helper6 a = ((λ j i → h ((lUnit (gluel (f a)) (~ j)) i))) + +-- helper3 : PathP {!!} (path1 (f a)) (λ i → d (gluel a i)) +-- helper3 = compPathP {!!} {!!} +-- helper4 : PathP {!!} (path1 (f a)) λ i → d (gluel a i) +-- helper4 = {!!} +-- hsection (gluer b i) = {!!} + + + +-- -- isConnectedSmashIdfun : (f : A' →∙ B') (nf nc : ℕ) +-- -- → isHLevelConnectedFun nf (fst f) +-- -- → isHLevelConnected (2 + nc) (fst C') +-- -- → isHLevelConnectedFun (1 + nf + nc) (f ⋀⃗ idfun∙ C') +-- -- isConnectedSmashIdfun {A' = (A , ptA)} {B' = (B , ptB)} {C' = (C , ptC)} (f , fpt) nf nc conf conC = {!isHLel!} +-- -- where +-- -- module _ (P : ((B , ptB) ⋀ (C , ptC)) → HLevel (ℓ-max (ℓ-max ℓ ℓ') ℓ'') (1 + nc + nf)) +-- -- (d : (x : (A , ptA) ⋀ (C , ptC)) → P (((f , fpt) ⋀⃗ idfun∙ (C , ptC) ) x) .fst) +-- -- where +-- -- F : (c : C) → _ +-- -- F c = λ(s : (b : B) → P (inr (b , c)) .fst) → s ∘ f + +-- -- step1 : (c : C) → isHLevelTruncatedFun (1 + nc) (F c) +-- -- step1 c = isOfHLevelPrecomposeConnected (1 + nc) nf ((λ b → P (inr (b , c)) .fst , P (inr (b , c)) .snd)) f conf + +-- -- codomFun : (c : C) (a : A) → P (inr ((f a) , c)) .fst +-- -- codomFun c = d ∘ λ a → inr (a , c) + +-- -- Q : C → HLevel _ (1 + nc) +-- -- Q c = fiber (F c) (codomFun c) , step1 c _ + + +-- -- helper : (a : A) → transport (λ i → P (push (inl (f a)) i) .fst) (d (inl tt)) +-- -- ≡ d (inr (a , ptC)) +-- -- helper a = transport (PathP≡Path (λ i → P (push (inl (f a)) i) .fst) (d (inl tt)) (d (inr (a , ptC)))) +-- -- (transport (λ j → PathP (λ i → P (rUnit (push (inl (f a))) (~ j) i) .fst) (d (inl tt)) (d (inr (a , ptC)))) +-- -- λ i → d (push (inl a) i)) + + +-- -- QptC : Q ptC .fst +-- -- QptC = (λ b → transport (λ i → P (push (inl b) i) .fst) (d (inl tt))) , +-- -- funExt helper +-- -- where + + +-- -- Q-constructor : (c : C) → Q c .fst +-- -- Q-constructor c = Iso.inv (elim.isIsoPrecompose (λ _ → ptC) (1 + nc) Q (isHLevelConnectedPoint _ conC ptC)) (λ ( _ : Unit) → QptC) c + +-- -- g : (b : B) (c : C) → P (inr (b , c)) .fst +-- -- g b c = Q-constructor c .fst b + +-- -- Q-constructor-β : (b : B) → Q-constructor ptC .fst b ≡ transport (λ i → P (push (inl b) i) .fst) (d (inl tt)) +-- -- Q-constructor-β b = ((λ i → (Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) (helper3 i)) .fst b)) ∙ +-- -- (λ i → transportRefl (QptC) i .fst b) + +-- -- where +-- -- helper3 : (isHLevelConnectedPoint _ conC ptC) ptC .fst ≡ ∣ tt , refl ∣ +-- -- helper3 = (isHLevelConnectedPoint _ conC ptC) ptC .snd ∣ tt , refl ∣ + +-- -- test : (b : B) (c : C) → PathP (λ i → P (push {!!} {!(~ i)!}) .fst) (Q-constructor ptC .fst b) (Q-constructor c .fst ptB) +-- -- test b = {!!} + +-- -- Q-constructor-β2 : (c : C) → Q-constructor c .fst ptB ≡ transport (λ i → P (push (inr c) i) .fst) (d (inl tt)) +-- -- Q-constructor-β2 c = ((λ i → (Trunc.rec (Q c .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) ((isHLevelConnectedPoint _ conC ptC) c .fst)) .fst ptB)) ∙ +-- -- {!isHLevelConnectedPoint _ conC ptC c .fst!} + +-- -- gid1 : (a : A) (c : C) → g (f a) c ≡ d (inr (a , c)) +-- -- gid1 a c i = (Q-constructor c .snd) i a + +-- -- gid1' : {!!} +-- -- gid1' = {!!} + +-- -- gid1Path : (c : C) → PathP (λ i → P (push (inr c) i) .fst) (d (inl tt)) (g ptB c) +-- -- gid1Path c i = +-- -- hcomp (λ k → λ {(i = i0) → {!!} ; (i = i1) → {!!}}) +-- -- {!!} + +-- -- compPath : (c : C) → PathP _ (d (inl tt)) (g ptB c) +-- -- compPath c = compPathP (λ i → d (push (inr c) i)) (compPathP (sym (gid1 ptA c)) (λ i → g (fpt i) c)) + + +-- -- compPathTransport : (c : C) → ((λ z → P ((push (inr c) ∙ +-- -- (λ i → inr (fpt (~ i) , c))) z) .fst) ∙ +-- -- (λ i → ((λ i₁ → P (inr (f ptA , c)) .fst) ∙ (λ i₁ → P (inr (fpt i₁ , c)) .fst)) i)) +-- -- ≡ λ i → P (push (inr c) i) .fst +-- -- compPathTransport c = (λ k → ((λ z → P ((push (inr c) ∙ +-- -- (λ i → inr (fpt (~ i) , c))) z) .fst) ∙ +-- -- ((lUnit (λ i₁ → P (inr (fpt i₁ , c)) .fst) (~ k) )))) +-- -- ∙ (λ k → ((λ z → P ((push (inr c) ∙ +-- -- (λ i → inr (fpt (~ i ∨ k) , c))) z) .fst) ∙ +-- -- (λ i₁ → P (inr (fpt (i₁ ∨ k) , c)) .fst) )) +-- -- ∙ (λ k → ((λ z → P ((push (inr c) ∙ refl) z) .fst) ∙ refl)) +-- -- ∙ (λ k → rUnit ((λ z → P ((rUnit (push (inr c)) (~ k)) z) .fst)) (~ k)) + +-- -- compPathTransport-ptC : compPathTransport ptC ≡ {!!} +-- -- compPathTransport-ptC = {!!} + +-- -- compPathTrue : (c : C) → PathP (λ i → P (push (inr c) i) .fst) (d (inl tt)) (g ptB c) +-- -- compPathTrue c = transport (λ i → PathP (λ j → compPathTransport c i j) (d (inl tt)) (g ptB c)) (compPath c) + +-- -- gid1transp : (c : C) → PathP {!!} (g ptB c) (transport (λ i → P (push (inr c) i) .fst) (d (inl tt))) +-- -- gid1transp c = compPathP (λ i → g (fpt (~ i)) c) (compPathP (gid1 ptA c) (compPathP (λ i → d (push (inr c) (~ i))) λ i → transp (λ j → P (push (inr c) (i ∧ j)) .fst) (~ i) (d (inl tt)))) + + +-- -- gid2 : (b : B) → g b ptC ≡ transport (λ i → P (push (inl b) i) .fst) (d (inl tt)) -- (λ i → P (push (inl b) i) .fst) (d (inl tt)) +-- -- gid2 b = Q-constructor-β b + + +-- -- gid1≡gid2 : (c : C) (a : A) → PathP {!gid1transp c !} (gid1 a ptC) (gid2 (f a)) +-- -- gid1≡gid2 c a = compPathP {!λ j → (Q-constructor ptC .snd) i a!} (compPathP {!!} {!!}) + + +-- -- gid2Path : (b : B) → PathP (λ i → P (push (inl b) i) .fst) (d (inl tt)) (g b ptC) +-- -- gid2Path b i = +-- -- hcomp (λ k → λ {(i = i0) → (d (inl tt)) ; (i = i1) → gid2 b (~ k)}) +-- -- (transp (λ j → P (push (inl b) (j ∧ i)) .fst) (~ i) (d (inl tt))) + +-- -- PathP2 : PathP {!!} ((gid2Path ptB)) (compPathTrue ptC) +-- -- PathP2 = compPathP {!gid2Path ptB!} (compPathP {!!} {!compPathTrue ptC!}) + +-- -- h : (x : (B , ptB) ⋀ (C , ptC)) → P x .fst +-- -- h (inl x) = d (inl tt) +-- -- h (inr (b , c)) = g b c +-- -- h (push (inl b) i) = gid2Path b i +-- -- h (push (inr x) i) = compPathTrue x i +-- -- h (push (push a i₁) i) = {!!} diff --git a/Cubical/ZCohomology/MayerVietorisReduced.agda b/Cubical/ZCohomology/MayerVietorisReduced.agda index 6979f236c..89571737a 100644 --- a/Cubical/ZCohomology/MayerVietorisReduced.agda +++ b/Cubical/ZCohomology/MayerVietorisReduced.agda @@ -21,6 +21,7 @@ open import Cubical.Data.Sigma hiding (_×_) open import Cubical.HITs.Susp open import Cubical.HITs.Wedge open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₋₁ ; ∣_∣ to ∣_∣₋₁) open import Cubical.HITs.Nullification open import Cubical.Data.Int hiding (_+_) open import Cubical.Data.Nat @@ -41,34 +42,6 @@ open import Cubical.Data.Sum.Base open import Cubical.Data.HomotopyGroup open import Cubical.Data.Bool hiding (_⊕_) - --- {- --- record AbGroup {ℓ} (A : Type ℓ) : Type ℓ where --- constructor abgr --- field --- noll : A --- _⊕_ : A → A → A --- associate : (a b c : A) → (a ⊕ b) ⊕ c ≡ a ⊕ (b ⊕ c) --- commute : (a b : A) → a ⊕ b ≡ (b ⊕ a) --- r-unit : (a : A) → a ⊕ noll ≡ a --- negate : (a : A) → Σ[ a⁻ ∈ A ] (a ⊕ a⁻ ≡ noll) --- -} --- {- --- grHom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (ϕ : A → B) --- → AbGroup A → AbGroup B → Type (ℓ-max ℓ ℓ') --- grHom {A = A} {B = B} ϕ (abgr 0A _+A_ as-A comm-A r-A neg-A) (abgr 0B _+B_ as-B comm-B r-B neg-B) = (x y : A) → ϕ (x +A y) ≡ (ϕ x +B ϕ y) --- ImHom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (ϕ : A → B) (A' : AbGroup A) (B' : AbGroup B) --- → grHom ϕ A' B' → Type (ℓ-max ℓ ℓ') --- ImHom {A = A} {B = B} ϕ Agr Bgr hom = Σ[ b ∈ B ] Σ[ a ∈ A ] ϕ a ≡ b --- KerHom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (ϕ : A → B) (A' : AbGroup A) (B' : AbGroup B) --- → grHom ϕ A' B' → Type (ℓ-max ℓ ℓ') --- KerHom {A = A} {B = B} ϕ Agr (abgr 0B _+B_ as-B comm-B r-B neg-B) hom = Σ[ a ∈ A ] ϕ a ≡ 0B --- -} - - --- --- - - private variable ℓ ℓ' ℓ'' : Level @@ -86,9 +59,6 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : D : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') D = Pushout f g - -- σ : typ A → typ (Ω (∙Susp (typ A))) - -- σ a = merid a ∙ merid (pt A) ⁻¹ - i : (n : ℕ) → coHom n D → coHom n A × coHom n B i zero = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) λ { δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀} @@ -96,14 +66,14 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : λ { δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀} Δ : (n : ℕ) → coHom n A × coHom n B → coHom n C - Δ n (α , β) = coHomFun n f α +ₕ -ₕ (coHomFun n g β) -- coHomFun∙2 n f α +ₕ (-ₕ (coHomFun n g β)) + Δ n (α , β) = coHomFun n f α +ₕ -ₕ (coHomFun n g β) d-pre : (n : ℕ) → (C → coHomK n) → D → coHomK (suc n) d-pre n γ (inl x) = 0ₖ d-pre n γ (inr x) = 0ₖ d-pre zero γ (push a i) = Kn→ΩKn+1 zero (γ a) i d-pre (suc n) γ (push a i) = Kn→ΩKn+1 (suc n) (γ a) i - + d : (n : ℕ) → coHom n C → coHom (suc n) D d n = sRec setTruncIsSet λ a → ∣ d-pre n a ∣₀ @@ -183,10 +153,185 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : dIsHom (suc n) = sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ f g i → ∣ funExt (λ x → dHomHelper (suc n) f g x) i ∣₀ - -- Long Exact Sequence - - Im-d⊂Ker-i : (n : ℕ) (x : coHom (suc n) D) → isInIm (coHomGr n C) (coHomGr (suc n) D) (d n) x → isInKer (coHomGr (suc n) D) (×coHomGr (suc n) A B) (i (suc n)) x + Im-d⊂Ker-i : (n : ℕ) (x : coHom (suc n) D) + → isInIm (coHomGr n C) (coHomGr (suc n) D) (d n) x + → isInKer (coHomGr (suc n) D) (×coHomGr (suc n) A B) (i (suc n)) x Im-d⊂Ker-i n = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) - λ a (δ , p) → {!!} + λ a → pElim (λ _ → isOfHLevelPath' 1 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + (sigmaElim (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ δ b → (λ i → sElim (λ _ → isOfHLevelProd 2 setTruncIsSet setTruncIsSet) + (λ { δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀ }) (b (~ i)))) + + Ker-i⊂Im-d : (n : ℕ) (x : coHom (suc n) D) + → isInKer (coHomGr (suc n) D) (×coHomGr (suc n) A B) (i (suc n)) x + → isInIm (coHomGr n C) (coHomGr (suc n) D) (d n) x + Ker-i⊂Im-d n = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ a p → pRec {A = (λ x → a (inl x)) ≡ λ _ → 0ₖ} (isOfHLevelΠ 1 (λ _ → propTruncIsProp)) + (λ p1 → pRec propTruncIsProp λ p2 → ∣ ∣ (λ c → ΩKn+1→Kn (sym (cong (λ F → F (f c)) p1) + ∙∙ cong a (push c) + ∙∙ cong (λ F → F (g c)) p2)) ∣₀ + , cong ∣_∣₀ (funExt (λ δ → helper n a p1 p2 δ)) ∣₋₁) + (Iso.fun (PathIdTrunc₀Iso) (cong proj₁ p)) + (Iso.fun (PathIdTrunc₀Iso) (cong proj₂ p)) + where + pushFiller : (n : ℕ) (F : D → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) + (p1 : Path (_ → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) (λ a₁ → F (inl a₁)) (λ _ → ∣ north ∣)) + (p2 : Path (_ → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) (λ a₁ → F (inr a₁)) (λ _ → ∣ north ∣)) → + (c : C) → I → I → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n) + pushFiller n F p1 p2 c i j = + hcomp (λ k → λ { (i = i1) → F (push c j) + ; (j = i0) → p1 (~ i ∧ k) (f c) + ; (j = i1) → p2 (~ i ∧ k) (g c)}) + (F (push c j)) + helper : (n : ℕ) (F : D → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) + (p1 : Path (_ → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) (λ a₁ → F (inl a₁)) (λ _ → ∣ north ∣)) + (p2 : Path (_ → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) (λ a₁ → F (inr a₁)) (λ _ → ∣ north ∣)) → + (δ : D) → d-pre n (λ c → ΩKn+1→Kn ((λ i₁ → p1 (~ i₁) (f c)) + ∙∙ cong F (push c) + ∙∙ cong (λ F → F (g c)) p2)) δ ≡ F δ + helper zero F p1 p2 (inl x) = sym (cong (λ f → f x) p1) + helper zero F p1 p2 (inr x) = sym (cong (λ f → f x) p2) + helper zero F p1 p2 (push a i) j = + hcomp (λ k → λ { (i = i0) → p1 (~ j) (f a) + ; (i = i1) → p2 (~ j) (g a) + ; (j = i0) → Iso.rightInv (Iso3-Kn-ΩKn+1 zero) ((λ i₁ → p1 (~ i₁) (f a)) + ∙∙ cong F (push a) + ∙∙ cong (λ F₁ → F₁ (g a)) p2) (~ k) i + ; (j = i1) → F (push a i)}) + (pushFiller zero F p1 p2 a j i) + helper (suc n) F p1 p2 (inl x) = sym (cong (λ f → f x) p1) + helper (suc n) F p1 p2 (inr x) = sym (cong (λ f → f x) p2) + helper (suc n) F p1 p2 (push a i) j = + hcomp (λ k → λ { (i = i0) → p1 (~ j) (f a) + ; (i = i1) → p2 (~ j) (g a) + ; (j = i0) → Iso.rightInv (Iso3-Kn-ΩKn+1 (suc n)) ((λ i₁ → p1 (~ i₁) (f a)) + ∙∙ cong F (push a) + ∙∙ cong (λ F₁ → F₁ (g a)) p2) (~ k) i + ; (j = i1) → F (push a i)}) + (pushFiller (suc n) F p1 p2 a j i) + + Im-i⊂Ker-Δ : (n : ℕ) (x : coHom n A × coHom n B) + → isInIm (coHomGr n D) (×coHomGr n A B) (i n) x + → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) x + Im-i⊂Ker-Δ n (Fa , Fb) = + sElim {B = λ Fa → (Fb : _) → isInIm (coHomGr n D) (×coHomGr n A B) (i n) (Fa , Fb) + → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) (Fa , Fb)} + (λ _ → isOfHLevelΠ 2 λ _ → (isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _)) + (λ Fa → sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ Fb → pRec (setTruncIsSet _ _) + (sigmaElim (λ x → isOfHLevelSuc 1 (setTruncIsSet _ _)) + λ Fd p → helper n Fa Fb Fd p)) + Fa + Fb + where + helper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) (Fd : D → coHomK n) + → (i n ∣ Fd ∣₀ ≡ (∣ Fa ∣₀ , ∣ Fb ∣₀)) + → (Δ n) (∣ Fa ∣₀ , ∣ Fb ∣₀) ≡ 0ₕ + helper zero Fa Fb Fd p = cong (Δ zero) (sym p) + ∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₀ +ₕ -ₕ ∣ (λ x → Fd (push x (~ i))) ∣₀) + ∙ rCancelₕ ∣ (λ x → Fd (inl (f x))) ∣₀ + helper (suc n) Fa Fb Fd p = cong (Δ (suc n)) (sym p) + ∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₀ +ₕ -ₕ ∣ (λ x → Fd (push x (~ i))) ∣₀) + ∙ rCancelₕ ∣ (λ x → Fd (inl (f x))) ∣₀ + + + + Ker-Δ⊂Im-i : (n : ℕ) (a : coHom n A × coHom n B) + → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) a + → isInIm (coHomGr n D) (×coHomGr n A B) (i n) a + Ker-Δ⊂Im-i n (Fa , Fb) = + sElim {B = λ Fa → (Fb : _) → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) (Fa , Fb) + → isInIm (coHomGr n D) (×coHomGr n A B) (i n) (Fa , Fb)} + (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) + (λ Fa → sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ Fb p → pRec propTruncIsProp + (λ q → ∣ ∣ helpFun n Fa Fb (funExt⁻ q) ∣₀ + , anotherHelper n Fa Fb q ∣₋₁) + (helper n Fa Fb p)) + Fa + Fb + + where + helper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) + → (Δ n) (∣ Fa ∣₀ , ∣ Fb ∣₀) ≡ 0ₕ + → ∥ (Path (_ → _) (λ c → Fa (f c)) (λ c → Fb (g c))) ∥₋₁ + helper zero Fa Fb p = Iso.fun (PathIdTrunc₀Iso) + (sym (rUnitₕ (coHomFun zero f ∣ Fa ∣₀)) + ∙ (λ i → coHomFun zero f ∣ Fa ∣₀ +ₕ (lCancelₕ (coHomFun zero g ∣ Fb ∣₀) (~ i))) + ∙ sym (assocₕ (coHomFun zero f ∣ Fa ∣₀) (-ₕ (coHomFun zero g ∣ Fb ∣₀)) (coHomFun zero g ∣ Fb ∣₀)) + ∙ cong (λ x → x +ₕ (coHomFun zero g ∣ Fb ∣₀)) p + ∙ lUnitₕ (coHomFun zero g ∣ Fb ∣₀)) + helper (suc n) Fa Fb p = Iso.fun (PathIdTrunc₀Iso) + (sym (rUnitₕ (coHomFun (suc n) f ∣ Fa ∣₀)) + ∙ (λ i → coHomFun (suc n) f ∣ Fa ∣₀ +ₕ (lCancelₕ (coHomFun (suc n) g ∣ Fb ∣₀) (~ i))) + ∙ sym (assocₕ (coHomFun (suc n) f ∣ Fa ∣₀) (-ₕ (coHomFun (suc n) g ∣ Fb ∣₀)) (coHomFun (suc n) g ∣ Fb ∣₀)) + ∙ cong (λ x → x +ₕ (coHomFun (suc n) g ∣ Fb ∣₀)) p + ∙ lUnitₕ (coHomFun (suc n) g ∣ Fb ∣₀)) + + helpFun : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) + → ((c : C) → Fa (f c) ≡ Fb (g c)) + → D → coHomK n + helpFun n Fa Fb p (inl x) = Fa x + helpFun n Fa Fb p (inr x) = Fb x + helpFun n Fa Fb p (push a i) = p a i + + anotherHelper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) + → (q : Path (C → coHomK n) (λ c → Fa (f c)) (λ c → Fb (g c))) + → i n ∣ helpFun n Fa Fb (λ x i₁ → q i₁ x) ∣₀ ≡ (∣ Fa ∣₀ , ∣ Fb ∣₀) + anotherHelper zero Fa Fb q = refl + anotherHelper (suc n) Fa Fb q = refl + + + Ker-d⊂Im-Δ : (n : ℕ) (a : coHom n C) + → isInKer (coHomGr n C) (coHomGr (suc n) D) (d n) a + → isInIm (×coHomGr n A B) (coHomGr n C) (Δ n) a + Ker-d⊂Im-Δ n = + sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ Fa p → pRec propTruncIsProp (λ p → ∣ {!!} , {!!} ∣₋₁) (Iso.fun (PathIdTrunc₀Iso) p) + + where + + + -- Im-Δ⊂Ker-d : (n : ℕ) (a : coHom n C) + -- → isInIm (×coHomGr n A B) (coHomGr n C) (Δ n) a + -- → isInKer (coHomGr n C) (coHomGr (suc n) D) (d n) a + -- Im-Δ⊂Ker-d n = + -- sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + -- λ Fc → pRec (isOfHLevelPath' 1 setTruncIsSet _ _) + -- (sigmaProdElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + -- λ Fa Fb p → pRec (isOfHLevelPath' 1 setTruncIsSet _ _) + -- (λ q → ((λ i → d n ∣ (q (~ i)) ∣₀) ∙ dΔ-Id n Fa Fb)) + -- (Iso.fun (PathIdTrunc₀Iso) p)) + + -- where + + -- d-preLeftId : (n : ℕ) (Fa : A → coHomK n)(d : D) + -- → d-pre n (Fa ∘ f) d ≡ 0ₖ + -- d-preLeftId n Fa (inl x) = Kn→ΩKn+1 n (Fa x) + -- d-preLeftId n Fa (inr x) = refl + -- d-preLeftId zero Fa (push a i) j = Kn→ΩKn+1 zero (Fa (f a)) (j ∨ i) + -- d-preLeftId (suc n) Fa (push a i) j = Kn→ΩKn+1 (suc n) (Fa (f a)) (j ∨ i) + + -- d-preRightId : (n : ℕ) (Fb : B → coHomK n) (d : D) + -- → d-pre n (Fb ∘ g) d ≡ 0ₖ + -- d-preRightId n Fb (inl x) = refl + -- d-preRightId n Fb (inr x) = sym (Kn→ΩKn+1 n (Fb x)) + -- d-preRightId zero Fb (push a i) j = Kn→ΩKn+1 zero (Fb (g a)) (~ j ∧ i) + -- d-preRightId (suc n) Fb (push a i) j = Kn→ΩKn+1 (suc n) (Fb (g a)) (~ j ∧ i) + + -- dΔ-Id : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) + -- → d n (Δ n (∣ Fa ∣₀ , ∣ Fb ∣₀)) ≡ 0ₕ + -- dΔ-Id zero Fa Fb = + -- dIsHom zero ∣ (λ x → Fa (f x)) ∣₀ (-ₕ ∣ (λ x → Fb (g x)) ∣₀) + -- ∙ (λ i → d zero ∣ (λ x → Fa (f x)) ∣₀ +ₕ morphMinus (coHomGr zero C) (coHomGr 1 D) (d zero) (dIsHom zero) ∣ (λ x → Fb (g x)) ∣₀ i) + -- ∙ (λ i → d zero ∣ (λ x → Fa (f x)) ∣₀ +ₕ -ₕ (d zero ∣ (λ x → Fb (g x)) ∣₀)) + -- ∙ (λ i → ∣ funExt (d-preLeftId zero Fa) i ∣₀ +ₕ -ₕ ∣ funExt (d-preRightId zero Fb) i ∣₀) + -- ∙ rCancelₕ 0ₕ + -- dΔ-Id (suc n) Fa Fb = + -- dIsHom (suc n) ∣ (λ x → Fa (f x)) ∣₀ (-ₕ ∣ (λ x → Fb (g x)) ∣₀) + -- ∙ (λ i → d (suc n) ∣ (λ x → Fa (f x)) ∣₀ +ₕ morphMinus (coHomGr (suc n) C) (coHomGr (2 + n) D) (d (suc n)) (dIsHom (suc n)) ∣ (λ x → Fb (g x)) ∣₀ i) + -- ∙ (λ i → d (suc n) ∣ (λ x → Fa (f x)) ∣₀ +ₕ -ₕ (d (suc n) ∣ (λ x → Fb (g x)) ∣₀)) + -- ∙ (λ i → ∣ funExt (d-preLeftId (suc n) Fa) i ∣₀ +ₕ -ₕ ∣ funExt (d-preRightId (suc n) Fb) i ∣₀) + -- ∙ rCancelₕ 0ₕ diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index 4a8c42667..0052848bd 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -100,11 +100,33 @@ _+ₖ_ {n = n} x y = ΩKn+1→Kn (Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y) -ₖ_ : {n : ℕ} → coHomK n → coHomK n -ₖ_ {n = n} x = ΩKn+1→Kn (sym (Kn→ΩKn+1 n x)) - Kn→ΩKn+10ₖ : (n : ℕ) → Kn→ΩKn+1 n 0ₖ ≡ refl Kn→ΩKn+10ₖ zero = refl Kn→ΩKn+10ₖ (suc n) = (λ i → cong ∣_∣ (rCancel (merid north) i)) -- could also use refl for n = 1, but for computational reasons I don't want to expand the definition if not necessary. +-0ₖ : {n : ℕ} → -ₖ 0ₖ {n = n} ≡ 0ₖ +-0ₖ {n = n} = (λ i → ΩKn+1→Kn (sym (Kn→ΩKn+10ₖ n i))) + ∙∙ (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ n (~ i))) + ∙∙ Iso.leftInv (Iso3-Kn-ΩKn+1 n) 0ₖ + + + +-- cong-ₖ : {n : ℕ} (x : coHomK n) → PathP (λ i → -0ₖ {n = (suc n)} i ≡ -0ₖ i) (cong -ₖ_ (Kn→ΩKn+1 n x)) (Kn→ΩKn+1 n (-ₖ x)) +-- cong-ₖ {n = n} x j i = +-- hcomp (λ k → λ{ (i = i0) → -0ₖ j +-- ; (i = i1) → -0ₖ j +-- ; (j = i0) → ΩKn+1→Kn (sym (Kn→ΩKn+1 (suc n) (Kn→ΩKn+1 n x i))) +-- ; (j = i1) → Iso.rightInv (Iso3-Kn-ΩKn+1 n) (sym (Kn→ΩKn+1 n x)) (~ k) i}) +-- (hcomp (λ k → λ{ (j = i0) → {!!} +-- ; (j = i1) → {!!}}) +-- {!!}) +-- {- +-- j = i0 ⊢ ΩKn+1→Kn (sym (Kn→ΩKn+1 (suc n) (Kn→ΩKn+1 n x i))) +-- j = i1 ⊢ Kn→ΩKn+1 x (~ i) -- Kn→ΩKn+1 n (-ₖ x) i +-- i = i0 ⊢ -0ₖ j +-- i = i1 ⊢ -0ₖ j +-- -} + +ₖ→∙ : (n : ℕ) (a b : coHomK n) → Kn→ΩKn+1 n (a +ₖ b) ≡ Kn→ΩKn+1 n a ∙ Kn→ΩKn+1 n b +ₖ→∙ n a b = Iso.rightInv (Iso3-Kn-ΩKn+1 n) (Kn→ΩKn+1 n a ∙ Kn→ΩKn+1 n b) @@ -194,6 +216,7 @@ assocₖ {n = n} x y z = (λ i → ΩKn+1→Kn (Kn→ΩKn+1 n (ΩKn+1→Kn (Kn -- Group structure of cohomology groups --- + private § : isSet ∥ A ∥₀ § = setTruncIsSet From cdeb64daafbd190fd242e3df206fd5673bf7e5b1 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Wed, 13 May 2020 12:55:27 +0200 Subject: [PATCH 33/89] MV done --- Cubical/Foundations/Prelude.agda | 37 +++++- Cubical/ZCohomology/MayerVietorisReduced.agda | 115 ++++++++++-------- 2 files changed, 100 insertions(+), 52 deletions(-) diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index 2a0850ad0..a0cfb5058 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -362,6 +362,37 @@ record Lift {i j} (A : Type i) : Type (ℓ-max i j) where open Lift public - -liftExt : ∀ {A : Type ℓ} {a b : Lift {ℓ} {ℓ'} A} → (lower a ≡ lower b) → a ≡ b -liftExt x i = lift (x i) +doubleCompPath-filler∙ : {a b c d : A} (p : a ≡ b) (q : b ≡ c) (r : c ≡ d) + → PathP (λ i → p i ≡ r (~ i)) (p ∙ q ∙ r) q +doubleCompPath-filler∙ {A = A} {b = b} p q r j i = + hcomp (λ k → λ { (i = i0) → p j + ; (i = i1) → side j k + ; (j = i1) → q (i ∧ k)}) + (p (j ∨ i)) + where + side : I → I → A + side i j = + hcomp (λ k → λ { (i = i1) → q j + ; (j = i0) → b + ; (j = i1) → r (~ i ∧ k)}) + (q j) + +PathP→compPathL : {a b c d : A} {p : a ≡ c} {q : b ≡ d} {r : a ≡ b} {s : c ≡ d} + → PathP (λ i → p i ≡ q i) r s + → sym p ∙ r ∙ q ≡ s +PathP→compPathL {p = p} {q = q} {r = r} {s = s} P j i = + hcomp (λ k → λ { (i = i0) → p (j ∨ k) + ; (i = i1) → q (j ∨ k) + ; (j = i0) → doubleCompPath-filler∙ (sym p) r q (~ k) i + ; (j = i1) → s i }) + (P j i) + +PathP→compPathR : {a b c d : A} {p : a ≡ c} {q : b ≡ d} {r : a ≡ b} {s : c ≡ d} + → PathP (λ i → p i ≡ q i) r s + → r ≡ p ∙ s ∙ sym q +PathP→compPathR {p = p} {q = q} {r = r} {s = s} P j i = + hcomp (λ k → λ { (i = i0) → p (j ∧ (~ k)) + ; (i = i1) → q (j ∧ (~ k)) + ; (j = i0) → r i + ; (j = i1) → doubleCompPath-filler∙ p s (sym q) (~ k) i}) + (P j i) diff --git a/Cubical/ZCohomology/MayerVietorisReduced.agda b/Cubical/ZCohomology/MayerVietorisReduced.agda index 89571737a..def107b80 100644 --- a/Cubical/ZCohomology/MayerVietorisReduced.agda +++ b/Cubical/ZCohomology/MayerVietorisReduced.agda @@ -190,9 +190,10 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : (p2 : Path (_ → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) (λ a₁ → F (inr a₁)) (λ _ → ∣ north ∣)) → (δ : D) → d-pre n (λ c → ΩKn+1→Kn ((λ i₁ → p1 (~ i₁) (f c)) ∙∙ cong F (push c) - ∙∙ cong (λ F → F (g c)) p2)) δ ≡ F δ - helper zero F p1 p2 (inl x) = sym (cong (λ f → f x) p1) - helper zero F p1 p2 (inr x) = sym (cong (λ f → f x) p2) + ∙∙ cong (λ F → F (g c)) p2)) δ + ≡ F δ + helper n F p1 p2 (inl x) = sym (cong (λ f → f x) p1) + helper n F p1 p2 (inr x) = sym (cong (λ f → f x) p2) helper zero F p1 p2 (push a i) j = hcomp (λ k → λ { (i = i0) → p1 (~ j) (f a) ; (i = i1) → p2 (~ j) (g a) @@ -201,8 +202,6 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : ∙∙ cong (λ F₁ → F₁ (g a)) p2) (~ k) i ; (j = i1) → F (push a i)}) (pushFiller zero F p1 p2 a j i) - helper (suc n) F p1 p2 (inl x) = sym (cong (λ f → f x) p1) - helper (suc n) F p1 p2 (inr x) = sym (cong (λ f → f x) p2) helper (suc n) F p1 p2 (push a i) j = hcomp (λ k → λ { (i = i0) → p1 (~ j) (f a) ; (i = i1) → p2 (~ j) (g a) @@ -289,49 +288,67 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : → isInIm (×coHomGr n A B) (coHomGr n C) (Δ n) a Ker-d⊂Im-Δ n = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) - λ Fa p → pRec propTruncIsProp (λ p → ∣ {!!} , {!!} ∣₋₁) (Iso.fun (PathIdTrunc₀Iso) p) + λ Fc p → pRec propTruncIsProp (λ p → ∣ (∣ (λ a → ΩKn+1→Kn (cong (λ f → f (inl a)) p)) ∣₀ , + ∣ (λ b → ΩKn+1→Kn (cong (λ f → f (inr b)) p)) ∣₀) , + Iso.inv (PathIdTrunc₀Iso) ∣ funExt (λ c → helper2 n Fc p c) ∣₋₁ ∣₋₁) + (Iso.fun (PathIdTrunc₀Iso) p) where - - - -- Im-Δ⊂Ker-d : (n : ℕ) (a : coHom n C) - -- → isInIm (×coHomGr n A B) (coHomGr n C) (Δ n) a - -- → isInKer (coHomGr n C) (coHomGr (suc n) D) (d n) a - -- Im-Δ⊂Ker-d n = - -- sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) - -- λ Fc → pRec (isOfHLevelPath' 1 setTruncIsSet _ _) - -- (sigmaProdElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) - -- λ Fa Fb p → pRec (isOfHLevelPath' 1 setTruncIsSet _ _) - -- (λ q → ((λ i → d n ∣ (q (~ i)) ∣₀) ∙ dΔ-Id n Fa Fb)) - -- (Iso.fun (PathIdTrunc₀Iso) p)) - - -- where - - -- d-preLeftId : (n : ℕ) (Fa : A → coHomK n)(d : D) - -- → d-pre n (Fa ∘ f) d ≡ 0ₖ - -- d-preLeftId n Fa (inl x) = Kn→ΩKn+1 n (Fa x) - -- d-preLeftId n Fa (inr x) = refl - -- d-preLeftId zero Fa (push a i) j = Kn→ΩKn+1 zero (Fa (f a)) (j ∨ i) - -- d-preLeftId (suc n) Fa (push a i) j = Kn→ΩKn+1 (suc n) (Fa (f a)) (j ∨ i) - - -- d-preRightId : (n : ℕ) (Fb : B → coHomK n) (d : D) - -- → d-pre n (Fb ∘ g) d ≡ 0ₖ - -- d-preRightId n Fb (inl x) = refl - -- d-preRightId n Fb (inr x) = sym (Kn→ΩKn+1 n (Fb x)) - -- d-preRightId zero Fb (push a i) j = Kn→ΩKn+1 zero (Fb (g a)) (~ j ∧ i) - -- d-preRightId (suc n) Fb (push a i) j = Kn→ΩKn+1 (suc n) (Fb (g a)) (~ j ∧ i) - - -- dΔ-Id : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) - -- → d n (Δ n (∣ Fa ∣₀ , ∣ Fb ∣₀)) ≡ 0ₕ - -- dΔ-Id zero Fa Fb = - -- dIsHom zero ∣ (λ x → Fa (f x)) ∣₀ (-ₕ ∣ (λ x → Fb (g x)) ∣₀) - -- ∙ (λ i → d zero ∣ (λ x → Fa (f x)) ∣₀ +ₕ morphMinus (coHomGr zero C) (coHomGr 1 D) (d zero) (dIsHom zero) ∣ (λ x → Fb (g x)) ∣₀ i) - -- ∙ (λ i → d zero ∣ (λ x → Fa (f x)) ∣₀ +ₕ -ₕ (d zero ∣ (λ x → Fb (g x)) ∣₀)) - -- ∙ (λ i → ∣ funExt (d-preLeftId zero Fa) i ∣₀ +ₕ -ₕ ∣ funExt (d-preRightId zero Fb) i ∣₀) - -- ∙ rCancelₕ 0ₕ - -- dΔ-Id (suc n) Fa Fb = - -- dIsHom (suc n) ∣ (λ x → Fa (f x)) ∣₀ (-ₕ ∣ (λ x → Fb (g x)) ∣₀) - -- ∙ (λ i → d (suc n) ∣ (λ x → Fa (f x)) ∣₀ +ₕ morphMinus (coHomGr (suc n) C) (coHomGr (2 + n) D) (d (suc n)) (dIsHom (suc n)) ∣ (λ x → Fb (g x)) ∣₀ i) - -- ∙ (λ i → d (suc n) ∣ (λ x → Fa (f x)) ∣₀ +ₕ -ₕ (d (suc n) ∣ (λ x → Fb (g x)) ∣₀)) - -- ∙ (λ i → ∣ funExt (d-preLeftId (suc n) Fa) i ∣₀ +ₕ -ₕ ∣ funExt (d-preRightId (suc n) Fb) i ∣₀) - -- ∙ rCancelₕ 0ₕ + distrHelper : (n : ℕ) (p q : _) → ΩKn+1→Kn {n = n} p +ₖ (-ₖ ΩKn+1→Kn q) ≡ ΩKn+1→Kn (p ∙ sym q) + distrHelper n p q i = ΩKn+1→Kn (Iso.rightInv (Iso3-Kn-ΩKn+1 n) p i ∙ Iso.rightInv (Iso3-Kn-ΩKn+1 n) (sym (Iso.rightInv (Iso3-Kn-ΩKn+1 n) q i)) i) + + + helper2 : (n : ℕ) (Fc : C → coHomK n) (p : d-pre n Fc ≡ (λ _ → ∣ north ∣)) (c : C) + → ΩKn+1→Kn (λ i₁ → p i₁ (inl (f c))) +ₖ (-ₖ ΩKn+1→Kn (λ i₁ → p i₁ (inr (g c)))) ≡ Fc c + helper2 zero Fc p c = distrHelper zero (cong (λ F → F (inl (f c))) p) (cong (λ F → F (inr (g c))) p) + ∙ cong ΩKn+1→Kn (sym ((PathP→compPathR (cong (λ f → cong f (push c)) p)) + ∙ (λ i → (λ i₁ → p i₁ (inl (f c))) + ∙ (lUnit (sym (λ i₁ → p i₁ (inr (g c)))) (~ i))))) + ∙ Iso.leftInv (Iso3-Kn-ΩKn+1 zero) (Fc c) + helper2 (suc n) Fc p c = distrHelper (suc n) (cong (λ F → F (inl (f c))) p) (cong (λ F → F (inr (g c))) p) + ∙ cong ΩKn+1→Kn (sym ((PathP→compPathR (cong (λ f → cong f (push c)) p)) + ∙ (λ i → (λ i₁ → p i₁ (inl (f c))) + ∙ (lUnit (sym (λ i₁ → p i₁ (inr (g c)))) (~ i))))) + ∙ Iso.leftInv (Iso3-Kn-ΩKn+1 (suc n)) (Fc c) + + Im-Δ⊂Ker-d : (n : ℕ) (a : coHom n C) + → isInIm (×coHomGr n A B) (coHomGr n C) (Δ n) a + → isInKer (coHomGr n C) (coHomGr (suc n) D) (d n) a + Im-Δ⊂Ker-d n = + sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ Fc → pRec (isOfHLevelPath' 1 setTruncIsSet _ _) + (sigmaProdElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ Fa Fb p → pRec (isOfHLevelPath' 1 setTruncIsSet _ _) + (λ q → ((λ i → d n ∣ (q (~ i)) ∣₀) ∙ dΔ-Id n Fa Fb)) + (Iso.fun (PathIdTrunc₀Iso) p)) + + where + + d-preLeftId : (n : ℕ) (Fa : A → coHomK n)(d : D) + → d-pre n (Fa ∘ f) d ≡ 0ₖ + d-preLeftId n Fa (inl x) = Kn→ΩKn+1 n (Fa x) + d-preLeftId n Fa (inr x) = refl + d-preLeftId zero Fa (push a i) j = Kn→ΩKn+1 zero (Fa (f a)) (j ∨ i) + d-preLeftId (suc n) Fa (push a i) j = Kn→ΩKn+1 (suc n) (Fa (f a)) (j ∨ i) + + d-preRightId : (n : ℕ) (Fb : B → coHomK n) (d : D) + → d-pre n (Fb ∘ g) d ≡ 0ₖ + d-preRightId n Fb (inl x) = refl + d-preRightId n Fb (inr x) = sym (Kn→ΩKn+1 n (Fb x)) + d-preRightId zero Fb (push a i) j = Kn→ΩKn+1 zero (Fb (g a)) (~ j ∧ i) + d-preRightId (suc n) Fb (push a i) j = Kn→ΩKn+1 (suc n) (Fb (g a)) (~ j ∧ i) + + dΔ-Id : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) + → d n (Δ n (∣ Fa ∣₀ , ∣ Fb ∣₀)) ≡ 0ₕ + dΔ-Id zero Fa Fb = + dIsHom zero ∣ (λ x → Fa (f x)) ∣₀ (-ₕ ∣ (λ x → Fb (g x)) ∣₀) + ∙ (λ i → d zero ∣ (λ x → Fa (f x)) ∣₀ +ₕ morphMinus (coHomGr zero C) (coHomGr 1 D) (d zero) (dIsHom zero) ∣ (λ x → Fb (g x)) ∣₀ i) + ∙ (λ i → d zero ∣ (λ x → Fa (f x)) ∣₀ +ₕ -ₕ (d zero ∣ (λ x → Fb (g x)) ∣₀)) + ∙ (λ i → ∣ funExt (d-preLeftId zero Fa) i ∣₀ +ₕ -ₕ ∣ funExt (d-preRightId zero Fb) i ∣₀) + ∙ rCancelₕ 0ₕ + dΔ-Id (suc n) Fa Fb = + dIsHom (suc n) ∣ (λ x → Fa (f x)) ∣₀ (-ₕ ∣ (λ x → Fb (g x)) ∣₀) + ∙ (λ i → d (suc n) ∣ (λ x → Fa (f x)) ∣₀ +ₕ morphMinus (coHomGr (suc n) C) (coHomGr (2 + n) D) (d (suc n)) (dIsHom (suc n)) ∣ (λ x → Fb (g x)) ∣₀ i) + ∙ (λ i → d (suc n) ∣ (λ x → Fa (f x)) ∣₀ +ₕ -ₕ (d (suc n) ∣ (λ x → Fb (g x)) ∣₀)) + ∙ (λ i → ∣ funExt (d-preLeftId (suc n) Fa) i ∣₀ +ₕ -ₕ ∣ funExt (d-preRightId (suc n) Fb) i ∣₀) + ∙ rCancelₕ 0ₕ From 9633b5256a700f3f8db0118ba95c58b9387b386f Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 19 May 2020 17:17:46 +0200 Subject: [PATCH 34/89] finalbackup --- Cubical/Foundations/GroupoidLaws.agda | 34 +- Cubical/Foundations/Prelude.agda | 47 ++ Cubical/HITs/Truncation/Properties.agda | 2 +- Cubical/Homotopy/Connected.agda | 552 ++++++++++++++++-------- 4 files changed, 445 insertions(+), 190 deletions(-) diff --git a/Cubical/Foundations/GroupoidLaws.agda b/Cubical/Foundations/GroupoidLaws.agda index 2f2c9a10d..257943b26 100644 --- a/Cubical/Foundations/GroupoidLaws.agda +++ b/Cubical/Foundations/GroupoidLaws.agda @@ -280,11 +280,30 @@ hcomp-cong : ∀ {ℓ} {A : Set ℓ} {φ} → (u : I → Partial φ A) → (u0 : hcomp-cong u u0 u' u0' ueq 0eq = inS (\ j → hcomp (\ i o → ueq i o j) (outS 0eq j)) --- - +{- congFunct : ∀ {ℓ} {B : Type ℓ} → (f : A → B) (p : x ≡ y) (q : y ≡ z) → cong f (p ∙ q) ≡ cong f p ∙ cong f q congFunct f p q i = hcomp (λ j → λ{(i = i0) → rUnit (cong f (p ∙ q)) (~ j) ; (i = i1) → cong f (rUnit p (~ j)) ∙ cong f q}) (cong f (p ∙ (λ k → q (k ∧ (~ i)))) ∙ cong f λ k → q ((~ i) ∨ k) ) +-} + +congFunct-filler : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y z : A} (f : A → B) (p : x ≡ y) (q : y ≡ z) + → I → I → I → B +congFunct-filler {x = x} f p q i j z = + hfill (λ k → λ { (i = i0) → f x + ; (i = i1) → f (q k) + ; (j = i0) → f (compPath-filler p q k i)}) + (inS (f (p i))) + z + +congFunct : ∀ {ℓ} {B : Type ℓ} (f : A → B) (p : x ≡ y) (q : y ≡ z) → cong f (p ∙ q) ≡ cong f p ∙ cong f q +congFunct f p q j i = congFunct-filler f p q i j i1 + + +-- congFunct for dependent types +congFunct-dep : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} {x y z : A} (f : (a : A) → B a) (p : x ≡ y) (q : y ≡ z) + → PathP (λ i → PathP (λ j → B (compPath-filler p q i j)) (f x) (f (q i))) (cong f p) (cong f (p ∙ q)) +congFunct-dep {B = B} {x = x} f p q i j = f (compPath-filler p q i j) cong₂Funct : (f : A → A → A) → (p : x ≡ y) → @@ -296,7 +315,20 @@ cong₂Funct {x = x} {y = y} f p {u = u} {v = v} q j i = ; (j = i0) → f (p i) (q (i ∧ k))}) (f (p i) u) + +symDistr-filler : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) → I → I → I → A +symDistr-filler {A = A} {z = z} p q i j k = + hfill (λ k → λ { (i = i0) → q (k ∨ j) + ; (i = i1) → p (~ k ∧ j) }) + (inS (invSides-filler q (sym p) i j)) + k + +symDistr : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) → sym (p ∙ q) ≡ sym q ∙ sym p +symDistr p q i j = symDistr-filler p q j i i1 + +{- symDistr : (p : x ≡ y) (q : y ≡ z) → sym (p ∙ q) ≡ sym q ∙ sym p symDistr p q i = hcomp (λ j → λ{(i = i0) → rUnit (sym (p ∙ q)) (~ j) ; (i = i1) → sym (lUnit q (~ j)) ∙ sym p}) (sym ((λ k → p (k ∨ i)) ∙ q) ∙ sym λ k → p (i ∧ k)) +-} diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index a0cfb5058..330a07fe0 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -156,6 +156,17 @@ compPathP {A = A} {x = x} {B = B} p q i = (i = i1) → q j }) (p i) + +compPathP' : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} {x y z : A} {x' : B x} {y' : B y} {z' : B z} {p : x ≡ y} {q : y ≡ z} + (P : PathP (λ i → B (p i)) x' y') (Q : PathP (λ i → B (q i)) y' z') + → PathP (λ i → B ((p ∙ q) i)) x' z' +compPathP' {B = B} {x' = x'} {p = p} {q = q} P Q i = + comp (λ j → B (compPath-filler p q j i)) + (λ j → λ { (i = i0) → x' ; + (i = i1) → Q j }) + (P i) + + compPathP-filler : ∀ {A : I → Type ℓ} {x : A i0} {y : A i1} {B_i1 : Type ℓ} {B : A i1 ≡ B_i1} {z : B i1} → (p : PathP A x y) (q : PathP (λ i → B i) y z) → PathP (λ j → PathP (λ k → (compPath-filler (λ i → A i) B j k)) x (q j)) p (compPathP p q) @@ -165,6 +176,17 @@ compPathP-filler {A = A} {x = x} {B = B} p q j i = (i = i1) → q j }) (inS (p i)) j + +compPathP'-filler : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} {x y z : A} {x' : B x} {y' : B y} {z' : B z} {p : x ≡ y} {q : y ≡ z} + (P : PathP (λ i → B (p i)) x' y') (Q : PathP (λ i → B (q i)) y' z') + → PathP (λ j → PathP (λ i → B (compPath-filler p q j i)) x' (Q j)) P (compPathP' {x = x} {y = y} {x' = x'} {y' = y'} P Q) +compPathP'-filler {B = B} {x' = x'} {p = p} {q = q} P Q j i = + fill (λ j → B (compPath-filler p q j i)) + (λ j → λ { (i = i0) → x' ; + (i = i1) → Q j }) + (inS (P i)) + j + _≡⟨_⟩_ : (x : A) → x ≡ y → y ≡ z → x ≡ z _ ≡⟨ x≡y ⟩ y≡z = x≡y ∙ y≡z @@ -362,6 +384,31 @@ record Lift {i j} (A : Type i) : Type (ℓ-max i j) where open Lift public + +{- filler of + q⁻¹ + z — — — > x + ^ ^ + q | | p⁻¹ + | | + x — — — > y + p +-} + +invSides-filler : {x y z : A} (p : x ≡ y) (q : x ≡ z) → PathP (λ j → q j ≡ p (~ j)) p (sym q) +invSides-filler {A = A} {x = x} p q i j = cube-filler i j i1 + where + cube-filler : I → I → I → A + cube-filler i j z = + hfill (λ k → λ { (i = i0) → p (k ∧ j) + ; (i = i1) → q (~ j ∧ k) + ; (j = i0) → q (i ∧ k) + ; (j = i1) → p (~ i ∧ k)}) + (inS x) + z + + + doubleCompPath-filler∙ : {a b c d : A} (p : a ≡ b) (q : b ≡ c) (r : c ≡ d) → PathP (λ i → p i ≡ r (~ i)) (p ∙ q ∙ r) q doubleCompPath-filler∙ {A = A} {b = b} p q r j i = diff --git a/Cubical/HITs/Truncation/Properties.agda b/Cubical/HITs/Truncation/Properties.agda index ef2058b69..ca41f3593 100644 --- a/Cubical/HITs/Truncation/Properties.agda +++ b/Cubical/HITs/Truncation/Properties.agda @@ -120,7 +120,7 @@ rec : {n : ℕ} (isOfHLevel n B) → (g : (a : A) → B) → (hLevelTrunc n A → B) -rec {B = B} h = Null.elim {B = λ _ → B} λ x → isOfHLevel→isSnNull h +rec {B = B} h = Null.rec (isOfHLevel→isSnNull h) elim : {n : ℕ} {B : hLevelTrunc n A → Type ℓ'} diff --git a/Cubical/Homotopy/Connected.agda b/Cubical/Homotopy/Connected.agda index 8251b83e4..fa224f91b 100644 --- a/Cubical/Homotopy/Connected.agda +++ b/Cubical/Homotopy/Connected.agda @@ -25,6 +25,7 @@ open import Cubical.Data.Unit open import Cubical.Data.NatMinusTwo.Base + isHLevelConnected : ∀ {ℓ} (n : ℕ) (A : Type ℓ) → Type ℓ isHLevelConnected n A = isContr (hLevelTrunc n A) @@ -53,9 +54,7 @@ private Iso.rightInv typeToFiberIso b i = fst b , (isOfHLevelSuc 1 (isPropUnit) tt tt (snd b) refl) i Iso.leftInv typeToFiberIso a = refl - module elim {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} (f : A → B) (n : ℕ) where - isIsoPrecompose : ∀ {ℓ'''} (P : B → HLevel ℓ''' n) → isHLevelConnectedFun n f → Iso ((b : B) → P b .fst) ((a : A) → P (f a) .fst) @@ -78,6 +77,9 @@ module elim {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} (f : A → B) (n : (P b .snd) (λ {(a , p) → subst (fst ∘ P) p (t a)}) + + + isEquivPrecompose : ∀ {ℓ'''} (P : B → HLevel ℓ''' n) → isHLevelConnectedFun n f → isEquiv λ(s : (b : B) → P b .fst) → s ∘ f @@ -165,6 +167,15 @@ isHLevelConnectedPath n connA a₀ a₁ = subst isContr (PathIdTrunc _) (isContr→isContrPath connA _ _) +isHLevelConnectedPath2 : ∀ {ℓ} (n : ℕ) {A : Type ℓ} + → isHLevelConnected (suc n) A + → (a₀ a₁ : A) → isHLevelConnected n (a₀ ≡ a₁) +isHLevelConnectedPath2 n connA a₀ a₁ = + Iso.fun (ΩTrunc.IsoFinal _ ∣ a₀ ∣ ∣ a₁ ∣) (sym (connA .snd ∣ a₀ ∣) ∙ (connA .snd ∣ a₁ ∣)) + , λ y → sym (subst isContr (PathIdTrunc _) + (isContr→isContrPath connA _ _) .snd _) ∙ (subst isContr (PathIdTrunc _) + (isContr→isContrPath connA _ _) .snd y) + isHLevelConnectedRetract : ∀ {ℓ ℓ'} (n : ℕ) {A : Type ℓ} {B : Type ℓ'} (f : A → B) (g : B → A) @@ -186,6 +197,15 @@ isHLevelConnectedPoint n connA a₀ a = snd (_ ,_) (λ _ → refl) (isHLevelConnectedPath n connA a₀ a) +isHLevelConnectedPoint' : ∀ {ℓ} (n : ℕ) {A : Type ℓ} + → isHLevelConnected (suc n) A + → (a : A) → isHLevelConnectedFun n (λ(_ : Unit) → a) +isHLevelConnectedPoint' n connA a₀ a = + isHLevelConnectedRetract n + snd (_ ,_) (λ _ → refl) + (isHLevelConnectedPath2 n connA a₀ a) + + isHLevelConnectedPoint2 : ∀ {ℓ} (n : ℕ) {A : Type ℓ} (a : A) → isHLevelConnectedFun n (λ(_ : Unit) → a) @@ -323,228 +343,384 @@ private --- isConnectedSmashIdfun2 : {A B C : Pointed₀} --- → (f : A →∙ B) (nf nc : ℕ) --- → isHLevelConnectedFun nf (fst f) --- → isHLevelConnected (2 + nc) (fst C) --- → isHLevelConnectedFun (1 + nc + nf) (Smash-map f (idfun∙ C)) --- isConnectedSmashIdfun2 {A = A , ptA} {B = B , ptB} {C = C , ptC} (f , fpt) nf nc conf conC = λ b → {!!} , {!!} --- where --- module proof {ℓ : Level} (P : (Smash (B , ptB) (C , ptC)) → HLevel ℓ (1 + nc + nf)) --- (d : (x : Smash (A , ptA) (C , ptC)) → P ((Smash-map (f , fpt) (idfun∙ (C , ptC)) ) x) .fst) --- where --- F : (c : C) → _ --- F c = λ(s : (b : B) → P (proj b c) .fst) → s ∘ f - --- step1 : (c : C) → isHLevelTruncatedFun (1 + nc) (F c) --- step1 c = isOfHLevelPrecomposeConnected (1 + nc) nf ((λ b → P (proj b c) .fst , P (proj b c) .snd)) f conf - --- codomFun : (c : C) (a : A) → P (proj (f a) c) .fst --- codomFun c = d ∘ λ a → proj a c - --- Q : C → HLevel _ (1 + nc) --- Q c = fiber (F c) (codomFun c) , step1 c _ - - --- helper : (a : A) → transport (λ i → P (gluel (f a) (~ i)) .fst) (d (basel)) --- ≡ d (proj a ptC) --- helper a = transport (PathP≡Path (λ i → P (gluel (f a) (~ i)) .fst) (d basel) (d (proj a ptC))) --- (transport (λ j → PathP (λ i → P (lUnit (gluel (f a)) (~ j) (~ i)) .fst) (d basel) (d (proj a ptC))) --- λ i → d (gluel a (~ i))) - --- QptC : Q ptC .fst --- QptC = (λ b → transport (λ i → P (gluel b (~ i)) .fst) (d basel)) , --- funExt helper - - --- Q-constructor : (c : C) → Q c .fst --- Q-constructor c = Iso.inv (elim.isIsoPrecompose (λ (_ : Unit) → ptC) (1 + nc) Q (isHLevelConnectedPoint _ conC ptC)) (λ _ → QptC) c - --- g : (b : B) (c : C) → P (proj b c) .fst --- g b c = Q-constructor c .fst b - --- Q-constructor-β : (b : B) → Q-constructor ptC .fst b ≡ transport (λ i → P (gluel b (~ i)) .fst) (d basel) --- Q-constructor-β b = ((λ i → (Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) (helper3 i)) .fst b)) ∙ --- (λ i → transportRefl (QptC) i .fst b) - --- where --- helper3 : (isHLevelConnectedPoint _ conC ptC) ptC .fst ≡ ∣ tt , refl ∣ --- helper3 = (isHLevelConnectedPoint _ conC ptC) ptC .snd ∣ tt , refl ∣ - --- moveTransport : ∀ {ℓ ℓ'} → {A : Type ℓ} {B : A → Type ℓ'} {x y : A} (p : x ≡ y) (x1 : B x) (y1 : B y) --- → transport (λ i → (B (p i))) x1 ≡ y1 --- → transport⁻ (λ i → (B (p i))) y1 ≡ x1 --- moveTransport {B = B} p x1 y1 transp = ((λ i → transport⁻ (λ i → B (p i)) (transp (~ i))) ∙ (transport⁻Transport (λ i → B (p i)) x1)) - --- otherway : (b : B) → transport (λ i → P (gluel b i) .fst) (g b ptC) ≡ d basel --- otherway b = (λ i → transport (λ i → P (gluel b i) .fst) (Q-constructor-β b i)) ∙ transportTransport⁻ (λ i → P (gluel b i) .fst) (d basel) - --- path1 : (b : B) → PathP (λ i → P (gluel b i) .fst) (g b ptC) (d basel) --- path1 b i = hcomp (λ k → λ {(i = i0) → g b ptC --- ; (i = i1) → otherway b k}) --- (transp (λ j → P (gluel b (i ∧ j)) .fst) (~ i) (g b ptC)) - - --- where --- massive : (c : C) → PathP _ (g ptB c) (d baser) --- massive c = compPathP (λ i → g (fpt (~ i)) c) (compPathP (λ i → Q-constructor c .snd i (ptA)) (λ i → d (gluer c i))) - --- pathPCancel : (c : C) --- → ((λ z → P (proj (fpt (~ z)) c) .fst) ∙ (λ i → ((λ _ → P (proj (f ptA) c) .fst) ∙ (λ i₁ → P (((λ j₁ → proj (fpt j₁) c) ∙ gluer c) i₁) .fst)) i)) --- ≡ (λ i → P (gluer c i) .fst) --- pathPCancel c = (λ i → (λ z → P (proj (fpt (~ z)) c) .fst) ∙ lUnit ((λ i₁ → P (((λ j₁ → proj (fpt j₁) c) ∙ gluer c) i₁) .fst)) (~ i)) --- ∙ (λ i → (λ z → P (proj (fpt (~ z ∨ i)) c) .fst) ∙ λ i₁ → P (((λ j₁ → proj (fpt (j₁ ∨ i)) c) ∙ gluer c) i₁) .fst ) --- ∙ (λ i → lUnit (λ i₁ → P (lUnit (gluer c) (~ i) i₁) .fst) (~ i)) - --- path2 : (c : C) → PathP (λ i → P (gluer c i) .fst) (g ptB c) (d baser) --- path2 c = transport (λ i → PathP (λ j → pathPCancel c i j) (g ptB c) (d baser)) (massive c) - --- h : (x : Smash (B , ptB) (C , ptC)) → P x .fst --- h basel = d basel --- h baser = d baser --- h (proj x y) = g x y --- h (gluel b i) = {!!} -- path1 b i --- h (gluer c i) = {!!} -- path2 c i - --- hsection : (x : Smash (A , ptA) (C , ptC)) → h (Smash-map (f , fpt) (idfun∙ (C , ptC)) x) ≡ d x --- hsection basel = refl --- hsection baser = refl --- hsection (proj x y) i = Q-constructor y .snd i x --- hsection (gluel a i) j = hcomp (λ k → λ { (i = i0) → transportRefl {!!} k --- ; (i = i1) → transportRefl (d basel) k --- ; (j = i0) → {!!} --- ; (j = i1) → {!!}}) --- {!hcomp!} -- hcomp {!!} {!!} --- where --- helper2 : Path ((P (proj (f a) ptC) .fst) ≡ (P basel .fst)) --- (λ i → P (Smash-map (f , fpt) (idfun∙ (C , ptC)) (gluel a i)) .fst) --- λ i → P (gluel (f a) i) .fst --- helper2 = (λ j i → P ((lUnit (gluel (f a)) (~ j)) i) .fst) --- helper6 : (a : A) → PathP ((λ j → PathP (λ i → P {!Q-constructor ? .snd i ?!} .fst) {!P (proj (f a) ptC) .fst!} {!!})) (λ i → h (Smash-map (f , fpt) (idfun∙ (C , ptC)) (gluel a i))) (path1 (f a)) --- helper6 a = ((λ j i → h ((lUnit (gluel (f a)) (~ j)) i))) - --- helper3 : PathP {!!} (path1 (f a)) (λ i → d (gluel a i)) --- helper3 = compPathP {!!} {!!} --- helper4 : PathP {!!} (path1 (f a)) λ i → d (gluel a i) --- helper4 = {!!} --- hsection (gluer b i) = {!!} +isConnectedSmashIdfun2 : {A B C : Pointed₀} + → (f : A →∙ B) (nf nc : ℕ) + → isHLevelConnectedFun nf (fst f) + → isHLevelConnected (2 + nc) (fst C) + → isHLevelConnectedFun (1 + nc + nf) (Smash-map f (idfun∙ C)) +isConnectedSmashIdfun2 {A = A , ptA} {B = B , ptB} {C = C , ptC} (f , fpt) nf nc conf conC = λ b → {!!} , {!!} + where + module proof {ℓ : Level} (P : (Smash (B , ptB) (C , ptC)) → HLevel ℓ (1 + nc + nf)) + (d : (x : Smash (A , ptA) (C , ptC)) → P ((Smash-map (f , fpt) (idfun∙ (C , ptC)) ) x) .fst) + where + F : (c : C) → _ + F c = λ(s : (b : B) → P (proj b c) .fst) → s ∘ f + step1 : (c : C) → isHLevelTruncatedFun (1 + nc) (F c) + step1 c = isOfHLevelPrecomposeConnected (1 + nc) nf ((λ b → P (proj b c) .fst , P (proj b c) .snd)) f conf --- -- isConnectedSmashIdfun : (f : A' →∙ B') (nf nc : ℕ) --- -- → isHLevelConnectedFun nf (fst f) --- -- → isHLevelConnected (2 + nc) (fst C') --- -- → isHLevelConnectedFun (1 + nf + nc) (f ⋀⃗ idfun∙ C') --- -- isConnectedSmashIdfun {A' = (A , ptA)} {B' = (B , ptB)} {C' = (C , ptC)} (f , fpt) nf nc conf conC = {!isHLel!} --- -- where --- -- module _ (P : ((B , ptB) ⋀ (C , ptC)) → HLevel (ℓ-max (ℓ-max ℓ ℓ') ℓ'') (1 + nc + nf)) --- -- (d : (x : (A , ptA) ⋀ (C , ptC)) → P (((f , fpt) ⋀⃗ idfun∙ (C , ptC) ) x) .fst) --- -- where --- -- F : (c : C) → _ --- -- F c = λ(s : (b : B) → P (inr (b , c)) .fst) → s ∘ f + codomFun : (c : C) (a : A) → P (proj (f a) c) .fst + codomFun c = d ∘ λ a → proj a c --- -- step1 : (c : C) → isHLevelTruncatedFun (1 + nc) (F c) --- -- step1 c = isOfHLevelPrecomposeConnected (1 + nc) nf ((λ b → P (inr (b , c)) .fst , P (inr (b , c)) .snd)) f conf + Q : C → HLevel _ (1 + nc) + Q c = fiber (F c) (codomFun c) , step1 c _ --- -- codomFun : (c : C) (a : A) → P (inr ((f a) , c)) .fst --- -- codomFun c = d ∘ λ a → inr (a , c) --- -- Q : C → HLevel _ (1 + nc) --- -- Q c = fiber (F c) (codomFun c) , step1 c _ + helper : (a : A) → transport (λ i → P (gluel (f a) (~ i)) .fst) (d (basel)) + ≡ d (proj a ptC) + helper a i = + hcomp (λ k → λ {(i = i0) → transport (λ j → P (compPath-filler refl (gluel (f a)) (~ j) (~ j ∨ k)) .fst) (d basel) + ; (i = i1) → d (proj a ptC)}) + (transp (λ j → P (compPath-filler refl (gluel (f a)) (~ j) (~ i ∧ (~ j))) .fst) i (d (gluel a (~ i)))) +{- +i = i0 ⊢ transport (λ i₁ → P (gluel (f a) (~ i₁)) .fst) (d basel) +i = i1 ⊢ d (proj a ptC) +-} + {- transport (PathP≡Path (λ i → P (gluel (f a) (~ i)) .fst) (d basel) (d (proj a ptC))) + (transport (λ j → PathP (λ i → P (lUnit (gluel (f a)) (~ j) (~ i)) .fst) (d basel) (d (proj a ptC))) + λ i → d (gluel a (~ i))) -} --- -- helper : (a : A) → transport (λ i → P (push (inl (f a)) i) .fst) (d (inl tt)) --- -- ≡ d (inr (a , ptC)) --- -- helper a = transport (PathP≡Path (λ i → P (push (inl (f a)) i) .fst) (d (inl tt)) (d (inr (a , ptC)))) --- -- (transport (λ j → PathP (λ i → P (rUnit (push (inl (f a))) (~ j) i) .fst) (d (inl tt)) (d (inr (a , ptC)))) --- -- λ i → d (push (inl a) i)) + QptC : Q ptC .fst + QptC = (λ b → transport (λ i → P (gluel b (~ i)) .fst) (d basel)) , + λ i a → hcomp (λ k → λ {(i = i0) → transport (λ j → P (compPath-filler refl (gluel (f a)) (~ j) (~ j ∨ k)) .fst) (d basel) + ; (i = i1) → d (proj a ptC)}) + (transp (λ j → P (compPath-filler refl (gluel (f a)) (~ j) (~ i ∧ (~ j))) .fst) i (d (gluel a (~ i)))) --- -- QptC : Q ptC .fst --- -- QptC = (λ b → transport (λ i → P (push (inl b) i) .fst) (d (inl tt))) , --- -- funExt helper --- -- where + Q-constructor : (c : C) → Q c .fst + Q-constructor c = Iso.inv (elim.isIsoPrecompose (λ (_ : Unit) → ptC) (1 + nc) Q (isHLevelConnectedPoint _ conC ptC)) (λ _ → QptC) c + Q-constructor' : Q-constructor ptC .snd ≡ (Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) ((isHLevelConnectedPoint _ conC ptC) ptC .fst)) .snd + Q-constructor' = refl --- -- Q-constructor : (c : C) → Q c .fst --- -- Q-constructor c = Iso.inv (elim.isIsoPrecompose (λ _ → ptC) (1 + nc) Q (isHLevelConnectedPoint _ conC ptC)) (λ ( _ : Unit) → QptC) c + Q-test : isHLevelConnected (suc nc) (fiber (λ (_ : Unit) → ptC) ptC) + Q-test = ∣ tt , refl ∣ , λ y → sym ((isHLevelConnectedPoint _ conC ptC) ptC .snd _) ∙ (isHLevelConnectedPoint _ conC ptC) ptC .snd y --- -- g : (b : B) (c : C) → P (inr (b , c)) .fst --- -- g b c = Q-constructor c .fst b + hLevelContr : isContr (isHLevelConnected (suc nc) (fiber (λ (_ : Unit) → ptC) ptC)) + hLevelContr = Q-test , isPropIsOfHLevel 0 Q-test --- -- Q-constructor-β : (b : B) → Q-constructor ptC .fst b ≡ transport (λ i → P (push (inl b) i) .fst) (d (inl tt)) --- -- Q-constructor-β b = ((λ i → (Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) (helper3 i)) .fst b)) ∙ --- -- (λ i → transportRefl (QptC) i .fst b) + helper3 : (isHLevelConnectedPoint _ conC ptC) ptC .fst ≡ ∣ tt , refl ∣ + helper3 = (isHLevelConnectedPoint _ conC ptC) ptC .snd ∣ tt , refl ∣ --- -- where --- -- helper3 : (isHLevelConnectedPoint _ conC ptC) ptC .fst ≡ ∣ tt , refl ∣ --- -- helper3 = (isHLevelConnectedPoint _ conC ptC) ptC .snd ∣ tt , refl ∣ + g : (b : B) (c : C) → P (proj b c) .fst + g b c = Q-constructor c .fst b --- -- test : (b : B) (c : C) → PathP (λ i → P (push {!!} {!(~ i)!}) .fst) (Q-constructor ptC .fst b) (Q-constructor c .fst ptB) --- -- test b = {!!} + Q-constructor-β : (b : B) → Q-constructor ptC .fst b ≡ QptC .fst b + Q-constructor-β b = sym ((λ i → transportRefl (QptC) (~ i) .fst b) ∙ + ((λ i → (Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) (helper3 (~ i))) .fst b))) --- -- Q-constructor-β2 : (c : C) → Q-constructor c .fst ptB ≡ transport (λ i → P (push (inr c) i) .fst) (d (inl tt)) --- -- Q-constructor-β2 c = ((λ i → (Trunc.rec (Q c .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) ((isHLevelConnectedPoint _ conC ptC) c .fst)) .fst ptB)) ∙ --- -- {!isHLevelConnectedPoint _ conC ptC c .fst!} + Q-constructor23 : PathP (λ i → _) (Q-constructor ptC .snd) ((Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) (Q-test .fst)) .snd) + Q-constructor23 i = ((Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) (helper3 i)) .snd) --- -- gid1 : (a : A) (c : C) → g (f a) c ≡ d (inr (a , c)) --- -- gid1 a c i = (Q-constructor c .snd) i a + Q-constructor24 : (a : A) → PathP (λ i → Q-constructor-β (f a) i ≡ d (proj a ptC)) (λ i → Q-constructor ptC .snd i a) λ i → QptC .snd i a + Q-constructor24 a i j = + hcomp (λ k → λ { (i = i0) → Q-constructor23 (~ k) j a + ; (i = i1) → QptC .snd j a + ; (j = i1) → d (proj a ptC)}) + (transportRefl (QptC) i .snd j a) --- -- gid1' : {!!} --- -- gid1' = {!!} + Q-constructor25 : (a : A) → PathP (λ i → d (proj a ptC) ≡ Q-constructor-β (f a) i) ((λ i → Q-constructor ptC .snd (~ i) a)) λ i → QptC .snd (~ i) a + Q-constructor25 a i j = Q-constructor24 a i (~ j) --- -- gid1Path : (c : C) → PathP (λ i → P (push (inr c) i) .fst) (d (inl tt)) (g ptB c) --- -- gid1Path c i = --- -- hcomp (λ k → λ {(i = i0) → {!!} ; (i = i1) → {!!}}) --- -- {!!} --- -- compPath : (c : C) → PathP _ (d (inl tt)) (g ptB c) --- -- compPath c = compPathP (λ i → d (push (inr c) i)) (compPathP (sym (gid1 ptA c)) (λ i → g (fpt i) c)) +{- +i = i0 ⊢ +i = i1 ⊢ QptC .snd j a +j = i0 ⊢ +j = i1 ⊢ d (proj a ptC) +-} + + otherway-filler : (b : B) → I → I → P basel .fst + otherway-filler b i j = + fill (λ _ → P basel .fst) + (λ k → λ {(i = i0) → transport (λ i → P (gluel b i) .fst) (g b ptC) + ; (i = i1) → transportTransport⁻ (λ i → P (gluel b i) .fst) (d basel) k}) + (inS (transport (λ i → P (gluel b i) .fst) (Q-constructor-β b i))) + j + + otherway : (b : B) → transport (λ i → P (gluel b i) .fst) (g b ptC) ≡ d basel + otherway b i = otherway-filler b i i1 + + otherway2 : (b : B) → transport (λ i → P (gluel b i) .fst) (g b ptC) ≡ d basel + otherway2 b = (λ i → transport (λ i → P (gluel b i) .fst) (Q-constructor-β b i)) + ∙ transportTransport⁻ (λ i → P (gluel b i) .fst) (d basel) + + {- (λ i → transport (λ i → P (gluel b i) .fst) (Q-constructor-β b i)) ∙ + transportTransport⁻ (λ i → P (gluel b i) .fst) (d basel) -} + + + otherwaygluel : (i : I) (a : A) → {!!} + otherwaygluel = {!!} + + + path1-filler : (b : B) → I → ((i : I) → P (gluel b i) .fst) + path1-filler b j i = + hfill + (λ k → λ {(i = i0) → g b ptC + ; (i = i1) → otherway2 b k}) + (inS (transp (λ j → P (gluel b (i ∧ j)) .fst) (~ i) (g b ptC))) + j + + path1 : (b : B) → PathP (λ i → P (gluel b i) .fst) (g b ptC) (d basel) + path1 b i = path1-filler b i1 i + + path1* : (a : A) → PathP (λ i → P (gluel (f a) i) .fst) (g (f a) ptC) (d basel) + path1* a i = + comp (λ k → P (compPath-filler refl (gluel (f a)) i k) .fst) + (λ k → λ {(i = i0) → g (f a) ptC + ; (i = i1) → d (gluel a k)}) + ((cong (λ F → F a) (Q-constructor ptC .snd)) i) + + +{- +Goal: P (gluel (f a) i) .fst +———— Boundary —————————————————————————————————————————————— +i = i0 ⊢ g (f a) ptC +i = i1 ⊢ d basel +-} + + path1*-filler : (a : A) (i : I) (j : I) → P (compPath-filler (λ _ → proj (f a) ptC) (gluel (f a)) i j) .fst + path1*-filler a i j = + fill (λ k → P (compPath-filler refl (gluel (f a)) i k) .fst) + (λ k → λ {(i = i0) → g (f a) ptC + ; (i = i1) → d (gluel a k)}) + (inS ( ((Q-constructor ptC .snd i a)))) + j + + path2 : (c : C) → PathP (λ i → P (gluer c i) .fst) (g ptB c) (d baser) + path2 c i = + comp (λ k → P (compPath-filler (λ i → proj (fpt i) c) (gluer c) i k) .fst) + (λ k → λ { (i = i0) → g (fpt k) c + ; (i = i1) → d (gluer c k) }) + (Q-constructor c .snd i ptA) + + + + + +{- +j = i0 ⊢ path1 (f a) i +j = i1 ⊢ path1* a i +i = i0 ⊢ g (f a) ptC +i = i1 ⊢ d basel +-} + + + h : (x : Smash (B , ptB) (C , ptC)) → P x .fst + h basel = d basel + h baser = d baser + h (proj x y) = g x y + h (gluel b i) = path1 b i + h (gluer c i) = path2 c i + + hsection : (x : Smash (A , ptA) (C , ptC)) → h (Smash-map (f , fpt) (idfun∙ (C , ptC)) x) ≡ d x + hsection basel = refl + hsection baser = refl + hsection (proj x y) i = Q-constructor y .snd i x + hsection (gluel a i) j = + comp (λ k → P (lUnit (gluel (f a)) (j ∨ k) i) .fst) + (λ k → λ { (i = i0) → Q-constructor ptC .snd j a + ; (i = i1) → d basel + ; (j = i0) → h (lUnit (gluel (f a)) k i) + ; (j = i1) → {!d (gluel a i)!}}) + {!!} + {- + comp (λ k → P (compPath-filler refl (gluel (f a)) (j ∨ k) i) .fst) + (λ k → λ { (i = i0) → Q-constructor ptC .snd j a + ; (i = i1) → path1 (f a) (k ∨ j) + ; (j = i0) → h (compPath-filler refl (gluel (f a)) k i) + ; (j = i1) → d (gluel a i)}) + {!!} -} + {- + comp (λ k → P (compPath-filler refl (gluel (f a)) (k ∨ j) i) .fst) + (λ k → λ { (i = i0) → Q-constructor ptC .snd j a + ; (i = i1) → {!path1 (f a) ?!} -- path1-filler (f a) k (k ∨ j) -- path1 (f a) (k ∨ j) + ; (j = i0) → h (compPath-filler refl (gluel (f a)) k i) + ; (j = i1) → {!!}}) -- d (gluel a i) + -- transp (λ r → P (compPath-filler refl (gluel (f a)) (r ∨ k) i) .fst) k ((transp (λ r → P (compPath-filler refl (gluel (f a)) (~ r ∨ k) i) .fst)) k (d (gluel a i))) + {!Q-constructor-β (f a)!} + -} + where + test2 : I → (i : I) → P ((refl ∙ gluel (f a)) i) .fst + test2 j i = + comp (λ k → P (compPath-filler refl (gluel (f a)) (j ∨ k) i) .fst) + (λ k → λ { (i = i0) → d (proj a ptC) -- g (f a) ptC + ; (i = i1) → transp (λ r → P (gluel (f a) (r ∧ (k ∨ j))) .fst) {!j!} (Q-constructor-β (f a) j) -- transp (λ r → P (gluel (f a) (r ∧ k)) .fst) (~ k) (Q-constructor-β (f a) j) + ; (j = i0) → compPathP'-filler (λ i → Q-constructor ptC .snd (~ i) a) (λ i → transp (λ r → P (gluel (f a) (r ∧ i)) .fst) (~ i) (g (f a) ptC)) k i + ; (j = i1) → d (gluel a i) }) -- d (gluel a i)}) + {!!} + + + {- + (comp (λ k → P (compPath-filler refl (gluel (f a)) (j ∧ k) i) .fst) + (λ k → λ { (i = i0) → Q-constructor ptC .snd j a + ; (i = i1) → path1-filler (f a) k (j ∧ k) + ; (j = i0) → g (f a) ptC + ; (j = i1) → side-filler i k}) + {!compPath-filler refl (gluel (f a)) i0 j!}) -} + + side-filler : (i : I) (k : I) → P (compPath-filler (λ _ → proj (f a) (pt (C , ptC))) (gluel (f a)) k i) .fst + side-filler i k = + comp (λ l → P (compPath-filler refl (gluel (f a)) k {!!}) .fst) + (λ l → λ { (i = i0) → {!transport (λ i₂ → P (gluel (f a) ₂) .fst) (g (f a) ptC)!} + ; (i = i1) → {!!} + ; (k = i0) → {!!} + ; (k = i1) → {!!} }) + {!otherway (f a) ?!} + hsection (gluer c i) j = {!transp (λ r → P (((refl ∙ gluel (f a)) ((r ∨ ?))) .fst)) l (transp (λ r → P (((refl ∙ gluel (f a)) ((~ r) ∨ ?))) .fst) l (d (gluel a i))) !} + {- + comp (λ k → P (compPath-filler (λ i → proj (fpt i) c) (gluer c) (k ∨ j) i) .fst) + (λ k → λ { (i = i0) → Q-constructor c .snd j ptA + ; (i = i1) → path2 c (k ∨ j) + ; (j = i0) → h (compPath-filler (λ i → proj (fpt i) c) (gluer c) k i) + ; (j = i1) → d (gluer c i) + }) + (comp (λ k → P (compPath-filler (λ i₁ → proj (fpt i₁) c) (gluer c) (j ∧ k) (i ∧ k)) .fst) + (λ k → λ { (i = i0) → Q-constructor c .snd j ptA + ; (j = i0) → h (proj (fpt (i ∧ k)) c) + ; (j = i1) → d (gluer c (i ∧ k)) + }) + (Q-constructor c .snd j ptA)) + -} + +-- isConnectedSmashIdfun : (f : A' →∙ B') (nf nc : ℕ) +-- → isHLevelConnectedFun nf (fst f) +-- → isHLevelConnected (2 + nc) (fst C') +-- → isHLevelConnectedFun (1 + nf + nc) (f ⋀⃗ idfun∙ C') +-- isConnectedSmashIdfun {A' = (A , ptA)} {B' = (B , ptB)} {C' = (C , ptC)} (f , fpt) nf nc conf conC = {!isHLel!} +-- where +-- module _ (P : ((B , ptB) ⋀ (C , ptC)) → HLevel (ℓ-max (ℓ-max ℓ ℓ') ℓ'') (1 + nc + nf)) +-- (d : (x : (A , ptA) ⋀ (C , ptC)) → P (((f , fpt) ⋀⃗ idfun∙ (C , ptC) ) x) .fst) +-- where +-- F : (c : C) → _ +-- F c = λ(s : (b : B) → P (inr (b , c)) .fst) → s ∘ f + +-- step1 : (c : C) → isHLevelTruncatedFun (1 + nc) (F c) +-- step1 c = isOfHLevelPrecomposeConnected (1 + nc) nf ((λ b → P (inr (b , c)) .fst , P (inr (b , c)) .snd)) f conf --- -- compPathTransport : (c : C) → ((λ z → P ((push (inr c) ∙ --- -- (λ i → inr (fpt (~ i) , c))) z) .fst) ∙ --- -- (λ i → ((λ i₁ → P (inr (f ptA , c)) .fst) ∙ (λ i₁ → P (inr (fpt i₁ , c)) .fst)) i)) --- -- ≡ λ i → P (push (inr c) i) .fst --- -- compPathTransport c = (λ k → ((λ z → P ((push (inr c) ∙ --- -- (λ i → inr (fpt (~ i) , c))) z) .fst) ∙ --- -- ((lUnit (λ i₁ → P (inr (fpt i₁ , c)) .fst) (~ k) )))) --- -- ∙ (λ k → ((λ z → P ((push (inr c) ∙ --- -- (λ i → inr (fpt (~ i ∨ k) , c))) z) .fst) ∙ --- -- (λ i₁ → P (inr (fpt (i₁ ∨ k) , c)) .fst) )) --- -- ∙ (λ k → ((λ z → P ((push (inr c) ∙ refl) z) .fst) ∙ refl)) --- -- ∙ (λ k → rUnit ((λ z → P ((rUnit (push (inr c)) (~ k)) z) .fst)) (~ k)) +-- codomFun : (c : C) (a : A) → P (inr ((f a) , c)) .fst +-- codomFun c = d ∘ λ a → inr (a , c) --- -- compPathTransport-ptC : compPathTransport ptC ≡ {!!} --- -- compPathTransport-ptC = {!!} +-- Q : C → HLevel _ (1 + nc) +-- Q c = fiber (F c) (codomFun c) , step1 c _ --- -- compPathTrue : (c : C) → PathP (λ i → P (push (inr c) i) .fst) (d (inl tt)) (g ptB c) --- -- compPathTrue c = transport (λ i → PathP (λ j → compPathTransport c i j) (d (inl tt)) (g ptB c)) (compPath c) --- -- gid1transp : (c : C) → PathP {!!} (g ptB c) (transport (λ i → P (push (inr c) i) .fst) (d (inl tt))) --- -- gid1transp c = compPathP (λ i → g (fpt (~ i)) c) (compPathP (gid1 ptA c) (compPathP (λ i → d (push (inr c) (~ i))) λ i → transp (λ j → P (push (inr c) (i ∧ j)) .fst) (~ i) (d (inl tt)))) +-- helper : (a : A) → transport (λ i → P (push (inl (f a)) i) .fst) (d (inl tt)) +-- ≡ d (inr (a , ptC)) +-- helper a = transport (PathP≡Path (λ i → P (push (inl (f a)) i) .fst) (d (inl tt)) (d (inr (a , ptC)))) +-- (transport (λ j → PathP (λ i → P (rUnit (push (inl (f a))) (~ j) i) .fst) (d (inl tt)) (d (inr (a , ptC)))) +-- λ i → d (push (inl a) i)) --- -- gid2 : (b : B) → g b ptC ≡ transport (λ i → P (push (inl b) i) .fst) (d (inl tt)) -- (λ i → P (push (inl b) i) .fst) (d (inl tt)) --- -- gid2 b = Q-constructor-β b +-- QptC : Q ptC .fst +-- QptC = (λ b → transport (λ i → P (push (inl b) i) .fst) (d (inl tt))) , +-- funExt helper +-- where --- -- gid1≡gid2 : (c : C) (a : A) → PathP {!gid1transp c !} (gid1 a ptC) (gid2 (f a)) --- -- gid1≡gid2 c a = compPathP {!λ j → (Q-constructor ptC .snd) i a!} (compPathP {!!} {!!}) +-- Q-constructor : (c : C) → Q c .fst +-- Q-constructor c = Iso.inv (elim.isIsoPrecompose (λ _ → ptC) (1 + nc) Q (isHLevelConnectedPoint _ conC ptC)) (λ ( _ : Unit) → QptC) c +-- g : (b : B) (c : C) → P (inr (b , c)) .fst +-- g b c = Q-constructor c .fst b --- -- gid2Path : (b : B) → PathP (λ i → P (push (inl b) i) .fst) (d (inl tt)) (g b ptC) --- -- gid2Path b i = --- -- hcomp (λ k → λ {(i = i0) → (d (inl tt)) ; (i = i1) → gid2 b (~ k)}) --- -- (transp (λ j → P (push (inl b) (j ∧ i)) .fst) (~ i) (d (inl tt))) +-- Q-constructor-β : (b : B) → Q-constructor ptC .fst b ≡ transport (λ i → P (push (inl b) i) .fst) (d (inl tt)) +-- Q-constructor-β b = ((λ i → (Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) (helper3 i)) .fst b)) ∙ +-- (λ i → transportRefl (QptC) i .fst b) --- -- PathP2 : PathP {!!} ((gid2Path ptB)) (compPathTrue ptC) --- -- PathP2 = compPathP {!gid2Path ptB!} (compPathP {!!} {!compPathTrue ptC!}) +-- where +-- helper3 : (isHLevelConnectedPoint _ conC ptC) ptC .fst ≡ ∣ tt , refl ∣ +-- helper3 = (isHLevelConnectedPoint _ conC ptC) ptC .snd ∣ tt , refl ∣ +-- gid1 : (a : A) (c : C) → g (f a) c ≡ d (inr (a , c)) +-- gid1 a c i = (Q-constructor c .snd) i a + +-- gid2 : (b : B) → g b ptC ≡ transport (λ i → P (push (inl b) i) .fst) (d (inl tt)) +-- gid2 b = Q-constructor-β b + +-- {- +-- compPath : (c : C) → PathP _ (d (inl tt)) (g ptB c) +-- compPath c = compPathP (λ i → d (push (inr c) i)) (compPathP (sym (gid1 ptA c)) (λ i → g (fpt i) c)) + + +-- compPathTransport : (c : C) → ((λ z → P ((push (inr c) ∙ +-- (λ i → inr (fpt (~ i) , c))) z) .fst) ∙ +-- (λ i → ((λ i₁ → P (inr (f ptA , c)) .fst) ∙ (λ i₁ → P (inr (fpt i₁ , c)) .fst)) i)) +-- ≡ λ i → P (push (inr c) i) .fst +-- compPathTransport c = (λ k → ((λ z → P ((push (inr c) ∙ +-- (λ i → inr (fpt (~ i) , c))) z) .fst) ∙ +-- ((lUnit (λ i₁ → P (inr (fpt i₁ , c)) .fst) (~ k) )))) +-- ∙ (λ k → ((λ z → P ((push (inr c) ∙ +-- (λ i → inr (fpt (~ i ∨ k) , c))) z) .fst) ∙ +-- (λ i₁ → P (inr (fpt (i₁ ∨ k) , c)) .fst) )) +-- ∙ (λ k → ((λ z → P ((push (inr c) ∙ refl) z) .fst) ∙ refl)) +-- ∙ (λ k → rUnit ((λ z → P ((rUnit (push (inr c)) (~ k)) z) .fst)) (~ k)) +-- -} +-- compP-filler : (c : C) → I → I → I → (B , ptB) ⋀ (C , ptC) +-- compP-filler c i j z = +-- hfill (λ k → λ { (i = i0) → inr (fpt (~ j ∧ ~ k) , c) +-- ; (i = i1) → push (inr c) j +-- ; (j = i1) → inr (fpt i , c)}) +-- (inS (invSides-filler (λ i → inr (fpt (~ i) , c)) +-- (sym (push (inr c))) i j)) +-- z + + +-- gid2-filler : C → I → I → {!!} +-- gid2-filler c i j = +-- fill (λ l → P (compP-filler c l i i1) .fst) +-- (λ k → λ { (i = i0) → d (push (inr c) (~ k)) +-- ; (i = i1) → g (fpt k) c }) +-- (inS (gid1 ptA c (~ i))) +-- {!j!} + +-- gid1Path : (c : C) → PathP (λ i → P (push (inr c) i) .fst) (d (inl tt)) (g ptB c) +-- gid1Path c i = comp (λ j → P (compP-filler c j i i1) .fst) +-- (λ k → λ { (i = i0) → d (push (inr c) (~ k)) +-- ; (i = i1) → g (fpt k) c }) +-- (gid1 ptA c (~ i)) + + + +-- gid2Path : (b : B) → PathP (λ i → P (push (inl b) i) .fst) (d (inl tt)) (g b ptC) +-- gid2Path b i = +-- comp (λ _ → P (push (inl b) i) .fst) +-- (λ k → λ {(i = i0) → (d (inl tt)) +-- ; (i = i1) → gid2 b (~ k)}) +-- (transp (λ j → P (push (inl b) (j ∧ i)) .fst) (~ i) (d (inl tt))) + +-- gid1Path≡gid2Path : PathP (λ j → PathP (λ i → P (push (push tt (~ j)) i) .fst) (d (inl tt)) (g ptB ptC)) (gid1Path ptC) (gid2Path ptB) +-- gid1Path≡gid2Path j i = +-- comp (λ k → P {!Q-constructor-β ptB!} .fst) +-- (λ k → λ { (i = i0) → {!!} -- Q-constructor-β (fpt (~ k)) j +-- ; (i = i1) → {!!} -- d (push (inr ptC) (~ k ∨ j)) +-- ; (j = i0) → {!!} +-- ; (j = i1) → {!!} }) +-- {!gid2!} + +-- -- -- {- +-- -- -- Goal: P (push (push tt (~ j)) i) .fst +-- -- -- ———— Boundary —————————————————————————————————————————————— +-- -- -- j = i0 ⊢ gid1path ptC i +-- -- -- j = i1 ⊢ gid2Path ptB i +-- -- -- i = i0 ⊢ d (inl tt) +-- -- -- i = i1 ⊢ g ptB ptC +-- -- -- -} + -- -- h : (x : (B , ptB) ⋀ (C , ptC)) → P x .fst -- -- h (inl x) = d (inl tt) -- -- h (inr (b , c)) = g b c -- -- h (push (inl b) i) = gid2Path b i --- -- h (push (inr x) i) = compPathTrue x i --- -- h (push (push a i₁) i) = {!!} +-- -- h (push (inr x) i) = gid1Path x i +-- -- h (push (push tt i) j) = gid1Path≡gid2Path (~ i) j + +-- -- sect-h : (x : (A , ptA) ⋀ (C , ptC)) → h (((f , fpt) ⋀⃗ idfun∙ (C , ptC)) x) ≡ d x +-- -- sect-h (inl x) = refl +-- -- sect-h (inr (x , x₁)) i = gid1 x x₁ i -- (Q-constructor x₁ .snd) i x +-- -- sect-h (push (inl x) i) j = {!!} +-- -- sect-h (push (inr x) i) = {!!} +-- -- sect-h (push (push a i) j) k = {!!} From 3201417538d3bbf5fa6cf87d76f7f12e6002744b Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 26 May 2020 13:42:21 +0200 Subject: [PATCH 35/89] backup before merge --- Cubical/Data/Group/Base.agda | 110 ++-- Cubical/Data/Group/GroupLibrary.agda | 217 +++++++ Cubical/Data/Unit/Properties.agda | 4 + Cubical/HITs/S1/Properties.agda | 4 + Cubical/HITs/SetQuotients/Properties.agda | 26 + Cubical/HITs/Truncation/Properties.agda | 4 +- Cubical/Homotopy/Connected.agda | 703 +++++++++++----------- Cubical/ZCohomology/KcompPrelims.agda | 6 +- Cubical/ZCohomology/Sn/Sn.agda | 211 +++++++ 9 files changed, 894 insertions(+), 391 deletions(-) create mode 100644 Cubical/Data/Group/GroupLibrary.agda create mode 100644 Cubical/ZCohomology/Sn/Sn.agda diff --git a/Cubical/Data/Group/Base.agda b/Cubical/Data/Group/Base.agda index 9c6f08f0b..fb3caa724 100644 --- a/Cubical/Data/Group/Base.agda +++ b/Cubical/Data/Group/Base.agda @@ -11,6 +11,8 @@ import Cubical.Foundations.Isomorphism as I import Cubical.Foundations.Equiv as E import Cubical.Foundations.Equiv.HalfAdjoint as HAE +open import Cubical.HITs.SetQuotients as sq + record isGroup {ℓ} (A : Type ℓ) : Type ℓ where constructor group-struct field @@ -38,6 +40,20 @@ isMorph (group G Gset (group-struct _ _ _⊙_ _ _ _ _ _)) morph : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → Type (ℓ-max ℓ ℓ') morph G H = Σ (Group.type G → Group.type H) (isMorph G H) + +open import Cubical.Data.Sigma hiding (_×_) + +isInIm : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (Group.type G → Group.type H) + → Group.type H → Type _ +isInIm G H ϕ h = ∃[ g ∈ Group.type G ] ϕ g ≡ h + +isInKer : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (Group.type G → Group.type H) + → Group.type G → Type _ +isInKer G H ϕ g = ϕ g ≡ isGroup.id (Group.groupStruc H) + + + +{- morphisms takes id to id -} morph0→0 : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (f : (Group.type G → Group.type H)) → isMorph G H f → f (isGroup.id (Group.groupStruc G)) ≡ isGroup.id (Group.groupStruc H) @@ -50,6 +66,7 @@ morph0→0 (group G strucG (group-struct idG invG compG lUnitG rUnitG assocG lCa compH (f idG) (invH (f idG)) ≡⟨ rCancelH (f idG) ⟩ idH ∎ +{- a morphism ϕ satisfies ϕ(- a) = - ϕ(a) -} morphMinus : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (f : (Group.type G → Group.type H)) → isMorph G H f → (g : Group.type G) → f (isGroup.inv (Group.groupStruc G) g) ≡ isGroup.inv (Group.groupStruc H) (f g) @@ -74,10 +91,6 @@ morphMinus G H f morph g = compH idH (invH (f g)) ≡⟨ lUnitH (invH (f g)) ⟩ invH (f g) ∎ - - --- sym (rUnitH (f (invG g))) ∙ (λ i → compH (f (invG g)) (rCancelH (f g) (~ i))) ∙ sym (assocH (f (invG g)) (f g) (invH (f g))) ∙ cong (λ x → compH x (invH (f g))) helper ∙ lUnitH (invH (f g)) - record Iso {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') : Type (ℓ-max ℓ ℓ') where constructor iso field @@ -92,6 +105,13 @@ record Iso' {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') : Type (ℓ-max ℓ ℓ' isoSet : I.Iso (Group.type G) (Group.type H) isoSetMorph : isMorph G H (I.Iso.fun isoSet) +record Iso'' {ℓ ℓ'} (A : Group ℓ) (B : Group ℓ') : Type (ℓ-max ℓ ℓ') where + constructor iso'' + field + ϕ : morph A B + inj : (x : Group.type A) → isInKer A B (fst ϕ) x → x ≡ isGroup.id (Group.groupStruc A) + surj : (x : Group.type B) → isInIm A B (fst ϕ) x + _≃_ : ∀ {ℓ ℓ'} (A : Group ℓ) (B : Group ℓ') → Type (ℓ-max ℓ ℓ') A ≃ B = Σ (morph A B) \ f → (E.isEquiv (f .fst)) @@ -138,7 +158,6 @@ Iso'→Iso {G = group G Gset Ggroup} {H = group H Hset Hgroup} i = iso (fun , fu (fun (inv h0)) ∘ (fun (inv h1)) ≡⟨ sym (funMorph (inv h0) (inv h1)) ⟩ fun ((inv h0) ⊙ (inv h1)) ∎ ) - Equiv→Iso' : ∀ {ℓ ℓ'} {G : Group ℓ} {H : Group ℓ'} → G ≃ H → Iso' G H Equiv→Iso' {G = group G Gset Ggroup} {H = group H Hset Hgroup} @@ -204,34 +223,53 @@ compIso {ℓ} {F} {G} {H} g→f (f→g f) ≡⟨ gf _ ⟩ f ∎ - -isInIm : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (Group.type G → Group.type H) - → Group.type H → Type _ -isInIm G H ϕ h = ∃[ g ∈ Group.type G ] ϕ g ≡ h - -isInKer : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (Group.type G → Group.type H) - → Group.type G → Type _ -isInKer G H ϕ g = ϕ g ≡ isGroup.id (Group.groupStruc H) - - - -{- direct products of groups -} -dirProd : ∀ {ℓ ℓ'} (A : Group ℓ) (B : Group ℓ') → Group (ℓ-max ℓ ℓ') -Group.type (dirProd (group A setA strucA) (group B setB strucB)) = A × B -Group.setStruc (dirProd (group A setA strucA) (group B setB strucB)) = isOfHLevelProd 2 setA setB -isGroup.id (Group.groupStruc (dirProd (group A setA strucA) (group B setB strucB))) = - (isGroup.id strucA) , (isGroup.id strucB) -isGroup.inv (Group.groupStruc (dirProd (group A setA strucA) (group B setB strucB))) = - map (isGroup.inv strucA) (isGroup.inv strucB) -isGroup.comp (Group.groupStruc (dirProd (group A setA strucA) (group B setB strucB))) (a1 , b1) (a2 , b2) = - (isGroup.comp strucA a1 a2) , isGroup.comp strucB b1 b2 -isGroup.lUnit (Group.groupStruc (dirProd (group A setA strucA) (group B setB strucB))) (a , b) i = - (isGroup.lUnit strucA a i) , (isGroup.lUnit strucB b i) -isGroup.rUnit (Group.groupStruc (dirProd (group A setA strucA) (group B setB strucB))) (a , b) i = - (isGroup.rUnit strucA a i) , (isGroup.rUnit strucB b i) -isGroup.assoc (Group.groupStruc (dirProd (group A setA strucA) (group B setB strucB))) (a1 , b1) (a2 , b2) (a3 , b3) i = - (isGroup.assoc strucA a1 a2 a3 i) , (isGroup.assoc strucB b1 b2 b3 i) -isGroup.lCancel (Group.groupStruc (dirProd (group A setA strucA) (group B setB strucB))) (a , b) i = - (isGroup.lCancel strucA a i) , (isGroup.lCancel strucB b i) -isGroup.rCancel (Group.groupStruc (dirProd (group A setA strucA) (group B setB strucB))) (a , b) i = - (isGroup.rCancel strucA a i) , (isGroup.rCancel strucB b i) +idIso : ∀ {ℓ} (G : Group ℓ) → Iso G G +Iso.fun (idIso G) = (λ x → x) , λ _ _ → refl +Iso.inv (idIso G) = (λ x → x) , λ _ _ → refl +Iso.rightInv (idIso G) _ = refl +Iso.leftInv (idIso G) _ = refl + +Iso''→Iso : ∀ {ℓ ℓ'} (A : Group ℓ) (B : Group ℓ') → Iso'' A B → Iso A B +Iso''→Iso A B (iso'' ϕ inj surj) = + Iso'→Iso + (iso' + (I.iso + (fst ϕ) + (λ b → rec (helper b) (λ a → a) (surj b) .fst) + (λ b → rec (helper b) (λ a → a) (surj b) .snd) + λ b i → rec (helper (fst ϕ b)) (λ a → a) (propTruncIsProp (surj (fst ϕ b)) ∣ b , refl ∣ i) .fst) + (snd ϕ)) + where + helper : (b : _) → isProp (Σ (Group.type A) (λ a → ϕ .fst a ≡ b)) + helper b (a1 , pf1) (a2 , pf2) = + ΣProp≡ (λ _ → isOfHLevelPath' 1 (Group.setStruc B) _ _) + fstId + where + fstIdHelper : isGroup.comp (Group.groupStruc A) a1 (isGroup.inv (Group.groupStruc A) a2) + ≡ isGroup.id (Group.groupStruc A) + fstIdHelper = + let -A = isGroup.inv (Group.groupStruc A) + -B = isGroup.inv (Group.groupStruc B) + rCancelB = isGroup.rCancel (Group.groupStruc B) + _+B_ = isGroup.comp (Group.groupStruc B) + in inj _ + (ϕ .snd a1 (-A a2) + ∙ (λ i → (ϕ .fst a1 +B (morphMinus A B (ϕ .fst) (ϕ .snd) a2 i) )) + ∙ cong (λ x → x +B (-B (ϕ .fst a2))) (pf1 ∙ sym pf2) + ∙ rCancelB (ϕ .fst a2)) + fstId : a1 ≡ a2 + fstId = + let - = isGroup.inv (Group.groupStruc A) + id = isGroup.id (Group.groupStruc A) + assoc = isGroup.assoc (Group.groupStruc A) + rUnit = isGroup.rUnit (Group.groupStruc A) + lUnit = isGroup.lUnit (Group.groupStruc A) + lCancel = isGroup.lCancel (Group.groupStruc A) + assoc = isGroup.assoc (Group.groupStruc A) + _+_ = isGroup.comp (Group.groupStruc A) + in a1 ≡⟨ sym (rUnit a1) ⟩ + (a1 + id) ≡⟨ cong (λ x → a1 + x) (sym (lCancel a2)) ⟩ + (a1 + (- a2 + a2)) ≡⟨ sym (assoc a1 (- a2) a2) ⟩ + ((a1 + - a2) + a2) ≡⟨ cong (λ x → x + a2) fstIdHelper ⟩ + (id + a2) ≡⟨ lUnit a2 ⟩ + a2 ∎ diff --git a/Cubical/Data/Group/GroupLibrary.agda b/Cubical/Data/Group/GroupLibrary.agda new file mode 100644 index 000000000..3db758d1f --- /dev/null +++ b/Cubical/Data/Group/GroupLibrary.agda @@ -0,0 +1,217 @@ +{-# OPTIONS --cubical --safe #-} + +module Cubical.Data.Group.GroupLibrary where + +open import Cubical.Foundations.Prelude hiding (comp) +open import Cubical.Data.Group.Base +open import Cubical.Foundations.HLevels +open import Cubical.Data.Prod +open import Cubical.HITs.PropositionalTruncation hiding (map) + +import Cubical.Foundations.Isomorphism as I +import Cubical.Foundations.Equiv as E +import Cubical.Foundations.Equiv.HalfAdjoint as HAE + +open import Cubical.HITs.SetQuotients as sq + + +-- imGroup : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → morph G H +-- → Group (ℓ-max ℓ ℓ') +-- Group.type (imGroup G H (ϕ , mf)) = Σ[ x ∈ Group.type H ] isInIm G H ϕ x +-- Group.setStruc (imGroup G H (ϕ , mf)) = isOfHLevelΣ 2 (Group.setStruc H) λ _ → isOfHLevelSuc 1 propTruncIsProp +-- isGroup.id (Group.groupStruc (imGroup G H (ϕ , mf))) = +-- let idH = isGroup.id (Group.groupStruc H) +-- idG = isGroup.id (Group.groupStruc G) +-- invG = isGroup.inv (Group.groupStruc G) +-- invH = isGroup.inv (Group.groupStruc H) +-- lUnit = isGroup.lUnit (Group.groupStruc H) +-- rCancel = isGroup.rCancel (Group.groupStruc H) +-- in idH , ∣ idG , morph0→0 G H ϕ mf ∣ +-- isGroup.inv (Group.groupStruc (imGroup G H (ϕ , mf))) (h , hinim) = +-- let idH = isGroup.id (Group.groupStruc H) +-- idG = isGroup.id (Group.groupStruc G) +-- invG = isGroup.inv (Group.groupStruc G) +-- invH = isGroup.inv (Group.groupStruc H) +-- lUnit = isGroup.lUnit (Group.groupStruc H) +-- rCancel = isGroup.rCancel (Group.groupStruc H) +-- in invH h , rec propTruncIsProp +-- (λ a → ∣ (invG (fst a)) +-- , morphMinus G H ϕ mf (fst a) ∙ cong (λ x → invH x) (snd a) ∣) +-- hinim +-- isGroup.comp (Group.groupStruc (imGroup G H (ϕ , mf))) (h1 , h1inim) (h2 , h2inim) = +-- let idH = isGroup.id (Group.groupStruc H) +-- idG = isGroup.id (Group.groupStruc G) +-- invG = isGroup.inv (Group.groupStruc G) +-- invH = isGroup.inv (Group.groupStruc H) +-- compH = isGroup.comp (Group.groupStruc H) +-- compG = isGroup.comp (Group.groupStruc G) +-- lUnit = isGroup.lUnit (Group.groupStruc H) +-- rCancel = isGroup.rCancel (Group.groupStruc H) +-- in compH h1 h2 , rec propTruncIsProp +-- (λ p1 → rec propTruncIsProp +-- (λ p2 → ∣ (compG (fst p1) (fst p2)) +-- , mf (fst p1) (fst p2) ∙ cong₂ compH (snd p1) (snd p2) ∣) +-- h2inim) +-- h1inim +-- isGroup.lUnit (Group.groupStruc (imGroup G H (ϕ , mf))) (h , _) = +-- let lUnit = isGroup.lUnit (Group.groupStruc H) +-- in ΣProp≡ (λ _ → propTruncIsProp) +-- (lUnit h) +-- isGroup.rUnit (Group.groupStruc (imGroup G H (ϕ , mf))) (h , _) = +-- let rUnit = isGroup.rUnit (Group.groupStruc H) +-- in ΣProp≡ (λ _ → propTruncIsProp) +-- (rUnit h) +-- isGroup.assoc (Group.groupStruc (imGroup G H (ϕ , mf))) (h1 , _) (h2 , _) (h3 , _) = +-- let assoc = isGroup.assoc (Group.groupStruc H) +-- in ΣProp≡ (λ _ → propTruncIsProp) +-- (assoc h1 h2 h3) +-- isGroup.lCancel (Group.groupStruc (imGroup G H (ϕ , mf))) (h , _) = +-- let lCancel = isGroup.lCancel (Group.groupStruc H) +-- in ΣProp≡ (λ _ → propTruncIsProp) +-- (lCancel h) +-- isGroup.rCancel (Group.groupStruc (imGroup G H (ϕ , mf))) (h , _) = +-- let rCancel = isGroup.rCancel (Group.groupStruc H) +-- in ΣProp≡ (λ _ → propTruncIsProp) +-- (rCancel h) + + +-- kerGroup : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → morph G H +-- → Group (ℓ-max ℓ ℓ') +-- Group.type (kerGroup G H (ϕ , pf)) = Σ[ x ∈ Group.type G ] isInKer G H ϕ x +-- Group.setStruc (kerGroup G H (ϕ , pf)) = isOfHLevelΣ 2 (Group.setStruc G) λ _ → isOfHLevelPath 2 (Group.setStruc H) _ _ +-- isGroup.id (Group.groupStruc (kerGroup G H (ϕ , pf))) = isGroup.id (Group.groupStruc G) , morph0→0 G H ϕ pf +-- isGroup.inv (Group.groupStruc (kerGroup G H (ϕ , pf))) (g , inker) = +-- let +-- idH = isGroup.id (Group.groupStruc H) +-- invG = isGroup.inv (Group.groupStruc G) +-- invH = isGroup.inv (Group.groupStruc H) +-- lUnit = isGroup.lUnit (Group.groupStruc H) +-- rCancel = isGroup.rCancel (Group.groupStruc H) +-- in invG g , morphMinus G H ϕ pf g ∙ cong (λ f → invH f) inker ∙ sym (lUnit (invH idH)) ∙ rCancel idH +-- isGroup.comp (Group.groupStruc (kerGroup G H (ϕ , pf))) (g1 , inkerg1) (g2 , inkerg2) = +-- let +-- idH = isGroup.id (Group.groupStruc H) +-- compG = isGroup.comp (Group.groupStruc G) +-- compH = isGroup.comp (Group.groupStruc H) +-- lUnit = isGroup.lUnit (Group.groupStruc H) +-- in (compG g1 g2) , pf g1 g2 ∙ (λ i → compH (inkerg1 i) (inkerg2 i)) ∙ lUnit idH +-- isGroup.lUnit (Group.groupStruc (kerGroup G H (ϕ , pf))) (g , _) = +-- let lUnit = isGroup.lUnit (Group.groupStruc G) +-- in ΣProp≡ (λ _ → isOfHLevelPath' 1 (Group.setStruc H) _ _) (lUnit g) +-- isGroup.rUnit (Group.groupStruc (kerGroup G H (ϕ , pf))) (g , _) = +-- let rUnit = isGroup.rUnit (Group.groupStruc G) +-- in ΣProp≡ (λ _ → isOfHLevelPath' 1 (Group.setStruc H) _ _) (rUnit g) +-- isGroup.assoc (Group.groupStruc (kerGroup G H (ϕ , pf))) (g1 , _) (g2 , _) (g3 , _) = +-- let assoc = isGroup.assoc (Group.groupStruc G) +-- in ΣProp≡ (λ _ → isOfHLevelPath' 1 (Group.setStruc H) _ _) (assoc g1 g2 g3) +-- isGroup.lCancel (Group.groupStruc (kerGroup G H (ϕ , pf))) (g1 , _) = +-- let lCancel = isGroup.lCancel (Group.groupStruc G) +-- in ΣProp≡ (λ _ → isOfHLevelPath' 1 (Group.setStruc H) _ _) (lCancel g1) +-- isGroup.rCancel (Group.groupStruc (kerGroup G H (ϕ , pf))) (g1 , _) = +-- let rCancel = isGroup.rCancel (Group.groupStruc G) +-- in ΣProp≡ (λ _ → isOfHLevelPath' 1 (Group.setStruc H) _ _) (rCancel g1) + +open import Cubical.Data.Unit + +trivialGroup : Group ℓ-zero +Group.type trivialGroup = Unit +Group.setStruc trivialGroup = isOfHLevelSuc 1 isPropUnit +isGroup.id (Group.groupStruc trivialGroup) = tt +isGroup.inv (Group.groupStruc trivialGroup) = λ _ → tt +isGroup.comp (Group.groupStruc trivialGroup) = λ _ _ → tt +isGroup.lUnit (Group.groupStruc trivialGroup) = λ _ → refl +isGroup.rUnit (Group.groupStruc trivialGroup) = λ _ → refl +isGroup.assoc (Group.groupStruc trivialGroup) = λ _ _ _ → refl +isGroup.lCancel (Group.groupStruc trivialGroup) = λ _ → refl +isGroup.rCancel (Group.groupStruc trivialGroup) = λ _ → refl + +open import Cubical.Data.Int + +intGroup : Group ℓ-zero +Group.type intGroup = Int +Group.setStruc intGroup = isSetInt +isGroup.id (Group.groupStruc intGroup) = pos 0 +isGroup.inv (Group.groupStruc intGroup) = pos 0 -_ +isGroup.comp (Group.groupStruc intGroup) = _+_ +isGroup.lUnit (Group.groupStruc intGroup) = λ a → +-comm (pos 0) a +isGroup.rUnit (Group.groupStruc intGroup) = λ _ → refl +isGroup.assoc (Group.groupStruc intGroup) = λ a b c → sym (+-assoc a b c) +isGroup.lCancel (Group.groupStruc intGroup) = λ a → minusPlus a 0 +isGroup.rCancel (Group.groupStruc intGroup) = λ a → +-comm a (pos 0 - a) ∙ minusPlus a 0 + + + +-- Short exact sequences without req for first group to be definitionally equal to the trivial group -- +record SES {ℓ ℓ' ℓ'' ℓ'''} (A : Group ℓ) (B : Group ℓ') (leftGr : Group ℓ'') (rightGr : Group ℓ''') : Type (ℓ-suc (ℓ-max ℓ (ℓ-max ℓ' (ℓ-max ℓ'' ℓ''')))) where + constructor ses + field + isTrivialLeft : Iso leftGr trivialGroup + isTrivialRight : Iso rightGr trivialGroup + + left : morph leftGr A + right : morph B rightGr + ϕ : morph A B + + Ker-ϕ⊂Im-left : (x : Group.type A) -- + → isInKer A B (fst ϕ) x + → isInIm leftGr A (fst left) x + Ker-right⊂Im-ϕ : (x : Group.type B) -- + → isInKer B rightGr (fst right) x + → isInIm A B (fst ϕ) x + +SES→inj-surj-morph : ∀ {ℓ ℓ' ℓ'' ℓ'''} (A : Group ℓ) (B : Group ℓ') (leftGr : Group ℓ'') (rightGr : Group ℓ''') + → SES A B leftGr rightGr + → inj-surj-morph A B +inj-surj-morph.ϕ (SES→inj-surj-morph _ _ _ _ (ses _ _ _ _ ϕ _ _)) = ϕ +inj-surj-morph.inj + (SES→inj-surj-morph A _ lGr _ + (ses (iso ψ ξ rinv linv) _ left _ ϕ Ker-ϕ⊂Im-left _)) x inKer = + rec (isOfHLevelPath' 1 (Group.setStruc A) _ _) + (λ (a , pf) → sym pf ∙∙ (λ i → fst left (helper a i)) ∙∙ morph0→0 lGr A (fst left) (snd left)) + (Ker-ϕ⊂Im-left x inKer) + where + helper : (a : Group.type lGr) → a ≡ isGroup.id (Group.groupStruc lGr) + helper a = + a ≡⟨ sym (linv a) ⟩ + ξ .fst (ψ .fst a) ≡⟨ refl ⟩ + ξ .fst tt ≡⟨ morph0→0 trivialGroup lGr (fst ξ) (snd ξ) ⟩ + isGroup.id (Group.groupStruc lGr) ∎ + +inj-surj-morph.surj (SES→inj-surj-morph _ _ _ rGr + (ses _ (iso ψ ξ rinv linv) _ right ϕ _ Ker-right⊂Im-ϕ)) b = + Ker-right⊂Im-ϕ b (helper (right .fst b)) + where + helper : (a : Group.type rGr) → a ≡ isGroup.id (Group.groupStruc rGr) + helper a = + a ≡⟨ sym (linv a) ⟩ + ξ .fst (ψ .fst a) ≡⟨ refl ⟩ + ξ .fst tt ≡⟨ morph0→0 trivialGroup rGr (fst ξ) (snd ξ) ⟩ + isGroup.id (Group.groupStruc rGr) ∎ + +SES→Iso : ∀ {ℓ ℓ' ℓ'' ℓ'''} (A : Group ℓ) (B : Group ℓ') (leftGr : Group ℓ'') (rightGr : Group ℓ''') + → SES A B leftGr rightGr + → Iso A B +SES→Iso A B lGr rGr seq = inj-surj-morph→Iso A B (SES→inj-surj-morph A B lGr rGr seq) + + +{- direct products of groups -} +dirProd : ∀ {ℓ ℓ'} (A : Group ℓ) (B : Group ℓ') → Group (ℓ-max ℓ ℓ') +Group.type (dirProd (group A _ _) (group B _ _)) = A × B +Group.setStruc (dirProd (group _ setA _) (group _ setB _)) = isOfHLevelProd 2 setA setB +isGroup.id (Group.groupStruc (dirProd (group _ _ strucA) (group _ _ strucB))) = + (isGroup.id strucA) , (isGroup.id strucB) +isGroup.inv (Group.groupStruc (dirProd (group _ _ strucA) (group _ _ strucB))) = + map (isGroup.inv strucA) (isGroup.inv strucB) +isGroup.comp (Group.groupStruc (dirProd (group _ _ strucA) (group _ _ strucB))) (a1 , b1) (a2 , b2) = + (isGroup.comp strucA a1 a2) , isGroup.comp strucB b1 b2 +isGroup.lUnit (Group.groupStruc (dirProd (group _ _ strucA) (group _ _ strucB))) (a , b) i = + (isGroup.lUnit strucA a i) , (isGroup.lUnit strucB b i) +isGroup.rUnit (Group.groupStruc (dirProd (group _ _ strucA) (group _ _ strucB))) (a , b) i = + (isGroup.rUnit strucA a i) , (isGroup.rUnit strucB b i) +isGroup.assoc (Group.groupStruc (dirProd (group _ _ strucA) (group _ _ strucB))) (a1 , b1) (a2 , b2) (a3 , b3) i = + (isGroup.assoc strucA a1 a2 a3 i) , (isGroup.assoc strucB b1 b2 b3 i) +isGroup.lCancel (Group.groupStruc (dirProd (group _ _ strucA) (group _ _ strucB))) (a , b) i = + (isGroup.lCancel strucA a i) , (isGroup.lCancel strucB b i) +isGroup.rCancel (Group.groupStruc (dirProd (group _ _ strucA) (group _ _ strucB))) (a , b) i = + (isGroup.rCancel strucA a i) , (isGroup.rCancel strucB b i) + diff --git a/Cubical/Data/Unit/Properties.agda b/Cubical/Data/Unit/Properties.agda index 606a88c78..16e6da17f 100644 --- a/Cubical/Data/Unit/Properties.agda +++ b/Cubical/Data/Unit/Properties.agda @@ -5,6 +5,7 @@ open import Cubical.Core.Everything open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism open import Cubical.Data.Nat open import Cubical.Data.Unit.Base @@ -17,3 +18,6 @@ isPropUnit _ _ i = tt -- definitionally equal to: isContr→isProp isContrUnit isOfHLevelUnit : (n : ℕ) → isOfHLevel n Unit isOfHLevelUnit n = isContr→isOfHLevel n isContrUnit + +UnitToTypeId : ∀ {ℓ} (A : Type ℓ) → (Unit → A) ≡ A +UnitToTypeId A = isoToPath (iso (λ f → f tt) (λ a _ → a) (λ _ → refl) λ _ → refl) diff --git a/Cubical/HITs/S1/Properties.agda b/Cubical/HITs/S1/Properties.agda index 1806d1a65..2813e905b 100644 --- a/Cubical/HITs/S1/Properties.agda +++ b/Cubical/HITs/S1/Properties.agda @@ -10,6 +10,10 @@ open import Cubical.Foundations.Univalence open import Cubical.HITs.S1.Base open import Cubical.HITs.PropositionalTruncation as PropTrunc +open import Cubical.HITs.SetTruncation as s +open import Cubical.Data.Int +open import Cubical.Data.Nat +open import Cubical.Data.Prod isConnectedS¹ : (s : S¹) → ∥ base ≡ s ∥ isConnectedS¹ base = ∣ refl ∣ diff --git a/Cubical/HITs/SetQuotients/Properties.agda b/Cubical/HITs/SetQuotients/Properties.agda index cd1101b39..0df03cd2a 100644 --- a/Cubical/HITs/SetQuotients/Properties.agda +++ b/Cubical/HITs/SetQuotients/Properties.agda @@ -35,6 +35,10 @@ private R : A → A → Type ℓ' B : A / R → Type ℓ'' +setQuotientIsSet : isSet (A / R) +setQuotientIsSet a b p q = squash/ a b p q + + elimEq/ : (Bprop : (x : A / R ) → isProp (B x)) {x y : A / R} (eq : x ≡ y) @@ -74,6 +78,27 @@ elim Bset f feq (squash/ x y p q i j) = where g = elim Bset f feq +-- elim2 : {B : A / R → A / R → Type ℓ} → +-- (Bset : (x y : A / R) → isSet (B x y)) → +-- (f : (a : A) (b : A) → (B [ a ] [ b ])) → +-- (feq1 : (a a₁ b : A) (r : R a₁ b) → +-- PathP (λ i → B [ a ] (eq/ a₁ b r i)) (f a a₁) (f a b)) → +-- (feq2 : (a a₁ b : A) (r : R a₁ a) → +-- PathP (λ i → B (eq/ a₁ a r i) [ b ]) (f a₁ b) (f a b)) → +-- {!(a b c d : A) (r1 : R a b) (r2 : R a b) → !} → +-- (x y : A / R) → B x y + + +-- elim2 {A = A} {B = B} Bset f feq1 feq2 _ = +-- elim (λ _ → isOfHLevelΠ 2 λ _ → Bset _ _) +-- (λ a → elim (λ _ → Bset _ _) (λ b → f a b) (feq1 _)) +-- λ p q r i → elim (Bset (eq/ p q r i)) (λ b → feq2 q p b r i) (λ a b r2 → {!isOfHLevel→isOfHLevelDep 2 {A = A} {λ x y → B [ y ] [ y ]}!}) +-- {- +-- Goal: PathP (λ i → (y : A / R) → B (eq/ p q r i) y) +-- (elim (λ z → Bset [ p ] z) (λ b → f p b) (feq p)) +-- (elim (λ z → Bset [ q ] z) (λ b → f q b) (feq q)) +-- -} + setQuotUniversal : {B : Type ℓ} (Bset : isSet B) → (A / R → B) ≃ (Σ[ f ∈ (A → B) ] ((a b : A) → R a b → f a ≡ f b)) @@ -145,3 +170,4 @@ discreteSetQuotients {A = A} {R = R} Adis Rprop Req Rdec = J (λ b ab → ∀ k → PathP (λ i → (y : A / R) → Dec (ab i ≡ y)) (discreteSetQuotients' a) k) (λ k → funExt (λ x → isPropDec (squash/ _ _) _ _)) (eq/ a b ab) (discreteSetQuotients' b) + diff --git a/Cubical/HITs/Truncation/Properties.agda b/Cubical/HITs/Truncation/Properties.agda index ca41f3593..d4c400f6c 100644 --- a/Cubical/HITs/Truncation/Properties.agda +++ b/Cubical/HITs/Truncation/Properties.agda @@ -120,7 +120,7 @@ rec : {n : ℕ} (isOfHLevel n B) → (g : (a : A) → B) → (hLevelTrunc n A → B) -rec {B = B} h = Null.rec (isOfHLevel→isSnNull h) +rec {B = B} h = Null.elim (λ _ → isOfHLevel→isSnNull (h)) --Null.rec (isOfHLevel→isSnNull h) elim : {n : ℕ} {B : hLevelTrunc n A → Type ℓ'} @@ -217,7 +217,7 @@ Iso.leftInv setTruncTrunc0Iso = SetTrunc.elim (λ _ → isOfHLevelPath 2 squash setTrunc≃Trunc0 : ∥ A ∥₀ ≃ ∥ A ∥ 0 setTrunc≃Trunc0 = isoToEquiv setTruncTrunc0Iso -propTrunc≡Trunc0 : ∥ A ∥₀ ≡ ∥ A ∥ -0 +propTrunc≡Trunc0 : ∥ A ∥₀ ≡ ∥ A ∥ 0 propTrunc≡Trunc0 = ua setTrunc≃Trunc0 -- 1 -- diff --git a/Cubical/Homotopy/Connected.agda b/Cubical/Homotopy/Connected.agda index fa224f91b..a226742c6 100644 --- a/Cubical/Homotopy/Connected.agda +++ b/Cubical/Homotopy/Connected.agda @@ -35,14 +35,17 @@ isHLevelConnectedFun n f = ∀ b → isHLevelConnected n (fiber f b) isHLevelTruncatedFun : ∀ {ℓ ℓ'} (n : ℕ) {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Type (ℓ-max ℓ ℓ') isHLevelTruncatedFun n f = ∀ b → isOfHLevel n (fiber f b) -isConnectedSubtr : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n m : ℕ) (f : A → B) +isConnectedSubtr : ∀ {ℓ} {A : Type ℓ} (n m : ℕ) + → isHLevelConnected (m + n) A + → isHLevelConnected n A +isConnectedSubtr {A = A} n m iscon = + transport (λ i → isContr (ua (truncOfTruncEq {A = A} n m) (~ i))) + (∣ iscon .fst ∣ , Trunc.elim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → cong ∣_∣ (iscon .snd a)) + +isConnectedFunSubtr : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n m : ℕ) (f : A → B) → isHLevelConnectedFun (m + n) f → isHLevelConnectedFun n f -isConnectedSubtr n m f iscon b = - transport (λ i → isContr (ua (truncOfTruncEq {A = fiber f b} n m) (~ i) )) - (∣ iscon b .fst ∣ , - Trunc.elim (λ x → isOfHLevelPath n (isOfHLevelTrunc n) _ _) - λ a → cong ∣_∣ (iscon b .snd a)) +isConnectedFunSubtr n m f iscon b = isConnectedSubtr n m (iscon b) private typeToFiber : ∀ {ℓ} (A : Type ℓ) (b : Unit) → A ≡ fiber (λ (x : A) → b) b @@ -274,7 +277,7 @@ connectedTruncIso2 : ∀ {ℓ} {A B : Type ℓ} (n m : ℕ) (f : A → B) → isHLevelConnectedFun m f → Iso (hLevelTrunc n A) (hLevelTrunc n B) connectedTruncIso2 {A = A} {B = B} n m f (x , pf) con = - connectedTruncIso n f (isConnectedSubtr n x f (transport (λ i → isHLevelConnectedFun (pf (~ i)) f) con)) + connectedTruncIso n f (isConnectedFunSubtr n x f (transport (λ i → isHLevelConnectedFun (pf (~ i)) f) con)) connectedTruncEquiv : ∀ {ℓ} {A B : Type ℓ} (n : ℕ) (f : A → B) → isHLevelConnectedFun n f @@ -332,395 +335,395 @@ sphereConnected (suc n) = Iso.rightInv (fibIso) b = refl -private - variable - ℓ ℓ' ℓ'' : Level - A : Type ℓ - B : Type ℓ' - A' : Pointed ℓ - B' : Pointed ℓ' - C' : Pointed ℓ'' +-- private +-- variable +-- ℓ ℓ' ℓ'' : Level +-- A : Type ℓ +-- B : Type ℓ' +-- A' : Pointed ℓ +-- B' : Pointed ℓ' +-- C' : Pointed ℓ'' -isConnectedSmashIdfun2 : {A B C : Pointed₀} - → (f : A →∙ B) (nf nc : ℕ) - → isHLevelConnectedFun nf (fst f) - → isHLevelConnected (2 + nc) (fst C) - → isHLevelConnectedFun (1 + nc + nf) (Smash-map f (idfun∙ C)) -isConnectedSmashIdfun2 {A = A , ptA} {B = B , ptB} {C = C , ptC} (f , fpt) nf nc conf conC = λ b → {!!} , {!!} - where - module proof {ℓ : Level} (P : (Smash (B , ptB) (C , ptC)) → HLevel ℓ (1 + nc + nf)) - (d : (x : Smash (A , ptA) (C , ptC)) → P ((Smash-map (f , fpt) (idfun∙ (C , ptC)) ) x) .fst) - where - F : (c : C) → _ - F c = λ(s : (b : B) → P (proj b c) .fst) → s ∘ f +-- isConnectedSmashIdfun2 : {A B C : Pointed₀} +-- → (f : A →∙ B) (nf nc : ℕ) +-- → isHLevelConnectedFun nf (fst f) +-- → isHLevelConnected (2 + nc) (fst C) +-- → isHLevelConnectedFun (1 + nc + nf) (Smash-map f (idfun∙ C)) +-- isConnectedSmashIdfun2 {A = A , ptA} {B = B , ptB} {C = C , ptC} (f , fpt) nf nc conf conC = λ b → {!!} , {!!} +-- where +-- module proof {ℓ : Level} (P : (Smash (B , ptB) (C , ptC)) → HLevel ℓ (1 + nc + nf)) +-- (d : (x : Smash (A , ptA) (C , ptC)) → P ((Smash-map (f , fpt) (idfun∙ (C , ptC)) ) x) .fst) +-- where +-- F : (c : C) → _ +-- F c = λ(s : (b : B) → P (proj b c) .fst) → s ∘ f - step1 : (c : C) → isHLevelTruncatedFun (1 + nc) (F c) - step1 c = isOfHLevelPrecomposeConnected (1 + nc) nf ((λ b → P (proj b c) .fst , P (proj b c) .snd)) f conf +-- step1 : (c : C) → isHLevelTruncatedFun (1 + nc) (F c) +-- step1 c = isOfHLevelPrecomposeConnected (1 + nc) nf ((λ b → P (proj b c) .fst , P (proj b c) .snd)) f conf - codomFun : (c : C) (a : A) → P (proj (f a) c) .fst - codomFun c = d ∘ λ a → proj a c +-- codomFun : (c : C) (a : A) → P (proj (f a) c) .fst +-- codomFun c = d ∘ λ a → proj a c - Q : C → HLevel _ (1 + nc) - Q c = fiber (F c) (codomFun c) , step1 c _ +-- Q : C → HLevel _ (1 + nc) +-- Q c = fiber (F c) (codomFun c) , step1 c _ - helper : (a : A) → transport (λ i → P (gluel (f a) (~ i)) .fst) (d (basel)) - ≡ d (proj a ptC) - helper a i = - hcomp (λ k → λ {(i = i0) → transport (λ j → P (compPath-filler refl (gluel (f a)) (~ j) (~ j ∨ k)) .fst) (d basel) - ; (i = i1) → d (proj a ptC)}) - (transp (λ j → P (compPath-filler refl (gluel (f a)) (~ j) (~ i ∧ (~ j))) .fst) i (d (gluel a (~ i)))) +-- helper : (a : A) → transport (λ i → P (gluel (f a) (~ i)) .fst) (d (basel)) +-- ≡ d (proj a ptC) +-- helper a i = +-- hcomp (λ k → λ {(i = i0) → transport (λ j → P (compPath-filler refl (gluel (f a)) (~ j) (~ j ∨ k)) .fst) (d basel) +-- ; (i = i1) → d (proj a ptC)}) +-- (transp (λ j → P (compPath-filler refl (gluel (f a)) (~ j) (~ i ∧ (~ j))) .fst) i (d (gluel a (~ i)))) -{- -i = i0 ⊢ transport (λ i₁ → P (gluel (f a) (~ i₁)) .fst) (d basel) -i = i1 ⊢ d (proj a ptC) --} - {- transport (PathP≡Path (λ i → P (gluel (f a) (~ i)) .fst) (d basel) (d (proj a ptC))) - (transport (λ j → PathP (λ i → P (lUnit (gluel (f a)) (~ j) (~ i)) .fst) (d basel) (d (proj a ptC))) - λ i → d (gluel a (~ i))) -} +-- {- +-- i = i0 ⊢ transport (λ i₁ → P (gluel (f a) (~ i₁)) .fst) (d basel) +-- i = i1 ⊢ d (proj a ptC) +-- -} +-- {- transport (PathP≡Path (λ i → P (gluel (f a) (~ i)) .fst) (d basel) (d (proj a ptC))) +-- (transport (λ j → PathP (λ i → P (lUnit (gluel (f a)) (~ j) (~ i)) .fst) (d basel) (d (proj a ptC))) +-- λ i → d (gluel a (~ i))) -} - QptC : Q ptC .fst - QptC = (λ b → transport (λ i → P (gluel b (~ i)) .fst) (d basel)) , - λ i a → hcomp (λ k → λ {(i = i0) → transport (λ j → P (compPath-filler refl (gluel (f a)) (~ j) (~ j ∨ k)) .fst) (d basel) - ; (i = i1) → d (proj a ptC)}) - (transp (λ j → P (compPath-filler refl (gluel (f a)) (~ j) (~ i ∧ (~ j))) .fst) i (d (gluel a (~ i)))) +-- QptC : Q ptC .fst +-- QptC = (λ b → transport (λ i → P (gluel b (~ i)) .fst) (d basel)) , +-- λ i a → hcomp (λ k → λ {(i = i0) → transport (λ j → P (compPath-filler refl (gluel (f a)) (~ j) (~ j ∨ k)) .fst) (d basel) +-- ; (i = i1) → d (proj a ptC)}) +-- (transp (λ j → P (compPath-filler refl (gluel (f a)) (~ j) (~ i ∧ (~ j))) .fst) i (d (gluel a (~ i)))) - Q-constructor : (c : C) → Q c .fst - Q-constructor c = Iso.inv (elim.isIsoPrecompose (λ (_ : Unit) → ptC) (1 + nc) Q (isHLevelConnectedPoint _ conC ptC)) (λ _ → QptC) c +-- Q-constructor : (c : C) → Q c .fst +-- Q-constructor c = Iso.inv (elim.isIsoPrecompose (λ (_ : Unit) → ptC) (1 + nc) Q (isHLevelConnectedPoint _ conC ptC)) (λ _ → QptC) c - Q-constructor' : Q-constructor ptC .snd ≡ (Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) ((isHLevelConnectedPoint _ conC ptC) ptC .fst)) .snd - Q-constructor' = refl +-- Q-constructor' : Q-constructor ptC .snd ≡ (Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) ((isHLevelConnectedPoint _ conC ptC) ptC .fst)) .snd +-- Q-constructor' = refl - Q-test : isHLevelConnected (suc nc) (fiber (λ (_ : Unit) → ptC) ptC) - Q-test = ∣ tt , refl ∣ , λ y → sym ((isHLevelConnectedPoint _ conC ptC) ptC .snd _) ∙ (isHLevelConnectedPoint _ conC ptC) ptC .snd y +-- Q-test : isHLevelConnected (suc nc) (fiber (λ (_ : Unit) → ptC) ptC) +-- Q-test = ∣ tt , refl ∣ , λ y → sym ((isHLevelConnectedPoint _ conC ptC) ptC .snd _) ∙ (isHLevelConnectedPoint _ conC ptC) ptC .snd y - hLevelContr : isContr (isHLevelConnected (suc nc) (fiber (λ (_ : Unit) → ptC) ptC)) - hLevelContr = Q-test , isPropIsOfHLevel 0 Q-test +-- hLevelContr : isContr (isHLevelConnected (suc nc) (fiber (λ (_ : Unit) → ptC) ptC)) +-- hLevelContr = Q-test , isPropIsOfHLevel 0 Q-test - helper3 : (isHLevelConnectedPoint _ conC ptC) ptC .fst ≡ ∣ tt , refl ∣ - helper3 = (isHLevelConnectedPoint _ conC ptC) ptC .snd ∣ tt , refl ∣ +-- helper3 : (isHLevelConnectedPoint _ conC ptC) ptC .fst ≡ ∣ tt , refl ∣ +-- helper3 = (isHLevelConnectedPoint _ conC ptC) ptC .snd ∣ tt , refl ∣ - g : (b : B) (c : C) → P (proj b c) .fst - g b c = Q-constructor c .fst b +-- g : (b : B) (c : C) → P (proj b c) .fst +-- g b c = Q-constructor c .fst b - Q-constructor-β : (b : B) → Q-constructor ptC .fst b ≡ QptC .fst b - Q-constructor-β b = sym ((λ i → transportRefl (QptC) (~ i) .fst b) ∙ - ((λ i → (Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) (helper3 (~ i))) .fst b))) +-- Q-constructor-β : (b : B) → Q-constructor ptC .fst b ≡ QptC .fst b +-- Q-constructor-β b = sym ((λ i → transportRefl (QptC) (~ i) .fst b) ∙ +-- ((λ i → (Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) (helper3 (~ i))) .fst b))) - Q-constructor23 : PathP (λ i → _) (Q-constructor ptC .snd) ((Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) (Q-test .fst)) .snd) - Q-constructor23 i = ((Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) (helper3 i)) .snd) +-- Q-constructor23 : PathP (λ i → _) (Q-constructor ptC .snd) ((Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) (Q-test .fst)) .snd) +-- Q-constructor23 i = ((Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) (helper3 i)) .snd) - Q-constructor24 : (a : A) → PathP (λ i → Q-constructor-β (f a) i ≡ d (proj a ptC)) (λ i → Q-constructor ptC .snd i a) λ i → QptC .snd i a - Q-constructor24 a i j = - hcomp (λ k → λ { (i = i0) → Q-constructor23 (~ k) j a - ; (i = i1) → QptC .snd j a - ; (j = i1) → d (proj a ptC)}) - (transportRefl (QptC) i .snd j a) +-- Q-constructor24 : (a : A) → PathP (λ i → Q-constructor-β (f a) i ≡ d (proj a ptC)) (λ i → Q-constructor ptC .snd i a) λ i → QptC .snd i a +-- Q-constructor24 a i j = +-- hcomp (λ k → λ { (i = i0) → Q-constructor23 (~ k) j a +-- ; (i = i1) → QptC .snd j a +-- ; (j = i1) → d (proj a ptC)}) +-- (transportRefl (QptC) i .snd j a) - Q-constructor25 : (a : A) → PathP (λ i → d (proj a ptC) ≡ Q-constructor-β (f a) i) ((λ i → Q-constructor ptC .snd (~ i) a)) λ i → QptC .snd (~ i) a - Q-constructor25 a i j = Q-constructor24 a i (~ j) +-- Q-constructor25 : (a : A) → PathP (λ i → d (proj a ptC) ≡ Q-constructor-β (f a) i) ((λ i → Q-constructor ptC .snd (~ i) a)) λ i → QptC .snd (~ i) a +-- Q-constructor25 a i j = Q-constructor24 a i (~ j) -{- -i = i0 ⊢ -i = i1 ⊢ QptC .snd j a -j = i0 ⊢ -j = i1 ⊢ d (proj a ptC) --} +-- {- +-- i = i0 ⊢ +-- i = i1 ⊢ QptC .snd j a +-- j = i0 ⊢ +-- j = i1 ⊢ d (proj a ptC) +-- -} - otherway-filler : (b : B) → I → I → P basel .fst - otherway-filler b i j = - fill (λ _ → P basel .fst) - (λ k → λ {(i = i0) → transport (λ i → P (gluel b i) .fst) (g b ptC) - ; (i = i1) → transportTransport⁻ (λ i → P (gluel b i) .fst) (d basel) k}) - (inS (transport (λ i → P (gluel b i) .fst) (Q-constructor-β b i))) - j +-- otherway-filler : (b : B) → I → I → P basel .fst +-- otherway-filler b i j = +-- fill (λ _ → P basel .fst) +-- (λ k → λ {(i = i0) → transport (λ i → P (gluel b i) .fst) (g b ptC) +-- ; (i = i1) → transportTransport⁻ (λ i → P (gluel b i) .fst) (d basel) k}) +-- (inS (transport (λ i → P (gluel b i) .fst) (Q-constructor-β b i))) +-- j - otherway : (b : B) → transport (λ i → P (gluel b i) .fst) (g b ptC) ≡ d basel - otherway b i = otherway-filler b i i1 +-- otherway : (b : B) → transport (λ i → P (gluel b i) .fst) (g b ptC) ≡ d basel +-- otherway b i = otherway-filler b i i1 - otherway2 : (b : B) → transport (λ i → P (gluel b i) .fst) (g b ptC) ≡ d basel - otherway2 b = (λ i → transport (λ i → P (gluel b i) .fst) (Q-constructor-β b i)) - ∙ transportTransport⁻ (λ i → P (gluel b i) .fst) (d basel) +-- otherway2 : (b : B) → transport (λ i → P (gluel b i) .fst) (g b ptC) ≡ d basel +-- otherway2 b = (λ i → transport (λ i → P (gluel b i) .fst) (Q-constructor-β b i)) +-- ∙ transportTransport⁻ (λ i → P (gluel b i) .fst) (d basel) - {- (λ i → transport (λ i → P (gluel b i) .fst) (Q-constructor-β b i)) ∙ - transportTransport⁻ (λ i → P (gluel b i) .fst) (d basel) -} +-- {- (λ i → transport (λ i → P (gluel b i) .fst) (Q-constructor-β b i)) ∙ +-- transportTransport⁻ (λ i → P (gluel b i) .fst) (d basel) -} - otherwaygluel : (i : I) (a : A) → {!!} - otherwaygluel = {!!} +-- otherwaygluel : (i : I) (a : A) → {!!} +-- otherwaygluel = {!!} - path1-filler : (b : B) → I → ((i : I) → P (gluel b i) .fst) - path1-filler b j i = - hfill - (λ k → λ {(i = i0) → g b ptC - ; (i = i1) → otherway2 b k}) - (inS (transp (λ j → P (gluel b (i ∧ j)) .fst) (~ i) (g b ptC))) - j +-- path1-filler : (b : B) → I → ((i : I) → P (gluel b i) .fst) +-- path1-filler b j i = +-- hfill +-- (λ k → λ {(i = i0) → g b ptC +-- ; (i = i1) → otherway2 b k}) +-- (inS (transp (λ j → P (gluel b (i ∧ j)) .fst) (~ i) (g b ptC))) +-- j - path1 : (b : B) → PathP (λ i → P (gluel b i) .fst) (g b ptC) (d basel) - path1 b i = path1-filler b i1 i +-- path1 : (b : B) → PathP (λ i → P (gluel b i) .fst) (g b ptC) (d basel) +-- path1 b i = path1-filler b i1 i - path1* : (a : A) → PathP (λ i → P (gluel (f a) i) .fst) (g (f a) ptC) (d basel) - path1* a i = - comp (λ k → P (compPath-filler refl (gluel (f a)) i k) .fst) - (λ k → λ {(i = i0) → g (f a) ptC - ; (i = i1) → d (gluel a k)}) - ((cong (λ F → F a) (Q-constructor ptC .snd)) i) +-- path1* : (a : A) → PathP (λ i → P (gluel (f a) i) .fst) (g (f a) ptC) (d basel) +-- path1* a i = +-- comp (λ k → P (compPath-filler refl (gluel (f a)) i k) .fst) +-- (λ k → λ {(i = i0) → g (f a) ptC +-- ; (i = i1) → d (gluel a k)}) +-- ((cong (λ F → F a) (Q-constructor ptC .snd)) i) -{- -Goal: P (gluel (f a) i) .fst -———— Boundary —————————————————————————————————————————————— -i = i0 ⊢ g (f a) ptC -i = i1 ⊢ d basel --} +-- {- +-- Goal: P (gluel (f a) i) .fst +-- ———— Boundary —————————————————————————————————————————————— +-- i = i0 ⊢ g (f a) ptC +-- i = i1 ⊢ d basel +-- -} - path1*-filler : (a : A) (i : I) (j : I) → P (compPath-filler (λ _ → proj (f a) ptC) (gluel (f a)) i j) .fst - path1*-filler a i j = - fill (λ k → P (compPath-filler refl (gluel (f a)) i k) .fst) - (λ k → λ {(i = i0) → g (f a) ptC - ; (i = i1) → d (gluel a k)}) - (inS ( ((Q-constructor ptC .snd i a)))) - j +-- path1*-filler : (a : A) (i : I) (j : I) → P (compPath-filler (λ _ → proj (f a) ptC) (gluel (f a)) i j) .fst +-- path1*-filler a i j = +-- fill (λ k → P (compPath-filler refl (gluel (f a)) i k) .fst) +-- (λ k → λ {(i = i0) → g (f a) ptC +-- ; (i = i1) → d (gluel a k)}) +-- (inS ( ((Q-constructor ptC .snd i a)))) +-- j - path2 : (c : C) → PathP (λ i → P (gluer c i) .fst) (g ptB c) (d baser) - path2 c i = - comp (λ k → P (compPath-filler (λ i → proj (fpt i) c) (gluer c) i k) .fst) - (λ k → λ { (i = i0) → g (fpt k) c - ; (i = i1) → d (gluer c k) }) - (Q-constructor c .snd i ptA) +-- path2 : (c : C) → PathP (λ i → P (gluer c i) .fst) (g ptB c) (d baser) +-- path2 c i = +-- comp (λ k → P (compPath-filler (λ i → proj (fpt i) c) (gluer c) i k) .fst) +-- (λ k → λ { (i = i0) → g (fpt k) c +-- ; (i = i1) → d (gluer c k) }) +-- (Q-constructor c .snd i ptA) -{- -j = i0 ⊢ path1 (f a) i -j = i1 ⊢ path1* a i -i = i0 ⊢ g (f a) ptC -i = i1 ⊢ d basel --} - - - h : (x : Smash (B , ptB) (C , ptC)) → P x .fst - h basel = d basel - h baser = d baser - h (proj x y) = g x y - h (gluel b i) = path1 b i - h (gluer c i) = path2 c i - - hsection : (x : Smash (A , ptA) (C , ptC)) → h (Smash-map (f , fpt) (idfun∙ (C , ptC)) x) ≡ d x - hsection basel = refl - hsection baser = refl - hsection (proj x y) i = Q-constructor y .snd i x - hsection (gluel a i) j = - comp (λ k → P (lUnit (gluel (f a)) (j ∨ k) i) .fst) - (λ k → λ { (i = i0) → Q-constructor ptC .snd j a - ; (i = i1) → d basel - ; (j = i0) → h (lUnit (gluel (f a)) k i) - ; (j = i1) → {!d (gluel a i)!}}) - {!!} - {- - comp (λ k → P (compPath-filler refl (gluel (f a)) (j ∨ k) i) .fst) - (λ k → λ { (i = i0) → Q-constructor ptC .snd j a - ; (i = i1) → path1 (f a) (k ∨ j) - ; (j = i0) → h (compPath-filler refl (gluel (f a)) k i) - ; (j = i1) → d (gluel a i)}) - {!!} -} - {- - comp (λ k → P (compPath-filler refl (gluel (f a)) (k ∨ j) i) .fst) - (λ k → λ { (i = i0) → Q-constructor ptC .snd j a - ; (i = i1) → {!path1 (f a) ?!} -- path1-filler (f a) k (k ∨ j) -- path1 (f a) (k ∨ j) - ; (j = i0) → h (compPath-filler refl (gluel (f a)) k i) - ; (j = i1) → {!!}}) -- d (gluel a i) - -- transp (λ r → P (compPath-filler refl (gluel (f a)) (r ∨ k) i) .fst) k ((transp (λ r → P (compPath-filler refl (gluel (f a)) (~ r ∨ k) i) .fst)) k (d (gluel a i))) - {!Q-constructor-β (f a)!} - -} - where - test2 : I → (i : I) → P ((refl ∙ gluel (f a)) i) .fst - test2 j i = - comp (λ k → P (compPath-filler refl (gluel (f a)) (j ∨ k) i) .fst) - (λ k → λ { (i = i0) → d (proj a ptC) -- g (f a) ptC - ; (i = i1) → transp (λ r → P (gluel (f a) (r ∧ (k ∨ j))) .fst) {!j!} (Q-constructor-β (f a) j) -- transp (λ r → P (gluel (f a) (r ∧ k)) .fst) (~ k) (Q-constructor-β (f a) j) - ; (j = i0) → compPathP'-filler (λ i → Q-constructor ptC .snd (~ i) a) (λ i → transp (λ r → P (gluel (f a) (r ∧ i)) .fst) (~ i) (g (f a) ptC)) k i - ; (j = i1) → d (gluel a i) }) -- d (gluel a i)}) - {!!} - - - {- - (comp (λ k → P (compPath-filler refl (gluel (f a)) (j ∧ k) i) .fst) - (λ k → λ { (i = i0) → Q-constructor ptC .snd j a - ; (i = i1) → path1-filler (f a) k (j ∧ k) - ; (j = i0) → g (f a) ptC - ; (j = i1) → side-filler i k}) - {!compPath-filler refl (gluel (f a)) i0 j!}) -} - - side-filler : (i : I) (k : I) → P (compPath-filler (λ _ → proj (f a) (pt (C , ptC))) (gluel (f a)) k i) .fst - side-filler i k = - comp (λ l → P (compPath-filler refl (gluel (f a)) k {!!}) .fst) - (λ l → λ { (i = i0) → {!transport (λ i₂ → P (gluel (f a) ₂) .fst) (g (f a) ptC)!} - ; (i = i1) → {!!} - ; (k = i0) → {!!} - ; (k = i1) → {!!} }) - {!otherway (f a) ?!} - hsection (gluer c i) j = {!transp (λ r → P (((refl ∙ gluel (f a)) ((r ∨ ?))) .fst)) l (transp (λ r → P (((refl ∙ gluel (f a)) ((~ r) ∨ ?))) .fst) l (d (gluel a i))) !} - {- - comp (λ k → P (compPath-filler (λ i → proj (fpt i) c) (gluer c) (k ∨ j) i) .fst) - (λ k → λ { (i = i0) → Q-constructor c .snd j ptA - ; (i = i1) → path2 c (k ∨ j) - ; (j = i0) → h (compPath-filler (λ i → proj (fpt i) c) (gluer c) k i) - ; (j = i1) → d (gluer c i) - }) - (comp (λ k → P (compPath-filler (λ i₁ → proj (fpt i₁) c) (gluer c) (j ∧ k) (i ∧ k)) .fst) - (λ k → λ { (i = i0) → Q-constructor c .snd j ptA - ; (j = i0) → h (proj (fpt (i ∧ k)) c) - ; (j = i1) → d (gluer c (i ∧ k)) - }) - (Q-constructor c .snd j ptA)) - -} - --- isConnectedSmashIdfun : (f : A' →∙ B') (nf nc : ℕ) --- → isHLevelConnectedFun nf (fst f) --- → isHLevelConnected (2 + nc) (fst C') --- → isHLevelConnectedFun (1 + nf + nc) (f ⋀⃗ idfun∙ C') --- isConnectedSmashIdfun {A' = (A , ptA)} {B' = (B , ptB)} {C' = (C , ptC)} (f , fpt) nf nc conf conC = {!isHLel!} --- where --- module _ (P : ((B , ptB) ⋀ (C , ptC)) → HLevel (ℓ-max (ℓ-max ℓ ℓ') ℓ'') (1 + nc + nf)) --- (d : (x : (A , ptA) ⋀ (C , ptC)) → P (((f , fpt) ⋀⃗ idfun∙ (C , ptC) ) x) .fst) --- where --- F : (c : C) → _ --- F c = λ(s : (b : B) → P (inr (b , c)) .fst) → s ∘ f - --- step1 : (c : C) → isHLevelTruncatedFun (1 + nc) (F c) --- step1 c = isOfHLevelPrecomposeConnected (1 + nc) nf ((λ b → P (inr (b , c)) .fst , P (inr (b , c)) .snd)) f conf - --- codomFun : (c : C) (a : A) → P (inr ((f a) , c)) .fst --- codomFun c = d ∘ λ a → inr (a , c) - --- Q : C → HLevel _ (1 + nc) --- Q c = fiber (F c) (codomFun c) , step1 c _ - - --- helper : (a : A) → transport (λ i → P (push (inl (f a)) i) .fst) (d (inl tt)) --- ≡ d (inr (a , ptC)) --- helper a = transport (PathP≡Path (λ i → P (push (inl (f a)) i) .fst) (d (inl tt)) (d (inr (a , ptC)))) --- (transport (λ j → PathP (λ i → P (rUnit (push (inl (f a))) (~ j) i) .fst) (d (inl tt)) (d (inr (a , ptC)))) --- λ i → d (push (inl a) i)) - - --- QptC : Q ptC .fst --- QptC = (λ b → transport (λ i → P (push (inl b) i) .fst) (d (inl tt))) , --- funExt helper --- where - - --- Q-constructor : (c : C) → Q c .fst --- Q-constructor c = Iso.inv (elim.isIsoPrecompose (λ _ → ptC) (1 + nc) Q (isHLevelConnectedPoint _ conC ptC)) (λ ( _ : Unit) → QptC) c - --- g : (b : B) (c : C) → P (inr (b , c)) .fst --- g b c = Q-constructor c .fst b +-- {- +-- j = i0 ⊢ path1 (f a) i +-- j = i1 ⊢ path1* a i +-- i = i0 ⊢ g (f a) ptC +-- i = i1 ⊢ d basel +-- -} --- Q-constructor-β : (b : B) → Q-constructor ptC .fst b ≡ transport (λ i → P (push (inl b) i) .fst) (d (inl tt)) --- Q-constructor-β b = ((λ i → (Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) (helper3 i)) .fst b)) ∙ --- (λ i → transportRefl (QptC) i .fst b) +-- h : (x : Smash (B , ptB) (C , ptC)) → P x .fst +-- h basel = d basel +-- h baser = d baser +-- h (proj x y) = g x y +-- h (gluel b i) = path1 b i +-- h (gluer c i) = path2 c i + +-- hsection : (x : Smash (A , ptA) (C , ptC)) → h (Smash-map (f , fpt) (idfun∙ (C , ptC)) x) ≡ d x +-- hsection basel = refl +-- hsection baser = refl +-- hsection (proj x y) i = Q-constructor y .snd i x +-- hsection (gluel a i) j = +-- comp (λ k → P (lUnit (gluel (f a)) (j ∨ k) i) .fst) +-- (λ k → λ { (i = i0) → Q-constructor ptC .snd j a +-- ; (i = i1) → d basel +-- ; (j = i0) → h (lUnit (gluel (f a)) k i) +-- ; (j = i1) → {!d (gluel a i)!}}) +-- {!!} +-- {- +-- comp (λ k → P (compPath-filler refl (gluel (f a)) (j ∨ k) i) .fst) +-- (λ k → λ { (i = i0) → Q-constructor ptC .snd j a +-- ; (i = i1) → path1 (f a) (k ∨ j) +-- ; (j = i0) → h (compPath-filler refl (gluel (f a)) k i) +-- ; (j = i1) → d (gluel a i)}) +-- {!!} -} +-- {- +-- comp (λ k → P (compPath-filler refl (gluel (f a)) (k ∨ j) i) .fst) +-- (λ k → λ { (i = i0) → Q-constructor ptC .snd j a +-- ; (i = i1) → {!path1 (f a) ?!} -- path1-filler (f a) k (k ∨ j) -- path1 (f a) (k ∨ j) +-- ; (j = i0) → h (compPath-filler refl (gluel (f a)) k i) +-- ; (j = i1) → {!!}}) -- d (gluel a i) +-- -- transp (λ r → P (compPath-filler refl (gluel (f a)) (r ∨ k) i) .fst) k ((transp (λ r → P (compPath-filler refl (gluel (f a)) (~ r ∨ k) i) .fst)) k (d (gluel a i))) +-- {!Q-constructor-β (f a)!} +-- -} -- where --- helper3 : (isHLevelConnectedPoint _ conC ptC) ptC .fst ≡ ∣ tt , refl ∣ --- helper3 = (isHLevelConnectedPoint _ conC ptC) ptC .snd ∣ tt , refl ∣ - --- gid1 : (a : A) (c : C) → g (f a) c ≡ d (inr (a , c)) --- gid1 a c i = (Q-constructor c .snd) i a - --- gid2 : (b : B) → g b ptC ≡ transport (λ i → P (push (inl b) i) .fst) (d (inl tt)) --- gid2 b = Q-constructor-β b - --- {- --- compPath : (c : C) → PathP _ (d (inl tt)) (g ptB c) --- compPath c = compPathP (λ i → d (push (inr c) i)) (compPathP (sym (gid1 ptA c)) (λ i → g (fpt i) c)) - - --- compPathTransport : (c : C) → ((λ z → P ((push (inr c) ∙ --- (λ i → inr (fpt (~ i) , c))) z) .fst) ∙ --- (λ i → ((λ i₁ → P (inr (f ptA , c)) .fst) ∙ (λ i₁ → P (inr (fpt i₁ , c)) .fst)) i)) --- ≡ λ i → P (push (inr c) i) .fst --- compPathTransport c = (λ k → ((λ z → P ((push (inr c) ∙ --- (λ i → inr (fpt (~ i) , c))) z) .fst) ∙ --- ((lUnit (λ i₁ → P (inr (fpt i₁ , c)) .fst) (~ k) )))) --- ∙ (λ k → ((λ z → P ((push (inr c) ∙ --- (λ i → inr (fpt (~ i ∨ k) , c))) z) .fst) ∙ --- (λ i₁ → P (inr (fpt (i₁ ∨ k) , c)) .fst) )) --- ∙ (λ k → ((λ z → P ((push (inr c) ∙ refl) z) .fst) ∙ refl)) --- ∙ (λ k → rUnit ((λ z → P ((rUnit (push (inr c)) (~ k)) z) .fst)) (~ k)) --- -} --- compP-filler : (c : C) → I → I → I → (B , ptB) ⋀ (C , ptC) --- compP-filler c i j z = --- hfill (λ k → λ { (i = i0) → inr (fpt (~ j ∧ ~ k) , c) --- ; (i = i1) → push (inr c) j --- ; (j = i1) → inr (fpt i , c)}) --- (inS (invSides-filler (λ i → inr (fpt (~ i) , c)) --- (sym (push (inr c))) i j)) --- z - - --- gid2-filler : C → I → I → {!!} --- gid2-filler c i j = --- fill (λ l → P (compP-filler c l i i1) .fst) --- (λ k → λ { (i = i0) → d (push (inr c) (~ k)) --- ; (i = i1) → g (fpt k) c }) --- (inS (gid1 ptA c (~ i))) --- {!j!} - --- gid1Path : (c : C) → PathP (λ i → P (push (inr c) i) .fst) (d (inl tt)) (g ptB c) --- gid1Path c i = comp (λ j → P (compP-filler c j i i1) .fst) --- (λ k → λ { (i = i0) → d (push (inr c) (~ k)) --- ; (i = i1) → g (fpt k) c }) --- (gid1 ptA c (~ i)) - - - --- gid2Path : (b : B) → PathP (λ i → P (push (inl b) i) .fst) (d (inl tt)) (g b ptC) --- gid2Path b i = --- comp (λ _ → P (push (inl b) i) .fst) --- (λ k → λ {(i = i0) → (d (inl tt)) --- ; (i = i1) → gid2 b (~ k)}) --- (transp (λ j → P (push (inl b) (j ∧ i)) .fst) (~ i) (d (inl tt))) - --- gid1Path≡gid2Path : PathP (λ j → PathP (λ i → P (push (push tt (~ j)) i) .fst) (d (inl tt)) (g ptB ptC)) (gid1Path ptC) (gid2Path ptB) --- gid1Path≡gid2Path j i = --- comp (λ k → P {!Q-constructor-β ptB!} .fst) --- (λ k → λ { (i = i0) → {!!} -- Q-constructor-β (fpt (~ k)) j --- ; (i = i1) → {!!} -- d (push (inr ptC) (~ k ∨ j)) --- ; (j = i0) → {!!} --- ; (j = i1) → {!!} }) --- {!gid2!} - --- -- -- {- --- -- -- Goal: P (push (push tt (~ j)) i) .fst --- -- -- ———— Boundary —————————————————————————————————————————————— --- -- -- j = i0 ⊢ gid1path ptC i --- -- -- j = i1 ⊢ gid2Path ptB i --- -- -- i = i0 ⊢ d (inl tt) --- -- -- i = i1 ⊢ g ptB ptC --- -- -- -} +-- test2 : I → (i : I) → P ((refl ∙ gluel (f a)) i) .fst +-- test2 j i = +-- comp (λ k → P (compPath-filler refl (gluel (f a)) (j ∨ k) i) .fst) +-- (λ k → λ { (i = i0) → d (proj a ptC) -- g (f a) ptC +-- ; (i = i1) → transp (λ r → P (gluel (f a) (r ∧ (k ∨ j))) .fst) {!j!} (Q-constructor-β (f a) j) -- transp (λ r → P (gluel (f a) (r ∧ k)) .fst) (~ k) (Q-constructor-β (f a) j) +-- ; (j = i0) → compPathP'-filler (λ i → Q-constructor ptC .snd (~ i) a) (λ i → transp (λ r → P (gluel (f a) (r ∧ i)) .fst) (~ i) (g (f a) ptC)) k i +-- ; (j = i1) → d (gluel a i) }) -- d (gluel a i)}) +-- {!!} + + +-- {- +-- (comp (λ k → P (compPath-filler refl (gluel (f a)) (j ∧ k) i) .fst) +-- (λ k → λ { (i = i0) → Q-constructor ptC .snd j a +-- ; (i = i1) → path1-filler (f a) k (j ∧ k) +-- ; (j = i0) → g (f a) ptC +-- ; (j = i1) → side-filler i k}) +-- {!compPath-filler refl (gluel (f a)) i0 j!}) -} + +-- side-filler : (i : I) (k : I) → P (compPath-filler (λ _ → proj (f a) (pt (C , ptC))) (gluel (f a)) k i) .fst +-- side-filler i k = +-- comp (λ l → P (compPath-filler refl (gluel (f a)) k {!!}) .fst) +-- (λ l → λ { (i = i0) → {!transport (λ i₂ → P (gluel (f a) ₂) .fst) (g (f a) ptC)!} +-- ; (i = i1) → {!!} +-- ; (k = i0) → {!!} +-- ; (k = i1) → {!!} }) +-- {!otherway (f a) ?!} +-- hsection (gluer c i) j = {!transp (λ r → P (((refl ∙ gluel (f a)) ((r ∨ ?))) .fst)) l (transp (λ r → P (((refl ∙ gluel (f a)) ((~ r) ∨ ?))) .fst) l (d (gluel a i))) !} +-- {- +-- comp (λ k → P (compPath-filler (λ i → proj (fpt i) c) (gluer c) (k ∨ j) i) .fst) +-- (λ k → λ { (i = i0) → Q-constructor c .snd j ptA +-- ; (i = i1) → path2 c (k ∨ j) +-- ; (j = i0) → h (compPath-filler (λ i → proj (fpt i) c) (gluer c) k i) +-- ; (j = i1) → d (gluer c i) +-- }) +-- (comp (λ k → P (compPath-filler (λ i₁ → proj (fpt i₁) c) (gluer c) (j ∧ k) (i ∧ k)) .fst) +-- (λ k → λ { (i = i0) → Q-constructor c .snd j ptA +-- ; (j = i0) → h (proj (fpt (i ∧ k)) c) +-- ; (j = i1) → d (gluer c (i ∧ k)) +-- }) +-- (Q-constructor c .snd j ptA)) +-- -} + +-- -- isConnectedSmashIdfun : (f : A' →∙ B') (nf nc : ℕ) +-- -- → isHLevelConnectedFun nf (fst f) +-- -- → isHLevelConnected (2 + nc) (fst C') +-- -- → isHLevelConnectedFun (1 + nf + nc) (f ⋀⃗ idfun∙ C') +-- -- isConnectedSmashIdfun {A' = (A , ptA)} {B' = (B , ptB)} {C' = (C , ptC)} (f , fpt) nf nc conf conC = {!isHLel!} +-- -- where +-- -- module _ (P : ((B , ptB) ⋀ (C , ptC)) → HLevel (ℓ-max (ℓ-max ℓ ℓ') ℓ'') (1 + nc + nf)) +-- -- (d : (x : (A , ptA) ⋀ (C , ptC)) → P (((f , fpt) ⋀⃗ idfun∙ (C , ptC) ) x) .fst) +-- -- where +-- -- F : (c : C) → _ +-- -- F c = λ(s : (b : B) → P (inr (b , c)) .fst) → s ∘ f + +-- -- step1 : (c : C) → isHLevelTruncatedFun (1 + nc) (F c) +-- -- step1 c = isOfHLevelPrecomposeConnected (1 + nc) nf ((λ b → P (inr (b , c)) .fst , P (inr (b , c)) .snd)) f conf + +-- -- codomFun : (c : C) (a : A) → P (inr ((f a) , c)) .fst +-- -- codomFun c = d ∘ λ a → inr (a , c) + +-- -- Q : C → HLevel _ (1 + nc) +-- -- Q c = fiber (F c) (codomFun c) , step1 c _ + + +-- -- helper : (a : A) → transport (λ i → P (push (inl (f a)) i) .fst) (d (inl tt)) +-- -- ≡ d (inr (a , ptC)) +-- -- helper a = transport (PathP≡Path (λ i → P (push (inl (f a)) i) .fst) (d (inl tt)) (d (inr (a , ptC)))) +-- -- (transport (λ j → PathP (λ i → P (rUnit (push (inl (f a))) (~ j) i) .fst) (d (inl tt)) (d (inr (a , ptC)))) +-- -- λ i → d (push (inl a) i)) + + +-- -- QptC : Q ptC .fst +-- -- QptC = (λ b → transport (λ i → P (push (inl b) i) .fst) (d (inl tt))) , +-- -- funExt helper +-- -- where + + +-- -- Q-constructor : (c : C) → Q c .fst +-- -- Q-constructor c = Iso.inv (elim.isIsoPrecompose (λ _ → ptC) (1 + nc) Q (isHLevelConnectedPoint _ conC ptC)) (λ ( _ : Unit) → QptC) c + +-- -- g : (b : B) (c : C) → P (inr (b , c)) .fst +-- -- g b c = Q-constructor c .fst b + +-- -- Q-constructor-β : (b : B) → Q-constructor ptC .fst b ≡ transport (λ i → P (push (inl b) i) .fst) (d (inl tt)) +-- -- Q-constructor-β b = ((λ i → (Trunc.rec (Q ptC .snd) (λ { (a , p) → subst (fst ∘ Q) p QptC}) (helper3 i)) .fst b)) ∙ +-- -- (λ i → transportRefl (QptC) i .fst b) + +-- -- where +-- -- helper3 : (isHLevelConnectedPoint _ conC ptC) ptC .fst ≡ ∣ tt , refl ∣ +-- -- helper3 = (isHLevelConnectedPoint _ conC ptC) ptC .snd ∣ tt , refl ∣ + +-- -- gid1 : (a : A) (c : C) → g (f a) c ≡ d (inr (a , c)) +-- -- gid1 a c i = (Q-constructor c .snd) i a + +-- -- gid2 : (b : B) → g b ptC ≡ transport (λ i → P (push (inl b) i) .fst) (d (inl tt)) +-- -- gid2 b = Q-constructor-β b + +-- -- {- +-- -- compPath : (c : C) → PathP _ (d (inl tt)) (g ptB c) +-- -- compPath c = compPathP (λ i → d (push (inr c) i)) (compPathP (sym (gid1 ptA c)) (λ i → g (fpt i) c)) + + +-- -- compPathTransport : (c : C) → ((λ z → P ((push (inr c) ∙ +-- -- (λ i → inr (fpt (~ i) , c))) z) .fst) ∙ +-- -- (λ i → ((λ i₁ → P (inr (f ptA , c)) .fst) ∙ (λ i₁ → P (inr (fpt i₁ , c)) .fst)) i)) +-- -- ≡ λ i → P (push (inr c) i) .fst +-- -- compPathTransport c = (λ k → ((λ z → P ((push (inr c) ∙ +-- -- (λ i → inr (fpt (~ i) , c))) z) .fst) ∙ +-- -- ((lUnit (λ i₁ → P (inr (fpt i₁ , c)) .fst) (~ k) )))) +-- -- ∙ (λ k → ((λ z → P ((push (inr c) ∙ +-- -- (λ i → inr (fpt (~ i ∨ k) , c))) z) .fst) ∙ +-- -- (λ i₁ → P (inr (fpt (i₁ ∨ k) , c)) .fst) )) +-- -- ∙ (λ k → ((λ z → P ((push (inr c) ∙ refl) z) .fst) ∙ refl)) +-- -- ∙ (λ k → rUnit ((λ z → P ((rUnit (push (inr c)) (~ k)) z) .fst)) (~ k)) +-- -- -} +-- -- compP-filler : (c : C) → I → I → I → (B , ptB) ⋀ (C , ptC) +-- -- compP-filler c i j z = +-- -- hfill (λ k → λ { (i = i0) → inr (fpt (~ j ∧ ~ k) , c) +-- -- ; (i = i1) → push (inr c) j +-- -- ; (j = i1) → inr (fpt i , c)}) +-- -- (inS (invSides-filler (λ i → inr (fpt (~ i) , c)) +-- -- (sym (push (inr c))) i j)) +-- -- z + + +-- -- gid2-filler : C → I → I → {!!} +-- -- gid2-filler c i j = +-- -- fill (λ l → P (compP-filler c l i i1) .fst) +-- -- (λ k → λ { (i = i0) → d (push (inr c) (~ k)) +-- -- ; (i = i1) → g (fpt k) c }) +-- -- (inS (gid1 ptA c (~ i))) +-- -- {!j!} + +-- -- gid1Path : (c : C) → PathP (λ i → P (push (inr c) i) .fst) (d (inl tt)) (g ptB c) +-- -- gid1Path c i = comp (λ j → P (compP-filler c j i i1) .fst) +-- -- (λ k → λ { (i = i0) → d (push (inr c) (~ k)) +-- -- ; (i = i1) → g (fpt k) c }) +-- -- (gid1 ptA c (~ i)) + + + +-- -- gid2Path : (b : B) → PathP (λ i → P (push (inl b) i) .fst) (d (inl tt)) (g b ptC) +-- -- gid2Path b i = +-- -- comp (λ _ → P (push (inl b) i) .fst) +-- -- (λ k → λ {(i = i0) → (d (inl tt)) +-- -- ; (i = i1) → gid2 b (~ k)}) +-- -- (transp (λ j → P (push (inl b) (j ∧ i)) .fst) (~ i) (d (inl tt))) + +-- -- gid1Path≡gid2Path : PathP (λ j → PathP (λ i → P (push (push tt (~ j)) i) .fst) (d (inl tt)) (g ptB ptC)) (gid1Path ptC) (gid2Path ptB) +-- -- gid1Path≡gid2Path j i = +-- -- comp (λ k → P {!Q-constructor-β ptB!} .fst) +-- -- (λ k → λ { (i = i0) → {!!} -- Q-constructor-β (fpt (~ k)) j +-- -- ; (i = i1) → {!!} -- d (push (inr ptC) (~ k ∨ j)) +-- -- ; (j = i0) → {!!} +-- -- ; (j = i1) → {!!} }) +-- -- {!gid2!} + +-- -- -- -- {- +-- -- -- -- Goal: P (push (push tt (~ j)) i) .fst +-- -- -- -- ———— Boundary —————————————————————————————————————————————— +-- -- -- -- j = i0 ⊢ gid1path ptC i +-- -- -- -- j = i1 ⊢ gid2Path ptB i +-- -- -- -- i = i0 ⊢ d (inl tt) +-- -- -- -- i = i1 ⊢ g ptB ptC +-- -- -- -- -} --- -- h : (x : (B , ptB) ⋀ (C , ptC)) → P x .fst --- -- h (inl x) = d (inl tt) --- -- h (inr (b , c)) = g b c --- -- h (push (inl b) i) = gid2Path b i --- -- h (push (inr x) i) = gid1Path x i --- -- h (push (push tt i) j) = gid1Path≡gid2Path (~ i) j - --- -- sect-h : (x : (A , ptA) ⋀ (C , ptC)) → h (((f , fpt) ⋀⃗ idfun∙ (C , ptC)) x) ≡ d x --- -- sect-h (inl x) = refl --- -- sect-h (inr (x , x₁)) i = gid1 x x₁ i -- (Q-constructor x₁ .snd) i x --- -- sect-h (push (inl x) i) j = {!!} --- -- sect-h (push (inr x) i) = {!!} --- -- sect-h (push (push a i) j) k = {!!} +-- -- -- h : (x : (B , ptB) ⋀ (C , ptC)) → P x .fst +-- -- -- h (inl x) = d (inl tt) +-- -- -- h (inr (b , c)) = g b c +-- -- -- h (push (inl b) i) = gid2Path b i +-- -- -- h (push (inr x) i) = gid1Path x i +-- -- -- h (push (push tt i) j) = gid1Path≡gid2Path (~ i) j + +-- -- -- sect-h : (x : (A , ptA) ⋀ (C , ptC)) → h (((f , fpt) ⋀⃗ idfun∙ (C , ptC)) x) ≡ d x +-- -- -- sect-h (inl x) = refl +-- -- -- sect-h (inr (x , x₁)) i = gid1 x x₁ i -- (Q-constructor x₁ .snd) i x +-- -- -- sect-h (push (inl x) i) j = {!!} +-- -- -- sect-h (push (inr x) i) = {!!} +-- -- -- sect-h (push (push a i) j) k = {!!} diff --git a/Cubical/ZCohomology/KcompPrelims.agda b/Cubical/ZCohomology/KcompPrelims.agda index f8d9f2219..98ab6f08e 100644 --- a/Cubical/ZCohomology/KcompPrelims.agda +++ b/Cubical/ZCohomology/KcompPrelims.agda @@ -50,9 +50,9 @@ private ϕ pt a = (merid a) ∙ sym (merid pt) {- To define the map for n=0 we use the λ k → loopᵏ map for S₊ 1. The loop is given by ϕ south north -} -private - loop* : Path (S₊ 1) north north - loop* = ϕ north south + +loop* : Path (S₊ 1) north north +loop* = ϕ north south looper : Int → Path (S₊ 1) north north looper (pos zero) = refl diff --git a/Cubical/ZCohomology/Sn/Sn.agda b/Cubical/ZCohomology/Sn/Sn.agda new file mode 100644 index 000000000..56b1cb0cd --- /dev/null +++ b/Cubical/ZCohomology/Sn/Sn.agda @@ -0,0 +1,211 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.Sn.Sn where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.S1.S1 +open import Cubical.HITs.S1 +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.MayerVietorisReduced +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence +open import Cubical.Data.NatMinusTwo.Base +open import Cubical.Data.Empty +open import Cubical.Data.Sigma hiding (_×_) +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₋₁ ; ∣_∣ to ∣_∣₋₁) +open import Cubical.HITs.Nullification +open import Cubical.Data.Int renaming (_+_ to _+ℤ_; +-comm to +ℤ-comm ; +-assoc to +ℤ-assoc) +open import Cubical.Data.Nat +open import Cubical.Data.Prod +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Freudenthal +open import Cubical.HITs.SmashProduct.Base +open import Cubical.Data.Unit +open import Cubical.Data.Group.Base renaming (Iso to grIso) + + +open import Cubical.ZCohomology.KcompPrelims + + +open import Cubical.HITs.Pushout +open import Cubical.Data.Sum.Base +open import Cubical.Data.HomotopyGroup +open import Cubical.Data.Bool hiding (_⊕_) + +pushoutSn : (n : ℕ) → Iso (S₊ (suc n)) (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt) +Iso.fun (pushoutSn n) north = inl tt +Iso.fun (pushoutSn n) south = inr tt +Iso.fun (pushoutSn n) (merid a i) = push a i +Iso.inv (pushoutSn n) (inl x) = north +Iso.inv (pushoutSn n) (inr x) = south +Iso.inv (pushoutSn n) (push a i) = merid a i +Iso.rightInv (pushoutSn n) (inl x) = refl +Iso.rightInv (pushoutSn n) (inr x) = refl +Iso.rightInv (pushoutSn n) (push a i) = refl +Iso.leftInv (pushoutSn n) north = refl +Iso.leftInv (pushoutSn n) south = refl +Iso.leftInv (pushoutSn n) (merid a i) = refl + +Sn≡Pushout : (n : ℕ) → (S₊ (suc n)) ≡ (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt) +Sn≡Pushout n = isoToPath (pushoutSn n) + +isContr→≡Unit : {A : Type₀} → isContr A → A ≡ Unit +isContr→≡Unit contr = isoToPath (iso (λ _ → tt) (λ _ → fst contr) (λ _ → refl) λ _ → snd contr _) + +isContr→isContrTrunc : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → isContr A → isContr (hLevelTrunc n A) +isContr→isContrTrunc n contr = ∣ fst contr ∣ , (trElim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → cong ∣_∣ (snd contr a)) +isContr→isContrSetTrunc : ∀ {ℓ} {A : Type ℓ} → isContr A → isContr (∥ A ∥₀) +isContr→isContrSetTrunc contr = ∣ fst contr ∣₀ , sElim (λ _ → isOfHLevelPath 2 (setTruncIsSet) _ _) λ a → cong ∣_∣₀ (snd contr a) + +-- coHom Unit + + +lemma10000 : (n : ℕ) → looper (pos (suc n)) ≡ loop* ∙ looper (pos n) +lemma10000 zero = sym (lUnit loop*) ∙ rUnit loop* +lemma10000 (suc n) = (λ i → (lemma10000 n i) ∙ loop*) ∙ sym (assoc loop* (looper (pos n)) loop*) + +lemma10001 : (a : Int) → looper (sucInt a) ≡ loop* ∙ looper a +lemma10001 (pos n) = lemma10000 n +lemma10001 (negsuc zero) = sym (rCancel loop*) +lemma10001 (negsuc (suc zero)) = lUnit (sym loop*) + ∙ (λ i → rCancel loop* (~ i) ∙ (sym loop*)) + ∙ sym (assoc loop* (sym loop*) (sym loop*)) +lemma10001 (negsuc (suc (suc n))) = cong (λ x → x ∙ (sym loop*)) (lemma10001 (negsuc (suc n))) + ∙ sym (assoc loop* (looper (negsuc n) ∙ sym loop*) (sym loop*)) + +lemma10002 : (a : Int) → looper (predInt a) ≡ sym loop* ∙ looper a +lemma10002 (pos zero) = rUnit (sym loop*) +lemma10002 (pos (suc n)) = lUnit (looper (pos n)) + ∙ (λ i → lCancel loop* (~ i) ∙ (looper (pos n))) + ∙ sym (assoc (sym loop*) loop* (looper (pos n))) + ∙ λ i → sym loop* ∙ lemma10000 n (~ i) +lemma10002 (negsuc zero) = refl +lemma10002 (negsuc (suc n)) = (λ i → (lemma10002 (negsuc n) i) ∙ sym loop*) ∙ sym (assoc (sym loop*) (looper (negsuc n)) (sym loop*)) + +-- loop*comm : (a : Int) → looper a ∙ loop* ≡ loop* ∙ looper a +-- loop*comm (pos zero) = sym (lUnit loop*) ∙ rUnit loop* +-- loop*comm (pos (suc n)) = (λ i → loop*comm (pos n) i ∙ loop*) ∙ sym (assoc loop* (looper (pos n)) loop*) +-- loop*comm (negsuc zero) = lCancel loop* ∙ sym (rCancel loop*) +-- loop*comm (negsuc (suc n)) = sym (assoc (looper (negsuc n)) (sym loop*) loop*) +-- ∙ (λ i → looper (negsuc n) ∙ lCancel loop* i) +-- ∙ (λ i → looper (negsuc n) ∙ rCancel loop* (~ i)) +-- ∙ assoc (looper (negsuc n)) loop* (sym loop*) +-- ∙ (λ i → loop*comm (negsuc n) i ∙ sym loop*) +-- ∙ sym (assoc loop* (looper (negsuc n)) (sym loop*)) +-- symloop*comm : (a : Int) → looper a ∙ sym loop* ≡ sym loop* ∙ looper a +-- symloop*comm (pos zero) = sym (lUnit (sym loop*)) ∙ rUnit (sym loop*) +-- symloop*comm (pos (suc n)) = sym (assoc (looper (pos n)) loop* (sym loop*)) +-- ∙ (λ i → looper (pos n) ∙ (rCancel loop* i)) +-- ∙ (λ i → looper (pos n) ∙ lCancel loop* (~ i)) +-- ∙ assoc (looper (pos n)) (sym loop*) loop* +-- ∙ (λ i → symloop*comm (pos n) i ∙ loop*) +-- ∙ sym (assoc (sym loop*) (looper (pos n)) loop*) +-- symloop*comm (negsuc zero) = refl +-- symloop*comm (negsuc (suc n)) = (λ i → symloop*comm (negsuc n) i ∙ sym loop*) +-- ∙ sym (assoc (sym loop*) (looper (negsuc n)) (sym loop*)) + + +sucIntPredInt : (a : Int) → sucInt (predInt a) ≡ a +sucIntPredInt (pos zero) = refl +sucIntPredInt (pos (suc n)) = refl +sucIntPredInt (negsuc n) = refl + +looperId : (a b : Int) → looper a ∙ looper b ≡ looper (a +ℤ b) +looperId (pos zero) b = sym (lUnit (looper b)) ∙ λ i → looper (pos0+ b i) +looperId (pos (suc n)) (pos zero) = sym (rUnit (looper (pos (suc n)))) +looperId (pos (suc n)) (pos (suc m)) = (λ i → (lemma10000 n i) ∙ looper (pos (suc m))) + ∙ sym (assoc loop* (looper (pos n)) (looper (pos (suc m)))) + ∙ (λ i → loop* ∙ looperId (pos n) (pos (suc m)) i) + ∙ sym (lemma10001 (sucInt (pos n +pos m))) + ∙ λ i → looper (sucInt (sucInt+pos m (pos n) i)) +looperId (pos (suc n)) (negsuc zero) = sym (assoc (looper (pos n)) loop* (sym loop*)) ∙ (λ i → looper (pos n) ∙ (rCancel loop* i)) ∙ sym (rUnit (looper (pos n))) +looperId (pos (suc n)) (negsuc (suc m)) = (λ i → (looper (pos n) ∙ loop*) ∙ lemma10002 (negsuc m) i) + ∙ sym (assoc (looper (pos n)) loop* (sym loop* ∙ looper (negsuc m))) + ∙ (λ i → looper (pos n) ∙ assoc loop* (sym loop*) (looper (negsuc m)) i) + ∙ (λ i → looper (pos n) ∙ rCancel loop* i ∙ looper (negsuc m)) + ∙ (λ i → looper (pos n) ∙ lUnit (looper (negsuc m)) (~ i)) + ∙ looperId (pos n) (negsuc m) + ∙ λ i → looper (predInt+negsuc m (pos (suc n)) (~ i)) +looperId (negsuc zero) (pos zero) = sym (rUnit (sym loop*)) +looperId (negsuc zero) (pos (suc n)) = (λ i → (sym loop*) ∙ lemma10001 (pos n) i) + ∙ assoc (sym loop*) loop* (looper (pos n)) + ∙ (λ i → lCancel loop* i ∙ looper (pos n)) + ∙ sym (lUnit (looper (pos n))) + ∙ (λ i → looper (sucIntPredInt (pos n) (~ i))) + ∙ λ i → looper (sucInt (negsuc0+ (pos n) i)) +looperId (negsuc zero) (negsuc zero) = refl +looperId (negsuc zero) (negsuc (suc n)) = assoc (sym loop*) (looper (negsuc n)) (sym loop*) ∙ + (λ i → looperId (negsuc zero) (negsuc n) i ∙ sym loop*) ∙ + (λ i → looper (negsuc zero +negsuc n) ∙ sym loop*) ∙ + (λ i → looper (negsuc0+ (negsuc n) (~ i)) ∙ sym loop*) ∙ + λ i → looper (predInt (negsuc0+ (negsuc n) i)) +looperId (negsuc (suc n)) b = (λ i → lemma10002 (negsuc n) i ∙ looper b) + ∙ sym (assoc (sym loop*) (looper (negsuc n)) (looper b)) + ∙ (λ i → sym loop* ∙ looperId (negsuc n) b i) + ∙ sym (lemma10002 (negsuc n +ℤ b)) + ∙ λ i → looper (predInt+ (negsuc n) b i) + +addLemma : (a b : Int) → a +ₖ b ≡ (a +ℤ b) +addLemma a b = (λ i → ΩKn+1→Kn (cong ∣_∣ (looper a) ∙ cong ∣_∣ (looper b))) + ∙ (λ i → ΩKn+1→Kn (congFunct ∣_∣ (looper a) (looper b) (~ i))) + ∙ (λ i → ΩKn+1→Kn (cong ∣_∣ (looperId a b i))) + ∙ Iso.leftInv (Iso3-Kn-ΩKn+1 zero) (a +ℤ b) + + +coHomGrUnit0 : grIso (coHomGr 0 Unit) intGroup +grIso.fun coHomGrUnit0 = sRec isSetInt (λ f → f tt) , + sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) λ a b → addLemma (a tt) (b tt) +grIso.inv coHomGrUnit0 = (λ a → ∣ (λ _ → a) ∣₀) , λ a b i → ∣ (λ _ → addLemma a b (~ i)) ∣₀ +grIso.rightInv coHomGrUnit0 a = refl +grIso.leftInv coHomGrUnit0 = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ a → refl + +coHomGrUnit≥1 : (n : ℕ) → grIso (coHomGr (suc n) Unit) trivialGroup +grIso.fun (coHomGrUnit≥1 n) = (λ _ → tt) , (λ _ _ → refl) +grIso.inv (coHomGrUnit≥1 n) = (λ _ → ∣ (λ _ → ∣ north ∣) ∣₀) , λ _ _ → sym (rUnitₕ 0ₕ) +grIso.rightInv (coHomGrUnit≥1 n) a = refl +grIso.leftInv (coHomGrUnit≥1 n) a = sym (helper2 .snd 0ₕ) ∙ helper2 .snd a + where + helper : (n : ℕ) → isContr (∥ coHomK (suc n) ∥₀) + helper n = subst isContr ((isoToPath (truncOfTruncIso {A = S₊ (1 + n)} 2 (1 + n))) ∙ sym propTrunc≡Trunc0 ∙ λ i → ∥ hLevelTrunc (suc (+-comm n 2 i)) (S₊ (1 + n)) ∥₀) + (isConnectedSubtr 2 (helper2 n .fst) (subst (λ x → isHLevelConnected x (S₊ (suc n))) (sym (helper2 n .snd)) (sphereConnected (suc n))) ) + where + helper2 : (n : ℕ) → Σ[ m ∈ ℕ ] m + 2 ≡ 2 + n + helper2 zero = 0 , refl + helper2 (suc n) = (suc n) , λ i → suc (+-comm n 2 i) + + helper2 : isContr (coHom (suc n) Unit) + helper2 = subst isContr (λ i → ∥ UnitToTypeId (coHomK (suc n)) (~ i) ∥₀) (helper n) + +coHom0-S0 : grIso (coHomGr 0 (S₊ 0)) (dirProd intGroup intGroup) +coHom0-S0 = + Iso'→Iso + (iso' + (iso (sRec (isOfHLevelProd 2 isSetInt isSetInt) + λ f → (f north) , (f south)) + (λ {(a , b) → ∣ (λ {north → a ; south → b}) ∣₀}) + (λ { (a , b) → refl}) + (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ f → cong ∣_∣₀ (funExt (λ {north → refl ; south → refl})))) + (sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 isSetInt isSetInt) _ _) + λ a b i → addLemma (a north) (b north) i , addLemma (a south) (b south) i)) + +coHomGr0-S1 : grIso (coHomGr 1 (S₊ 1)) intGroup +coHomGr0-S1 = + Iso'→Iso + (iso' + (iso (sRec isSetInt {!!}) + {!!} + {!!} + {!!}) + {!!}) From 3e3ac2a13c6cc8bd4d522ff24a8f84a47b320165 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Sun, 31 May 2020 03:12:00 +0200 Subject: [PATCH 36/89] injectivity --- Cubical/Data/Group/GroupLibrary.agda | 121 +-- Cubical/HITs/S1/Base.agda | 107 ++- Cubical/HITs/Truncation/Properties.agda | 17 +- Cubical/Homotopy/Connected.agda | 10 +- Cubical/Homotopy/Freudenthal.agda | 6 +- Cubical/ZCohomology/KcompPrelims.agda | 10 +- Cubical/ZCohomology/MayerVietorisReduced.agda | 2 +- Cubical/ZCohomology/Properties.agda | 2 +- Cubical/ZCohomology/Sn/Sn.agda | 801 ++++++++++++++++-- 9 files changed, 856 insertions(+), 220 deletions(-) diff --git a/Cubical/Data/Group/GroupLibrary.agda b/Cubical/Data/Group/GroupLibrary.agda index 65a758adb..d515f0ab1 100644 --- a/Cubical/Data/Group/GroupLibrary.agda +++ b/Cubical/Data/Group/GroupLibrary.agda @@ -12,67 +12,72 @@ import Cubical.Foundations.Isomorphism as I import Cubical.Foundations.Equiv as E import Cubical.Foundations.Equiv.HalfAdjoint as HAE +open import Cubical.Data.Sigma hiding (_×_) + open import Cubical.HITs.SetQuotients as sq --- imGroup : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → morph G H --- → Group (ℓ-max ℓ ℓ') --- Group.type (imGroup G H (ϕ , mf)) = Σ[ x ∈ Group.type H ] isInIm G H ϕ x --- Group.setStruc (imGroup G H (ϕ , mf)) = isOfHLevelΣ 2 (Group.setStruc H) λ _ → isOfHLevelSuc 1 propTruncIsProp --- isGroup.id (Group.groupStruc (imGroup G H (ϕ , mf))) = --- let idH = isGroup.id (Group.groupStruc H) --- idG = isGroup.id (Group.groupStruc G) --- invG = isGroup.inv (Group.groupStruc G) --- invH = isGroup.inv (Group.groupStruc H) --- lUnit = isGroup.lUnit (Group.groupStruc H) --- rCancel = isGroup.rCancel (Group.groupStruc H) --- in idH , ∣ idG , morph0→0 G H ϕ mf ∣ --- isGroup.inv (Group.groupStruc (imGroup G H (ϕ , mf))) (h , hinim) = --- let idH = isGroup.id (Group.groupStruc H) --- idG = isGroup.id (Group.groupStruc G) --- invG = isGroup.inv (Group.groupStruc G) --- invH = isGroup.inv (Group.groupStruc H) --- lUnit = isGroup.lUnit (Group.groupStruc H) --- rCancel = isGroup.rCancel (Group.groupStruc H) --- in invH h , rec propTruncIsProp --- (λ a → ∣ (invG (fst a)) --- , morphMinus G H ϕ mf (fst a) ∙ cong (λ x → invH x) (snd a) ∣) --- hinim --- isGroup.comp (Group.groupStruc (imGroup G H (ϕ , mf))) (h1 , h1inim) (h2 , h2inim) = --- let idH = isGroup.id (Group.groupStruc H) --- idG = isGroup.id (Group.groupStruc G) --- invG = isGroup.inv (Group.groupStruc G) --- invH = isGroup.inv (Group.groupStruc H) --- compH = isGroup.comp (Group.groupStruc H) --- compG = isGroup.comp (Group.groupStruc G) --- lUnit = isGroup.lUnit (Group.groupStruc H) --- rCancel = isGroup.rCancel (Group.groupStruc H) --- in compH h1 h2 , rec propTruncIsProp --- (λ p1 → rec propTruncIsProp --- (λ p2 → ∣ (compG (fst p1) (fst p2)) --- , mf (fst p1) (fst p2) ∙ cong₂ compH (snd p1) (snd p2) ∣) --- h2inim) --- h1inim --- isGroup.lUnit (Group.groupStruc (imGroup G H (ϕ , mf))) (h , _) = --- let lUnit = isGroup.lUnit (Group.groupStruc H) --- in ΣProp≡ (λ _ → propTruncIsProp) --- (lUnit h) --- isGroup.rUnit (Group.groupStruc (imGroup G H (ϕ , mf))) (h , _) = --- let rUnit = isGroup.rUnit (Group.groupStruc H) --- in ΣProp≡ (λ _ → propTruncIsProp) --- (rUnit h) --- isGroup.assoc (Group.groupStruc (imGroup G H (ϕ , mf))) (h1 , _) (h2 , _) (h3 , _) = --- let assoc = isGroup.assoc (Group.groupStruc H) --- in ΣProp≡ (λ _ → propTruncIsProp) --- (assoc h1 h2 h3) --- isGroup.lCancel (Group.groupStruc (imGroup G H (ϕ , mf))) (h , _) = --- let lCancel = isGroup.lCancel (Group.groupStruc H) --- in ΣProp≡ (λ _ → propTruncIsProp) --- (lCancel h) --- isGroup.rCancel (Group.groupStruc (imGroup G H (ϕ , mf))) (h , _) = --- let rCancel = isGroup.rCancel (Group.groupStruc H) --- in ΣProp≡ (λ _ → propTruncIsProp) --- (rCancel h) + + +imGroup : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → morph G H + + → Group (ℓ-max ℓ ℓ') +Group.type (imGroup G H (ϕ , mf)) = Σ[ x ∈ Group.type H ] isInIm G H ϕ x +Group.setStruc (imGroup G H (ϕ , mf)) = isOfHLevelΣ 2 (Group.setStruc H) λ _ → isOfHLevelSuc 1 propTruncIsProp +isGroup.id (Group.groupStruc (imGroup G H (ϕ , mf))) = + let idH = isGroup.id (Group.groupStruc H) + idG = isGroup.id (Group.groupStruc G) + invG = isGroup.inv (Group.groupStruc G) + invH = isGroup.inv (Group.groupStruc H) + lUnit = isGroup.lUnit (Group.groupStruc H) + rCancel = isGroup.rCancel (Group.groupStruc H) + in idH , ∣ idG , morph0→0 G H ϕ mf ∣ +isGroup.inv (Group.groupStruc (imGroup G H (ϕ , mf))) (h , hinim) = + let idH = isGroup.id (Group.groupStruc H) + idG = isGroup.id (Group.groupStruc G) + invG = isGroup.inv (Group.groupStruc G) + invH = isGroup.inv (Group.groupStruc H) + lUnit = isGroup.lUnit (Group.groupStruc H) + rCancel = isGroup.rCancel (Group.groupStruc H) + in invH h , rec propTruncIsProp + (λ a → ∣ (invG (fst a)) + , morphMinus G H ϕ mf (fst a) ∙ cong (λ x → invH x) (snd a) ∣) + hinim +isGroup.comp (Group.groupStruc (imGroup G H (ϕ , mf))) (h1 , h1inim) (h2 , h2inim) = + let idH = isGroup.id (Group.groupStruc H) + idG = isGroup.id (Group.groupStruc G) + invG = isGroup.inv (Group.groupStruc G) + invH = isGroup.inv (Group.groupStruc H) + compH = isGroup.comp (Group.groupStruc H) + compG = isGroup.comp (Group.groupStruc G) + lUnit = isGroup.lUnit (Group.groupStruc H) + rCancel = isGroup.rCancel (Group.groupStruc H) + in compH h1 h2 , rec propTruncIsProp + (λ p1 → rec propTruncIsProp + (λ p2 → ∣ (compG (fst p1) (fst p2)) + , mf (fst p1) (fst p2) ∙ cong₂ compH (snd p1) (snd p2) ∣) + h2inim) + h1inim +isGroup.lUnit (Group.groupStruc (imGroup G H (ϕ , mf))) (h , _) = + let lUnit = isGroup.lUnit (Group.groupStruc H) + in ΣProp≡ (λ _ → propTruncIsProp) + (lUnit h) +isGroup.rUnit (Group.groupStruc (imGroup G H (ϕ , mf))) (h , _) = + let rUnit = isGroup.rUnit (Group.groupStruc H) + in ΣProp≡ (λ _ → propTruncIsProp) + (rUnit h) +isGroup.assoc (Group.groupStruc (imGroup G H (ϕ , mf))) (h1 , _) (h2 , _) (h3 , _) = + let assoc = isGroup.assoc (Group.groupStruc H) + in ΣProp≡ (λ _ → propTruncIsProp) + (assoc h1 h2 h3) +isGroup.lCancel (Group.groupStruc (imGroup G H (ϕ , mf))) (h , _) = + let lCancel = isGroup.lCancel (Group.groupStruc H) + in ΣProp≡ (λ _ → propTruncIsProp) + (lCancel h) +isGroup.rCancel (Group.groupStruc (imGroup G H (ϕ , mf))) (h , _) = + let rCancel = isGroup.rCancel (Group.groupStruc H) + in ΣProp≡ (λ _ → propTruncIsProp) + (rCancel h) -- kerGroup : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → morph G H diff --git a/Cubical/HITs/S1/Base.agda b/Cubical/HITs/S1/Base.agda index c2cf8337e..c9103a2fc 100644 --- a/Cubical/HITs/S1/Base.agda +++ b/Cubical/HITs/S1/Base.agda @@ -183,60 +183,59 @@ basedΩS¹→ΩS¹-isequiv i = isoToIsEquiv (iso (basedΩS¹→ΩS¹ i) (ΩS¹ -- now extend the basechange so that both ends match -- (and therefore we get a basechange for any x : S¹) -private - loop-conjugation : basedΩS¹→ΩS¹ i1 ≡ λ x → x - loop-conjugation i x = - let p = (doubleCompPath-elim loop x (sym loop)) - ∙ (λ i → (lUnit loop i ∙ x) ∙ sym loop) - in - ((sym (decodeEncode base (basedΩS¹→ΩS¹ i1 x))) - ∙ (λ t → intLoop (winding (p t))) - ∙ (λ t → intLoop (winding-hom (intLoop (pos (suc zero)) ∙ x) - (intLoop (negsuc zero)) t)) - ∙ (λ t → intLoop ((winding-hom (intLoop (pos (suc zero))) x t) - + (windingIntLoop (negsuc zero) t))) - ∙ (λ t → intLoop (((windingIntLoop (pos (suc zero)) t) + (winding x)) + (negsuc zero))) - ∙ (λ t → intLoop ((+-comm (pos (suc zero)) (winding x) t) + (negsuc zero))) - ∙ (λ t → intLoop (+-assoc (winding x) (pos (suc zero)) (negsuc zero) (~ t))) - ∙ (decodeEncode base x)) i - - refl-conjugation : basedΩS¹→ΩS¹ i0 ≡ λ x → x - refl-conjugation i x j = - hfill (λ t → λ { (j = i0) → base - ; (j = i1) → base }) - (inS (x j)) (~ i) - - basechange : (x : S¹) → basedΩS¹ x → ΩS¹ - basechange base y = y - basechange (loop i) y = - hcomp (λ t → λ { (i = i0) → refl-conjugation t y - ; (i = i1) → loop-conjugation t y }) - (basedΩS¹→ΩS¹ i y) - - -- for any loop i, the old basechange is equal to the new one - basedΩS¹→ΩS¹≡basechange : (i : I) → basedΩS¹→ΩS¹ i ≡ basechange (loop i) - basedΩS¹→ΩS¹≡basechange i j y = - hfill (λ t → λ { (i = i0) → refl-conjugation t y - ; (i = i1) → loop-conjugation t y }) - (inS (basedΩS¹→ΩS¹ i y)) j - - -- so for any loop i, the extended basechange is an equivalence - basechange-isequiv-aux : (i : I) → isEquiv (basechange (loop i)) - basechange-isequiv-aux i = - transport (λ j → isEquiv (basedΩS¹→ΩS¹≡basechange i j)) (basedΩS¹→ΩS¹-isequiv i) - - - -- as being an equivalence is contractible, basechange is an equivalence for all x : S¹ - basechange-isequiv : (x : S¹) → isEquiv (basechange x) - basechange-isequiv base = basechange-isequiv-aux i0 - basechange-isequiv (loop i) = - hcomp (λ t → λ { (i = i0) → basechange-isequiv-aux i0 - ; (i = i1) → isPropIsEquiv (basechange base) (basechange-isequiv-aux i1) - (basechange-isequiv-aux i0) t }) - (basechange-isequiv-aux i) - - basedΩS¹≡ΩS¹ : (x : S¹) → basedΩS¹ x ≡ ΩS¹ - basedΩS¹≡ΩS¹ x = ua (basechange x , basechange-isequiv x) +loop-conjugation : basedΩS¹→ΩS¹ i1 ≡ λ x → x +loop-conjugation i x = + let p = (doubleCompPath-elim loop x (sym loop)) + ∙ (λ i → (lUnit loop i ∙ x) ∙ sym loop) + in + ((sym (decodeEncode base (basedΩS¹→ΩS¹ i1 x))) + ∙ (λ t → intLoop (winding (p t))) + ∙ (λ t → intLoop (winding-hom (intLoop (pos (suc zero)) ∙ x) + (intLoop (negsuc zero)) t)) + ∙ (λ t → intLoop ((winding-hom (intLoop (pos (suc zero))) x t) + + (windingIntLoop (negsuc zero) t))) + ∙ (λ t → intLoop (((windingIntLoop (pos (suc zero)) t) + (winding x)) + (negsuc zero))) + ∙ (λ t → intLoop ((+-comm (pos (suc zero)) (winding x) t) + (negsuc zero))) + ∙ (λ t → intLoop (+-assoc (winding x) (pos (suc zero)) (negsuc zero) (~ t))) + ∙ (decodeEncode base x)) i + +refl-conjugation : basedΩS¹→ΩS¹ i0 ≡ λ x → x +refl-conjugation i x j = + hfill (λ t → λ { (j = i0) → base + ; (j = i1) → base }) + (inS (x j)) (~ i) + +basechange : (x : S¹) → basedΩS¹ x → ΩS¹ +basechange base y = y +basechange (loop i) y = + hcomp (λ t → λ { (i = i0) → refl-conjugation t y + ; (i = i1) → loop-conjugation t y }) + (basedΩS¹→ΩS¹ i y) + +-- for any loop i, the old basechange is equal to the new one +basedΩS¹→ΩS¹≡basechange : (i : I) → basedΩS¹→ΩS¹ i ≡ basechange (loop i) +basedΩS¹→ΩS¹≡basechange i j y = + hfill (λ t → λ { (i = i0) → refl-conjugation t y + ; (i = i1) → loop-conjugation t y }) + (inS (basedΩS¹→ΩS¹ i y)) j + +-- so for any loop i, the extended basechange is an equivalence +basechange-isequiv-aux : (i : I) → isEquiv (basechange (loop i)) +basechange-isequiv-aux i = + transport (λ j → isEquiv (basedΩS¹→ΩS¹≡basechange i j)) (basedΩS¹→ΩS¹-isequiv i) + + +-- as being an equivalence is contractible, basechange is an equivalence for all x : S¹ +basechange-isequiv : (x : S¹) → isEquiv (basechange x) +basechange-isequiv base = basechange-isequiv-aux i0 +basechange-isequiv (loop i) = + hcomp (λ t → λ { (i = i0) → basechange-isequiv-aux i0 + ; (i = i1) → isPropIsEquiv (basechange base) (basechange-isequiv-aux i1) + (basechange-isequiv-aux i0) t }) + (basechange-isequiv-aux i) + +basedΩS¹≡ΩS¹ : (x : S¹) → basedΩS¹ x ≡ ΩS¹ +basedΩS¹≡ΩS¹ x = ua (basechange x , basechange-isequiv x) basedΩS¹≡Int : (x : S¹) → basedΩS¹ x ≡ Int basedΩS¹≡Int x = (basedΩS¹≡ΩS¹ x) ∙ ΩS¹≡Int diff --git a/Cubical/HITs/Truncation/Properties.agda b/Cubical/HITs/Truncation/Properties.agda index be04e9173..64726af9e 100644 --- a/Cubical/HITs/Truncation/Properties.agda +++ b/Cubical/HITs/Truncation/Properties.agda @@ -111,6 +111,7 @@ isSnNull→isOfHLevel {n = suc n} nA = isSphereFilled→isOfHLevelSuc (λ f → isOfHLevelTrunc : (n : ℕ) → isOfHLevel n (hLevelTrunc n A) isOfHLevelTrunc zero = hub ⊥.rec , λ _ → ≡hub ⊥.rec isOfHLevelTrunc (suc n) = isSphereFilled→isOfHLevelSuc isSphereFilledTrunc + -- isOfHLevelTrunc n = isSnNull→isOfHLevel isNull-Null -- hLevelTrunc n is a modality @@ -118,13 +119,13 @@ isOfHLevelTrunc (suc n) = isSphereFilled→isOfHLevelSuc isSphereFilledTrunc -- This more direct definition should behave better than recElim -- below. Commented for now as it breaks some cohomology code if we -- use it instead of recElim. --- rec : {n : ℕ} --- {B : Type ℓ'} → --- isOfHLevel n B → --- (A → B) → --- hLevelTrunc n A → --- B --- rec h = Null.rec (isOfHLevel→isSnNull h) +rec : {n : ℕ} + {B : Type ℓ'} → + isOfHLevel n B → + (A → B) → + hLevelTrunc n A → + B +rec h = Null.rec (isOfHLevel→isSnNull h) -- TODO: remove this recElim : {n : ℕ} @@ -192,7 +193,7 @@ Iso.leftInv (univTrunc n {B , lev}) b = funExt (elim (λ x → isOfHLevelPath _ map : {n : ℕ} {B : Type ℓ'} (g : A → B) → hLevelTrunc n A → hLevelTrunc n B -map g = recElim (isOfHLevelTrunc _) (λ a → ∣ g a ∣) +map g = rec (isOfHLevelTrunc _) (λ a → ∣ g a ∣) mapCompIso : {n : ℕ} {B : Type ℓ'} → (Iso A B) → Iso (hLevelTrunc n A) (hLevelTrunc n B) Iso.fun (mapCompIso g) = recElim (isOfHLevelTrunc _) λ a → ∣ Iso.fun g a ∣ diff --git a/Cubical/Homotopy/Connected.agda b/Cubical/Homotopy/Connected.agda index efe5f0fee..257c9c181 100644 --- a/Cubical/Homotopy/Connected.agda +++ b/Cubical/Homotopy/Connected.agda @@ -17,7 +17,7 @@ open import Cubical.Data.Sigma hiding (_×_) open import Cubical.HITs.Nullification open import Cubical.HITs.Susp open import Cubical.HITs.SmashProduct -open import Cubical.HITs.Truncation as Trunc +open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec) open import Cubical.Homotopy.Loopspace open import Cubical.HITs.Pushout open import Cubical.HITs.Sn.Base @@ -76,7 +76,7 @@ module elim {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} (f : A → B) (n : where inv : ((a : A) → P (f a) .fst) → (b : B) → hLevelTrunc n (fiber f b) → P b .fst inv t b = - Trunc.recElim + trRec (P b .snd) (λ {(a , p) → subst (fst ∘ P) p (t a)}) @@ -231,9 +231,9 @@ connectedTruncIso {A = A} {B = B} (suc n) f con = g back y = map fst ((con y) .fst) backSection : (b : B) → Path (hLevelTrunc (suc n) B) - (Trunc.recElim (isOfHLevelTrunc (suc n)) + (trRec (isOfHLevelTrunc (suc n)) (λ a → ∣ f a ∣) - (Trunc.recElim {n = suc n } + (trRec {n = suc n } {B = hLevelTrunc (suc n) A} (isOfHLevelTrunc (suc n)) back ∣ b ∣)) ∣ b ∣ backSection b = helper (λ p → map f p ≡ ∣ b ∣) @@ -254,7 +254,7 @@ connectedTruncIso {A = A} {B = B} (suc n) f con = g g : Iso (hLevelTrunc (suc n) A) (hLevelTrunc (suc n) B) Iso.fun g = map f - Iso.inv g = Trunc.recElim (isOfHLevelTrunc _) back + Iso.inv g = trRec (isOfHLevelTrunc _) back Iso.leftInv g = Trunc.elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc _) _ _) λ a → cong (map fst) (con (f a) .snd ∣ a , refl ∣) Iso.rightInv g = Trunc.elim (λ x → isOfHLevelPath (suc n) (isOfHLevelTrunc _) _ _) diff --git a/Cubical/Homotopy/Freudenthal.agda b/Cubical/Homotopy/Freudenthal.agda index d5b1ec645..690d363fd 100644 --- a/Cubical/Homotopy/Freudenthal.agda +++ b/Cubical/Homotopy/Freudenthal.agda @@ -13,7 +13,7 @@ import Cubical.Data.NatMinusTwo as ℕ₋₂ open import Cubical.Data.Sigma open import Cubical.HITs.Nullification open import Cubical.HITs.Susp -open import Cubical.HITs.Truncation as Trunc +open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec) open import Cubical.Homotopy.Connected open import Cubical.Homotopy.WedgeConnectivity open import Cubical.Homotopy.Loopspace @@ -42,7 +42,7 @@ module _ {ℓ} (n : ℕ) {A : Pointed ℓ} (connA : isHLevelConnected (suc (suc fwd : (p : north ≡ north) (a : typ A) → hLevelTrunc 2n+2 (fiber σ p) → hLevelTrunc 2n+2 (fiber (λ x → merid x ∙ merid a ⁻¹) p) - fwd p a = Trunc.recElim (isOfHLevelTrunc 2n+2) (uncurry (WC.extension p a)) + fwd p a = trRec (isOfHLevelTrunc 2n+2) (uncurry (WC.extension p a)) isEquivFwd : (p : north ≡ north) (a : typ A) → isEquiv (fwd p a) isEquivFwd p a .equiv-proof = @@ -58,7 +58,7 @@ module _ {ℓ} (n : ℕ) {A : Pointed ℓ} (connA : isHLevelConnected (suc (suc (λ _ → isProp→isOfHLevelSuc (n + suc n) isPropIsContr) (λ fib → subst (λ k → isContr (fiber k ∣ fib ∣)) - (cong (Trunc.recElim (isOfHLevelTrunc 2n+2) ∘ uncurry) + (cong (trRec (isOfHLevelTrunc 2n+2) ∘ uncurry) (funExt (WC.right p) ⁻¹)) (subst isEquiv (funExt (Trunc.mapId) ⁻¹) diff --git a/Cubical/ZCohomology/KcompPrelims.agda b/Cubical/ZCohomology/KcompPrelims.agda index 1b56be826..f19a3468a 100644 --- a/Cubical/ZCohomology/KcompPrelims.agda +++ b/Cubical/ZCohomology/KcompPrelims.agda @@ -7,7 +7,7 @@ open import Cubical.HITs.Hopf open import Cubical.Homotopy.Freudenthal hiding (encode) open import Cubical.HITs.Sn open import Cubical.HITs.S1 -open import Cubical.HITs.Truncation renaming (elim to trElim ; recElim to trRec ; map to trMap) +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec) open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels @@ -171,7 +171,7 @@ private d-iso2 : Iso (hLevelTrunc 3 (typ (Ω (Susp S¹ , north)))) (hLevelTrunc 3 S¹) d-iso2 = connectedTruncIso _ d-map is1Connected-dmap - d-iso : isIso {A = ∥ typ (Ω (Susp S¹ , north)) ∥ (ℕ→ℕ₋₂ 1)} {B = ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (trElim (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣ ) + d-iso : isIso {A = ∥ typ (Ω (Susp S¹ , north)) ∥ (ℕ→ℕ₋₂ 1)} {B = ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} (trRec (isOfHLevelTrunc 3) λ x → ∣ d-map x ∣ ) d-iso = (Iso.inv (connectedTruncIso _ d-map is1Connected-dmap)) , (Iso.rightInv (connectedTruncIso _ d-map is1Connected-dmap) , Iso.leftInv (connectedTruncIso _ d-map is1Connected-dmap)) @@ -179,10 +179,8 @@ private d-mapId2 : (λ (x : hLevelTrunc 3 S¹) → (trRec {n = 3} {B = hLevelTrunc 3 S¹} (isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) (trRec (isOfHLevelTrunc 3) (λ a → ∣ ϕ base a ∣) x)) ≡ λ x → x - d-mapId2 = funExt (trElim (λ x → isOfHLevelSuc 2 (isOfHLevelTrunc 3 ((trElim {n = 3} - {B = λ _ → ∥ S¹ ∥ (ℕ→ℕ₋₂ 1)} - (λ x → isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) - (trElim (λ _ → isOfHLevelTrunc 3) + d-mapId2 = funExt (trElim (λ x → isOfHLevelSuc 2 (isOfHLevelTrunc 3 ((trRec (isOfHLevelTrunc 3) λ x → ∣ d-map x ∣) + (trRec (isOfHLevelTrunc 3) (λ a → ∣ ϕ base a ∣) x)) x)) λ a i → ∣ d-mapId a i ∣) diff --git a/Cubical/ZCohomology/MayerVietorisReduced.agda b/Cubical/ZCohomology/MayerVietorisReduced.agda index 5d357038c..def107b80 100644 --- a/Cubical/ZCohomology/MayerVietorisReduced.agda +++ b/Cubical/ZCohomology/MayerVietorisReduced.agda @@ -26,7 +26,7 @@ open import Cubical.HITs.Nullification open import Cubical.Data.Int hiding (_+_) open import Cubical.Data.Nat open import Cubical.Data.Prod -open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; recElim to trRec ; elim3 to trElim3) +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) open import Cubical.Homotopy.Loopspace open import Cubical.Homotopy.Connected open import Cubical.Homotopy.Freudenthal diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index 36e908601..220dd93fd 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -22,7 +22,7 @@ open import Cubical.HITs.Nullification open import Cubical.Data.Int hiding (_+_) open import Cubical.Data.Nat open import Cubical.Data.Prod -open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; recElim to trRec ; elim3 to trElim3) +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) open import Cubical.Homotopy.Loopspace open import Cubical.Homotopy.Connected open import Cubical.Homotopy.Freudenthal diff --git a/Cubical/ZCohomology/Sn/Sn.agda b/Cubical/ZCohomology/Sn/Sn.agda index e806b1f31..f2c40350a 100644 --- a/Cubical/ZCohomology/Sn/Sn.agda +++ b/Cubical/ZCohomology/Sn/Sn.agda @@ -27,14 +27,15 @@ open import Cubical.HITs.Nullification open import Cubical.Data.Int renaming (_+_ to _+ℤ_; +-comm to +ℤ-comm ; +-assoc to +ℤ-assoc) open import Cubical.Data.Nat open import Cubical.Data.Prod -open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; recElim to trRec ; elim3 to trElim3) +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; recElim to trRec ; rec to trRec2 ; elim3 to trElim3) open import Cubical.Homotopy.Loopspace open import Cubical.Homotopy.Connected open import Cubical.Homotopy.Freudenthal open import Cubical.HITs.SmashProduct.Base open import Cubical.Data.Unit -open import Cubical.Data.Group.Base renaming (Iso to grIso) +open import Cubical.Data.Group.Base renaming (Iso to grIso ; compIso to compGrIso) open import Cubical.Data.Group.GroupLibrary +open import Cubical.ZCohomology.coHomZero-map open import Cubical.ZCohomology.KcompPrelims @@ -45,6 +46,43 @@ open import Cubical.Data.Sum.Base open import Cubical.Data.HomotopyGroup open import Cubical.Data.Bool hiding (_⊕_) + +congDep : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {a : A} (p q r s : a ≡ a) (f : (a ≡ a) → B) → PathP (λ i → p i ≡ q i) r s + → PathP (λ i → {!B!}) (f r) (f s) -- PathP (λ i → f (p i) ≡ f (q i)) (cong f r) (cong f s) +congDep = {!!} + +moveDownPathP : {!!} +moveDownPathP = {!!} + +congFunct2 : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A} (f : A → A → B) (j₁ : x ≡ y) {u v : A} (q : u ≡ v) → cong₂ f j₁ q ≡ cong (λ x₁ → f x₁ u) j₁ ∙ cong (f y) q +congFunct2 = {!!} + +prodId : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A × B} → proj₁ x ≡ proj₁ y → proj₂ x ≡ proj₂ y → x ≡ y +prodId {x = (a , b)} {y = (c , d)} id1 id2 i = (id1 i) , (id2 i) + +indIntGroup : ∀ {ℓ} {G : Group ℓ} → (ϕ : Int → Group.type G) + → ϕ 0 ≡ isGroup.id (Group.groupStruc G) + → ((a : Int) → ϕ (a +ℤ 1) ≡ isGroup.comp (Group.groupStruc G) (ϕ a) (ϕ 1)) + → ((n : Int) → ϕ (predInt n) ≡ isGroup.comp (Group.groupStruc G) (ϕ n) (ϕ (negsuc zero))) + → isMorph intGroup G ϕ +indIntGroup {G = group G gSet (group-struct _ _ _+G_ _ rUnit₁ _ _ _)} ϕ zeroId _ _ n (pos zero) = + sym (rUnit₁ (ϕ n)) ∙ cong (λ x → ϕ n +G x) (sym zeroId) +indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId oneId minOneId n (pos (suc m)) = + (λ i → ϕ ((n +pos m) +ℤ 1)) + ∙ oneId (n +pos m) + ∙ cong (λ x → x +G ϕ (pos 1)) + (indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} + ϕ zeroId oneId minOneId n (pos m)) + ∙ assoc₁ (ϕ n) (ϕ (pos m)) (ϕ (pos 1)) + ∙ sym (cong (λ x → ϕ n +G x) (oneId (pos m))) +indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId _ minOneId n (negsuc zero) = minOneId n +indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId a minOneId n (negsuc (suc m)) = + (λ i → ϕ ((n +negsuc m) +ℤ (negsuc zero))) + ∙ minOneId (n +negsuc m) + ∙ cong (λ x → x +G ϕ (negsuc zero)) (indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId a minOneId n (negsuc m)) + ∙ assoc₁ (ϕ n) (ϕ (negsuc m)) (ϕ (negsuc zero)) + ∙ cong (λ x → ϕ n +G x) (sym (minOneId (negsuc m))) + pushoutSn : (n : ℕ) → Iso (S₊ (suc n)) (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt) Iso.fun (pushoutSn n) north = inl tt Iso.fun (pushoutSn n) south = inr tt @@ -70,78 +108,6 @@ isContr→isContrTrunc n contr = ∣ fst contr ∣ , (trElim (λ _ → isOfHLeve isContr→isContrSetTrunc : ∀ {ℓ} {A : Type ℓ} → isContr A → isContr (∥ A ∥₀) isContr→isContrSetTrunc contr = ∣ fst contr ∣₀ , sElim (λ _ → isOfHLevelPath 2 (setTruncIsSet) _ _) λ a → cong ∣_∣₀ (snd contr a) --- coHom Unit - - -lemma10000 : (n : ℕ) → looper (pos (suc n)) ≡ loop* ∙ looper (pos n) -lemma10000 zero = sym (lUnit loop*) ∙ rUnit loop* -lemma10000 (suc n) = (λ i → (lemma10000 n i) ∙ loop*) ∙ sym (assoc loop* (looper (pos n)) loop*) - -lemma10001 : (a : Int) → looper (sucInt a) ≡ loop* ∙ looper a -lemma10001 (pos n) = lemma10000 n -lemma10001 (negsuc zero) = sym (rCancel loop*) -lemma10001 (negsuc (suc zero)) = lUnit (sym loop*) - ∙ (λ i → rCancel loop* (~ i) ∙ (sym loop*)) - ∙ sym (assoc loop* (sym loop*) (sym loop*)) -lemma10001 (negsuc (suc (suc n))) = cong (λ x → x ∙ (sym loop*)) (lemma10001 (negsuc (suc n))) - ∙ sym (assoc loop* (looper (negsuc n) ∙ sym loop*) (sym loop*)) - -lemma10002 : (a : Int) → looper (predInt a) ≡ sym loop* ∙ looper a -lemma10002 (pos zero) = rUnit (sym loop*) -lemma10002 (pos (suc n)) = lUnit (looper (pos n)) - ∙ (λ i → lCancel loop* (~ i) ∙ (looper (pos n))) - ∙ sym (assoc (sym loop*) loop* (looper (pos n))) - ∙ λ i → sym loop* ∙ lemma10000 n (~ i) -lemma10002 (negsuc zero) = refl -lemma10002 (negsuc (suc n)) = (λ i → (lemma10002 (negsuc n) i) ∙ sym loop*) ∙ sym (assoc (sym loop*) (looper (negsuc n)) (sym loop*)) - -sucIntPredInt : (a : Int) → sucInt (predInt a) ≡ a -sucIntPredInt (pos zero) = refl -sucIntPredInt (pos (suc n)) = refl -sucIntPredInt (negsuc n) = refl - -looperId : (a b : Int) → looper a ∙ looper b ≡ looper (a +ℤ b) -looperId (pos zero) b = sym (lUnit (looper b)) ∙ λ i → looper (pos0+ b i) -looperId (pos (suc n)) (pos zero) = sym (rUnit (looper (pos (suc n)))) -looperId (pos (suc n)) (pos (suc m)) = (λ i → (lemma10000 n i) ∙ looper (pos (suc m))) - ∙ sym (assoc loop* (looper (pos n)) (looper (pos (suc m)))) - ∙ (λ i → loop* ∙ looperId (pos n) (pos (suc m)) i) - ∙ sym (lemma10001 (sucInt (pos n +pos m))) - ∙ λ i → looper (sucInt (sucInt+pos m (pos n) i)) -looperId (pos (suc n)) (negsuc zero) = sym (assoc (looper (pos n)) loop* (sym loop*)) ∙ (λ i → looper (pos n) ∙ (rCancel loop* i)) ∙ sym (rUnit (looper (pos n))) -looperId (pos (suc n)) (negsuc (suc m)) = (λ i → (looper (pos n) ∙ loop*) ∙ lemma10002 (negsuc m) i) - ∙ sym (assoc (looper (pos n)) loop* (sym loop* ∙ looper (negsuc m))) - ∙ (λ i → looper (pos n) ∙ assoc loop* (sym loop*) (looper (negsuc m)) i) - ∙ (λ i → looper (pos n) ∙ rCancel loop* i ∙ looper (negsuc m)) - ∙ (λ i → looper (pos n) ∙ lUnit (looper (negsuc m)) (~ i)) - ∙ looperId (pos n) (negsuc m) - ∙ λ i → looper (predInt+negsuc m (pos (suc n)) (~ i)) -looperId (negsuc zero) (pos zero) = sym (rUnit (sym loop*)) -looperId (negsuc zero) (pos (suc n)) = (λ i → (sym loop*) ∙ lemma10001 (pos n) i) - ∙ assoc (sym loop*) loop* (looper (pos n)) - ∙ (λ i → lCancel loop* i ∙ looper (pos n)) - ∙ sym (lUnit (looper (pos n))) - ∙ (λ i → looper (sucIntPredInt (pos n) (~ i))) - ∙ λ i → looper (sucInt (negsuc0+ (pos n) i)) -looperId (negsuc zero) (negsuc zero) = refl -looperId (negsuc zero) (negsuc (suc n)) = assoc (sym loop*) (looper (negsuc n)) (sym loop*) ∙ - (λ i → looperId (negsuc zero) (negsuc n) i ∙ sym loop*) ∙ - (λ i → looper (negsuc zero +negsuc n) ∙ sym loop*) ∙ - (λ i → looper (negsuc0+ (negsuc n) (~ i)) ∙ sym loop*) ∙ - λ i → looper (predInt (negsuc0+ (negsuc n) i)) -looperId (negsuc (suc n)) b = (λ i → lemma10002 (negsuc n) i ∙ looper b) - ∙ sym (assoc (sym loop*) (looper (negsuc n)) (looper b)) - ∙ (λ i → sym loop* ∙ looperId (negsuc n) b i) - ∙ sym (lemma10002 (negsuc n +ℤ b)) - ∙ λ i → looper (predInt+ (negsuc n) b i) - -addLemma : (a b : Int) → a +ₖ b ≡ (a +ℤ b) -addLemma a b = (λ i → ΩKn+1→Kn (cong ∣_∣ (looper a) ∙ cong ∣_∣ (looper b))) - ∙ (λ i → ΩKn+1→Kn (congFunct ∣_∣ (looper a) (looper b) (~ i))) - ∙ (λ i → ΩKn+1→Kn (cong ∣_∣ (looperId a b i))) - ∙ Iso.leftInv (Iso3-Kn-ΩKn+1 zero) (a +ℤ b) - - coHomGrUnit0 : grIso (coHomGr 0 Unit) intGroup grIso.fun coHomGrUnit0 = sRec isSetInt (λ f → f tt) , sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) λ a b → addLemma (a tt) (b tt) @@ -166,24 +132,691 @@ grIso.leftInv (coHomGrUnit≥1 n) a = sym (helper2 .snd 0ₕ) ∙ helper2 .snd a helper2 : isContr (coHom (suc n) Unit) helper2 = subst isContr (λ i → ∥ UnitToTypeId (coHomK (suc n)) (~ i) ∥₀) (helper n) +S0→Int : (a : Int × Int) → S₊ 0 → Int +S0→Int a north = proj₁ a +S0→Int a south = proj₂ a + coHom0-S0 : grIso (coHomGr 0 (S₊ 0)) (dirProd intGroup intGroup) coHom0-S0 = Iso'→Iso (iso' (iso (sRec (isOfHLevelProd 2 isSetInt isSetInt) λ f → (f north) , (f south)) - (λ {(a , b) → ∣ (λ {north → a ; south → b}) ∣₀}) + (λ a → ∣ S0→Int a ∣₀) (λ { (a , b) → refl}) (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ f → cong ∣_∣₀ (funExt (λ {north → refl ; south → refl})))) (sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 isSetInt isSetInt) _ _) λ a b i → addLemma (a north) (b north) i , addLemma (a south) (b south) i)) -coHomGr0-S1 : grIso (coHomGr 1 (S₊ 1)) intGroup -coHomGr0-S1 = - Iso'→Iso - (iso' - (iso (sRec isSetInt {!!}) - {!!} - {!!} - {!!}) - {!!}) +×morph : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group ℓ} {B : Group ℓ'} {C : Group ℓ''} {D : Group ℓ'''} → morph A B → morph C D → morph (dirProd A C) (dirProd B D) +×morph mf1 mf2 = (λ {(a , b) → (mf1 .fst a) , mf2 .fst b}) , λ {(a , b) (c , d) i → mf1 .snd a c i , mf2 .snd b d i} + + +helper1 : Iso (S₊ 1) S¹ +Iso.fun helper1 north = base +Iso.fun helper1 south = base +Iso.fun helper1 (merid a i) = loop i +Iso.inv helper1 base = north +Iso.inv helper1 (loop i) = loop* i +Iso.rightInv helper1 base = refl +Iso.rightInv helper1 (loop i) j = {!helper loop!} +Iso.leftInv helper1 north = refl +Iso.leftInv helper1 south = merid north +Iso.leftInv helper1 (merid north i) j = {!compPath-filler (merid south) (sym (merid north))!} +Iso.leftInv helper1 (merid south i) j = {!!} + +isGroupoidS1 : isGroupoid (S₊ 1) +isGroupoidS1 = transport (λ i → isGroupoid (S¹≡S1 (~ i))) isGroupoidS¹ +isConnectedS1 : (x : S₊ 1) → ∥ north ≡ x ∥₋₁ +isConnectedS1 x = pRec propTruncIsProp + (λ p → ∣ cong (transport (sym (S¹≡S1))) p ∙ transport⁻Transport (S¹≡S1) x ∣₋₁) + (isConnectedS¹ (transport S¹≡S1 x)) +isConnectedS1Loop : (x : S₊ 1) (p : x ≡ x) → ∥ refl ≡ p ∥₋₁ +isConnectedS1Loop x p = {!∥ ? ∥!} + +-- basechange* : (x y : S¹) → x ≡ y → x ≡ y → ΩS¹ +-- basechange* x y = J (λ y p → (x ≡ y) → ΩS¹) (basechange x) + + +-- test1 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S₊ 1 × Int) +-- Iso.fun test1 f = (trRec isGroupoidS1 (λ a → a) (f north)) +-- , winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (f (loop* i))))) +-- Iso.inv test1 (north , b) x = ∣ x ∣ +-- Iso.inv test1 (south , b) x = ∣ x ∣ +-- Iso.inv test1 (merid a i , b) x = {!!} +-- Iso.rightInv test1 = {!!} +-- Iso.leftInv test1 = {!!} + +-- funRec : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) {C : (A → hLevelTrunc n B) → Type ℓ''} +-- → isOfHLevel n B +-- → ((f : A → B) → C (λ a → ∣ f a ∣)) +-- → (f : A → hLevelTrunc n B) → C f +-- funRec {A = A} {B = B} n {C = C} hLev ind f = subst C (helper f) (ind (λ a → trRec hLev (λ x → x) (f a))) +-- where +-- helper : retract {A = A → hLevelTrunc n B} {B = A → B} (λ f₁ a → trRec hLev (λ x → x) (f₁ a)) (λ f₁ a → ∣ f₁ a ∣) +-- helper f = funExt λ a → helper2 (f a) +-- where +-- helper2 : (x : hLevelTrunc n B) → ∣ trRec hLev (λ x → x) x ∣ ≡ x +-- helper2 = trElim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → refl + +-- invMapSurj : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (ϕ : morph G H) → ((x : Group.type H) → isInIm G H (fst ϕ) x) +-- → morph H G +-- fst (invMapSurj G H (ϕ , pf) surj) a = {!pRec!} +-- snd (invMapSurj G H (ϕ , pf) surj) = {!!} + +-- ImIso : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (ϕ : morph G H) → ((x : Group.type H) → isInIm G H (fst ϕ) x) +-- → grIso H (imGroup G H ϕ) +-- ImIso G H (ϕ , mf) surj = +-- let idH = isGroup.id (Group.groupStruc H) +-- idG = isGroup.id (Group.groupStruc G) +-- _+G_ = isGroup.comp (Group.groupStruc G) +-- _+H_ = isGroup.comp (Group.groupStruc H) +-- _+Im_ = isGroup.comp (Group.groupStruc (imGroup G H (ϕ , mf))) +-- invG = isGroup.inv (Group.groupStruc G) +-- invH = isGroup.inv (Group.groupStruc H) +-- lUnit = isGroup.lUnit (Group.groupStruc H) +-- lCancel = isGroup.rCancel (Group.groupStruc H) +-- in +-- Iso''→Iso _ _ +-- (iso'' ((λ x → x , pRec propTruncIsProp (λ (a , b) → ∣ a , b ∣₋₁) (surj x)) +-- , λ a b → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) +-- (λ surja → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) +-- (λ surjb → +-- pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) +-- (λ surja+b → +-- (λ i → (a +H b) , (pRec (propTruncIsProp) +-- (λ (a , b) → ∣ a , b ∣₋₁) +-- (propTruncIsProp (surj (isGroup.comp (Group.groupStruc H) a b)) ∣ surja+b ∣₋₁ i))) ∙ +-- (λ i → (a +H b) , ∣ (fst surja+b) , (snd surja+b) ∣₋₁) ∙ +-- ΣProp≡ (λ _ → propTruncIsProp) refl ∙ +-- λ i → (a +H b) , pRec (propTruncIsProp) +-- (λ p1 → +-- pRec (λ x y → squash x y) +-- (λ p2 → +-- ∣ +-- isGroup.comp (Group.groupStruc G) (fst p1) (fst p2) , +-- mf (fst p1) (fst p2) ∙ +-- cong₂ (isGroup.comp (Group.groupStruc H)) (snd p1) (snd p2) +-- ∣₋₁) +-- (pRec (propTruncIsProp) +-- ∣_∣₋₁ (propTruncIsProp ∣ surjb ∣₋₁ (surj b) i))) +-- (pRec (propTruncIsProp) +-- ∣_∣₋₁ (propTruncIsProp ∣ surja ∣₋₁ (surj a) i ))) +-- (surj (isGroup.comp (Group.groupStruc H) a b))) +-- (surj b)) +-- (surj a)) +-- (λ x inker → cong fst inker) +-- {- +-- pElim (λ _ → Group.setStruc H _ _) +-- (λ xinim → + +-- (surj x)) -} +-- λ x → pRec propTruncIsProp (λ inimx → ∣ (ϕ (inimx .fst)) , ΣProp≡ (λ _ → propTruncIsProp) (inimx .snd) ∣₋₁) (snd x)) + + +coHom1Iso : grIso intGroup + ((imGroup (coHomGr 0 (S₊ 0)) + (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) + (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 + , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0))) +coHom1Iso = + Iso''→Iso _ _ + (iso'' ((λ x → ∣ theFuns x ∣₀ , d-surj ∣ theFuns x ∣₀) + , (λ a b → ΣProp≡ (λ _ → propTruncIsProp) + (cong ∣_∣₀ (funExt λ x → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 1) _) + ∙ cong (ΩKn+1→Kn) (sym (theFunsId2 a b x)))))) + (λ x inker → pRec (isSetInt _ _) (inj x) (Iso.fun PathIdTrunc₀Iso (cong fst inker))) + finIm) + where + ImGr : _ + ImGr = (imGroup (coHomGr 0 (S₊ 0)) + (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) + (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 + , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) + + d : _ + d = MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 + + Δ : _ + Δ = MV.Δ _ _ (S₊ 1) (λ _ → tt) (λ _ → tt) 0 + + + d-surj : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) + → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x + d-surj x = {!!} + + theFuns : (a : Int) → Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 + theFuns a (inl x) = 0ₖ + theFuns a (inr x) = 0ₖ + theFuns a (push north i) = Kn→ΩKn+1 0 a i + theFuns a (push south i) = 0ₖ + + + inj : (a : Int) → theFuns a ≡ (λ _ → ∣ north ∣) → a ≡ pos 0 + inj a id = + pRec (isSetInt _ _) + (sigmaElim (λ _ → isOfHLevelPath 2 isSetInt _ _) + (λ a p → pRec (isSetInt _ _) + (λ id2 → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 0) _) + ∙ cong (ΩKn+1→Kn) + (PathP→compPathR + (cong (λ f → cong f (push north)) id) + ∙ test)) + (Iso.fun PathIdTrunc₀Iso p))) (d-surj ∣ theFuns a ∣₀) + where + + test : (λ i → id i (inl tt)) ∙ (λ i → ∣ north ∣) ∙ sym (λ i → id i (inr tt)) ≡ refl + test = (λ i → cong (λ f → f (inl tt)) id + ∙ lUnit (sym (cong (λ f → f (inr tt)) id)) (~ i)) + ∙ (λ i → cong (λ f → f (push south i)) id + ∙ sym (cong (λ f → f (inr tt)) id)) + ∙ rCancel (cong (λ f → f (inr tt)) id) + + isInImTheFuns : (x : Int) → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout (λ _ → tt) (λ _ → tt))) d ∣ theFuns x ∣₀ + isInImTheFuns x = {!!} + + consMember : (a : Int) → Group.type ImGr + consMember a = d ∣ (λ _ → a) ∣₀ , ∣ ∣ (λ _ → a) ∣₀ , refl ∣₋₁ + + consMember≡0 : (a : Int) → consMember a ≡ isGroup.id (Group.groupStruc ImGr) + consMember≡0 a = + ΣProp≡ (λ _ → propTruncIsProp) + (MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ (λ _ → a) ∣₀ ∣ + (∣ (λ _ → a) ∣₀ , ∣ (λ _ → 0) ∣₀) + , cong ∣_∣₀ (λ i x → (rUnitₖ a i)) ∣₋₁) + + f+consMember : (f : Group.type ImGr) → ∃[ x ∈ Int ] (isGroup.comp (Group.groupStruc ImGr) f (isGroup.inv (Group.groupStruc ImGr) (consMember x)) ≡ (∣ theFuns x ∣₀ , d-surj ∣ theFuns x ∣₀)) + f+consMember f = + let null = isGroup.id (Group.groupStruc ImGr) + - = isGroup.inv (Group.groupStruc ImGr) + _+_ = isGroup.comp (Group.groupStruc ImGr) + in + sigmaElim {C = λ (f : Group.type ImGr) + → ∃[ x ∈ Int ] (isGroup.comp (Group.groupStruc ImGr) f (isGroup.inv (Group.groupStruc ImGr) (consMember x)) + ≡ (∣ theFuns x ∣₀ , d-surj ∣ theFuns x ∣₀)) } + (λ _ → isOfHLevelSuc 1 propTruncIsProp) + (λ f idf → pRec propTruncIsProp + (sigmaElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ g id → ∣ (g south) , (ΣProp≡ (λ _ → propTruncIsProp) + (pRec (setTruncIsSet _ _) + (λ id → {! (λ i → ∣ id i ∣₀ +ₕ -ₕ (consMember (g south) .fst)) !} + ∙ {!!} + ∙ {!!}) + (Iso.fun PathIdTrunc₀Iso id))) ∣₋₁) + idf) + f + where + subtrLemma : (a b : Int) → a - b ≡ a +ₖ (-ₖ b) + subtrLemma = {!!} + + funId : (g : S₊ 0 → Int) + → (x : Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) + → MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g x +ₖ (-ₖ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) x) ≡ + theFuns (g north - g south) x + funId g (inl x) = cong (λ x → (0ₖ {n = 1} ) +ₖ x) (-0ₖ {n = 1}) ∙ rUnitₖ (0ₖ {n = 1}) + funId g (inr x) = cong (λ x → (0ₖ {n = 1} ) +ₖ x) (-0ₖ {n = 1}) ∙ rUnitₖ (0ₖ {n = 1}) + funId g (push north i) j = + {!!} + funId g (push south i) j = {!!} + + finIm : (f : Group.type ImGr) → ∃[ x ∈ Int ] ((∣ theFuns x ∣₀ , d-surj ∣ theFuns x ∣₀) ≡ f) + finIm f = + let null = isGroup.id (Group.groupStruc ImGr) + - = isGroup.inv (Group.groupStruc ImGr) + _+_ = isGroup.comp (Group.groupStruc ImGr) + rUnit = isGroup.rUnit (Group.groupStruc ImGr) + lCancel = isGroup.lCancel (Group.groupStruc ImGr) + in pRec propTruncIsProp (λ {(a , id) → ∣ a , (sym (sym (rUnit f) ∙ cong (λ x → f + x) (sym (lCancel null)) ∙ cong (λ x → f + x) (rUnit (- null)) ∙ cong (λ x → f + (- x)) (sym (consMember≡0 a)) ∙ id)) ∣₋₁}) (f+consMember f) -- (f+consMember f) + + + surjective : {!f + consMember f+consMember!} + surjective = {!!} + + + theFunsId3 : (a b : Int) → ∣ theFuns a ∣₀ ≡ ∣ theFuns b ∣₀ → a ≡ b + theFunsId3 = {!!} + + theFunsId2 : (a b : Int) (x : Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) + → Kn→ΩKn+1 1 (theFuns a x) ∙ Kn→ΩKn+1 1 (theFuns b x) ≡ Kn→ΩKn+1 1 (theFuns (a +ℤ b) x) + theFunsId2 a b (inl x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ sym (lUnit _) + theFunsId2 a b (inr x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ sym (lUnit _) + theFunsId2 a b (push north i) j = + hcomp (λ k → λ {(i = i0) → ((λ i₁ → + (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) + j + ; (i = i1) → ((λ i₁ → + (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) + j + ; (j = i0) → congFunct2 (λ p q → Kn→ΩKn+1 1 p ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b) (~ k) i + ; (j = i1) → (helper2 a b) k i }) + (hcomp (λ k → λ { (i = i1) → {!!} + ; (i = i0) → ((λ i₁ → + (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) + j}) + {!absRetr ? ?!}) + + + +{- +i = i0 ⊢ ((λ i₁ → + (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) + j +i = i1 ⊢ ((λ i₁ → + (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) + j +j = i0 ⊢ cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a) ∙ + cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) +j = i1 ⊢ (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) i +-} + + where + + + + + helper3 : sym ((λ i₁ → (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) ∙ cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a) ∙ cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) ∙ (λ i₁ → + (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁)) ≡ + cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b) + helper3 = {!λ i → !} + ∙ {!!} + ∙ {!!} + + helper2 : (a b : Int) → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b) ≡ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ℤ b)) + helper2 a b = + sym (congFunct (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b)) + ∙ (λ i → cong (Kn→ΩKn+1 1) (Iso.rightInv (Iso3-Kn-ΩKn+1 0) (Kn→ΩKn+1 0 a ∙ Kn→ΩKn+1 0 b) (~ i))) + ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ₖ b)) ) + ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (addLemma a b i))) + + theFunsId2 a b (push south i) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ sym (lUnit _) + + + theFunsId : (a b : Int) → ∣ theFuns (a +ℤ b) ∣₀ ≡ (∣ theFuns a ∣₀ +ₕ ∣ theFuns b ∣₀) + theFunsId a b = {!!} + ∙ {!!} + ∙ (λ i → ∣ (λ x → ΩKn+1→Kn {!congFunct ∣_∣ ? ? i!}) ∣₀) + ∙ (λ i → ∣ (λ x → ΩKn+1→Kn (Kn→ΩKn+1 1 (theFuns a x) ∙ Kn→ΩKn+1 1 (theFuns b x))) ∣₀) + ∙ λ i → ∣ (λ x → (theFuns a x) +ₖ theFuns b x) ∣₀ + + +-- coHomGrIm : grIso (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- (imGroup (coHomGr 0 (S₊ 0)) +-- (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 +-- , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) +-- coHomGrIm = +-- ImIso _ +-- _ +-- _ +-- {!!} + + +-- -- coHom1RedS1 : Iso (coHom 1 (S₊ 1)) (coHomRed 1 (S₊ 1 , north)) +-- -- Iso.fun coHom1RedS1 = sRec setTruncIsSet λ f → ∣ f , (pRec {!!} {!!} ((transport (λ i → (b : truncIdempotent 3 {!S₊ 1!} (~ i)) → ∥ (transp (λ j → truncIdempotent {!3!} {!!} (~ i ∨ (~ j))) (~ i) north) ≡ b ∥₋₁) isConnectedS1) (f north)) ) ∣₀ +-- -- Iso.inv coHom1RedS1 = {!!} +-- -- Iso.rightInv coHom1RedS1 = {!setTruncIdempotent!} +-- -- Iso.leftInv coHom1RedS1 = {!!} + +-- -- coHom1Red : ∀ {ℓ} (A : Pointed ℓ) → Iso (coHom 1 (typ A)) (coHomRed 1 A) +-- -- Iso.fun (coHom1Red A) = sRec setTruncIsSet λ f → ∣ f , {!!} ∣₀ +-- -- Iso.inv (coHom1Red A) = {!!} +-- -- Iso.rightInv (coHom1Red A) = {!!} +-- -- Iso.leftInv (coHom1Red A) = {!!} + +-- -- -- morphtest : morph (coHomGr 1 (S₊ 1)) intGroup +-- -- -- fst morphtest = sRec isSetInt λ f → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (f (loop* i))))) +-- -- -- snd morphtest = sElim2 {!!} +-- -- -- (funRec 3 isGroupoidS1 +-- -- -- λ f → funRec 3 isGroupoidS1 +-- -- -- λ g → pRec (isSetInt _ _) +-- -- -- (λ n=fn → +-- -- -- pRec (isSetInt _ _) +-- -- -- (λ n=gn → (λ j → winding (basechange (SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (∣ f (north) ∣ +ₖ ∣ n=gn (~ j) ∣)))) (λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (∣ f (loop* i) ∣ +ₖ ∣ transp (λ i → n=gn ((~ i) ∨ (~ j)) ≡ n=gn ((~ i) ∨ (~ j))) (~ j) (λ i → g (loop* i)) i ∣)))))) +-- -- -- ∙ {!!} +-- -- -- ∙ {!!}) +-- -- -- (isConnectedS1 (g north))) +-- -- -- (isConnectedS1 (f north))) + + +-- -- -- {- (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (∣ f (loop* i) ∣ +ₖ ∣ g (loop* i) ∣))))) +-- -- -- ∙ (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (Kn→ΩKn+1 1 ∣ f (loop* i) ∣ ∙ Kn→ΩKn+1 1 ∣ g (loop* i) ∣)))))) +-- -- -- ∙ (λ j → winding (basechange _ (cong (λ x → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (Kn→ΩKn+1 1 ∣ f x ∣ ∙ Kn→ΩKn+1 1 ∣ g x ∣))))) loop*)) ) +-- -- -- ∙ (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn ((cong ∣_∣ (merid (f (loop* i)) ∙ sym (merid north)) ∙ cong ∣_∣ (merid (g (loop* i)) ∙ sym (merid north))))))))) +-- -- -- ∙ (λ j → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (congFunct ∣_∣ (merid (f (loop* i)) ∙ sym (merid north)) (merid (g (loop* i)) ∙ sym (merid north)) (~ j))))))) +-- -- -- ∙ (λ j → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (cong ∣_∣ (({!!} ∙ {!!}) ∙ {!!}))))))) +-- -- -- ∙ {!!} +-- -- -- ∙ {!!} +-- -- -- ∙ {!!}) -} + +-- -- -- where +-- -- -- helper : ∀ {ℓ} {A : Type ℓ} (a : A) (f : A → S¹) (p q : a ≡ a) → winding (basechange (f a) (cong f (p ∙ q))) ≡ winding (basechange (f a) (cong f p ∙ cong f q)) +-- -- -- helper a f p q i = winding (basechange (f a) (congFunct f p q i)) +-- -- -- helper2 : (x : S¹) (p q : x ≡ x) → basechange x (p ∙ q) ≡ basechange x p ∙ basechange x q +-- -- -- helper2 base p q = refl +-- -- -- helper2 (loop i) p q = {!!} +-- -- -- helper4 : (x y z : S¹) (p : x ≡ z) (q r : x ≡ y) (s : y ≡ z) → basechange* x z p (q ∙ s) ≡ basechange* x y {!!} q ∙ {!!} +-- -- -- helper4 = {!!} +-- -- -- helper3 : (p q : ΩS¹) → winding (p ∙ q) ≡ (winding p +ℤ winding q) +-- -- -- helper3 = {!!} + + +-- -- -- -- fstmap : morph (dirProd intGroup intGroup) (coHomGr 0 (S₊ 0)) +-- -- -- -- fstmap = compMorph {F = dirProd intGroup intGroup} {G = ×coHomGr 0 Unit Unit} {H = coHomGr 0 (S₊ 0)} +-- -- -- -- (×morph (grIso.inv coHomGrUnit0) (grIso.inv coHomGrUnit0)) +-- -- -- -- (((MV.Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) zero)) , +-- -- -- -- {!MV.ΔIsHom _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) zero!}) + +-- -- -- -- fstMapId : (a : Int × Int) → fstmap .fst a ≡ ∣ (λ _ → proj₁ a +ℤ (0 - proj₂ a)) ∣₀ +-- -- -- -- fstMapId (a , b) = (λ i → ∣ (λ _ → a +ₖ (-ₖ b)) ∣₀) ∙ {!addLemma!} ∙ {!!} ∙ {!!} + +-- -- -- -- isoAgain : grIso intGroup (coHomGr 1 (S₊ 1)) +-- -- -- -- isoAgain = +-- -- -- -- Iso''→Iso _ _ +-- -- -- -- (iso'' ((λ a → ∣ (λ {north → 0ₖ ; south → 0ₖ ; (merid north i) → {!a!} ; (merid south i) → {!!}}) ∣₀) , {!!}) +-- -- -- -- {!!} +-- -- -- -- {!!}) + +-- -- -- -- -- -- test2 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → S¹) +-- -- -- -- -- -- Iso.fun test2 f = {!!} +-- -- -- -- -- -- Iso.inv test2 f north = ∣ transport (sym S¹≡S1) (f base) ∣ +-- -- -- -- -- -- Iso.inv test2 f south = ∣ transport (sym S¹≡S1) (f base) ∣ +-- -- -- -- -- -- Iso.inv test2 f (merid a i) = cong ∣_∣ {!transport (sym S¹≡S1) (f base)!} i +-- -- -- -- -- -- Iso.rightInv test2 = {!!} +-- -- -- -- -- -- Iso.leftInv test2 = {!!} + +-- -- -- -- -- -- F : winding (basechange base loop) ≡ 1 +-- -- -- -- -- -- F = refl + +-- -- -- -- -- -- another : (f g : Int) → winding (basechange {!!} {!!}) ≡ {!!} +-- -- -- -- -- -- another = {!!} + +-- -- -- -- -- -- test : Iso (S¹ → S¹) (S¹ × Int) +-- -- -- -- -- -- Iso.fun test f = f base , winding (basechange (f base) (cong f loop)) +-- -- -- -- -- -- Iso.inv test (x , int) base = x +-- -- -- -- -- -- Iso.inv test (x , int) (loop i) = {!!} +-- -- -- -- -- -- Iso.rightInv test = {!!} +-- -- -- -- -- -- Iso.leftInv test = {!!} + +-- -- -- -- -- -- -- test13 : Iso ∥ (S¹ → S¹) ∥₀ Int +-- -- -- -- -- -- -- Iso.fun test13 = sRec isSetInt λ f → winding (basechange (f base) (λ i → f (loop i))) +-- -- -- -- -- -- -- Iso.inv test13 a = ∣ (λ {base → {!!} ; (loop i) → {!!}}) ∣₀ +-- -- -- -- -- -- -- Iso.rightInv test13 = {!!} +-- -- -- -- -- -- -- Iso.leftInv test13 = {!!} + +-- -- -- -- -- -- -- test : Iso (S¹ → S¹) (S¹ × Int) +-- -- -- -- -- -- -- Iso.fun test f = (f base) , transport (basedΩS¹≡Int (f base)) λ i → f (loop i) +-- -- -- -- -- -- -- Iso.inv test (x , int) base = x +-- -- -- -- -- -- -- Iso.inv test (x , int) (loop i) = transport (sym (basedΩS¹≡Int x)) int i +-- -- -- -- -- -- -- Iso.rightInv test (x , int) i = (x , transportTransport⁻ (basedΩS¹≡Int x) int i) +-- -- -- -- -- -- -- Iso.leftInv test f = +-- -- -- -- -- -- -- funExt λ { base → refl +-- -- -- -- -- -- -- ; (loop i) j → transport⁻Transport (basedΩS¹≡Int (f base)) (λ i → f (loop i)) j i} + + +-- -- -- -- -- -- -- lem : S¹ ≡ hLevelTrunc 3 (S₊ 1) +-- -- -- -- -- -- -- lem = sym (truncIdempotent 3 isGroupoidS¹) ∙ λ i → hLevelTrunc 3 (S¹≡S1 (~ i)) + +-- -- -- -- -- -- -- prodId : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (a b : A × B) → proj₁ a ≡ proj₁ b → proj₂ a ≡ proj₂ b → a ≡ b +-- -- -- -- -- -- -- prodId (_ , _) (_ , _) id1 id2 i = (id1 i) , (id2 i) + +-- -- -- -- -- -- -- test22 : Iso (S₊ 1 → coHomK 1) (S₊ 1 × Int) +-- -- -- -- -- -- -- Iso.fun test22 f = {!f north!} , {!!} +-- -- -- -- -- -- -- Iso.inv test22 = {!!} +-- -- -- -- -- -- -- Iso.rightInv test22 = {!!} +-- -- -- -- -- -- -- Iso.leftInv test22 = {!!} + + + + + + +-- -- -- -- -- -- -- coHom1≃∥S1×ℤ∥₀ : Iso (coHom 1 (S₊ 1)) ∥ S₊ 1 × Int ∥₀ +-- -- -- -- -- -- -- coHom1≃∥S1×ℤ∥₀ = setTruncIso test2 +-- -- -- -- -- -- -- where +-- -- -- -- -- -- -- test2 : Iso (S₊ 1 → coHomK 1) (S₊ 1 × Int) +-- -- -- -- -- -- -- Iso.fun test2 f = transport (λ i → S¹≡S1 (~ i) × Int) (Iso.fun test (transport (λ i → (S¹≡S1 i → lem (~ i))) f)) +-- -- -- -- -- -- -- Iso.inv test2 x = transport (λ i → (S¹≡S1 (~ i) → lem i)) (Iso.inv test (transport (λ i → S¹≡S1 i × Int) x)) +-- -- -- -- -- -- -- Iso.rightInv test2 (s , int) = prodId _ _ {!!} {!!} +-- -- -- -- -- -- -- Iso.leftInv test2 f = {!!} ∙ {!!} ∙ {!!} + +-- -- -- -- -- -- -- test2Id : (a b : (S₊ 1 → coHomK 1)) → proj₂ (Iso.fun test2 (λ x → a x +ₖ b x)) ≡ (proj₂ (Iso.fun test2 a) +ₖ proj₂ (Iso.fun test2 a)) +-- -- -- -- -- -- -- test2Id a b = {! +-- -- -- -- -- -- -- transport +-- -- -- -- -- -- -- (basedΩS¹≡Int +-- -- -- -- -- -- -- (transport (λ i → S¹≡S1 i → lem (~ i)) (λ x → a x +ₖ b x) base)) +-- -- -- -- -- -- -- (λ i → +-- -- -- -- -- -- -- transport (λ i₁ → S¹≡S1 i₁ → lem (~ i₁)) (λ x → a x +ₖ b x) +-- -- -- -- -- -- -- (loop i))!} ∙ {!transport (λ i → S¹≡S1 i → lem (~ i)) (λ x → a x +ₖ b x) base!} + + +-- -- -- -- -- -- -- main : grIso intGroup (coHomGr 1 (S₊ 1)) +-- -- -- -- -- -- -- main = Iso'→Iso +-- -- -- -- -- -- -- (iso' {!!} +-- -- -- -- -- -- -- {!!}) + + +-- -- -- -- -- coHom1 : grIso (coHomGr 1 (S₊ 1)) intGroup +-- -- -- -- -- coHom1 = +-- -- -- -- -- Iso'→Iso +-- -- -- -- -- (iso' (iso ({!!} ∘ {!!} ∘ {!!} ∘ {!!}) +-- -- -- -- -- {!!} +-- -- -- -- -- {!!} +-- -- -- -- -- {!!}) +-- -- -- -- -- {!!}) + + + +-- -- -- -- -- schonf : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : (A × B) → Type ℓ''} +-- -- -- -- -- → ((a : A) (b : B) → C (a , b)) +-- -- -- -- -- → (x : A × B) → C x +-- -- -- -- -- schonf f (a , b) = f a b + +-- -- -- -- -- -- -- setTruncProdIso : ∀ {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') → Iso ∥ (A × B) ∥₀ (∥ A ∥₀ × ∥ B ∥₀) +-- -- -- -- -- -- -- Iso.fun (setTruncProdIso A B) = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) λ {(a , b) → ∣ a ∣₀ , ∣ b ∣₀} +-- -- -- -- -- -- -- Iso.inv (setTruncProdIso A B) (a , b) = sRec setTruncIsSet (λ a → sRec setTruncIsSet (λ b → ∣ a , b ∣₀) b) a +-- -- -- -- -- -- -- Iso.rightInv (setTruncProdIso A B) = +-- -- -- -- -- -- -- schonf (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- -- -- -- -- -- -- λ _ → sElim (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- -- -- -- -- -- -- λ _ → refl) +-- -- -- -- -- -- -- Iso.leftInv (setTruncProdIso A B) = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ {(a , b) → refl} + +-- -- -- -- -- -- -- setTruncProdLemma : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (a1 a2 : A) (b : B) → isHLevelConnected 2 A +-- -- -- -- -- -- -- → Path ∥ A × B ∥₀ ∣ a1 , b ∣₀ ∣ a2 , b ∣₀ +-- -- -- -- -- -- -- setTruncProdLemma {A = A} {B = B} a1 a2 b conA i = Iso.inv (setTruncProdIso A B) (Iso.inv setTruncTrunc0Iso ((sym (conA .snd ∣ a1 ∣) ∙ (conA .snd ∣ a2 ∣)) i) , ∣ b ∣₀) + +-- -- -- -- -- -- -- test3 : Iso ∥ S₊ 1 × Int ∥₀ Int +-- -- -- -- -- -- -- Iso.fun test3 = sRec isSetInt proj₂ +-- -- -- -- -- -- -- Iso.inv test3 a = ∣ north , a ∣₀ +-- -- -- -- -- -- -- Iso.rightInv test3 a = refl +-- -- -- -- -- -- -- Iso.leftInv test3 = +-- -- -- -- -- -- -- sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- -- -- -- -- -- -- λ {(s , int) → setTruncProdLemma north s int (sphereConnected 1)} + +-- -- -- -- -- -- -- coHomGr0-S1 : grIso intGroup (coHomGr 1 (S₊ 1)) +-- -- -- -- -- -- -- coHomGr0-S1 = +-- -- -- -- -- -- -- Iso'→Iso +-- -- -- -- -- -- -- (iso' (compIso (symIso test3) (symIso coHom1≃∥S1×ℤ∥₀)) +-- -- -- -- -- -- -- (indIntGroup {G = coHomGr 1 (S₊ 1)} +-- -- -- -- -- -- -- (Iso.fun (compIso (symIso test3) (symIso coHom1≃∥S1×ℤ∥₀))) +-- -- -- -- -- -- -- ((λ i → ∣ transport (λ i → (S¹≡S1 (~ i) → lem i)) (Iso.inv test (base , 0)) ∣₀) +-- -- -- -- -- -- -- ∙ (λ i → ∣ transport (λ i → (S¹≡S1 (~ i) → lem i)) (helper2 i) ∣₀) +-- -- -- -- -- -- -- ∙ cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → {!!}})) +-- -- -- -- -- -- -- {!!} +-- -- -- -- -- -- -- {!!})) +-- -- -- -- -- -- -- where +-- -- -- -- -- -- -- helper : basedΩS¹≡ΩS¹ base ≡ {!basechange!} +-- -- -- -- -- -- -- helper = {!substComposite!} + +-- -- -- -- -- -- -- substComposite2 : ∀ {ℓ} {A B C : Type ℓ} +-- -- -- -- -- -- -- (P : A ≡ B) (Q : B ≡ C) (a : A) +-- -- -- -- -- -- -- → transport (P ∙ Q) a ≡ transport Q (transport P a) +-- -- -- -- -- -- -- substComposite2 = {!!} + +-- -- -- -- -- -- -- helper1 : transport (λ i → S¹≡S1 i × Int) (north , 0) ≡ (base , 0) +-- -- -- -- -- -- -- helper1 = refl +-- -- -- -- -- -- -- helper3 : transport (sym (basedΩS¹≡Int base)) 0 ≡ refl +-- -- -- -- -- -- -- helper3 = (λ i → transport (symDistr (basedΩS¹≡ΩS¹ base) ΩS¹≡Int i) 0) +-- -- -- -- -- -- -- ∙ substComposite2 (sym ΩS¹≡Int) (sym (basedΩS¹≡ΩS¹ base)) 0 +-- -- -- -- -- -- -- ∙ (λ i → transport (λ i → basedΩS¹≡ΩS¹ base (~ i)) refl) -- +-- -- -- -- -- -- -- ∙ transportRefl ((equiv-proof (basechange-isequiv base) refl) .fst .fst) +-- -- -- -- -- -- -- ∙ (λ i → equiv-proof (transport (λ j → isEquiv (refl-conjugation j)) (basedΩS¹→ΩS¹-isequiv i0)) refl .fst .fst) +-- -- -- -- -- -- -- ∙ (λ i → {!equiv-proof (transport (λ j → isEquiv (refl-conjugation j)) (basedΩS¹→ΩS¹-isequiv i0)) refl .fst!}) +-- -- -- -- -- -- -- ∙ {!basedΩS¹→ΩS¹!} +-- -- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- -- helper4 : (x : S¹) → Iso.inv test (base , 0) x ≡ x +-- -- -- -- -- -- -- helper4 = {!!} +-- -- -- -- -- -- -- helper2 : Iso.inv test (transport (λ i → S¹≡S1 i × Int) (north , 0)) ≡ λ x → x +-- -- -- -- -- -- -- helper2 = (λ i → Iso.inv test (base , 0)) ∙ {!!} ∙ {!!} + +-- -- -- -- -- prodId : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A × B} → proj₁ x ≡ proj₁ y → proj₂ x ≡ proj₂ y → x ≡ y +-- -- -- -- -- prodId {x = (a , b)} {y = (c , d)} id1 id2 i = (id1 i) , (id2 i) + +-- -- -- -- -- fstFunHelper : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 _) (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x +-- -- -- -- -- fstFunHelper a = MV.Ker-i⊂Im-d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a +-- -- -- -- -- (sym (isContrH1Unit×H1Unit .snd _) ∙ (isContrH1Unit×H1Unit .snd _)) +-- -- -- -- -- where +-- -- -- -- -- isContrH1Unit×H1Unit : isContr (Group.type (×coHomGr 1 Unit Unit)) +-- -- -- -- -- isContrH1Unit×H1Unit = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) +-- -- -- -- -- , λ {(a , b) → sigmaProdElim {D = λ (x : Σ[ x ∈ Group.type (×coHomGr 1 Unit Unit)] Unit) → (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) ≡ fst x} +-- -- -- -- -- (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- -- -- -- -- (λ a b _ → prodId (grIso.leftInv (coHomGrUnit≥1 0) ∣ a ∣₀) (grIso.leftInv (coHomGrUnit≥1 0) ∣ b ∣₀)) ((a , b) , tt)} + + + +-- -- -- -- -- helperMorph : isMorph intGroup (dirProd intGroup intGroup) λ a → a , (0 - a) +-- -- -- -- -- helperMorph = +-- -- -- -- -- indIntGroup {G = dirProd intGroup intGroup} +-- -- -- -- -- (λ a → a , (0 - a)) +-- -- -- -- -- refl +-- -- -- -- -- (λ a → prodId refl (helper2 a)) +-- -- -- -- -- λ a → prodId refl (helper3 a) +-- -- -- -- -- where +-- -- -- -- -- helper1 : (a : Int) → predInt (sucInt a) ≡ a +-- -- -- -- -- helper1 (pos zero) = refl +-- -- -- -- -- helper1 (pos (suc n)) = refl +-- -- -- -- -- helper1 (negsuc zero) = refl +-- -- -- -- -- helper1 (negsuc (suc n)) = refl + +-- -- -- -- -- helper4 : (a : Int) → sucInt (predInt a) ≡ a +-- -- -- -- -- helper4 (pos zero) = refl +-- -- -- -- -- helper4 (pos (suc n)) = refl +-- -- -- -- -- helper4 (negsuc zero) = refl +-- -- -- -- -- helper4 (negsuc (suc n)) = refl + +-- -- -- -- -- helper2 : (a : Int) → (pos 0 - sucInt a) ≡ predInt (pos 0 - a) +-- -- -- -- -- helper2 (pos zero) = refl +-- -- -- -- -- helper2 (pos (suc n)) = refl +-- -- -- -- -- helper2 (negsuc zero) = refl +-- -- -- -- -- helper2 (negsuc (suc n)) = sym (helper1 _) + +-- -- -- -- -- helper3 : (a : Int) → (pos 0 - predInt a) ≡ sucInt (pos 0 - a) +-- -- -- -- -- helper3 (pos zero) = refl +-- -- -- -- -- helper3 (pos (suc zero)) = refl +-- -- -- -- -- helper3 (pos (suc (suc n))) = sym (helper4 _) +-- -- -- -- -- helper3 (negsuc zero) = refl +-- -- -- -- -- helper3 (negsuc (suc n)) = refl + +-- -- -- -- -- helper : (a b : Int) → (pos 0 - (a +ℤ b)) ≡ ((pos 0 - a) +ℤ (pos 0 - b)) +-- -- -- -- -- helper a (pos zero) = refl +-- -- -- -- -- helper (pos zero) (pos (suc n)) = +-- -- -- -- -- cong (λ x → pos 0 - sucInt x) (+ℤ-comm (pos zero) (pos n)) +-- -- -- -- -- ∙ +ℤ-comm (pos 0 +negsuc n) (pos zero) +-- -- -- -- -- helper (pos (suc n₁)) (pos (suc n)) = +-- -- -- -- -- {!!} +-- -- -- -- -- helper (negsuc n₁) (pos (suc n)) = {!!} +-- -- -- -- -- helper a (negsuc n) = {!!} + +-- -- -- -- -- fun : morph intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- fst fun = MV.d _ _ _ (λ _ → tt) (λ _ → tt) 0 ∘ grIso.inv coHom0-S0 .fst ∘ λ a → a , (0 - a) +-- -- -- -- -- snd fun = {!!} +-- -- -- -- -- {- compMorph {F = intGroup} {G = dirProd intGroup intGroup} {H = coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))} +-- -- -- -- -- ((λ a → a , a) , (λ a b → refl)) +-- -- -- -- -- (compMorph {F = dirProd intGroup intGroup} {G = coHomGr 0 (S₊ 0)} {H = coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))} (grIso.inv coHom0-S0) +-- -- -- -- -- (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 +-- -- -- -- -- , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) .snd -} +-- -- -- -- -- {- theIso : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- theIso = Iso''→Iso _ _ +-- -- -- -- -- (iso'' ((λ x → ∣ (λ {(inl tt) → 0ₖ ; (inr tt) → 0ₖ ; (push a i) → Kn→ΩKn+1 0 x i}) ∣₀) , {!!}) +-- -- -- -- -- {!!} +-- -- -- -- -- {!MV.d!}) +-- -- -- -- -- -} + + + +-- -- -- -- -- theIso : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- theIso = +-- -- -- -- -- Iso''→Iso _ _ +-- -- -- -- -- (iso'' fun +-- -- -- -- -- (λ x inker → {!MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 +-- -- -- -- -- (grIso.inv coHom0-S0 .fst (g , g))!}) +-- -- -- -- -- (sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- -- -- -- -- λ x → pRec propTruncIsProp +-- -- -- -- -- (λ {(a , b) → {!fun!} }) +-- -- -- -- -- (fstFunHelper (∣ x ∣₀)))) +-- -- -- -- -- where +-- -- -- -- -- whoKnows : (a : S₊ 0 → Int) (x : MV.D Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt)) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (λ _ → a north) x +-- -- -- -- -- ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 a x +-- -- -- -- -- whoKnows a (inl x) = refl +-- -- -- -- -- whoKnows a (inr x) = refl +-- -- -- -- -- whoKnows a (push north i) = refl +-- -- -- -- -- whoKnows a (push south i) j = {!!} + +-- -- -- -- -- helper : (a : Int) → (grIso.inv coHom0-S0 .fst (a , a)) ≡ ∣ S0→Int (a , a) ∣₀ +-- -- -- -- -- helper a = {!have : + +-- -- -- -- -- ∣ +-- -- -- -- -- MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 +-- -- -- -- -- (S0→Int (x , x)) +-- -- -- -- -- ∣₀ +-- -- -- -- -- ≡ ∣ (λ _ → ∣ north ∣) ∣₀!} + +-- -- -- -- -- helper2 : (a b : Int) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a)) ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (b , b)) +-- -- -- -- -- → a ≡ b +-- -- -- -- -- helper2 a b id = pRec (isSetInt a b) (λ {(pt , id) → {!!}}) (fstFunHelper ∣ (MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a))) ∣₀) + +-- -- -- -- -- idFun : (a : Int) → MV.D Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 +-- -- -- -- -- idFun a (inl x) = 0ₖ +-- -- -- -- -- idFun a (inr x) = 0ₖ +-- -- -- -- -- idFun a (push north i) = Kn→ΩKn+1 zero a i +-- -- -- -- -- idFun a (push south i) = Kn→ΩKn+1 zero a i + +-- -- -- -- -- helper3 : (a : Int) → (MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a))) ≡ idFun a +-- -- -- -- -- helper3 a = funExt λ {(inl x) → refl ; (inr x) → refl ; (push north i) → refl ; (push south i) → refl } + +-- -- -- -- -- helper4 : (a b : Int) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a)) ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (b , b)) +-- -- -- -- -- → a ≡ b +-- -- -- -- -- helper4 a b id = +-- -- -- -- -- {!!} +-- -- -- -- -- ∙ {!!} +-- -- -- -- -- ∙ {!!} +-- -- -- -- -- where +-- -- -- -- -- helper5 : {!!} --PathP (λ k → id k (inl tt) ≡ id k (inr tt)) (Kn→ΩKn+1 zero a) (Kn→ΩKn+1 zero a) +-- -- -- -- -- helper5 i j = {!id i!} + +-- -- -- -- -- -- fun : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) → coHom 0 (S₊ 0) +-- -- -- -- -- -- fun a = (pRec {P = Σ[ x ∈ coHom 0 (S₊ 0)] (MV.d _ _ _ (λ _ → tt) (λ _ → tt) 0 x ≡ a) × isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) (MV.Δ _ _ _ (λ _ → tt) (λ _ → tt) 0) x} +-- -- -- -- -- -- (λ {(a1 , b) (c , d) → ΣProp≡ (λ x → isOfHLevelProd 1 (setTruncIsSet _ _) propTruncIsProp) +-- -- -- -- -- -- (sElim2 {B = λ a1 c → (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a1 ≡ a) +-- -- -- -- -- -- → MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 c ≡ a +-- -- -- -- -- -- → isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) +-- -- -- -- -- -- (λ z → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 z) a1 +-- -- -- -- -- -- → isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) +-- -- -- -- -- -- (λ z → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 z) c → a1 ≡ c} +-- -- -- -- -- -- (λ _ _ → {!!}) +-- -- -- -- -- -- (λ a c b1 d1 → pElim (λ _ → isOfHLevelΠ 1 λ _ → setTruncIsSet _ _) +-- -- -- -- -- -- λ b2 → pElim (λ _ → setTruncIsSet _ _) +-- -- -- -- -- -- λ d2 → {!d2!}) +-- -- -- -- -- -- a1 c (proj₁ b) (proj₁ d) (proj₂ b) (proj₂ d))}) +-- -- -- -- -- -- (λ {(a , b) → a , b , MV.Ker-d⊂Im-Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a {!!}}) +-- -- -- -- -- -- (fstFunHelper a)) .fst -- pRec {!!} {!!} (fstFunHelper a) From 5e1a576d94433c0a1401cbfb3696d93898920876 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 1 Jun 2020 18:45:24 +0200 Subject: [PATCH 37/89] cohom1-S1 --- Cubical/Data/Group/Base.agda | 62 +++++- Cubical/Data/Group/GroupLibrary.agda | 4 +- Cubical/ZCohomology/Sn/Sn.agda | 287 +++++++++++++++------------ 3 files changed, 219 insertions(+), 134 deletions(-) diff --git a/Cubical/Data/Group/Base.agda b/Cubical/Data/Group/Base.agda index 8d6fd99c1..14e9ff1d9 100644 --- a/Cubical/Data/Group/Base.agda +++ b/Cubical/Data/Group/Base.agda @@ -51,6 +51,12 @@ isInKer : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (Group.type G → isInKer G H ϕ g = ϕ g ≡ isGroup.id (Group.groupStruc H) +isSurjective : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → morph G H → Type _ +isSurjective G H ϕ = (x : Group.type H) → isInIm G H (fst ϕ) x + +isInjective : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → morph G H → Type _ +isInjective G H ϕ = (x : Group.type G) → isInKer G H (fst ϕ) x → x ≡ isGroup.id (Group.groupStruc G) + {- morphisms takes id to id -} morph0→0 : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (f : (Group.type G → Group.type H)) @@ -192,6 +198,7 @@ leftist-group : ∀ {ℓ} {A : Type ℓ} (Aset : isSet A) leftist-group Aset id inv comp lUnit assoc lCancel = group _ Aset (leftist-group-struct id inv comp lUnit assoc lCancel) + record Iso {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') : Type (ℓ-max ℓ ℓ') where constructor iso field @@ -210,8 +217,8 @@ record Iso'' {ℓ ℓ'} (A : Group ℓ) (B : Group ℓ') : Type (ℓ-max ℓ ℓ constructor iso'' field ϕ : morph A B - inj : (x : Group.type A) → isInKer A B (fst ϕ) x → x ≡ isGroup.id (Group.groupStruc A) - surj : (x : Group.type B) → isInIm A B (fst ϕ) x + inj : isInjective A B ϕ + surj : isSurjective A B ϕ _≃_ : ∀ {ℓ ℓ'} (A : Group ℓ) (B : Group ℓ') → Type (ℓ-max ℓ ℓ') A ≃ B = Σ (morph A B) \ f → (E.isEquiv (f .fst)) @@ -374,3 +381,54 @@ Iso''→Iso A B (iso'' ϕ inj surj) = ((a1 + - a2) + a2) ≡⟨ cong (λ x → x + a2) fstIdHelper ⟩ (id + a2) ≡⟨ lUnit a2 ⟩ a2 ∎ + + +-- Injectivity and surjectivity in terms of exact sequences +exactInjectiveL : ∀ {ℓ ℓ' ℓ''} (E : Group ℓ'') (F : Group ℓ'') (G : Group ℓ) (H : Group ℓ') + (ϕ : morph E F) (ψ : morph F G) (ξ : morph G H) + → isSurjective E F ϕ + → ((x : Group.type G) → isInKer G H (fst ξ) x → (isInIm F G (fst ψ) x)) + → isProp (Group.type E) + → isInjective G H ξ +exactInjectiveL (group E _ (group-struct idE _ _ _ _ _ _ _)) F G H + (ϕ , mfϕ) (ψ , mfψ) _ surjϕ ker⊂im isPropE x inker = + rec (Group.setStruc G _ _) + (λ {(f , id) → sym id ∙ cong ψ (isPropF f (isGroup.id (Group.groupStruc F))) ∙ morph0→0 F G ψ mfψ}) + (ker⊂im x inker) + where + isPropF : isProp (Group.type F) + isPropF a b = rec (Group.setStruc F _ _) + (λ {(c , id-c) → rec (Group.setStruc F _ _) + (λ {(d , id-d) → sym id-c ∙ cong ϕ (isPropE c d) ∙ id-d}) + (surjϕ b) }) + (surjϕ a) + + +exactSurjectiveR : ∀ {ℓ ℓ' ℓ''} (E : Group ℓ'') (F : Group ℓ'') (G : Group ℓ) (H : Group ℓ') + (ϕ : morph E F) (ψ : morph F G) (ξ : morph G H) + → isInjective G H ξ + → ((x : Group.type F) → (isInKer F G (fst ψ) x) → isInIm E F (fst ϕ) x) + → isProp (Group.type H) + → isSurjective E F ϕ +exactSurjectiveR E F (group G GSet (group-struct idG -G _+G_ lUnitG rUnitG assocG lCancelG rCancelG)) + (group H _ (group-struct idH _ _ _ _ _ _ _)) + (ϕ , mfϕ) (ψ , mfψ) (ξ , mfξ) isInjξ ker⊂im isPropH x = ker⊂im x (isPropF (ψ x) idG) + + where + isPropF : isProp G + isPropF a b = sym (rUnitG a) + ∙ cong (a +G_) (sym (lCancelG b)) + ∙ sym (assocG a (-G b) b) + ∙ cong (λ x → x +G b) helper + ∙ lUnitG b + where + helper : a +G (-G b) ≡ idG + helper = isInjξ (a +G (-G b)) (isPropH (ξ (a +G -G b)) idH ) + +exactSurjectiveR' : ∀ {ℓ ℓ' ℓ''} (E : Group ℓ) (F : Group ℓ') (G : Group ℓ'') + (ϕ : morph E F) (ψ : morph F G) + → ((x : Group.type F) → (isInKer F G (fst ψ) x) → isInIm E F (fst ϕ) x) + → isProp (Group.type G) + → isSurjective E F ϕ +exactSurjectiveR' E F (group G GSet (group-struct idG -G _+G_ lUnitG rUnitG assocG lCancelG rCancelG)) + (ϕ , mfϕ) (ψ , mfψ) ker⊂im isPropG x = ker⊂im x (isPropG (ψ x) idG) diff --git a/Cubical/Data/Group/GroupLibrary.agda b/Cubical/Data/Group/GroupLibrary.agda index d515f0ab1..6cc2ec98a 100644 --- a/Cubical/Data/Group/GroupLibrary.agda +++ b/Cubical/Data/Group/GroupLibrary.agda @@ -16,9 +16,7 @@ open import Cubical.Data.Sigma hiding (_×_) open import Cubical.HITs.SetQuotients as sq - - - +-- The image of a morphism imGroup : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → morph G H → Group (ℓ-max ℓ ℓ') diff --git a/Cubical/ZCohomology/Sn/Sn.agda b/Cubical/ZCohomology/Sn/Sn.agda index f2c40350a..ed2cf24b6 100644 --- a/Cubical/ZCohomology/Sn/Sn.agda +++ b/Cubical/ZCohomology/Sn/Sn.agda @@ -47,15 +47,41 @@ open import Cubical.Data.HomotopyGroup open import Cubical.Data.Bool hiding (_⊕_) -congDep : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {a : A} (p q r s : a ≡ a) (f : (a ≡ a) → B) → PathP (λ i → p i ≡ q i) r s - → PathP (λ i → {!B!}) (f r) (f s) -- PathP (λ i → f (p i) ≡ f (q i)) (cong f r) (cong f s) -congDep = {!!} +cong₂Funct2 : ∀ {ℓ ℓ'} {A : Type ℓ} {x y : A} {B : Type ℓ'} (f : A → A → B) → + (p : x ≡ y) → + {u v : A} (q : u ≡ v) → + cong₂ f p q ≡ cong (λ x → f x u) p ∙ cong (f y) q +cong₂Funct2 {x = x} {y = y} f p {u = u} {v = v} q j i = + hcomp (λ k → λ { (i = i0) → f x u + ; (i = i1) → f y (q k) + ; (j = i0) → f (p i) (q (i ∧ k))}) + (f (p i) u) + +cong₂Sym : ∀ {ℓ} {A : Type ℓ} {x : A} + (p : x ≡ x) → + (refl ≡ p) → + (P : p ≡ p) → + cong₂ (λ x y → x ∙ y) P (sym P) ≡ refl +cong₂Sym {x = x} p = J (λ p _ → (P : p ≡ p) → cong₂ (λ x y → x ∙ y) P (sym P) ≡ refl) + λ P → cong₂Funct2 (λ x y → x ∙ y) P (sym P) + ∙ PathP→compPathR (λ j → cong (λ x → rUnit x (~ j)) P ∙ cong (λ x → lUnit x (~ j)) (sym P)) + ∙ (λ j → (sym (rUnit refl)) ∙ rCancel P j ∙ lUnit refl) + ∙ (λ j → sym (rUnit refl) ∙ lUnit (lUnit refl) (~ j) ) + ∙ rCancel (sym (rUnit refl)) + +abstract + cong₂Sym1 : ∀ {ℓ} {A : Type ℓ} {x : A} + (p : x ≡ x) → + (refl ≡ p) → + (P : p ≡ p) → + cong (λ x → x ∙ p) P ≡ cong (λ x → p ∙ x) P + cong₂Sym1 {x = x} p id P = rUnit _ + ∙ (λ i → cong (λ x → x ∙ p) P ∙ lCancel (λ i → p ∙ P i) (~ i)) -- cong (λ x → cong (λ x → x ∙ p) P ∙ x) {!!} + ∙ assoc _ _ _ + ∙ (λ j → cong₂Funct2 (λ x y → x ∙ y) P (sym P) (~ j) ∙ (λ i → p ∙ P i)) + ∙ (λ j → cong₂Sym p id P j ∙ (λ i → p ∙ P i)) + ∙ sym (lUnit _) -moveDownPathP : {!!} -moveDownPathP = {!!} - -congFunct2 : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A} (f : A → A → B) (j₁ : x ≡ y) {u v : A} (q : u ≡ v) → cong₂ f j₁ q ≡ cong (λ x₁ → f x₁ u) j₁ ∙ cong (f y) q -congFunct2 = {!!} prodId : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A × B} → proj₁ x ≡ proj₁ y → proj₂ x ≡ proj₂ y → x ≡ y prodId {x = (a , b)} {y = (c , d)} id1 id2 i = (id1 i) , (id2 i) @@ -152,18 +178,6 @@ coHom0-S0 = ×morph mf1 mf2 = (λ {(a , b) → (mf1 .fst a) , mf2 .fst b}) , λ {(a , b) (c , d) i → mf1 .snd a c i , mf2 .snd b d i} -helper1 : Iso (S₊ 1) S¹ -Iso.fun helper1 north = base -Iso.fun helper1 south = base -Iso.fun helper1 (merid a i) = loop i -Iso.inv helper1 base = north -Iso.inv helper1 (loop i) = loop* i -Iso.rightInv helper1 base = refl -Iso.rightInv helper1 (loop i) j = {!helper loop!} -Iso.leftInv helper1 north = refl -Iso.leftInv helper1 south = merid north -Iso.leftInv helper1 (merid north i) j = {!compPath-filler (merid south) (sym (merid north))!} -Iso.leftInv helper1 (merid south i) j = {!!} isGroupoidS1 : isGroupoid (S₊ 1) isGroupoidS1 = transport (λ i → isGroupoid (S¹≡S1 (~ i))) isGroupoidS¹ @@ -171,8 +185,6 @@ isConnectedS1 : (x : S₊ 1) → ∥ north ≡ x ∥₋₁ isConnectedS1 x = pRec propTruncIsProp (λ p → ∣ cong (transport (sym (S¹≡S1))) p ∙ transport⁻Transport (S¹≡S1) x ∣₋₁) (isConnectedS¹ (transport S¹≡S1 x)) -isConnectedS1Loop : (x : S₊ 1) (p : x ≡ x) → ∥ refl ≡ p ∥₋₁ -isConnectedS1Loop x p = {!∥ ? ∥!} -- basechange* : (x y : S¹) → x ≡ y → x ≡ y → ΩS¹ -- basechange* x y = J (λ y p → (x ≡ y) → ΩS¹) (basechange x) @@ -277,14 +289,37 @@ coHom1Iso = d : _ d = MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 + i : _ + i = MV.i _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1 + Δ : _ Δ = MV.Δ _ _ (S₊ 1) (λ _ → tt) (λ _ → tt) 0 d-surj : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x - d-surj x = {!!} - + d-surj = sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ x → exactSurjectiveR' (coHomGr 0 (S₊ 0)) + (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) + (×coHomGr 1 Unit Unit) + {!d , MV.dIsHom _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0!} + {!i , MV.iIsHom _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1!} + {!grIso.fun (coHomGrUnit≥1 0)!} + {!!} + ? + ∣ x ∣₀ + +{- exactSurjectiveR (coHomGr 0 (S₊ 0)) + (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) + (×coHomGr 1 Unit Unit) + trivialGroup + (d , MV.dIsHom _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) + {!i , M!} + (? , {!grIso.fun (coHomGrUnit≥1 0)!}) + {!!} + {!!} + {!!} +-} theFuns : (a : Int) → Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 theFuns a (inl x) = 0ₖ theFuns a (inr x) = 0ₖ @@ -292,6 +327,71 @@ coHom1Iso = theFuns a (push south i) = 0ₖ + theFunsId2 : (a b : Int) (x : Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) + → Kn→ΩKn+1 1 (theFuns a x) ∙ Kn→ΩKn+1 1 (theFuns b x) ≡ Kn→ΩKn+1 1 (theFuns (a +ℤ b) x) + theFunsId2 a b (inl x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) + theFunsId2 a b (inr x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) + theFunsId2 a b (push north i) j = + hcomp (λ k → λ {(i = i0) → ((λ i₁ → + (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) + j + ; (i = i1) → ((λ i₁ → + (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) + j + ; (j = i0) → cong₂Funct2 (λ p q → Kn→ΩKn+1 1 p ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b) (~ k) i + ; (j = i1) → (helper2 a b) k i }) + (hcomp (λ k → λ { (j = i0) → compPath-filler (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) k i + ; (j = i1) → compPath-filler (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) k i + ; (i = i1) → RHS-filler b j k + ; (i = i0) → ((λ i₁ → + (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) + j}) + (bottom-filler a j i)) + + where + + bottom-filler : (a : Int) → + PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) + (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) + ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) + j ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) + (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) + ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) + j) (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) + bottom-filler a j i = + hcomp (λ k → λ {(j = i0) → helper2 (~ k) i ; + (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 a) i}) + ((λ j₂ → ∣ rCancel (merid north) j j₂ ∣) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + + where + helper2 : Path (Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣ ≡ Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣) + (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i) ∙ Kn→ΩKn+1 1 ∣ north ∣) + (λ i → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + helper2 = cong₂Sym1 (Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) (~ i) j ∣) (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + + RHS-filler : (b : Int) → + PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j + ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j) + (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) + (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) + RHS-filler b j i = + hcomp (λ k → λ {(j = i0) → cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i ; + (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 b) i}) + (cong (λ q → (λ i → ∣ rCancel (merid north) j i ∣) ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i) + + helper2 : (a b : Int) → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b) ≡ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ℤ b)) + helper2 a b = + sym (congFunct (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b)) + ∙ (λ i → cong (Kn→ΩKn+1 1) (Iso.rightInv (Iso3-Kn-ΩKn+1 0) (Kn→ΩKn+1 0 a ∙ Kn→ΩKn+1 0 b) (~ i))) + ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ₖ b)) ) + ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (addLemma a b i))) + + theFunsId2 a b (push south i) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ sym (lUnit _) + inj : (a : Int) → theFuns a ≡ (λ _ → ∣ north ∣) → a ≡ pos 0 inj a id = pRec (isSetInt _ _) @@ -312,9 +412,6 @@ coHom1Iso = ∙ sym (cong (λ f → f (inr tt)) id)) ∙ rCancel (cong (λ f → f (inr tt)) id) - isInImTheFuns : (x : Int) → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout (λ _ → tt) (λ _ → tt))) d ∣ theFuns x ∣₀ - isInImTheFuns x = {!!} - consMember : (a : Int) → Group.type ImGr consMember a = d ∣ (λ _ → a) ∣₀ , ∣ ∣ (λ _ → a) ∣₀ , refl ∣₋₁ @@ -325,39 +422,49 @@ coHom1Iso = (∣ (λ _ → a) ∣₀ , ∣ (λ _ → 0) ∣₀) , cong ∣_∣₀ (λ i x → (rUnitₖ a i)) ∣₋₁) - f+consMember : (f : Group.type ImGr) → ∃[ x ∈ Int ] (isGroup.comp (Group.groupStruc ImGr) f (isGroup.inv (Group.groupStruc ImGr) (consMember x)) ≡ (∣ theFuns x ∣₀ , d-surj ∣ theFuns x ∣₀)) - f+consMember f = + f+consMember' : (f : Group.type ImGr) → ∃[ x ∈ Int × Int ] (isGroup.comp (Group.groupStruc ImGr) f (isGroup.inv (Group.groupStruc ImGr) (consMember (proj₁ x))) ≡ (∣ theFuns (proj₂ x) ∣₀ , d-surj ∣ theFuns (proj₂ x) ∣₀)) + f+consMember' f = let null = isGroup.id (Group.groupStruc ImGr) - = isGroup.inv (Group.groupStruc ImGr) _+_ = isGroup.comp (Group.groupStruc ImGr) in sigmaElim {C = λ (f : Group.type ImGr) - → ∃[ x ∈ Int ] (isGroup.comp (Group.groupStruc ImGr) f (isGroup.inv (Group.groupStruc ImGr) (consMember x)) - ≡ (∣ theFuns x ∣₀ , d-surj ∣ theFuns x ∣₀)) } + → ∃[ x ∈ Int × Int ] (isGroup.comp (Group.groupStruc ImGr) f (isGroup.inv (Group.groupStruc ImGr) (consMember (proj₁ x))) ≡ (∣ theFuns (proj₂ x) ∣₀ , d-surj ∣ theFuns (proj₂ x) ∣₀)) } (λ _ → isOfHLevelSuc 1 propTruncIsProp) (λ f idf → pRec propTruncIsProp - (sigmaElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) - λ g id → ∣ (g south) , (ΣProp≡ (λ _ → propTruncIsProp) - (pRec (setTruncIsSet _ _) - (λ id → {! (λ i → ∣ id i ∣₀ +ₕ -ₕ (consMember (g south) .fst)) !} - ∙ {!!} - ∙ {!!}) - (Iso.fun PathIdTrunc₀Iso id))) ∣₋₁) + ((sigmaElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ g id → ∣ ((g south) , ((g north) +ₖ (-ₖ g south))) + , (ΣProp≡ (λ _ → propTruncIsProp) + (pRec (setTruncIsSet _ _) + (λ id → (λ i → ∣ id (~ i) ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀) ∙ funId1 g) + (Iso.fun PathIdTrunc₀Iso id))) ∣₋₁)) idf) f + where - subtrLemma : (a b : Int) → a - b ≡ a +ₖ (-ₖ b) - subtrLemma = {!!} - - funId : (g : S₊ 0 → Int) - → (x : Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) - → MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g x +ₖ (-ₖ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) x) ≡ - theFuns (g north - g south) x - funId g (inl x) = cong (λ x → (0ₖ {n = 1} ) +ₖ x) (-0ₖ {n = 1}) ∙ rUnitₖ (0ₖ {n = 1}) - funId g (inr x) = cong (λ x → (0ₖ {n = 1} ) +ₖ x) (-0ₖ {n = 1}) ∙ rUnitₖ (0ₖ {n = 1}) - funId g (push north i) j = - {!!} - funId g (push south i) j = {!!} + funId1 : (g : S₊ 0 → Int) + → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀ ≡ + ∣ theFuns ((g north) +ₖ (-ₖ (g south))) ∣₀ + funId1 g = (λ i → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ + +ₕ (morphMinus (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) d + (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ∣ (λ _ → g south) ∣₀ (~ i))) + ∙ sym (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ g ∣₀ (-ₕ ∣ (λ _ → g south) ∣₀)) + ∙ (cong (λ x → d ∣ x ∣₀) g'Id) + ∙ cong ∣_∣₀ helper + where + g' : S₊ 0 → Int + g' north = (g north) +ₖ (-ₖ (g south)) + g' south = 0 + + g'Id : (λ x → g x +ₖ (-ₖ (g south))) ≡ g' + g'Id = funExt (λ {north → refl + ; south → rCancelₖ (g south)}) + + helper : MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g' ≡ theFuns (g north +ₖ (-ₖ g south)) + helper = funExt λ {(inl tt) → refl + ; (inr tt) → refl + ; (push north i) → refl + ; (push south i) → refl} finIm : (f : Group.type ImGr) → ∃[ x ∈ Int ] ((∣ theFuns x ∣₀ , d-surj ∣ theFuns x ∣₀) ≡ f) finIm f = @@ -366,86 +473,8 @@ coHom1Iso = _+_ = isGroup.comp (Group.groupStruc ImGr) rUnit = isGroup.rUnit (Group.groupStruc ImGr) lCancel = isGroup.lCancel (Group.groupStruc ImGr) - in pRec propTruncIsProp (λ {(a , id) → ∣ a , (sym (sym (rUnit f) ∙ cong (λ x → f + x) (sym (lCancel null)) ∙ cong (λ x → f + x) (rUnit (- null)) ∙ cong (λ x → f + (- x)) (sym (consMember≡0 a)) ∙ id)) ∣₋₁}) (f+consMember f) -- (f+consMember f) - - - surjective : {!f + consMember f+consMember!} - surjective = {!!} - - - theFunsId3 : (a b : Int) → ∣ theFuns a ∣₀ ≡ ∣ theFuns b ∣₀ → a ≡ b - theFunsId3 = {!!} - - theFunsId2 : (a b : Int) (x : Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) - → Kn→ΩKn+1 1 (theFuns a x) ∙ Kn→ΩKn+1 1 (theFuns b x) ≡ Kn→ΩKn+1 1 (theFuns (a +ℤ b) x) - theFunsId2 a b (inl x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) - ∙ sym (lUnit _) - theFunsId2 a b (inr x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) - ∙ sym (lUnit _) - theFunsId2 a b (push north i) j = - hcomp (λ k → λ {(i = i0) → ((λ i₁ → - (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) - ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) - j - ; (i = i1) → ((λ i₁ → - (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) - ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) - j - ; (j = i0) → congFunct2 (λ p q → Kn→ΩKn+1 1 p ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b) (~ k) i - ; (j = i1) → (helper2 a b) k i }) - (hcomp (λ k → λ { (i = i1) → {!!} - ; (i = i0) → ((λ i₁ → - (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) - ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) - j}) - {!absRetr ? ?!}) - - - -{- -i = i0 ⊢ ((λ i₁ → - (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) - ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) - j -i = i1 ⊢ ((λ i₁ → - (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) - ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) - j -j = i0 ⊢ cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a) ∙ - cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) -j = i1 ⊢ (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) i --} - - where - - - - - helper3 : sym ((λ i₁ → (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) ∙ cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a) ∙ cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) ∙ (λ i₁ → - (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) - ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁)) ≡ - cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b) - helper3 = {!λ i → !} - ∙ {!!} - ∙ {!!} - - helper2 : (a b : Int) → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b) ≡ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ℤ b)) - helper2 a b = - sym (congFunct (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b)) - ∙ (λ i → cong (Kn→ΩKn+1 1) (Iso.rightInv (Iso3-Kn-ΩKn+1 0) (Kn→ΩKn+1 0 a ∙ Kn→ΩKn+1 0 b) (~ i))) - ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ₖ b)) ) - ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (addLemma a b i))) - - theFunsId2 a b (push south i) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) - ∙ sym (lUnit _) - + in pRec propTruncIsProp (λ {((a , b) , id) → ∣ b , (ΣProp≡ (λ _ → propTruncIsProp) (cong fst (sym id) ∙ cong (λ x → f .fst +ₕ x) ((λ i → (-ₕ (consMember≡0 a i .fst))) ∙ sym (lUnitₕ (-ₕ 0ₕ)) ∙ rCancelₕ 0ₕ) ∙ (rUnitₕ (f .fst)))) ∣₋₁}) (f+consMember' f) - theFunsId : (a b : Int) → ∣ theFuns (a +ℤ b) ∣₀ ≡ (∣ theFuns a ∣₀ +ₕ ∣ theFuns b ∣₀) - theFunsId a b = {!!} - ∙ {!!} - ∙ (λ i → ∣ (λ x → ΩKn+1→Kn {!congFunct ∣_∣ ? ? i!}) ∣₀) - ∙ (λ i → ∣ (λ x → ΩKn+1→Kn (Kn→ΩKn+1 1 (theFuns a x) ∙ Kn→ΩKn+1 1 (theFuns b x))) ∣₀) - ∙ λ i → ∣ (λ x → (theFuns a x) +ₖ theFuns b x) ∣₀ -- coHomGrIm : grIso (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) From 39fca67040008b027ebf1c8beb2d68779aa57802 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 2 Jun 2020 14:30:42 +0200 Subject: [PATCH 38/89] S1-done --- Cubical/Data/Group/Base.agda | 6 + Cubical/Data/Group/GroupLibrary.agda | 77 ++-- Cubical/Homotopy/Loopspace.agda | 3 +- Cubical/ZCohomology/MayerVietorisReduced.agda | 13 + Cubical/ZCohomology/Sn/Sn.agda | 410 ++++++++++++------ 5 files changed, 330 insertions(+), 179 deletions(-) diff --git a/Cubical/Data/Group/Base.agda b/Cubical/Data/Group/Base.agda index 14e9ff1d9..ab3b691b3 100644 --- a/Cubical/Data/Group/Base.agda +++ b/Cubical/Data/Group/Base.agda @@ -337,6 +337,12 @@ Iso.inv (idIso G) = (λ x → x) , λ _ _ → refl Iso.rightInv (idIso G) _ = refl Iso.leftInv (idIso G) _ = refl +invIso : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → Iso G H → Iso H G +Iso.fun (invIso G H (iso _ inv _ _)) = inv +Iso.inv (invIso G H (iso fun _ _ _)) = fun +Iso.rightInv (invIso G H (iso _ _ _ leftInv)) = leftInv +Iso.leftInv (invIso G H (iso _ _ rightInv _)) = rightInv + Iso''→Iso : ∀ {ℓ ℓ'} (A : Group ℓ) (B : Group ℓ') → Iso'' A B → Iso A B Iso''→Iso A B (iso'' ϕ inj surj) = Iso'→Iso diff --git a/Cubical/Data/Group/GroupLibrary.agda b/Cubical/Data/Group/GroupLibrary.agda index 6cc2ec98a..933ca16b9 100644 --- a/Cubical/Data/Group/GroupLibrary.agda +++ b/Cubical/Data/Group/GroupLibrary.agda @@ -144,57 +144,34 @@ isGroup.rCancel (Group.groupStruc intGroup) = λ a → +-comm a (pos 0 - a) ∙ --- -- Short exact sequences without req for first group to be definitionally equal to the trivial group -- --- record SES {ℓ ℓ' ℓ'' ℓ'''} (A : Group ℓ) (B : Group ℓ') (leftGr : Group ℓ'') (rightGr : Group ℓ''') : Type (ℓ-suc (ℓ-max ℓ (ℓ-max ℓ' (ℓ-max ℓ'' ℓ''')))) where --- constructor ses --- field --- isTrivialLeft : Iso leftGr trivialGroup --- isTrivialRight : Iso rightGr trivialGroup - --- left : morph leftGr A --- right : morph B rightGr --- ϕ : morph A B + -- Short exact sequences without req for first group to be definitionally equal to the trivial group -- +record SES {ℓ ℓ' ℓ'' ℓ'''} (A : Group ℓ) (B : Group ℓ') (leftGr : Group ℓ'') (rightGr : Group ℓ''') : Type (ℓ-suc (ℓ-max ℓ (ℓ-max ℓ' (ℓ-max ℓ'' ℓ''')))) where + constructor ses + field + isTrivialLeft : isProp (Group.type leftGr) + isTrivialRight : isProp (Group.type rightGr) + + left : morph leftGr A + right : morph B rightGr + ϕ : morph A B --- Ker-ϕ⊂Im-left : (x : Group.type A) -- --- → isInKer A B (fst ϕ) x --- → isInIm leftGr A (fst left) x --- Ker-right⊂Im-ϕ : (x : Group.type B) -- --- → isInKer B rightGr (fst right) x --- → isInIm A B (fst ϕ) x - --- SES→inj-surj-morph : ∀ {ℓ ℓ' ℓ'' ℓ'''} (A : Group ℓ) (B : Group ℓ') (leftGr : Group ℓ'') (rightGr : Group ℓ''') --- → SES A B leftGr rightGr --- → inj-surj-morph A B --- inj-surj-morph.ϕ (SES→inj-surj-morph _ _ _ _ (ses _ _ _ _ ϕ _ _)) = ϕ --- inj-surj-morph.inj --- (SES→inj-surj-morph A _ lGr _ --- (ses (iso ψ ξ rinv linv) _ left _ ϕ Ker-ϕ⊂Im-left _)) x inKer = --- rec (isOfHLevelPath' 1 (Group.setStruc A) _ _) --- (λ (a , pf) → sym pf ∙∙ (λ i → fst left (helper a i)) ∙∙ morph0→0 lGr A (fst left) (snd left)) --- (Ker-ϕ⊂Im-left x inKer) --- where --- helper : (a : Group.type lGr) → a ≡ isGroup.id (Group.groupStruc lGr) --- helper a = --- a ≡⟨ sym (linv a) ⟩ --- ξ .fst (ψ .fst a) ≡⟨ refl ⟩ --- ξ .fst tt ≡⟨ morph0→0 trivialGroup lGr (fst ξ) (snd ξ) ⟩ --- isGroup.id (Group.groupStruc lGr) ∎ - --- inj-surj-morph.surj (SES→inj-surj-morph _ _ _ rGr --- (ses _ (iso ψ ξ rinv linv) _ right ϕ _ Ker-right⊂Im-ϕ)) b = --- Ker-right⊂Im-ϕ b (helper (right .fst b)) --- where --- helper : (a : Group.type rGr) → a ≡ isGroup.id (Group.groupStruc rGr) --- helper a = --- a ≡⟨ sym (linv a) ⟩ --- ξ .fst (ψ .fst a) ≡⟨ refl ⟩ --- ξ .fst tt ≡⟨ morph0→0 trivialGroup rGr (fst ξ) (snd ξ) ⟩ --- isGroup.id (Group.groupStruc rGr) ∎ - --- SES→Iso : ∀ {ℓ ℓ' ℓ'' ℓ'''} (A : Group ℓ) (B : Group ℓ') (leftGr : Group ℓ'') (rightGr : Group ℓ''') --- → SES A B leftGr rightGr --- → Iso A B --- SES→Iso A B lGr rGr seq = inj-surj-morph→Iso A B (SES→inj-surj-morph A B lGr rGr seq) + Ker-ϕ⊂Im-left : (x : Group.type A) -- + → isInKer A B (fst ϕ) x + → isInIm leftGr A (fst left) x + Ker-right⊂Im-ϕ : (x : Group.type B) -- + → isInKer B rightGr (fst right) x + → isInIm A B (fst ϕ) x +SES→Iso : ∀ {ℓ ℓ' ℓ'' ℓ'''} (A : Group ℓ) (B : Group ℓ') (leftGr : Group ℓ'') (rightGr : Group ℓ''') + → SES A B leftGr rightGr + → Iso A B +SES→Iso A B lGr rGr (ses PropL PropR (lfun , lmf) (rfun , rmf) (ϕ , mf) Ker-ϕ⊂Im-left Ker-right⊂Im-ϕ) = + Iso''→Iso _ _ + (iso'' + (ϕ , mf) + (λ a inker → rec (Group.setStruc A _ _) + (λ { (a , p) → sym p ∙ cong lfun (PropL a _) ∙ morph0→0 lGr A lfun lmf }) + (Ker-ϕ⊂Im-left a inker)) + λ a → Ker-right⊂Im-ϕ a (PropR _ _)) {- direct products of groups -} diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda index 06dab77b2..aba35f71b 100644 --- a/Cubical/Homotopy/Loopspace.agda +++ b/Cubical/Homotopy/Loopspace.agda @@ -43,8 +43,9 @@ Eckmann-Hilton {A = A} n α β i = (generalEH α β i) {- Homotopy group version -} -π-comp : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → ∥ typ ((Ω^ (suc n)) A) ∥₀ +π-comp : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → ∥ typ ((Ω^ (suc n)) A) ∥₀ → ∥ typ ((Ω^ (suc n)) A) ∥₀ + → ∥ typ ((Ω^ (suc n)) A) ∥₀ π-comp n = elim2 (λ _ _ → setTruncIsSet) λ p q → ∣ p ∙ q ∣₀ Eckmann-Hilton-π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p q : ∥ typ ((Ω^ (2 + n)) A) ∥₀) diff --git a/Cubical/ZCohomology/MayerVietorisReduced.agda b/Cubical/ZCohomology/MayerVietorisReduced.agda index def107b80..28270493c 100644 --- a/Cubical/ZCohomology/MayerVietorisReduced.agda +++ b/Cubical/ZCohomology/MayerVietorisReduced.agda @@ -352,3 +352,16 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : ∙ (λ i → d (suc n) ∣ (λ x → Fa (f x)) ∣₀ +ₕ -ₕ (d (suc n) ∣ (λ x → Fb (g x)) ∣₀)) ∙ (λ i → ∣ funExt (d-preLeftId (suc n) Fa) i ∣₀ +ₕ -ₕ ∣ funExt (d-preRightId (suc n) Fb) i ∣₀) ∙ rCancelₕ 0ₕ + + +d-morph : ∀ {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) (n : ℕ) + → morph (coHomGr n C) (coHomGr (suc n) (Pushout f g)) +d-morph A B C f g n = (MV.d A B C f g n) , (MV.dIsHom A B C f g n) + +i-morph : ∀ {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) (n : ℕ) + → morph (coHomGr n (Pushout f g)) (×coHomGr n A B) +i-morph A B C f g n = (MV.i A B C f g n) , (MV.iIsHom A B C f g n) + +Δ-morph : ∀ {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) (n : ℕ) + → morph (×coHomGr n A B) (coHomGr n C) +Δ-morph A B C f g n = MV.Δ A B C f g n , MV.ΔIsHom A B C f g n diff --git a/Cubical/ZCohomology/Sn/Sn.agda b/Cubical/ZCohomology/Sn/Sn.agda index ed2cf24b6..0ae8b5dee 100644 --- a/Cubical/ZCohomology/Sn/Sn.agda +++ b/Cubical/ZCohomology/Sn/Sn.agda @@ -33,7 +33,7 @@ open import Cubical.Homotopy.Connected open import Cubical.Homotopy.Freudenthal open import Cubical.HITs.SmashProduct.Base open import Cubical.Data.Unit -open import Cubical.Data.Group.Base renaming (Iso to grIso ; compIso to compGrIso) +open import Cubical.Data.Group.Base renaming (Iso to grIso ; compIso to compGrIso ; invIso to invGrIso ; idIso to idGrIso) open import Cubical.Data.Group.GroupLibrary open import Cubical.ZCohomology.coHomZero-map @@ -47,6 +47,7 @@ open import Cubical.Data.HomotopyGroup open import Cubical.Data.Bool hiding (_⊕_) + cong₂Funct2 : ∀ {ℓ ℓ'} {A : Type ℓ} {x y : A} {B : Type ℓ'} (f : A → A → B) → (p : x ≡ y) → {u v : A} (q : u ≡ v) → @@ -126,6 +127,19 @@ Iso.leftInv (pushoutSn n) (merid a i) = refl Sn≡Pushout : (n : ℕ) → (S₊ (suc n)) ≡ (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt) Sn≡Pushout n = isoToPath (pushoutSn n) + +coHomPushout≡coHomSn : (n m : ℕ) → grIso (coHomGr m (S₊ (suc n))) (coHomGr m (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt)) +grIso.fun (coHomPushout≡coHomSn n m) = (sRec setTruncIsSet λ f → ∣ (λ {(inl x) → f north ; (inr x) → f south ; (push a i) → f (merid a i)}) ∣₀) + , sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ a b → cong ∣_∣₀ (funExt λ {(inl x) → refl ; (inr x) → refl ; (push a i) → refl }) +grIso.inv (coHomPushout≡coHomSn n m) = sRec setTruncIsSet (λ f → ∣ (λ {north → f (inl tt) ; south → f (inr tt) ; (merid a i) → f (push a i)}) ∣₀) + , sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ a b → cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → refl }) +grIso.rightInv (coHomPushout≡coHomSn n m) = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ (funExt λ {(inl x) → refl ; (inr x) → refl ; (push a i) → refl }) +grIso.leftInv (coHomPushout≡coHomSn n m) = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → refl }) + isContr→≡Unit : {A : Type₀} → isContr A → A ≡ Unit isContr→≡Unit contr = isoToPath (iso (λ _ → tt) (λ _ → fst contr) (λ _ → refl) λ _ → snd contr _) @@ -141,11 +155,8 @@ grIso.inv coHomGrUnit0 = (λ a → ∣ (λ _ → a) ∣₀) , λ a b i → ∣ ( grIso.rightInv coHomGrUnit0 a = refl grIso.leftInv coHomGrUnit0 = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ a → refl -coHomGrUnit≥1 : (n : ℕ) → grIso (coHomGr (suc n) Unit) trivialGroup -grIso.fun (coHomGrUnit≥1 n) = (λ _ → tt) , (λ _ _ → refl) -grIso.inv (coHomGrUnit≥1 n) = (λ _ → ∣ (λ _ → ∣ north ∣) ∣₀) , λ _ _ → sym (rUnitₕ 0ₕ) -grIso.rightInv (coHomGrUnit≥1 n) a = refl -grIso.leftInv (coHomGrUnit≥1 n) a = sym (helper2 .snd 0ₕ) ∙ helper2 .snd a +isContrCohomUnit : (n : ℕ) → isContr (coHom (suc n) Unit) +isContrCohomUnit n = subst isContr (λ i → ∥ UnitToTypeId (coHomK (suc n)) (~ i) ∥₀) (helper n) where helper : (n : ℕ) → isContr (∥ coHomK (suc n) ∥₀) helper n = subst isContr ((isoToPath (truncOfTruncIso {A = S₊ (1 + n)} 2 (1 + n))) ∙ sym propTrunc≡Trunc0 ∙ λ i → ∥ hLevelTrunc (suc (+-comm n 2 i)) (S₊ (1 + n)) ∥₀) @@ -155,8 +166,11 @@ grIso.leftInv (coHomGrUnit≥1 n) a = sym (helper2 .snd 0ₕ) ∙ helper2 .snd a helper2 zero = 0 , refl helper2 (suc n) = (suc n) , λ i → suc (+-comm n 2 i) - helper2 : isContr (coHom (suc n) Unit) - helper2 = subst isContr (λ i → ∥ UnitToTypeId (coHomK (suc n)) (~ i) ∥₀) (helper n) +coHomGrUnit≥1 : (n : ℕ) → grIso (coHomGr (suc n) Unit) trivialGroup +grIso.fun (coHomGrUnit≥1 n) = (λ _ → tt) , (λ _ _ → refl) +grIso.inv (coHomGrUnit≥1 n) = (λ _ → ∣ (λ _ → ∣ north ∣) ∣₀) , λ _ _ → sym (rUnitₕ 0ₕ) +grIso.rightInv (coHomGrUnit≥1 n) a = refl +grIso.leftInv (coHomGrUnit≥1 n) a = sym (isContrCohomUnit n .snd 0ₕ) ∙ isContrCohomUnit n .snd a S0→Int : (a : Int × Int) → S₊ 0 → Int S0→Int a north = proj₁ a @@ -216,76 +230,57 @@ isConnectedS1 x = pRec propTruncIsProp -- fst (invMapSurj G H (ϕ , pf) surj) a = {!pRec!} -- snd (invMapSurj G H (ϕ , pf) surj) = {!!} --- ImIso : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (ϕ : morph G H) → ((x : Group.type H) → isInIm G H (fst ϕ) x) --- → grIso H (imGroup G H ϕ) --- ImIso G H (ϕ , mf) surj = --- let idH = isGroup.id (Group.groupStruc H) --- idG = isGroup.id (Group.groupStruc G) --- _+G_ = isGroup.comp (Group.groupStruc G) --- _+H_ = isGroup.comp (Group.groupStruc H) --- _+Im_ = isGroup.comp (Group.groupStruc (imGroup G H (ϕ , mf))) --- invG = isGroup.inv (Group.groupStruc G) --- invH = isGroup.inv (Group.groupStruc H) --- lUnit = isGroup.lUnit (Group.groupStruc H) --- lCancel = isGroup.rCancel (Group.groupStruc H) --- in --- Iso''→Iso _ _ --- (iso'' ((λ x → x , pRec propTruncIsProp (λ (a , b) → ∣ a , b ∣₋₁) (surj x)) --- , λ a b → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) --- (λ surja → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) --- (λ surjb → --- pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) --- (λ surja+b → --- (λ i → (a +H b) , (pRec (propTruncIsProp) --- (λ (a , b) → ∣ a , b ∣₋₁) --- (propTruncIsProp (surj (isGroup.comp (Group.groupStruc H) a b)) ∣ surja+b ∣₋₁ i))) ∙ --- (λ i → (a +H b) , ∣ (fst surja+b) , (snd surja+b) ∣₋₁) ∙ --- ΣProp≡ (λ _ → propTruncIsProp) refl ∙ --- λ i → (a +H b) , pRec (propTruncIsProp) --- (λ p1 → --- pRec (λ x y → squash x y) --- (λ p2 → --- ∣ --- isGroup.comp (Group.groupStruc G) (fst p1) (fst p2) , --- mf (fst p1) (fst p2) ∙ --- cong₂ (isGroup.comp (Group.groupStruc H)) (snd p1) (snd p2) --- ∣₋₁) --- (pRec (propTruncIsProp) --- ∣_∣₋₁ (propTruncIsProp ∣ surjb ∣₋₁ (surj b) i))) --- (pRec (propTruncIsProp) --- ∣_∣₋₁ (propTruncIsProp ∣ surja ∣₋₁ (surj a) i ))) --- (surj (isGroup.comp (Group.groupStruc H) a b))) --- (surj b)) --- (surj a)) --- (λ x inker → cong fst inker) --- {- --- pElim (λ _ → Group.setStruc H _ _) --- (λ xinim → - --- (surj x)) -} --- λ x → pRec propTruncIsProp (λ inimx → ∣ (ϕ (inimx .fst)) , ΣProp≡ (λ _ → propTruncIsProp) (inimx .snd) ∣₋₁) (snd x)) - - -coHom1Iso : grIso intGroup - ((imGroup (coHomGr 0 (S₊ 0)) - (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) - (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 - , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0))) -coHom1Iso = - Iso''→Iso _ _ - (iso'' ((λ x → ∣ theFuns x ∣₀ , d-surj ∣ theFuns x ∣₀) - , (λ a b → ΣProp≡ (λ _ → propTruncIsProp) - (cong ∣_∣₀ (funExt λ x → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 1) _) - ∙ cong (ΩKn+1→Kn) (sym (theFunsId2 a b x)))))) - (λ x inker → pRec (isSetInt _ _) (inj x) (Iso.fun PathIdTrunc₀Iso (cong fst inker))) - finIm) +ImIso : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (ϕ : morph G H) → ((x : Group.type H) → isInIm G H (fst ϕ) x) + → grIso H (imGroup G H ϕ) +ImIso G H (ϕ , mf) surj = + let idH = isGroup.id (Group.groupStruc H) + idG = isGroup.id (Group.groupStruc G) + _+G_ = isGroup.comp (Group.groupStruc G) + _+H_ = isGroup.comp (Group.groupStruc H) + _+Im_ = isGroup.comp (Group.groupStruc (imGroup G H (ϕ , mf))) + invG = isGroup.inv (Group.groupStruc G) + invH = isGroup.inv (Group.groupStruc H) + lUnit = isGroup.lUnit (Group.groupStruc H) + lCancel = isGroup.rCancel (Group.groupStruc H) + in + Iso''→Iso _ _ + (iso'' ((λ x → x , pRec propTruncIsProp (λ (a , b) → ∣ a , b ∣₋₁) (surj x)) + , λ a b → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) + (λ surja → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) + (λ surjb → + pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) + (λ surja+b → + (λ i → (a +H b) , (pRec (propTruncIsProp) + (λ (a , b) → ∣ a , b ∣₋₁) + (propTruncIsProp (surj (isGroup.comp (Group.groupStruc H) a b)) ∣ surja+b ∣₋₁ i))) ∙ + (λ i → (a +H b) , ∣ (fst surja+b) , (snd surja+b) ∣₋₁) ∙ + ΣProp≡ (λ _ → propTruncIsProp) refl ∙ + λ i → (a +H b) , pRec (propTruncIsProp) + (λ p1 → + pRec (λ x y → squash x y) + (λ p2 → + ∣ + isGroup.comp (Group.groupStruc G) (fst p1) (fst p2) , + mf (fst p1) (fst p2) ∙ + cong₂ (isGroup.comp (Group.groupStruc H)) (snd p1) (snd p2) + ∣₋₁) + (pRec (propTruncIsProp) + ∣_∣₋₁ (propTruncIsProp ∣ surjb ∣₋₁ (surj b) i))) + (pRec (propTruncIsProp) + ∣_∣₋₁ (propTruncIsProp ∣ surja ∣₋₁ (surj a) i ))) + (surj (isGroup.comp (Group.groupStruc H) a b))) + (surj b)) + (surj a)) + (λ x inker → cong fst inker) + λ x → pRec propTruncIsProp (λ inimx → ∣ (ϕ (inimx .fst)) , ΣProp≡ (λ _ → propTruncIsProp) (inimx .snd) ∣₋₁) (snd x)) +{- +H¹-S¹≃Int : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +H¹-S¹≃Int = + Iso''→Iso _ _ + (iso'' ((λ x → ∣ theFuns x ∣₀) , λ a b → cong ∣_∣₀ (funExt λ x → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 1) _) ∙ sym (cong (ΩKn+1→Kn) (theFunsId2 a b x)))) + (λ x inker → pRec (isSetInt _ _) (inj x) (Iso.fun PathIdTrunc₀Iso inker)) + finIm) where - ImGr : _ - ImGr = (imGroup (coHomGr 0 (S₊ 0)) - (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) - (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 - , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) - d : _ d = MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 @@ -299,27 +294,14 @@ coHom1Iso = d-surj : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x d-surj = sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) - λ x → exactSurjectiveR' (coHomGr 0 (S₊ 0)) - (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) - (×coHomGr 1 Unit Unit) - {!d , MV.dIsHom _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0!} - {!i , MV.iIsHom _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1!} - {!grIso.fun (coHomGrUnit≥1 0)!} - {!!} - ? - ∣ x ∣₀ - -{- exactSurjectiveR (coHomGr 0 (S₊ 0)) - (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) - (×coHomGr 1 Unit Unit) - trivialGroup - (d , MV.dIsHom _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) - {!i , M!} - (? , {!grIso.fun (coHomGrUnit≥1 0)!}) - {!!} - {!!} - {!!} --} + λ x → MV.Ker-i⊂Im-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ + (sym (isContrHelper .snd _)) + where + isContrHelper : isContr (Group.type (×coHomGr 1 Unit Unit)) + isContrHelper = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) + , λ y → prodId (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₁ y)) + (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₂ y)) + theFuns : (a : Int) → Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 theFuns a (inl x) = 0ₖ theFuns a (inr x) = 0ₖ @@ -412,35 +394,26 @@ coHom1Iso = ∙ sym (cong (λ f → f (inr tt)) id)) ∙ rCancel (cong (λ f → f (inr tt)) id) - consMember : (a : Int) → Group.type ImGr - consMember a = d ∣ (λ _ → a) ∣₀ , ∣ ∣ (λ _ → a) ∣₀ , refl ∣₋₁ - consMember≡0 : (a : Int) → consMember a ≡ isGroup.id (Group.groupStruc ImGr) + consMember : (a : Int) → coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) + consMember a = d ∣ (λ _ → a) ∣₀ + + consMember≡0 : (a : Int) → consMember a ≡ 0ₕ consMember≡0 a = - ΣProp≡ (λ _ → propTruncIsProp) (MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ (λ _ → a) ∣₀ ∣ (∣ (λ _ → a) ∣₀ , ∣ (λ _ → 0) ∣₀) , cong ∣_∣₀ (λ i x → (rUnitₖ a i)) ∣₋₁) - f+consMember' : (f : Group.type ImGr) → ∃[ x ∈ Int × Int ] (isGroup.comp (Group.groupStruc ImGr) f (isGroup.inv (Group.groupStruc ImGr) (consMember (proj₁ x))) ≡ (∣ theFuns (proj₂ x) ∣₀ , d-surj ∣ theFuns (proj₂ x) ∣₀)) - f+consMember' f = - let null = isGroup.id (Group.groupStruc ImGr) - - = isGroup.inv (Group.groupStruc ImGr) - _+_ = isGroup.comp (Group.groupStruc ImGr) - in - sigmaElim {C = λ (f : Group.type ImGr) - → ∃[ x ∈ Int × Int ] (isGroup.comp (Group.groupStruc ImGr) f (isGroup.inv (Group.groupStruc ImGr) (consMember (proj₁ x))) ≡ (∣ theFuns (proj₂ x) ∣₀ , d-surj ∣ theFuns (proj₂ x) ∣₀)) } - (λ _ → isOfHLevelSuc 1 propTruncIsProp) - (λ f idf → pRec propTruncIsProp - ((sigmaElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) - λ g id → ∣ ((g south) , ((g north) +ₖ (-ₖ g south))) - , (ΣProp≡ (λ _ → propTruncIsProp) - (pRec (setTruncIsSet _ _) - (λ id → (λ i → ∣ id (~ i) ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀) ∙ funId1 g) - (Iso.fun PathIdTrunc₀Iso id))) ∣₋₁)) - idf) - f - + f+consMember' : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int × Int ] (f +ₕ (-ₕ (consMember (proj₁ x))) ≡ ∣ theFuns (proj₂ x) ∣₀) + f+consMember' = + sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ f → pRec propTruncIsProp + (sigmaElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) + (λ g id → ∣ ((g south) , ((g north) +ₖ (-ₖ g south))) + , (pRec (setTruncIsSet _ _) + (λ id → (λ i → ∣ id (~ i) ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀) ∙ funId1 g) + (Iso.fun PathIdTrunc₀Iso id)) ∣₋₁)) + (d-surj ∣ f ∣₀) where funId1 : (g : S₊ 0 → Int) → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀ ≡ @@ -465,16 +438,197 @@ coHom1Iso = ; (inr tt) → refl ; (push north i) → refl ; (push south i) → refl} + finIm : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int ] (∣ theFuns x ∣₀ ≡ f) + finIm f = + pRec propTruncIsProp + (λ {((a , b) , id) → ∣ b , (sym id ∙ cong (λ x → f +ₕ x) ((λ i → (-ₕ (consMember≡0 a i))) ∙ sym (lUnitₕ (-ₕ 0ₕ)) ∙ rCancelₕ 0ₕ) ∙ (rUnitₕ f)) ∣₋₁}) + (f+consMember' f) +-} +Hⁿ-Sⁿ≃Int : (n : ℕ) → grIso intGroup (coHomGr (suc n) (S₊ (suc n))) +Hⁿ-Sⁿ≃Int zero = + compGrIso {F = intGroup} {G = {!!}} {H = {!coHomGr 1 (S₊ 1)!}} + (Iso''→Iso _ _ + (iso'' ((λ x → ∣ theFuns x ∣₀) , λ a b → cong ∣_∣₀ (funExt λ x → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 1) _) ∙ sym (cong (ΩKn+1→Kn) (theFunsId2 a b x)))) + (λ x inker → pRec (isSetInt _ _) (inj x) (Iso.fun PathIdTrunc₀Iso inker)) + finIm)) + {!invGrIso _ _ (coHomPushout≡coHomSn 0 1)!} + where + d : _ + d = MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 + + i : _ + i = MV.i _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1 + + Δ : _ + Δ = MV.Δ _ _ (S₊ 1) (λ _ → tt) (λ _ → tt) 0 + + + d-surj : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) + → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x + d-surj = sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ x → MV.Ker-i⊂Im-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ + (sym (isContrHelper .snd _)) + where + isContrHelper : isContr (Group.type (×coHomGr 1 Unit Unit)) + isContrHelper = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) + , λ y → prodId (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₁ y)) + (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₂ y)) + + theFuns : (a : Int) → Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 + theFuns a (inl x) = 0ₖ + theFuns a (inr x) = 0ₖ + theFuns a (push north i) = Kn→ΩKn+1 0 a i + theFuns a (push south i) = 0ₖ + + + theFunsId2 : (a b : Int) (x : Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) + → Kn→ΩKn+1 1 (theFuns a x) ∙ Kn→ΩKn+1 1 (theFuns b x) ≡ Kn→ΩKn+1 1 (theFuns (a +ℤ b) x) + theFunsId2 a b (inl x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) + theFunsId2 a b (inr x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) + theFunsId2 a b (push north i) j = + hcomp (λ k → λ {(i = i0) → ((λ i₁ → + (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) + j + ; (i = i1) → ((λ i₁ → + (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) + j + ; (j = i0) → cong₂Funct2 (λ p q → Kn→ΩKn+1 1 p ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b) (~ k) i + ; (j = i1) → (helper2 a b) k i }) + (hcomp (λ k → λ { (j = i0) → compPath-filler (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) k i + ; (j = i1) → compPath-filler (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) k i + ; (i = i1) → RHS-filler b j k + ; (i = i0) → ((λ i₁ → + (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) + j}) + (bottom-filler a j i)) + + where + + bottom-filler : (a : Int) → + PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) + (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) + ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) + j ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) + (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) + ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) + j) (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) + bottom-filler a j i = + hcomp (λ k → λ {(j = i0) → helper2 (~ k) i ; + (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 a) i}) + ((λ j₂ → ∣ rCancel (merid north) j j₂ ∣) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + + where + helper2 : Path (Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣ ≡ Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣) + (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i) ∙ Kn→ΩKn+1 1 ∣ north ∣) + (λ i → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + helper2 = cong₂Sym1 (Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) (~ i) j ∣) (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + + RHS-filler : (b : Int) → + PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j + ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j) + (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) + (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) + RHS-filler b j i = + hcomp (λ k → λ {(j = i0) → cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i ; + (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 b) i}) + (cong (λ q → (λ i → ∣ rCancel (merid north) j i ∣) ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i) + + helper2 : (a b : Int) → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b) ≡ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ℤ b)) + helper2 a b = + sym (congFunct (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b)) + ∙ (λ i → cong (Kn→ΩKn+1 1) (Iso.rightInv (Iso3-Kn-ΩKn+1 0) (Kn→ΩKn+1 0 a ∙ Kn→ΩKn+1 0 b) (~ i))) + ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ₖ b)) ) + ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (addLemma a b i))) + + theFunsId2 a b (push south i) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ sym (lUnit _) + + inj : (a : Int) → theFuns a ≡ (λ _ → ∣ north ∣) → a ≡ pos 0 + inj a id = + pRec (isSetInt _ _) + (sigmaElim (λ _ → isOfHLevelPath 2 isSetInt _ _) + (λ a p → pRec (isSetInt _ _) + (λ id2 → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 0) _) + ∙ cong (ΩKn+1→Kn) + (PathP→compPathR + (cong (λ f → cong f (push north)) id) + ∙ test)) + (Iso.fun PathIdTrunc₀Iso p))) (d-surj ∣ theFuns a ∣₀) + where - finIm : (f : Group.type ImGr) → ∃[ x ∈ Int ] ((∣ theFuns x ∣₀ , d-surj ∣ theFuns x ∣₀) ≡ f) + test : (λ i → id i (inl tt)) ∙ (λ i → ∣ north ∣) ∙ sym (λ i → id i (inr tt)) ≡ refl + test = (λ i → cong (λ f → f (inl tt)) id + ∙ lUnit (sym (cong (λ f → f (inr tt)) id)) (~ i)) + ∙ (λ i → cong (λ f → f (push south i)) id + ∙ sym (cong (λ f → f (inr tt)) id)) + ∙ rCancel (cong (λ f → f (inr tt)) id) + + + consMember : (a : Int) → coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) + consMember a = d ∣ (λ _ → a) ∣₀ + + consMember≡0 : (a : Int) → consMember a ≡ 0ₕ + consMember≡0 a = + (MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ (λ _ → a) ∣₀ ∣ + (∣ (λ _ → a) ∣₀ , ∣ (λ _ → 0) ∣₀) + , cong ∣_∣₀ (λ i x → (rUnitₖ a i)) ∣₋₁) + + f+consMember' : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int × Int ] (f +ₕ (-ₕ (consMember (proj₁ x))) ≡ ∣ theFuns (proj₂ x) ∣₀) + f+consMember' = + sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ f → pRec propTruncIsProp + (sigmaElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) + (λ g id → ∣ ((g south) , ((g north) +ₖ (-ₖ g south))) + , (pRec (setTruncIsSet _ _) + (λ id → (λ i → ∣ id (~ i) ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀) ∙ funId1 g) + (Iso.fun PathIdTrunc₀Iso id)) ∣₋₁)) + (d-surj ∣ f ∣₀) + where + funId1 : (g : S₊ 0 → Int) + → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀ ≡ + ∣ theFuns ((g north) +ₖ (-ₖ (g south))) ∣₀ + funId1 g = (λ i → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ + +ₕ (morphMinus (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) d + (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ∣ (λ _ → g south) ∣₀ (~ i))) + ∙ sym (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ g ∣₀ (-ₕ ∣ (λ _ → g south) ∣₀)) + ∙ (cong (λ x → d ∣ x ∣₀) g'Id) + ∙ cong ∣_∣₀ helper + where + g' : S₊ 0 → Int + g' north = (g north) +ₖ (-ₖ (g south)) + g' south = 0 + + g'Id : (λ x → g x +ₖ (-ₖ (g south))) ≡ g' + g'Id = funExt (λ {north → refl + ; south → rCancelₖ (g south)}) + + helper : MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g' ≡ theFuns (g north +ₖ (-ₖ g south)) + helper = funExt λ {(inl tt) → refl + ; (inr tt) → refl + ; (push north i) → refl + ; (push south i) → refl} + finIm : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int ] (∣ theFuns x ∣₀ ≡ f) finIm f = - let null = isGroup.id (Group.groupStruc ImGr) - - = isGroup.inv (Group.groupStruc ImGr) - _+_ = isGroup.comp (Group.groupStruc ImGr) - rUnit = isGroup.rUnit (Group.groupStruc ImGr) - lCancel = isGroup.lCancel (Group.groupStruc ImGr) - in pRec propTruncIsProp (λ {((a , b) , id) → ∣ b , (ΣProp≡ (λ _ → propTruncIsProp) (cong fst (sym id) ∙ cong (λ x → f .fst +ₕ x) ((λ i → (-ₕ (consMember≡0 a i .fst))) ∙ sym (lUnitₕ (-ₕ 0ₕ)) ∙ rCancelₕ 0ₕ) ∙ (rUnitₕ (f .fst)))) ∣₋₁}) (f+consMember' f) + pRec propTruncIsProp + (λ {((a , b) , id) → ∣ b , (sym id ∙ cong (λ x → f +ₕ x) ((λ i → (-ₕ (consMember≡0 a i))) ∙ sym (lUnitₕ (-ₕ 0ₕ)) ∙ rCancelₕ 0ₕ) ∙ (rUnitₕ f)) ∣₋₁}) + (f+consMember' f) +Hⁿ-Sⁿ≃Int (suc n) = + compGrIso (Hⁿ-Sⁿ≃Int n) + (transport (λ i → grIso {!!} {!coHomGr (suc (suc n)) (Pushout (λ _ → tt) (λ _ → tt))!}) {!!}) + + +{- +coHom1S1≃ℤ : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +grIso.fun coHom1S1≃ℤ = grIso.fun {!compGrIso coHom1Iso (invGrIso _ _ (ImIso _ _ (d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ?))!} +grIso.inv coHom1S1≃ℤ = {!!} +grIso.rightInv coHom1S1≃ℤ = {!!} +grIso.leftInv coHom1S1≃ℤ = {!!} +-} +-- compGrIso coHom1Iso (invGrIso _ _ (ImIso _ _ {!d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0!} {!!})) -- coHomGrIm : grIso (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) From 9cdb1da9d8a639d6cbe01d1f1ce067a1fbe3db2f Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Thu, 4 Jun 2020 04:52:03 +0200 Subject: [PATCH 39/89] basechange-morph --- Cubical/Data/Group/Base.agda | 372 ++- Cubical/Data/Group/GroupLibrary.agda | 218 +- Cubical/Data/Prod/Base.agda | 4 + Cubical/Data/Prod/Properties.agda | 3 + Cubical/HITs/S1/Base.agda | 73 + Cubical/HITs/S1/Properties.agda | 6 - Cubical/HITs/Sn/Properties.agda | 22 +- Cubical/ZCohomology/MayerVietorisReduced.agda | 6 +- Cubical/ZCohomology/Sn/Sn.agda | 2145 ++++++++++------- 9 files changed, 1701 insertions(+), 1148 deletions(-) diff --git a/Cubical/Data/Group/Base.agda b/Cubical/Data/Group/Base.agda index ab3b691b3..ab8dd4fca 100644 --- a/Cubical/Data/Group/Base.agda +++ b/Cubical/Data/Group/Base.agda @@ -2,7 +2,7 @@ module Cubical.Data.Group.Base where -open import Cubical.Foundations.Prelude hiding (comp) +open import Cubical.Foundations.Prelude renaming (comp to comp') open import Cubical.Foundations.HLevels open import Cubical.Data.Prod open import Cubical.HITs.PropositionalTruncation hiding (map) @@ -14,6 +14,7 @@ import Cubical.Foundations.Equiv.HalfAdjoint as HAE open import Cubical.HITs.SetQuotients as sq record isGroup {ℓ} (A : Type ℓ) : Type ℓ where + no-eta-equality constructor group-struct field id : A @@ -26,6 +27,7 @@ record isGroup {ℓ} (A : Type ℓ) : Type ℓ where rCancel : ∀ a → comp a (inv a) ≡ id record Group ℓ : Type (ℓ-suc ℓ) where + no-eta-equality constructor group field type : Type ℓ @@ -33,14 +35,18 @@ record Group ℓ : Type (ℓ-suc ℓ) where groupStruc : isGroup type isMorph : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (f : (Group.type G → Group.type H)) → Type (ℓ-max ℓ ℓ') -isMorph (group G Gset (group-struct _ _ _⊙_ _ _ _ _ _)) - (group H Hset (group-struct _ _ _∘_ _ _ _ _ _)) f - = (g0 g1 : G) → f (g0 ⊙ g1) ≡ (f g0) ∘ (f g1) +isMorph G H f = (g0 g1 : Group.type G) → f (isGroup.comp (Group.groupStruc G) g0 g1) ≡ isGroup.comp (Group.groupStruc H) (f g0) (f g1) -morph : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → Type (ℓ-max ℓ ℓ') -morph G H = Σ (Group.type G → Group.type H) (isMorph G H) +record morph {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') : Type (ℓ-max ℓ ℓ') where + no-eta-equality + constructor mph + field + fun : Group.type G → Group.type H + ismorph : isMorph G H fun + + +open import Cubical.Data.Sigma hiding (_×_ ; comp) -open import Cubical.Data.Sigma hiding (_×_) isInIm : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (Group.type G → Group.type H) → Group.type H → Type _ @@ -52,30 +58,42 @@ isInKer G H ϕ g = ϕ g ≡ isGroup.id (Group.groupStruc H) isSurjective : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → morph G H → Type _ -isSurjective G H ϕ = (x : Group.type H) → isInIm G H (fst ϕ) x +isSurjective G H ϕ = (x : Group.type H) → isInIm G H (morph.fun ϕ) x isInjective : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → morph G H → Type _ -isInjective G H ϕ = (x : Group.type G) → isInKer G H (fst ϕ) x → x ≡ isGroup.id (Group.groupStruc G) +isInjective G H ϕ = (x : Group.type G) → isInKer G H (morph.fun ϕ) x → x ≡ isGroup.id (Group.groupStruc G) + +-0≡0 : ∀ {ℓ} {G : Group ℓ} → isGroup.inv (Group.groupStruc G) (isGroup.id (Group.groupStruc G)) ≡ isGroup.id (Group.groupStruc G) +-0≡0 {G = G} = sym (isGroup.lUnit (Group.groupStruc G) _) ∙ isGroup.rCancel (Group.groupStruc G) _ {- morphisms takes id to id -} morph0→0 : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (f : (Group.type G → Group.type H)) → isMorph G H f → f (isGroup.id (Group.groupStruc G)) ≡ isGroup.id (Group.groupStruc H) -morph0→0 (group G strucG (group-struct idG invG compG lUnitG rUnitG assocG lCancelG rCancelG)) - (group H strucH (group-struct idH invH compH lUnitH rUnitH assocH lCancelH rCancelH)) f morph = - f idG ≡⟨ sym (rUnitH (f idG)) ⟩ - compH (f idG) idH ≡⟨ (λ i → compH (f idG) (rCancelH (f idG) (~ i))) ⟩ - compH (f idG) (compH (f idG) (invH (f idG))) ≡⟨ sym (assocH (f idG) (f idG) (invH (f idG))) ⟩ - compH (compH (f idG) (f idG)) (invH (f idG)) ≡⟨ sym (cong (λ x → compH x (invH (f idG))) (sym (cong f (lUnitG idG)) ∙ morph idG idG)) ⟩ - compH (f idG) (invH (f idG)) ≡⟨ rCancelH (f idG) ⟩ - idH ∎ +morph0→0 G' H' f ismorph = 0→0 + where + open Group + open isGroup + G = Group.groupStruc G' + H = Group.groupStruc H' + + 0→0 : f (id G) ≡ id H + 0→0 = + f (id G) ≡⟨ sym (rUnit H (f (id G))) ⟩ + comp H (f (id G)) (id H) ≡⟨ (λ i → comp H (f (id G)) (rCancel H (f (id G)) (~ i))) ⟩ + comp H (f (id G)) (comp H (f (id G)) (inv H (f (id G)))) ≡⟨ sym (assoc H (f (id G)) (f (id G)) (inv H (f (id G)))) ⟩ + comp H (comp H (f (id G)) (f (id G))) (inv H (f (id G))) ≡⟨ sym (cong (λ x → comp H x (inv H (f (id G)))) (sym (cong f (lUnit G (id G))) ∙ ismorph (id G) (id G))) ⟩ + comp H (f (id G)) (inv H (f (id G))) ≡⟨ rCancel H (f (id G)) ⟩ + id H ∎ + + {- a morphism ϕ satisfies ϕ(- a) = - ϕ(a) -} morphMinus : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (f : (Group.type G → Group.type H)) → isMorph G H f → (g : Group.type G) → f (isGroup.inv (Group.groupStruc G) g) ≡ isGroup.inv (Group.groupStruc H) (f g) -morphMinus G H f morph g = +morphMinus G H f mf g = let idG = isGroup.id (Group.groupStruc G) idH = isGroup.id (Group.groupStruc H) invG = isGroup.inv (Group.groupStruc G) @@ -88,7 +106,7 @@ morphMinus G H f morph g = compG = isGroup.comp (Group.groupStruc G) compH = isGroup.comp (Group.groupStruc H) helper : compH (f (invG g)) (f g) ≡ idH - helper = sym (morph (invG g) g) ∙ (λ i → f (lCancelG g i)) ∙ morph0→0 G H f morph + helper = sym (mf (invG g) g) ∙ (λ i → f (lCancelG g i)) ∙ morph0→0 G H f mf -- sym (morph (invG g) g) ∙ (λ i → f (lCancelG g i)) ∙ morph0→0 G H f morph in f (invG g) ≡⟨ sym (rUnitH (f (invG g))) ⟩ compH (f (invG g)) idH ≡⟨ (λ i → compH (f (invG g)) (rCancelH (f g) (~ i))) ⟩ compH (f (invG g)) (compH (f g) (invH (f g))) ≡⟨ sym (assocH (f (invG g)) (f g) (invH (f g))) ⟩ @@ -204,8 +222,8 @@ record Iso {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') : Type (ℓ-max ℓ ℓ') field fun : morph G H inv : morph H G - rightInv : I.section (fun .fst) (inv .fst) - leftInv : I.retract (fun .fst) (inv .fst) + rightInv : I.section (morph.fun fun) (morph.fun inv) + leftInv : I.retract (morph.fun fun) (morph.fun inv) record Iso' {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') : Type (ℓ-max ℓ ℓ') where constructor iso' @@ -221,220 +239,162 @@ record Iso'' {ℓ ℓ'} (A : Group ℓ) (B : Group ℓ') : Type (ℓ-max ℓ ℓ surj : isSurjective A B ϕ _≃_ : ∀ {ℓ ℓ'} (A : Group ℓ) (B : Group ℓ') → Type (ℓ-max ℓ ℓ') -A ≃ B = Σ (morph A B) \ f → (E.isEquiv (f .fst)) - -Iso'→Iso : ∀ {ℓ ℓ'} {G : Group ℓ} {H : Group ℓ'} → Iso' G H → Iso G H -Iso'→Iso {G = group G Gset Ggroup} {H = group H Hset Hgroup} i = iso (fun , funMorph) (inv , invMorph) rightInv leftInv +A ≃ B = Σ (morph A B) \ f → (E.isEquiv (morph.fun f)) + +Iso'→Iso : ∀ {ℓ ℓ'} {G : Group ℓ} {H : Group ℓ'} → Iso' G H → Iso G H +morph.fun (Iso.fun (Iso'→Iso {G = G} {H = H} iso1)) = I.Iso.fun (Iso'.isoSet iso1) +morph.ismorph (Iso.fun (Iso'→Iso {G = G} {H = H} iso1)) = Iso'.isoSetMorph iso1 +morph.fun (Iso.inv (Iso'→Iso {G = G} {H = H} iso1)) = I.Iso.inv (Iso'.isoSet iso1) +morph.ismorph (Iso.inv (Iso'→Iso {G = G} {H = H} iso1)) a b = + cong₂ (λ x y → ψ (comp H' x y)) + (sym (I.Iso.rightInv (Iso'.isoSet iso1) _)) + (sym (I.Iso.rightInv (Iso'.isoSet iso1) _)) + ∙ cong ψ (sym (Iso'.isoSetMorph iso1 _ _)) + ∙ I.Iso.leftInv (Iso'.isoSet iso1) _ where - G_ : Group _ - G_ = (group G Gset Ggroup) - - H_ : Group _ - H_ = (group H Hset Hgroup) - - open Iso' - - fun : G → H - fun = I.Iso.fun (isoSet i) - - inv : H → G - inv = I.Iso.inv (isoSet i) + open Iso' + open isGroup + H' = Group.groupStruc H + ψ = I.Iso.inv (Iso'.isoSet iso1) - rightInv : I.section fun inv - rightInv = I.Iso.rightInv (isoSet i) - - leftInv : I.retract fun inv - leftInv = I.Iso.leftInv (isoSet i) - - e' : G E.≃ H - e' = I.isoToEquiv (I.iso fun inv rightInv leftInv) - - funMorph : isMorph G_ H_ fun - funMorph = isoSetMorph i - - _∘_ : H → H → H - _∘_ = isGroup.comp Hgroup - - _⊙_ : G → G → G - _⊙_ = isGroup.comp Ggroup - - invMorph : isMorph H_ G_ inv - invMorph h0 h1 = E.invEq (HAE.congEquiv e') - (fun (inv (h0 ∘ h1)) ≡⟨ rightInv (h0 ∘ h1) ⟩ - h0 ∘ h1 ≡⟨ cong (λ x → x ∘ h1) (sym (rightInv h0)) ⟩ - (fun (inv h0)) ∘ h1 ≡⟨ cong (λ x → fun (inv h0) ∘ x) (sym (rightInv h1)) ⟩ - (fun (inv h0)) ∘ (fun (inv h1)) ≡⟨ sym (funMorph (inv h0) (inv h1)) ⟩ - fun ((inv h0) ⊙ (inv h1)) ∎ ) +Iso.rightInv (Iso'→Iso {G = G} {H = H} iso1) = I.Iso.rightInv (Iso'.isoSet iso1) +Iso.leftInv (Iso'→Iso {G = G} {H = H} iso1) = I.Iso.leftInv (Iso'.isoSet iso1) Equiv→Iso' : ∀ {ℓ ℓ'} {G : Group ℓ} {H : Group ℓ'} → G ≃ H → Iso' G H Equiv→Iso' {G = group G Gset Ggroup} {H = group H Hset Hgroup} - e = iso' i' (e .fst .snd) + e = iso' i' (morph.ismorph (e .fst)) where e' : G E.≃ H - e' = (e .fst .fst) , (e .snd) + e' = (morph.fun (e .fst)) , (e .snd) i' : I.Iso G H i' = E.equivToIso e' -compMorph : ∀ {ℓ} {F G H : Group ℓ} (I : morph F G) (J : morph G H) → morph F H -compMorph {ℓ} {group F Fset (group-struct _ _ _⊙_ _ _ _ _ _)} - {group G Gset (group-struct _ _ _∙_ _ _ _ _ _)} - {group H Hset (group-struct _ _ _∘_ _ _ _ _ _)} - (i , ic) (j , jc) = k , kc - where - k : F → H - k f = j (i f) - - kc : (g0 g1 : F) → k (g0 ⊙ g1) ≡ (k g0) ∘ (k g1) - kc g0 g1 = j (i (g0 ⊙ g1)) ≡⟨ cong j (ic _ _) ⟩ - j (i g0 ∙ i g1) ≡⟨ jc _ _ ⟩ - j (i g0) ∘ j (i g1) ∎ +compMorph : ∀ {ℓ ℓ' ℓ''} {F : Group ℓ} {G : Group ℓ'} {H : Group ℓ''} (I : morph F G) (J : morph G H) → morph F H +morph.fun (compMorph I J) x = morph.fun J (morph.fun I x) +morph.ismorph (compMorph {F = F} {G = G} {H = H} I J) g0 g1 = + cong (morph.fun J) (morph.ismorph I g0 g1) + ∙ morph.ismorph J (morph.fun I g0) (morph.fun I g1) compIso : ∀ {ℓ} {F G H : Group ℓ} (I : Iso F G) (J : Iso G H) → Iso F H -compIso {ℓ} {F} {G} {H} - (iso F→G G→F fg gf) (iso G→H H→G gh hg ) = iso F→H H→F sec ret - where - F→H : morph F H - F→H = compMorph {ℓ} {F} {G} {H} F→G G→H - - H→F : morph H F - H→F = compMorph {ℓ} {H} {G} {F} H→G G→F - - open Group - - f→h : F .type → H .type - f→h = F→H .fst - - f→g : F .type → G .type - f→g = F→G .fst - - g→h : G .type → H .type - g→h = G→H .fst - - h→f : H .type → F .type - h→f = H→F .fst - - h→g : H .type → G .type - h→g = H→G .fst - - g→f : G .type → F .type - g→f = G→F .fst - - sec : I.section f→h h→f - sec h = f→h (h→f h) ≡⟨ cong g→h (fg (h→g h)) ⟩ - g→h (h→g h) ≡⟨ gh _ ⟩ - h ∎ - - ret : I.retract f→h h→f - ret f = h→f (f→h f) ≡⟨ cong g→f (hg (f→g f)) ⟩ - g→f (f→g f) ≡⟨ gf _ ⟩ - f ∎ +Iso.fun (compIso {ℓ} {F} {G} {H} iso1 iso2) = compMorph (Iso.fun iso1) (Iso.fun iso2) +Iso.inv (compIso {ℓ} {F} {G} {H} iso1 iso2) = compMorph (Iso.inv iso2) (Iso.inv iso1) +Iso.rightInv (compIso {ℓ} {F} {G} {H} iso1 iso2) b = + cong (morph.fun (Iso.fun iso2)) (Iso.rightInv iso1 _) ∙ Iso.rightInv iso2 _ +Iso.leftInv (compIso {ℓ} {F} {G} {H} iso1 iso2) b = + cong (morph.fun (Iso.inv iso1)) (Iso.leftInv iso2 _) ∙ Iso.leftInv iso1 _ idIso : ∀ {ℓ} (G : Group ℓ) → Iso G G -Iso.fun (idIso G) = (λ x → x) , λ _ _ → refl -Iso.inv (idIso G) = (λ x → x) , λ _ _ → refl +Iso.fun (idIso G) = mph (λ x → x) (λ _ _ → refl) +Iso.inv (idIso G) = mph (λ x → x) (λ _ _ → refl) Iso.rightInv (idIso G) _ = refl Iso.leftInv (idIso G) _ = refl -invIso : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → Iso G H → Iso H G -Iso.fun (invIso G H (iso _ inv _ _)) = inv -Iso.inv (invIso G H (iso fun _ _ _)) = fun -Iso.rightInv (invIso G H (iso _ _ _ leftInv)) = leftInv -Iso.leftInv (invIso G H (iso _ _ rightInv _)) = rightInv +invIso : ∀ {ℓ ℓ'} {G : Group ℓ} {H : Group ℓ'} → Iso G H → Iso H G +Iso.fun (invIso is) = Iso.inv is +Iso.inv (invIso is) = Iso.fun is +Iso.rightInv (invIso is) = Iso.leftInv is +Iso.leftInv (invIso is) = Iso.rightInv is + -Iso''→Iso : ∀ {ℓ ℓ'} (A : Group ℓ) (B : Group ℓ') → Iso'' A B → Iso A B -Iso''→Iso A B (iso'' ϕ inj surj) = +Iso''→Iso : ∀ {ℓ ℓ'} {A : Group ℓ} {B : Group ℓ'} → Iso'' A B → Iso A B +Iso''→Iso {A = A} {B = B} is = Iso'→Iso - (iso' - (I.iso - (fst ϕ) - (λ b → rec (helper b) (λ a → a) (surj b) .fst) - (λ b → rec (helper b) (λ a → a) (surj b) .snd) - λ b i → rec (helper (fst ϕ b)) (λ a → a) (propTruncIsProp (surj (fst ϕ b)) ∣ b , refl ∣ i) .fst) - (snd ϕ)) + (iso' (I.iso (morph.fun (Iso''.ϕ is)) + (λ b → rec (helper b) (λ a → a) (Iso''.surj is b) .fst) + (λ b → rec (helper b) (λ a → a) (Iso''.surj is b) .snd) + λ b i → rec (helper (morph.fun (Iso''.ϕ is) b)) (λ a → a) (propTruncIsProp (Iso''.surj is (morph.fun (Iso''.ϕ is) b)) ∣ b , refl ∣ i) .fst) + (morph.ismorph (Iso''.ϕ is))) where - helper : (b : _) → isProp (Σ (Group.type A) (λ a → ϕ .fst a ≡ b)) + helper : (b : _) → isProp (Σ (Group.type A) (λ a → (morph.fun (Iso''.ϕ is)) a ≡ b)) helper b (a1 , pf1) (a2 , pf2) = ΣProp≡ (λ _ → isOfHLevelPath' 1 (Group.setStruc B) _ _) fstId where + open Group + open isGroup + open morph + A' = groupStruc A + B' = groupStruc B + ϕ = morph.fun (Iso''.ϕ is) + ϕmf = morph.ismorph (Iso''.ϕ is) + fstIdHelper : isGroup.comp (Group.groupStruc A) a1 (isGroup.inv (Group.groupStruc A) a2) ≡ isGroup.id (Group.groupStruc A) fstIdHelper = - let -A = isGroup.inv (Group.groupStruc A) - -B = isGroup.inv (Group.groupStruc B) - rCancelB = isGroup.rCancel (Group.groupStruc B) - _+B_ = isGroup.comp (Group.groupStruc B) - in inj _ - (ϕ .snd a1 (-A a2) - ∙ (λ i → (ϕ .fst a1 +B (morphMinus A B (ϕ .fst) (ϕ .snd) a2 i) )) - ∙ cong (λ x → x +B (-B (ϕ .fst a2))) (pf1 ∙ sym pf2) - ∙ rCancelB (ϕ .fst a2)) + let -A = isGroup.inv (Group.groupStruc A) + -B = isGroup.inv (Group.groupStruc B) + rCancelB = isGroup.rCancel (Group.groupStruc B) + _+B_ = isGroup.comp (Group.groupStruc B) + in Iso''.inj is _ + (ϕmf a1 (inv A' a2) + ∙ cong (λ x → comp B' (ϕ a1) x) (morphMinus A B ϕ ϕmf a2) + ∙ cong (λ x → comp B' x (inv B' (ϕ a2))) (pf1 ∙ sym pf2) + ∙ rCancel B' (ϕ a2)) fstId : a1 ≡ a2 fstId = - let - = isGroup.inv (Group.groupStruc A) - id = isGroup.id (Group.groupStruc A) - assoc = isGroup.assoc (Group.groupStruc A) - rUnit = isGroup.rUnit (Group.groupStruc A) - lUnit = isGroup.lUnit (Group.groupStruc A) - lCancel = isGroup.lCancel (Group.groupStruc A) - assoc = isGroup.assoc (Group.groupStruc A) - _+_ = isGroup.comp (Group.groupStruc A) - in a1 ≡⟨ sym (rUnit a1) ⟩ - (a1 + id) ≡⟨ cong (λ x → a1 + x) (sym (lCancel a2)) ⟩ - (a1 + (- a2 + a2)) ≡⟨ sym (assoc a1 (- a2) a2) ⟩ - ((a1 + - a2) + a2) ≡⟨ cong (λ x → x + a2) fstIdHelper ⟩ - (id + a2) ≡⟨ lUnit a2 ⟩ - a2 ∎ - - --- Injectivity and surjectivity in terms of exact sequences -exactInjectiveL : ∀ {ℓ ℓ' ℓ''} (E : Group ℓ'') (F : Group ℓ'') (G : Group ℓ) (H : Group ℓ') - (ϕ : morph E F) (ψ : morph F G) (ξ : morph G H) - → isSurjective E F ϕ - → ((x : Group.type G) → isInKer G H (fst ξ) x → (isInIm F G (fst ψ) x)) - → isProp (Group.type E) - → isInjective G H ξ -exactInjectiveL (group E _ (group-struct idE _ _ _ _ _ _ _)) F G H - (ϕ , mfϕ) (ψ , mfψ) _ surjϕ ker⊂im isPropE x inker = - rec (Group.setStruc G _ _) - (λ {(f , id) → sym id ∙ cong ψ (isPropF f (isGroup.id (Group.groupStruc F))) ∙ morph0→0 F G ψ mfψ}) - (ker⊂im x inker) - where - isPropF : isProp (Group.type F) - isPropF a b = rec (Group.setStruc F _ _) - (λ {(c , id-c) → rec (Group.setStruc F _ _) - (λ {(d , id-d) → sym id-c ∙ cong ϕ (isPropE c d) ∙ id-d}) - (surjϕ b) }) - (surjϕ a) - - -exactSurjectiveR : ∀ {ℓ ℓ' ℓ''} (E : Group ℓ'') (F : Group ℓ'') (G : Group ℓ) (H : Group ℓ') - (ϕ : morph E F) (ψ : morph F G) (ξ : morph G H) - → isInjective G H ξ - → ((x : Group.type F) → (isInKer F G (fst ψ) x) → isInIm E F (fst ϕ) x) - → isProp (Group.type H) - → isSurjective E F ϕ -exactSurjectiveR E F (group G GSet (group-struct idG -G _+G_ lUnitG rUnitG assocG lCancelG rCancelG)) - (group H _ (group-struct idH _ _ _ _ _ _ _)) - (ϕ , mfϕ) (ψ , mfψ) (ξ , mfξ) isInjξ ker⊂im isPropH x = ker⊂im x (isPropF (ψ x) idG) + a1 ≡⟨ sym (rUnit A' a1) ⟩ + comp A' a1 (id A') ≡⟨ cong (λ x → comp A' a1 x) (sym (lCancel A' a2)) ⟩ + comp A' a1 (comp A' (inv A' a2) a2) ≡⟨ sym (assoc A' a1 (inv A' a2) a2) ⟩ + comp A' (comp A' a1 (inv A' a2)) a2 ≡⟨ cong (λ x → comp A' x a2) fstIdHelper ⟩ + comp A' (id A') a2 ≡⟨ lUnit A' a2 ⟩ + a2 ∎ + + +-- -- Injectivity and surjectivity in terms of exact sequences + +-- {- +-- exactInjectiveL : ∀ {ℓ ℓ' ℓ''} (E : Group ℓ'') (F : Group ℓ'') (G : Group ℓ) (H : Group ℓ') +-- (ϕ : morph E F) (ψ : morph F G) (ξ : morph G H) +-- → isSurjective E F ϕ +-- → ((x : Group.type G) → isInKer G H (fst ξ) x → (isInIm F G (fst ψ) x)) +-- → isProp (Group.type E) +-- → isInjective G H ξ +-- exactInjectiveL (group E _ (group-struct idE _ _ _ _ _ _ _)) F G H +-- (ϕ , mfϕ) (ψ , mfψ) _ surjϕ ker⊂im isPropE x inker = +-- rec (Group.setStruc G _ _) +-- (λ {(f , id) → sym id ∙ cong ψ (isPropF f (isGroup.id (Group.groupStruc F))) ∙ morph0→0 F G ψ mfψ}) +-- (ker⊂im x inker) +-- where +-- isPropF : isProp (Group.type F) +-- isPropF a b = rec (Group.setStruc F _ _) +-- (λ {(c , id-c) → rec (Group.setStruc F _ _) +-- (λ {(d , id-d) → sym id-c ∙ cong ϕ (isPropE c d) ∙ id-d}) +-- (surjϕ b) }) +-- (surjϕ a) + + +-- exactSurjectiveR : ∀ {ℓ ℓ' ℓ''} (E : Group ℓ'') (F : Group ℓ'') (G : Group ℓ) (H : Group ℓ') +-- (ϕ : morph E F) (ψ : morph F G) (ξ : morph G H) +-- → isInjective G H ξ +-- → ((x : Group.type F) → (isInKer F G (fst ψ) x) → isInIm E F (fst ϕ) x) +-- → isProp (Group.type H) +-- → isSurjective E F ϕ +-- exactSurjectiveR E F (group G GSet (group-struct idG -G _+G_ lUnitG rUnitG assocG lCancelG rCancelG)) +-- (group H _ (group-struct idH _ _ _ _ _ _ _)) +-- (ϕ , mfϕ) (ψ , mfψ) (ξ , mfξ) isInjξ ker⊂im isPropH x = ker⊂im x (isPropF (ψ x) idG) + +-- where +-- isPropF : isProp G +-- isPropF a b = sym (rUnitG a) +-- ∙ cong (a +G_) (sym (lCancelG b)) +-- ∙ sym (assocG a (-G b) b) +-- ∙ cong (λ x → x +G b) helper +-- ∙ lUnitG b +-- where +-- helper : a +G (-G b) ≡ idG +-- helper = isInjξ (a +G (-G b)) (isPropH (ξ (a +G -G b)) idH ) + +-- exactSurjectiveR' : ∀ {ℓ ℓ' ℓ''} (E : Group ℓ) (F : Group ℓ') (G : Group ℓ'') +-- (ϕ : morph E F) (ψ : morph F G) +-- → ((x : Group.type F) → (isInKer F G (fst ψ) x) → isInIm E F (fst ϕ) x) +-- → isProp (Group.type G) +-- → isSurjective E F ϕ +-- exactSurjectiveR' E F (group G GSet (group-struct idG -G _+G_ lUnitG rUnitG assocG lCancelG rCancelG)) +-- (ϕ , mfϕ) (ψ , mfψ) ker⊂im isPropG x = ker⊂im x (isPropG (ψ x) idG) +-- -} + + + - where - isPropF : isProp G - isPropF a b = sym (rUnitG a) - ∙ cong (a +G_) (sym (lCancelG b)) - ∙ sym (assocG a (-G b) b) - ∙ cong (λ x → x +G b) helper - ∙ lUnitG b - where - helper : a +G (-G b) ≡ idG - helper = isInjξ (a +G (-G b)) (isPropH (ξ (a +G -G b)) idH ) - -exactSurjectiveR' : ∀ {ℓ ℓ' ℓ''} (E : Group ℓ) (F : Group ℓ') (G : Group ℓ'') - (ϕ : morph E F) (ψ : morph F G) - → ((x : Group.type F) → (isInKer F G (fst ψ) x) → isInIm E F (fst ϕ) x) - → isProp (Group.type G) - → isSurjective E F ϕ -exactSurjectiveR' E F (group G GSet (group-struct idG -G _+G_ lUnitG rUnitG assocG lCancelG rCancelG)) - (ϕ , mfϕ) (ψ , mfψ) ker⊂im isPropG x = ker⊂im x (isPropG (ψ x) idG) diff --git a/Cubical/Data/Group/GroupLibrary.agda b/Cubical/Data/Group/GroupLibrary.agda index 933ca16b9..75ef82d0c 100644 --- a/Cubical/Data/Group/GroupLibrary.agda +++ b/Cubical/Data/Group/GroupLibrary.agda @@ -12,11 +12,12 @@ import Cubical.Foundations.Isomorphism as I import Cubical.Foundations.Equiv as E import Cubical.Foundations.Equiv.HalfAdjoint as HAE -open import Cubical.Data.Sigma hiding (_×_) +open import Cubical.Data.Sigma hiding (_×_ ; comp) open import Cubical.HITs.SetQuotients as sq -- The image of a morphism +{- imGroup : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → morph G H → Group (ℓ-max ℓ ℓ') @@ -77,7 +78,7 @@ isGroup.rCancel (Group.groupStruc (imGroup G H (ϕ , mf))) (h , _) = in ΣProp≡ (λ _ → propTruncIsProp) (rCancel h) - +-} -- kerGroup : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → morph G H -- → Group (ℓ-max ℓ ℓ') -- Group.type (kerGroup G H (ϕ , pf)) = Σ[ x ∈ Group.type G ] isInKer G H ϕ x @@ -156,42 +157,189 @@ record SES {ℓ ℓ' ℓ'' ℓ'''} (A : Group ℓ) (B : Group ℓ') (leftGr : Gr ϕ : morph A B Ker-ϕ⊂Im-left : (x : Group.type A) -- - → isInKer A B (fst ϕ) x - → isInIm leftGr A (fst left) x + → isInKer A B (morph.fun ϕ) x + → isInIm leftGr A (morph.fun left) x Ker-right⊂Im-ϕ : (x : Group.type B) -- - → isInKer B rightGr (fst right) x - → isInIm A B (fst ϕ) x -SES→Iso : ∀ {ℓ ℓ' ℓ'' ℓ'''} (A : Group ℓ) (B : Group ℓ') (leftGr : Group ℓ'') (rightGr : Group ℓ''') + → isInKer B rightGr (morph.fun right) x + → isInIm A B (morph.fun ϕ) x +SES→Iso : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group ℓ} {B : Group ℓ'} (leftGr : Group ℓ'') (rightGr : Group ℓ''') → SES A B leftGr rightGr → Iso A B -SES→Iso A B lGr rGr (ses PropL PropR (lfun , lmf) (rfun , rmf) (ϕ , mf) Ker-ϕ⊂Im-left Ker-right⊂Im-ϕ) = - Iso''→Iso _ _ - (iso'' - (ϕ , mf) - (λ a inker → rec (Group.setStruc A _ _) - (λ { (a , p) → sym p ∙ cong lfun (PropL a _) ∙ morph0→0 lGr A lfun lmf }) - (Ker-ϕ⊂Im-left a inker)) - λ a → Ker-right⊂Im-ϕ a (PropR _ _)) - +SES→Iso {A = A} {B = B} lGr rGr sess = + Iso''→Iso + (iso'' (SES.ϕ sess) + (λ a inker → rec (Group.setStruc A _ _) + (λ {(a , p) → sym p ∙ cong (morph.fun (SES.left sess)) (SES.isTrivialLeft sess a _) + ∙ morph0→0 lGr A (morph.fun (SES.left sess)) (morph.ismorph (SES.left sess))}) + (SES.Ker-ϕ⊂Im-left sess a inker)) + λ a → SES.Ker-right⊂Im-ϕ sess a (SES.isTrivialRight sess _ _)) {- direct products of groups -} dirProd : ∀ {ℓ ℓ'} (A : Group ℓ) (B : Group ℓ') → Group (ℓ-max ℓ ℓ') -Group.type (dirProd (group A _ _) (group B _ _)) = A × B -Group.setStruc (dirProd (group _ setA _) (group _ setB _)) = isOfHLevelProd 2 setA setB -isGroup.id (Group.groupStruc (dirProd (group _ _ strucA) (group _ _ strucB))) = - (isGroup.id strucA) , (isGroup.id strucB) -isGroup.inv (Group.groupStruc (dirProd (group _ _ strucA) (group _ _ strucB))) = - map (isGroup.inv strucA) (isGroup.inv strucB) -isGroup.comp (Group.groupStruc (dirProd (group _ _ strucA) (group _ _ strucB))) (a1 , b1) (a2 , b2) = - (isGroup.comp strucA a1 a2) , isGroup.comp strucB b1 b2 -isGroup.lUnit (Group.groupStruc (dirProd (group _ _ strucA) (group _ _ strucB))) (a , b) i = - (isGroup.lUnit strucA a i) , (isGroup.lUnit strucB b i) -isGroup.rUnit (Group.groupStruc (dirProd (group _ _ strucA) (group _ _ strucB))) (a , b) i = - (isGroup.rUnit strucA a i) , (isGroup.rUnit strucB b i) -isGroup.assoc (Group.groupStruc (dirProd (group _ _ strucA) (group _ _ strucB))) (a1 , b1) (a2 , b2) (a3 , b3) i = - (isGroup.assoc strucA a1 a2 a3 i) , (isGroup.assoc strucB b1 b2 b3 i) -isGroup.lCancel (Group.groupStruc (dirProd (group _ _ strucA) (group _ _ strucB))) (a , b) i = - (isGroup.lCancel strucA a i) , (isGroup.lCancel strucB b i) -isGroup.rCancel (Group.groupStruc (dirProd (group _ _ strucA) (group _ _ strucB))) (a , b) i = - (isGroup.rCancel strucA a i) , (isGroup.rCancel strucB b i) - +Group.type (dirProd A B) = Group.type A × Group.type B +Group.setStruc (dirProd A B) = isOfHLevelProd 2 (Group.setStruc A) (Group.setStruc B) +isGroup.id (Group.groupStruc (dirProd A B)) = + isGroup.id (Group.groupStruc A) , isGroup.id (Group.groupStruc B) +isGroup.inv (Group.groupStruc (dirProd A B)) (a1 , b1) = + (isGroup.inv (Group.groupStruc A) a1) , (isGroup.inv (Group.groupStruc B) b1) +isGroup.comp (Group.groupStruc (dirProd A B)) (a1 , b1) (a2 , b2) = + (isGroup.comp (Group.groupStruc A) a1 a2) , (isGroup.comp (Group.groupStruc B) b1 b2) +isGroup.lUnit (Group.groupStruc (dirProd A B)) (a1 , b1) i = + (isGroup.lUnit (Group.groupStruc A) a1 i) , (isGroup.lUnit (Group.groupStruc B) b1 i) +isGroup.rUnit (Group.groupStruc (dirProd A B)) (a1 , b1) i = + (isGroup.rUnit (Group.groupStruc A) a1 i) , (isGroup.rUnit (Group.groupStruc B) b1 i) +isGroup.assoc (Group.groupStruc (dirProd A B)) (a1 , b1) (a2 , b2) (a3 , b3) i = + (isGroup.assoc (Group.groupStruc A) a1 a2 a3 i) , (isGroup.assoc (Group.groupStruc B) b1 b2 b3 i) +isGroup.lCancel (Group.groupStruc (dirProd A B)) (a1 , b1) i = + (isGroup.lCancel (Group.groupStruc A) a1 i) , (isGroup.lCancel (Group.groupStruc B) b1 i) +isGroup.rCancel (Group.groupStruc (dirProd A B)) (a1 , b1) i = + (isGroup.rCancel (Group.groupStruc A) a1 i) , (isGroup.rCancel (Group.groupStruc B) b1 i) + + +dirProdIso : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group ℓ} {B : Group ℓ'} {C : Group ℓ''} {D : Group ℓ'''} + → Iso A C → Iso B D + → Iso (dirProd A B) (dirProd C D) +dirProdIso {A = A'} {B = B'} {C = C'} {D = D'} isoAC isoBD = + Iso'→Iso + (iso' + (I.iso + (λ a → (morph.fun (Iso.fun isoAC) (proj₁ a)) , (morph.fun (Iso.fun isoBD) (proj₂ a))) + (λ a → (morph.fun (Iso.inv isoAC) (proj₁ a)) , (morph.fun (Iso.inv isoBD) (proj₂ a))) + (λ a → ×≡ (Iso.rightInv isoAC (proj₁ a)) (Iso.rightInv isoBD (proj₂ a))) + λ a → ×≡ (Iso.leftInv isoAC (proj₁ a)) (Iso.leftInv isoBD (proj₂ a))) + λ {(_ , _) (_ , _) → ×≡ (morph.ismorph (Iso.fun isoAC) _ _) (morph.ismorph (Iso.fun isoBD) _ _)}) + + + +{- +Given the following diagram + a ↦ (a , 0) ψ ϕ + A --> A × A -------> B ---> C +If ψ is an isomorphism and ϕ is surjective with ker ϕ ≡ {ψ (a , a) ∣ a ∈ A}, then C ≃ B +-} + +diagonalIso1 : ∀ {ℓ ℓ' ℓ''} (A : Group ℓ) (B : Group ℓ') (C : Group ℓ'') + (ψ : Iso (dirProd A A) B) (ϕ : morph B C) + → isSurjective _ _ ϕ + → ((x : Group.type B) → isInKer B C (morph.fun ϕ) x + → ∃[ y ∈ Group.type A ] x ≡ morph.fun (Iso.fun ψ) (y , y)) + → ((x : Group.type B) → (∃[ y ∈ Group.type A ] x ≡ morph.fun (Iso.fun ψ) (y , y)) + → isInKer B C (morph.fun ϕ) x) + → Iso A C +diagonalIso1 A' B' C' ψ' ϕ' issurj ker→diag diag→ker = + Iso''→Iso + (iso'' (compMorph fstProj (compMorph (Iso.fun ψ') ϕ')) + (λ a inker + → rec (Group.setStruc A' _ _) + (λ {(a' , id') → cong proj₁ (sym (Iso.leftInv ψ' (a , id A)) + ∙ (cong (morph.fun (Iso.inv ψ')) id') + ∙ Iso.leftInv ψ' (a' , a')) + ∙ cong proj₂ (sym ((sym (Iso.leftInv ψ' (a , id A)) + ∙ (cong (morph.fun (Iso.inv ψ')) id') + ∙ Iso.leftInv ψ' (a' , a'))))}) + (ker→diag (ψ (a , id A)) inker)) + λ c → rec propTruncIsProp + (λ { (b , id') → ∣ comp A (proj₁ (ψ⁻ b)) (inv A (proj₂ (ψ⁻ b))) + , sym (rUnit C _) + ∙ cong (comp C _) (sym (diag→ker (ψ (proj₂ (ψ⁻ b) , proj₂ (ψ⁻ b))) + ∣ proj₂ (ψ⁻ b) , refl ∣)) + ∙ sym (ϕmf _ _) + ∙ cong ϕ (sym (Iso.rightInv ψ' _)) + ∙ cong (λ x → ϕ (ψ x)) (morph.ismorph (Iso.inv ψ') _ _) + ∙ cong (λ x → ϕ (ψ x)) (cong₂ (comp (groupStruc (dirProd A' A'))) + (Iso.leftInv ψ' _) + (Iso.leftInv ψ' _)) + ∙ cong (λ x → ϕ (ψ x)) (×≡ {a = _ , _} {b = _ , _} + (assoc A _ _ _) + (lUnit A _)) + ∙ cong (λ x → ϕ (ψ x)) (×≡ {a = _ , _} {b = _ , _} + (cong (comp A (proj₁ (ψ⁻ b))) (lCancel A _)) + refl) + ∙ cong (λ x → ϕ (ψ x)) (×≡ {a = _ , _} {b = _ , _} + (rUnit A _) + refl) + ∙ cong (λ x → ϕ (ψ x)) (sym (×-η (ψ⁻ b))) + ∙ cong ϕ (Iso.rightInv ψ' b) + ∙ id' ∣ }) + (issurj c)) + where + open Group + open isGroup + A = groupStruc A' + B = groupStruc B' + C = groupStruc C' + ϕ = morph.fun ϕ' + ϕmf = morph.ismorph ϕ' + ψ = morph.fun (Iso.fun ψ') + ψ⁻ = morph.fun (Iso.inv ψ') + + fstProj : morph A' (dirProd A' A') + morph.fun fstProj = λ a → a , (id A) + morph.ismorph fstProj = λ g0 g1 i → comp A g0 g1 , lUnit A (id A) (~ i) + +{- +Given the following diagram + + ϕ +A × A ⇉ B + ^ + | + | a ↦ (a , 0) + | + A ≃ C + +If ϕ is surjective with ker ϕ ≡ {(a , a) ∣ a ∈ A}, then C ≃ B +-} + +diagonalIso : ∀ {ℓ ℓ' ℓ''} (A : Group ℓ) (B : Group ℓ') (C : Group ℓ'') + → (ϕ : morph (dirProd A A) B) + → ((x : Group.type (dirProd A A)) → isInKer (dirProd A A) B (morph.fun ϕ) x → ∃[ y ∈ Group.type A ] x ≡ (y , y)) + → ((x : Group.type (dirProd A A)) → ∃[ y ∈ Group.type A ] x ≡ (y , y) → isInKer (dirProd A A) B (morph.fun ϕ) x) + → isSurjective (dirProd A A) B ϕ + → Iso C A + → Iso C B +diagonalIso A' B' C' ϕ' diagKer1 diagKer2 issurj I = + Iso''→Iso + (iso'' (compMorph (compMorph (Iso.fun I) fstProj) ϕ') + inj + surj) + where + open Group + open isGroup + A = groupStruc A' + B = groupStruc B' + C = groupStruc C' + ϕ = morph.fun ϕ' + ϕmf = morph.ismorph ϕ' + C→A = morph.fun (Iso.fun I) + C→A-mf = morph.ismorph (Iso.fun I) + + fstProj : morph A' (dirProd A' A') + morph.fun fstProj = λ a → a , (id A) + morph.ismorph fstProj = λ g0 g1 i → comp A g0 g1 , lUnit A (id A) (~ i) + + inj : (a : Group.type C') → isInKer C' B' (λ x → ϕ (morph.fun fstProj (C→A x))) a + → a ≡ id C + inj a inker = rec (Group.setStruc C' _ _) + (λ {(b , id') → sym (Iso.leftInv I a) + ∙ cong (morph.fun (Iso.inv I)) (cong proj₁ id') + ∙ cong (morph.fun (Iso.inv I)) (cong proj₂ (sym id')) + ∙ morph0→0 A' C' (morph.fun (Iso.inv I)) (morph.ismorph (Iso.inv I))}) + (diagKer1 (C→A a , isGroup.id A) inker) + + surj : (b : Group.type B') → ∃[ c ∈ Group.type C' ] ϕ (morph.fun fstProj (C→A c)) ≡ b + surj b = rec propTruncIsProp + (λ {(a , id') → ∣ morph.fun (Iso.inv I) a , cong ϕ (×≡ {a = _ , _} {b = _ , _} (Iso.rightInv I a) refl) ∙ id' ∣}) + (idhelper b) + where + idhelper : (b : Group.type B') → ∃[ a ∈ Group.type A' ] ϕ (a , (id A)) ≡ b + idhelper b = rec propTruncIsProp + (λ {((a1 , a2) , id') → ∣ (comp A a1 (inv A a2)) + , sym (rUnit B (ϕ (comp A a1 (inv A a2) , id A))) + ∙ cong (comp B _) (sym (diagKer2 (a2 , a2) ∣ a2 , refl ∣)) + ∙ sym (ϕmf _ _) + ∙ cong ϕ (×≡ {a = _ , _} {b = _ , _} + (assoc A a1 (inv A a2) a2 ∙ cong (comp A a1) (lCancel A a2) ∙ rUnit A a1) + (lUnit A a2)) + ∙ id' ∣}) + (issurj b) diff --git a/Cubical/Data/Prod/Base.agda b/Cubical/Data/Prod/Base.agda index 0e8c5cd93..094f4e649 100644 --- a/Cubical/Data/Prod/Base.agda +++ b/Cubical/Data/Prod/Base.agda @@ -45,3 +45,7 @@ intro f g a = f a , g a map : {B : Type ℓ} {D : B → Type ℓ'} → (∀ a → C a) → (∀ b → D b) → (x : A × B) → C (proj₁ x) × D (proj₂ x) map f g = intro (f ∘ proj₁) (g ∘ proj₂) + + +×-η : {A : Type ℓ} {B : Type ℓ'} (x : A × B) → x ≡ ((proj₁ x) , (proj₂ x)) +×-η (x , x₁) = refl diff --git a/Cubical/Data/Prod/Properties.agda b/Cubical/Data/Prod/Properties.agda index 8798a0c3e..af78ac624 100644 --- a/Cubical/Data/Prod/Properties.agda +++ b/Cubical/Data/Prod/Properties.agda @@ -20,6 +20,9 @@ private -- Swapping is an equivalence +×≡ : {a b : A × B} → proj₁ a ≡ proj₁ b → proj₂ a ≡ proj₂ b → a ≡ b +×≡ {a = (a1 , b1)} {b = (a2 , b2)} id1 id2 i = (id1 i) , (id2 i) + swap : A × B → B × A swap (x , y) = (y , x) diff --git a/Cubical/HITs/S1/Base.agda b/Cubical/HITs/S1/Base.agda index c9103a2fc..04f6e23f2 100644 --- a/Cubical/HITs/S1/Base.agda +++ b/Cubical/HITs/S1/Base.agda @@ -8,9 +8,11 @@ module Cubical.HITs.S1.Base where open import Cubical.Foundations.Prelude open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv open import Cubical.Foundations.Function open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Transport open import Cubical.Foundations.Univalence open import Cubical.Data.Nat @@ -30,6 +32,14 @@ module _ where comp (λ _ → S¹) u (outS u0) ≡ hcomp u (outS u0) compS1 φ u u0 = refl + + +toPropElim : ∀ {ℓ} {B : S¹ → Type ℓ} → ((s : S¹) → isProp (B s)) → B base → (s : S¹) → B s +toPropElim {B = B} isprop b base = b +toPropElim {B = B} isprop b (loop i) = hcomp (λ k → λ {(i = i0) → b + ; (i = i1) → isprop base (subst B (loop) b) b k}) + (transp (λ j → B (loop (i ∧ j))) (~ i) b) + helix : S¹ → Type₀ helix base = Int helix (loop i) = sucPathInt i @@ -49,6 +59,7 @@ intLoop (pos (suc n)) = intLoop (pos n) ∙ loop intLoop (negsuc zero) = sym loop intLoop (negsuc (suc n)) = intLoop (negsuc n) ∙ sym loop + decodeSquare : (n : Int) → PathP (λ i → base ≡ loop i) (intLoop (predInt n)) (intLoop n) decodeSquare (pos zero) i j = loop (i ∨ ~ j) decodeSquare (pos (suc n)) i j = hfill (λ k → λ { (j = i0) → base @@ -72,6 +83,7 @@ decode (loop i) y j = decodeEncode : (x : S¹) (p : base ≡ x) → decode x (encode x p) ≡ p decodeEncode x p = J (λ y q → decode y (encode y q) ≡ q) (λ _ → refl) p + isSetΩS¹ : isSet ΩS¹ isSetΩS¹ p q r s j i = hcomp (λ k → λ { (i = i0) → decodeEncode base p k @@ -80,6 +92,12 @@ isSetΩS¹ p q r s j i = ; (j = i1) → decodeEncode base (s i) k }) (decode base (isSetInt (winding p) (winding q) (cong winding r) (cong winding s) j i)) + +isSetΩx : (x : S¹) → isSet (x ≡ x) +isSetΩx = toPropElim + (λ s → isPropIsSet) + isSetΩS¹ + -- This proof does not rely on rewriting hcomp with empty systems in -- Int as ghcomp has been implemented! windingIntLoop : (n : Int) → winding (intLoop n) ≡ n @@ -127,6 +145,14 @@ winding-hom a b i = ; (i = i1) → windingIntLoop (winding a + winding b) t }) (winding (intLoop-hom (winding a) (winding b) i)) +-- Commutativity +comm-ΩS¹ : (p q : ΩS¹) → p ∙ q ≡ q ∙ p +comm-ΩS¹ p q = sym (cong₂ (_∙_) (decodeEncode base p) (decodeEncode base q)) + ∙ intLoop-hom (winding p) (winding q) + ∙ cong intLoop (+-comm (winding p) (winding q)) + ∙ sym (intLoop-hom (winding q) (winding p)) + ∙ (cong₂ (_∙_) (decodeEncode base q) (decodeEncode base p)) + -- Based homotopy group basedΩS¹ : (x : S¹) → Type₀ @@ -154,6 +180,8 @@ private basedΩS¹→ΩS¹ : (i : I) → basedΩS¹ (loop i) → ΩS¹ basedΩS¹→ΩS¹ i x j = basedΩS¹→ΩS¹-filler i1 i x j + + basedΩS¹→ΩS¹→basedΩS¹ : (i : I) → (x : basedΩS¹ (loop i)) → ΩS¹→basedΩS¹ i (basedΩS¹→ΩS¹ i x) ≡ x basedΩS¹→ΩS¹→basedΩS¹ i x j k = @@ -241,6 +269,48 @@ basedΩS¹≡Int : (x : S¹) → basedΩS¹ x ≡ Int basedΩS¹≡Int x = (basedΩS¹≡ΩS¹ x) ∙ ΩS¹≡Int + + +---- A shorter proof of the same thing ----- +basechange2 : (x : S¹) → ΩS¹ → (x ≡ x) +basechange2 base p = p +basechange2 (loop i) p = + hcomp (λ k → λ { (i = i0) → lUnit (rUnit p (~ k)) (~ k) + ; (i = i1) → (cong ((sym loop) ∙_) (comm-ΩS¹ p loop) + ∙ assoc (sym loop) loop p + ∙ cong (_∙ p) (lCancel loop) + ∙ sym (lUnit _)) k }) + ((λ j → loop (~ j ∧ i)) ∙ p ∙ λ j → loop (j ∧ i)) + +basechange2⁻ : (x : S¹) → (x ≡ x) → ΩS¹ +basechange2⁻ base p = p +basechange2⁻ (loop i) p = + hcomp (λ k → λ { (i = i0) → lUnit (rUnit p (~ k)) (~ k) + ; (i = i1) → (cong (loop ∙_) (comm-ΩS¹ p (sym loop)) + ∙ assoc loop (sym loop) p + ∙ cong (_∙ p) (rCancel loop) + ∙ sym (lUnit _)) k }) + ((λ j → loop (i ∧ j)) ∙ p ∙ λ j → loop (i ∧ (~ j))) + +basechange2-sect : (x : S¹) → section (basechange2 x) (basechange2⁻ x) +basechange2-sect = + toPropElim (λ _ → isOfHLevelΠ 1 λ _ → isSetΩx _ _ _ ) + λ _ → refl + +basechange2-retr : (x : S¹) → retract (basechange2 x) (basechange2⁻ x) +basechange2-retr = + toPropElim (λ s → isOfHLevelΠ 1 λ x → isSetΩx _ _ _) + λ _ → refl + + +-- baschange2⁻ is a morphism + +basechange2⁻-morph : (x : S¹) (p q : x ≡ x) → + basechange2⁻ x (p ∙ q) ≡ basechange2⁻ x p ∙ basechange2⁻ x q +basechange2⁻-morph = toPropElim {B = λ x → (p q : x ≡ x) → basechange2⁻ x (p ∙ q) ≡ basechange2⁻ x p ∙ basechange2⁻ x q} + (λ x → isOfHLevelΠ 1 λ p → isOfHLevelΠ 1 λ q → isSetΩS¹ _ _ ) + λ _ _ → refl + -- Some tests module _ where private @@ -376,3 +446,6 @@ rotInv-4 base base i = base rotInv-4 base (loop k) i = rotInv-aux-4 i0 k (~ i) i1 rotInv-4 (loop j) base i = loop j rotInv-4 (loop j) (loop k) i = rotInv-aux-4 j k (~ i) i1 + + + diff --git a/Cubical/HITs/S1/Properties.agda b/Cubical/HITs/S1/Properties.agda index 2813e905b..3feb2a306 100644 --- a/Cubical/HITs/S1/Properties.agda +++ b/Cubical/HITs/S1/Properties.agda @@ -30,9 +30,3 @@ isGroupoidS¹ s t = (isConnectedS¹ t))) (isConnectedS¹ s) - -toPropElim : ∀ {ℓ} {B : S¹ → Type ℓ} → ((s : S¹) → isProp (B s)) → B base → (s : S¹) → B s -toPropElim {B = B} isprop b base = b -toPropElim {B = B} isprop b (loop i) = hcomp (λ k → λ {(i = i0) → b - ; (i = i1) → isprop base (subst B (loop) b) b k}) - (transp (λ j → B (loop (i ∧ j))) (~ i) b) diff --git a/Cubical/HITs/Sn/Properties.agda b/Cubical/HITs/Sn/Properties.agda index de7d0b143..7f10e899b 100644 --- a/Cubical/HITs/Sn/Properties.agda +++ b/Cubical/HITs/Sn/Properties.agda @@ -50,20 +50,26 @@ S1→SuspBool south = south S1→SuspBool (merid north i) = merid true i S1→SuspBool (merid south i) = merid false i +SuspBool→S1-sect : section SuspBool→S1 S1→SuspBool +SuspBool→S1-sect north = refl +SuspBool→S1-sect south = refl +SuspBool→S1-sect (merid north i) = refl +SuspBool→S1-sect (merid south i) = refl + +SuspBool→S1-retr : retract SuspBool→S1 S1→SuspBool +SuspBool→S1-retr north = refl +SuspBool→S1-retr south = refl +SuspBool→S1-retr (merid false i) = refl +SuspBool→S1-retr (merid true i) = refl + SuspBool≃S1 : Susp Bool ≃ S₊ 1 SuspBool≃S1 = isoToEquiv iso1 where iso1 : Iso (Susp Bool) (S₊ 1) Iso.fun iso1 = SuspBool→S1 Iso.inv iso1 = S1→SuspBool - Iso.rightInv iso1 north = refl - Iso.rightInv iso1 south = refl - Iso.rightInv iso1 (merid north i) = refl - Iso.rightInv iso1 (merid south i) = refl - Iso.leftInv iso1 north = refl - Iso.leftInv iso1 south = refl - Iso.leftInv iso1 (merid false i) = refl - Iso.leftInv iso1 (merid true i) = refl + Iso.rightInv iso1 = SuspBool→S1-sect + Iso.leftInv iso1 = SuspBool→S1-retr -- map between S¹ ∧ A and Susp A. diff --git a/Cubical/ZCohomology/MayerVietorisReduced.agda b/Cubical/ZCohomology/MayerVietorisReduced.agda index 28270493c..580f79b21 100644 --- a/Cubical/ZCohomology/MayerVietorisReduced.agda +++ b/Cubical/ZCohomology/MayerVietorisReduced.agda @@ -356,12 +356,12 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : d-morph : ∀ {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) (n : ℕ) → morph (coHomGr n C) (coHomGr (suc n) (Pushout f g)) -d-morph A B C f g n = (MV.d A B C f g n) , (MV.dIsHom A B C f g n) +d-morph A B C f g n = mph (MV.d A B C f g n) (MV.dIsHom A B C f g n) i-morph : ∀ {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) (n : ℕ) → morph (coHomGr n (Pushout f g)) (×coHomGr n A B) -i-morph A B C f g n = (MV.i A B C f g n) , (MV.iIsHom A B C f g n) +i-morph A B C f g n = mph (MV.i A B C f g n) (MV.iIsHom A B C f g n) Δ-morph : ∀ {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) (n : ℕ) → morph (×coHomGr n A B) (coHomGr n C) -Δ-morph A B C f g n = MV.Δ A B C f g n , MV.ΔIsHom A B C f g n +Δ-morph A B C f g n = mph (MV.Δ A B C f g n) (MV.ΔIsHom A B C f g n) diff --git a/Cubical/ZCohomology/Sn/Sn.agda b/Cubical/ZCohomology/Sn/Sn.agda index 0ae8b5dee..a190b3300 100644 --- a/Cubical/ZCohomology/Sn/Sn.agda +++ b/Cubical/ZCohomology/Sn/Sn.agda @@ -36,6 +36,7 @@ open import Cubical.Data.Unit open import Cubical.Data.Group.Base renaming (Iso to grIso ; compIso to compGrIso ; invIso to invGrIso ; idIso to idGrIso) open import Cubical.Data.Group.GroupLibrary open import Cubical.ZCohomology.coHomZero-map +open import Cubical.HITs.Wedge open import Cubical.ZCohomology.KcompPrelims @@ -47,6 +48,24 @@ open import Cubical.Data.HomotopyGroup open import Cubical.Data.Bool hiding (_⊕_) +toProdElimSuspRec : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Susp A → Type ℓ'} (a : A) + → ((x : Susp A) → isProp (B x)) + → B north + → (x : Susp A) → B x +toProdElimSuspRec a isProp Bnorth north = Bnorth +toProdElimSuspRec {B = B} a isProp Bnorth south = subst B (merid a) Bnorth +toProdElimSuspRec {B = B} a isProp Bnorth (merid a₁ i) = + hcomp (λ k → λ {(i = i0) → Bnorth ; + (i = i1) → isProp south (subst B (merid a₁) Bnorth) (subst B (merid a) Bnorth) k}) + (transp (λ j → B (merid a₁ (j ∧ i))) (~ i) Bnorth) + +toProdElimSuspElim2 : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Susp A → Susp A → Type ℓ'} (a : A) + → ((x y : Susp A) → isProp (B x y)) + → B north north + → (x y : Susp A) → B x y +toProdElimSuspElim2 a isProp Bnorth = + toProdElimSuspRec a (λ x → isOfHLevelΠ 1 λ y → isProp x y) + (toProdElimSuspRec a (λ x → isProp north x) Bnorth) cong₂Funct2 : ∀ {ℓ ℓ'} {A : Type ℓ} {x y : A} {B : Type ℓ'} (f : A → A → B) → (p : x ≡ y) → @@ -87,110 +106,45 @@ abstract prodId : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A × B} → proj₁ x ≡ proj₁ y → proj₂ x ≡ proj₂ y → x ≡ y prodId {x = (a , b)} {y = (c , d)} id1 id2 i = (id1 i) , (id2 i) -indIntGroup : ∀ {ℓ} {G : Group ℓ} → (ϕ : Int → Group.type G) - → ϕ 0 ≡ isGroup.id (Group.groupStruc G) - → ((a : Int) → ϕ (a +ℤ 1) ≡ isGroup.comp (Group.groupStruc G) (ϕ a) (ϕ 1)) - → ((n : Int) → ϕ (predInt n) ≡ isGroup.comp (Group.groupStruc G) (ϕ n) (ϕ (negsuc zero))) - → isMorph intGroup G ϕ -indIntGroup {G = group G gSet (group-struct _ _ _+G_ _ rUnit₁ _ _ _)} ϕ zeroId _ _ n (pos zero) = - sym (rUnit₁ (ϕ n)) ∙ cong (λ x → ϕ n +G x) (sym zeroId) -indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId oneId minOneId n (pos (suc m)) = - (λ i → ϕ ((n +pos m) +ℤ 1)) - ∙ oneId (n +pos m) - ∙ cong (λ x → x +G ϕ (pos 1)) - (indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} - ϕ zeroId oneId minOneId n (pos m)) - ∙ assoc₁ (ϕ n) (ϕ (pos m)) (ϕ (pos 1)) - ∙ sym (cong (λ x → ϕ n +G x) (oneId (pos m))) -indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId _ minOneId n (negsuc zero) = minOneId n -indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId a minOneId n (negsuc (suc m)) = - (λ i → ϕ ((n +negsuc m) +ℤ (negsuc zero))) - ∙ minOneId (n +negsuc m) - ∙ cong (λ x → x +G ϕ (negsuc zero)) (indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId a minOneId n (negsuc m)) - ∙ assoc₁ (ϕ n) (ϕ (negsuc m)) (ϕ (negsuc zero)) - ∙ cong (λ x → ϕ n +G x) (sym (minOneId (negsuc m))) - -pushoutSn : (n : ℕ) → Iso (S₊ (suc n)) (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt) -Iso.fun (pushoutSn n) north = inl tt -Iso.fun (pushoutSn n) south = inr tt -Iso.fun (pushoutSn n) (merid a i) = push a i -Iso.inv (pushoutSn n) (inl x) = north -Iso.inv (pushoutSn n) (inr x) = south -Iso.inv (pushoutSn n) (push a i) = merid a i -Iso.rightInv (pushoutSn n) (inl x) = refl -Iso.rightInv (pushoutSn n) (inr x) = refl -Iso.rightInv (pushoutSn n) (push a i) = refl -Iso.leftInv (pushoutSn n) north = refl -Iso.leftInv (pushoutSn n) south = refl -Iso.leftInv (pushoutSn n) (merid a i) = refl - -Sn≡Pushout : (n : ℕ) → (S₊ (suc n)) ≡ (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt) -Sn≡Pushout n = isoToPath (pushoutSn n) - - -coHomPushout≡coHomSn : (n m : ℕ) → grIso (coHomGr m (S₊ (suc n))) (coHomGr m (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt)) -grIso.fun (coHomPushout≡coHomSn n m) = (sRec setTruncIsSet λ f → ∣ (λ {(inl x) → f north ; (inr x) → f south ; (push a i) → f (merid a i)}) ∣₀) - , sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) - λ a b → cong ∣_∣₀ (funExt λ {(inl x) → refl ; (inr x) → refl ; (push a i) → refl }) -grIso.inv (coHomPushout≡coHomSn n m) = sRec setTruncIsSet (λ f → ∣ (λ {north → f (inl tt) ; south → f (inr tt) ; (merid a i) → f (push a i)}) ∣₀) - , sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) - λ a b → cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → refl }) -grIso.rightInv (coHomPushout≡coHomSn n m) = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) - λ f → cong ∣_∣₀ (funExt λ {(inl x) → refl ; (inr x) → refl ; (push a i) → refl }) -grIso.leftInv (coHomPushout≡coHomSn n m) = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) - λ f → cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → refl }) - -isContr→≡Unit : {A : Type₀} → isContr A → A ≡ Unit -isContr→≡Unit contr = isoToPath (iso (λ _ → tt) (λ _ → fst contr) (λ _ → refl) λ _ → snd contr _) - -isContr→isContrTrunc : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → isContr A → isContr (hLevelTrunc n A) -isContr→isContrTrunc n contr = ∣ fst contr ∣ , (trElim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → cong ∣_∣ (snd contr a)) -isContr→isContrSetTrunc : ∀ {ℓ} {A : Type ℓ} → isContr A → isContr (∥ A ∥₀) -isContr→isContrSetTrunc contr = ∣ fst contr ∣₀ , sElim (λ _ → isOfHLevelPath 2 (setTruncIsSet) _ _) λ a → cong ∣_∣₀ (snd contr a) - -coHomGrUnit0 : grIso (coHomGr 0 Unit) intGroup -grIso.fun coHomGrUnit0 = sRec isSetInt (λ f → f tt) , - sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) λ a b → addLemma (a tt) (b tt) -grIso.inv coHomGrUnit0 = (λ a → ∣ (λ _ → a) ∣₀) , λ a b i → ∣ (λ _ → addLemma a b (~ i)) ∣₀ -grIso.rightInv coHomGrUnit0 a = refl -grIso.leftInv coHomGrUnit0 = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ a → refl - -isContrCohomUnit : (n : ℕ) → isContr (coHom (suc n) Unit) -isContrCohomUnit n = subst isContr (λ i → ∥ UnitToTypeId (coHomK (suc n)) (~ i) ∥₀) (helper n) - where - helper : (n : ℕ) → isContr (∥ coHomK (suc n) ∥₀) - helper n = subst isContr ((isoToPath (truncOfTruncIso {A = S₊ (1 + n)} 2 (1 + n))) ∙ sym propTrunc≡Trunc0 ∙ λ i → ∥ hLevelTrunc (suc (+-comm n 2 i)) (S₊ (1 + n)) ∥₀) - (isConnectedSubtr 2 (helper2 n .fst) (subst (λ x → isHLevelConnected x (S₊ (suc n))) (sym (helper2 n .snd)) (sphereConnected (suc n))) ) - where - helper2 : (n : ℕ) → Σ[ m ∈ ℕ ] m + 2 ≡ 2 + n - helper2 zero = 0 , refl - helper2 (suc n) = (suc n) , λ i → suc (+-comm n 2 i) - -coHomGrUnit≥1 : (n : ℕ) → grIso (coHomGr (suc n) Unit) trivialGroup -grIso.fun (coHomGrUnit≥1 n) = (λ _ → tt) , (λ _ _ → refl) -grIso.inv (coHomGrUnit≥1 n) = (λ _ → ∣ (λ _ → ∣ north ∣) ∣₀) , λ _ _ → sym (rUnitₕ 0ₕ) -grIso.rightInv (coHomGrUnit≥1 n) a = refl -grIso.leftInv (coHomGrUnit≥1 n) a = sym (isContrCohomUnit n .snd 0ₕ) ∙ isContrCohomUnit n .snd a - -S0→Int : (a : Int × Int) → S₊ 0 → Int -S0→Int a north = proj₁ a -S0→Int a south = proj₂ a - -coHom0-S0 : grIso (coHomGr 0 (S₊ 0)) (dirProd intGroup intGroup) -coHom0-S0 = - Iso'→Iso - (iso' - (iso (sRec (isOfHLevelProd 2 isSetInt isSetInt) - λ f → (f north) , (f south)) - (λ a → ∣ S0→Int a ∣₀) - (λ { (a , b) → refl}) - (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ f → cong ∣_∣₀ (funExt (λ {north → refl ; south → refl})))) - (sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 isSetInt isSetInt) _ _) - λ a b i → addLemma (a north) (b north) i , addLemma (a south) (b south) i)) -×morph : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group ℓ} {B : Group ℓ'} {C : Group ℓ''} {D : Group ℓ'''} → morph A B → morph C D → morph (dirProd A C) (dirProd B D) -×morph mf1 mf2 = (λ {(a , b) → (mf1 .fst a) , mf2 .fst b}) , λ {(a , b) (c , d) i → mf1 .snd a c i , mf2 .snd b d i} +basechange2⁻Gen : (x y : S¹) → x ≡ y → x ≡ y → ΩS¹ +basechange2⁻Gen x y = J (λ y p → x ≡ y → ΩS¹) (basechange2⁻ x) + +basechange2⁻Gen-morphIsh : (x y z : S¹) → (P : x ≡ y) (Q : y ≡ z) (p : x ≡ y) (q : y ≡ z) → basechange2⁻Gen x z (P ∙ Q) (p ∙ q) ≡ basechange2⁻Gen x y P p ∙ basechange2⁻Gen y z Q q +basechange2⁻Gen-morphIsh x y z = + J (λ y P → (Q : y ≡ z) (p : x ≡ y) (q : y ≡ z) → basechange2⁻Gen x z (P ∙ Q) (p ∙ q) ≡ basechange2⁻Gen x y P p ∙ basechange2⁻Gen y z Q q) + (J (λ z Q → (p : x ≡ x) (q : x ≡ z) → basechange2⁻Gen x z (refl ∙ Q) (p ∙ q) ≡ basechange2⁻Gen x x refl p ∙ basechange2⁻Gen x z Q q) + λ p q → cong (λ y → basechange2⁻Gen x x y (p ∙ q)) (sym (lUnit refl)) + ∙ (λ i → JRefl (λ y p → x ≡ y → ΩS¹) (basechange2⁻ x) i (p ∙ q)) + ∙ basechange2⁻-morph x p q + ∙ λ i → JRefl (λ y p → x ≡ y → ΩS¹) (basechange2⁻ x) (~ i) p ∙ JRefl (λ y p → x ≡ y → ΩS¹) (basechange2⁻ x) (~ i) q) + +S1→S¹ : S₊ 1 → S¹ +S1→S¹ x = SuspBool→S¹ (S1→SuspBool x) +S¹→S1 : S¹ → S₊ 1 +S¹→S1 x = SuspBool→S1 (S¹→SuspBool x) + +S1→S¹-sect : section S1→S¹ S¹→S1 +S1→S¹-sect x = + cong SuspBool→S¹ (SuspBool→S1-retr (S¹→SuspBool x)) + ∙ S¹→SuspBool→S¹ x + +S1→S¹-retr : retract S1→S¹ S¹→S1 +S1→S¹-retr x = + cong SuspBool→S1 (SuspBool→S¹→SuspBool (S1→SuspBool x)) + ∙ SuspBool→S1-sect x + + +prodElim : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : ∥ A ∥₀ × ∥ B ∥₀ → Type ℓ''} + → ((x : ∥ A ∥₀ × ∥ B ∥₀) → isOfHLevel 2 (C x)) + → ((a : A) (b : B) → C (∣ a ∣₀ , ∣ b ∣₀)) + → (x : ∥ A ∥₀ × ∥ B ∥₀) → C x +prodElim {A = A} {B = B} {C = C} hlevel ind (a , b) = schonf a b + where + schonf : (a : ∥ A ∥₀) (b : ∥ B ∥₀) → C (a , b) + schonf = sElim (λ x → isOfHLevelΠ 2 λ y → hlevel (_ , _)) λ a → sElim (λ x → hlevel (_ , _)) + λ b → ind a b isGroupoidS1 : isGroupoid (S₊ 1) @@ -200,806 +154,1217 @@ isConnectedS1 x = pRec propTruncIsProp (λ p → ∣ cong (transport (sym (S¹≡S1))) p ∙ transport⁻Transport (S¹≡S1) x ∣₋₁) (isConnectedS¹ (transport S¹≡S1 x)) --- basechange* : (x y : S¹) → x ≡ y → x ≡ y → ΩS¹ --- basechange* x y = J (λ y p → (x ≡ y) → ΩS¹) (basechange x) +private + S¹map : hLevelTrunc 3 S¹ → S¹ + S¹map = trElim (λ _ → isGroupoidS¹) (idfun S¹) + + S¹map-id : (x : hLevelTrunc 3 S¹) → Path (hLevelTrunc 3 S¹) ∣ S¹map x ∣ x + S¹map-id = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → refl + + S1map : hLevelTrunc 3 (S₊ 1) → (S₊ 1) + S1map = trElim (λ _ → isGroupoidS1) (idfun _) + +S¹→S¹ : Iso (S¹ → hLevelTrunc 3 S¹) (S¹ × Int) +Iso.fun S¹→S¹ f = S¹map (f base) + , winding (basechange2⁻ (S¹map (f base)) λ i → S¹map (f (loop i))) +Iso.inv S¹→S¹ (s , int) base = ∣ s ∣ +Iso.inv S¹→S¹ (s , int) (loop i) = ∣ basechange2 s (intLoop int) i ∣ +Iso.rightInv S¹→S¹ (s , int) = ×≡ refl ((λ i → winding (basechange2-retr s (λ i → intLoop int i) i)) + ∙ windingIntLoop int) +Iso.leftInv S¹→S¹ f = funExt λ { base → S¹map-id (f base) + ; (loop i) j → helper2 j i} + where + helper2 : PathP (λ i → S¹map-id (f base) i ≡ S¹map-id (f base) i) + (λ i → ∣ basechange2 (S¹map (f base)) (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁)))))) i ∣) + (cong f loop) + helper2 i j = + hcomp (λ k → λ { (i = i0) → cong ∣_∣ (basechange2 (S¹map (f base)) (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁))))))) j + ; (i = i1) → S¹map-id (f (loop j)) k + ; (j = i0) → S¹map-id (f base) (i ∧ k) + ; (j = i1) → S¹map-id (f base) (i ∧ k)}) + (helper4 i j) + where + helper4 : Path (Path (hLevelTrunc 3 _) _ _) + (cong ∣_∣ (basechange2 (S¹map (f base)) + (intLoop + (winding + (basechange2⁻ (S¹map (f base)) + (λ i₁ → S¹map (f (loop i₁)))))))) + λ i → ∣ S¹map (f (loop i)) ∣ + helper4 i = + cong ∣_∣ + ((cong (basechange2 (S¹map (f base))) + (decodeEncode base (basechange2⁻ (S¹map (f base)) + (λ i₁ → S¹map (f (loop i₁))))) + ∙ basechange2-sect (S¹map (f base)) + (λ i → S¹map (f (loop i)))) i) + + + +S¹→S¹≡S1→S1 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → hLevelTrunc 3 S¹) +Iso.fun S¹→S¹≡S1→S1 f x = trMap S1→S¹ (f (S¹→S1 x)) +Iso.inv S¹→S¹≡S1→S1 f x = trMap S¹→S1 (f (S1→S¹ x)) +Iso.rightInv S¹→S¹≡S1→S1 F = funExt λ x i → helper (F (S1→S¹-sect x i)) i + where + helper : (x : hLevelTrunc 3 S¹) → trMap S1→S¹ (trMap S¹→S1 x) ≡ x + helper = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → cong ∣_∣ (S1→S¹-sect a) +Iso.leftInv S¹→S¹≡S1→S1 F = funExt λ x i → helper (F (S1→S¹-retr x i)) i + where + helper : (x : hLevelTrunc 3 (S₊ 1)) → trMap S¹→S1 (trMap S1→S¹ x) ≡ x + helper = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → cong ∣_∣ (S1→S¹-retr a) + + + + +basechange-lemma : ∀ {ℓ} {A : Type ℓ} {a : A} (x y : S¹) (F : a ≡ a → S¹) (f : S¹ → a ≡ a) (g : S¹ → a ≡ a) + → (f base ≡ refl) + → (g base ≡ refl) + → basechange2⁻ (F (f base ∙ g base)) (cong₂ {A = S¹} {B = λ x → S¹} (λ x y → F (f x ∙ g y)) loop loop) + ≡ basechange2⁻ (F (f base)) (cong (λ x → F (f x)) loop) ∙ basechange2⁻ (F (g base)) (cong (λ x → F (g x)) loop) +basechange-lemma x y F f g frefl grefl = + (λ i → basechange2⁻ (F (f base ∙ g base)) (cong₂Funct2 (λ x y → F (f x ∙ g y)) loop loop i)) + ∙ (λ i → basechange2⁻ (F (f base ∙ g base)) (cong (λ x₁ → F (f x₁ ∙ g base)) loop ∙ cong (λ y₁ → F (f base ∙ g y₁)) loop)) + ∙ basechange2⁻-morph (F (f base ∙ g base)) _ _ + ∙ (λ j → basechange2⁻ (F (f base ∙ grefl j)) + (λ i → F (f (loop i) ∙ grefl j)) + ∙ basechange2⁻ (F (frefl j ∙ g base)) + (λ i → F (frefl j ∙ g (loop i)))) + ∙ ((λ j → basechange2⁻ (F (rUnit (f base) (~ j))) + (λ i → F (rUnit (f (loop i)) (~ j))) + ∙ basechange2⁻ (F (lUnit (g base) (~ j))) + (λ i → F (lUnit (g (loop i)) (~ j))))) + + +coHom1S1≃ℤ : grIso (coHomGr 1 (S₊ 1)) intGroup +coHom1S1≃ℤ = + Iso'→Iso + (iso' + (iso + (sRec isSetInt (λ f → proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) + (λ a → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , a)) ∣₀) + (λ a → (λ i → proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 (Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , a)))))) + ∙ (λ i → proj₂ (Iso.fun S¹→S¹ (Iso.rightInv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , a)) i))) + ∙ λ i → proj₂ (Iso.rightInv S¹→S¹ (base , a) i)) + (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → (λ i → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) ∣₀) + ∙ (λ i → ∣ Iso.inv S¹→S¹≡S1→S1 {!!} ∣₀) + ∙ {!!} + ∙ {!!} + ∙ {!!})) + {!!}) + where + helper : (f : _) (x : S¹) → ∣ (Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f))))) ∣₀ ≡ ∣ (Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) ∣₀ + helper f x = sym ({!!} + {!!} ∙ + {!!} ∙ + {!!} ∙ + {!!} ∙ + {!!}) +{- + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) + λ f g → (λ i → winding (basechange2⁻ (S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)) ∙ Kn→ΩKn+1 1 (g (S¹→S1 base)))))) + (λ i → S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 (loop i))) ∙ Kn→ΩKn+1 1 (g (S¹→S1 (loop i))))))))) + ∙ cong winding (helper (f (S¹→S1 base)) (g (S¹→S1 base)) f g refl refl) + ∙ winding-hom ((basechange2⁻ (S¹map (trMap S1→S¹ (f north))) + (λ i → S¹map (trMap S1→S¹ (f (S¹→S1 (loop i))))))) + ((basechange2⁻ (S¹map (trMap S1→S¹ (g north))) + (λ i → S¹map (trMap S1→S¹ (g (S¹→S1 (loop i))))))))) --- test1 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S₊ 1 × Int) --- Iso.fun test1 f = (trRec isGroupoidS1 (λ a → a) (f north)) --- , winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (f (loop* i))))) --- Iso.inv test1 (north , b) x = ∣ x ∣ --- Iso.inv test1 (south , b) x = ∣ x ∣ --- Iso.inv test1 (merid a i , b) x = {!!} --- Iso.rightInv test1 = {!!} --- Iso.leftInv test1 = {!!} --- funRec : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) {C : (A → hLevelTrunc n B) → Type ℓ''} --- → isOfHLevel n B --- → ((f : A → B) → C (λ a → ∣ f a ∣)) --- → (f : A → hLevelTrunc n B) → C f --- funRec {A = A} {B = B} n {C = C} hLev ind f = subst C (helper f) (ind (λ a → trRec hLev (λ x → x) (f a))) --- where --- helper : retract {A = A → hLevelTrunc n B} {B = A → B} (λ f₁ a → trRec hLev (λ x → x) (f₁ a)) (λ f₁ a → ∣ f₁ a ∣) --- helper f = funExt λ a → helper2 (f a) --- where --- helper2 : (x : hLevelTrunc n B) → ∣ trRec hLev (λ x → x) x ∣ ≡ x --- helper2 = trElim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → refl - --- invMapSurj : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (ϕ : morph G H) → ((x : Group.type H) → isInIm G H (fst ϕ) x) --- → morph H G --- fst (invMapSurj G H (ϕ , pf) surj) a = {!pRec!} --- snd (invMapSurj G H (ϕ , pf) surj) = {!!} - -ImIso : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (ϕ : morph G H) → ((x : Group.type H) → isInIm G H (fst ϕ) x) - → grIso H (imGroup G H ϕ) -ImIso G H (ϕ , mf) surj = - let idH = isGroup.id (Group.groupStruc H) - idG = isGroup.id (Group.groupStruc G) - _+G_ = isGroup.comp (Group.groupStruc G) - _+H_ = isGroup.comp (Group.groupStruc H) - _+Im_ = isGroup.comp (Group.groupStruc (imGroup G H (ϕ , mf))) - invG = isGroup.inv (Group.groupStruc G) - invH = isGroup.inv (Group.groupStruc H) - lUnit = isGroup.lUnit (Group.groupStruc H) - lCancel = isGroup.rCancel (Group.groupStruc H) - in - Iso''→Iso _ _ - (iso'' ((λ x → x , pRec propTruncIsProp (λ (a , b) → ∣ a , b ∣₋₁) (surj x)) - , λ a b → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) - (λ surja → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) - (λ surjb → - pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) - (λ surja+b → - (λ i → (a +H b) , (pRec (propTruncIsProp) - (λ (a , b) → ∣ a , b ∣₋₁) - (propTruncIsProp (surj (isGroup.comp (Group.groupStruc H) a b)) ∣ surja+b ∣₋₁ i))) ∙ - (λ i → (a +H b) , ∣ (fst surja+b) , (snd surja+b) ∣₋₁) ∙ - ΣProp≡ (λ _ → propTruncIsProp) refl ∙ - λ i → (a +H b) , pRec (propTruncIsProp) - (λ p1 → - pRec (λ x y → squash x y) - (λ p2 → - ∣ - isGroup.comp (Group.groupStruc G) (fst p1) (fst p2) , - mf (fst p1) (fst p2) ∙ - cong₂ (isGroup.comp (Group.groupStruc H)) (snd p1) (snd p2) - ∣₋₁) - (pRec (propTruncIsProp) - ∣_∣₋₁ (propTruncIsProp ∣ surjb ∣₋₁ (surj b) i))) - (pRec (propTruncIsProp) - ∣_∣₋₁ (propTruncIsProp ∣ surja ∣₋₁ (surj a) i ))) - (surj (isGroup.comp (Group.groupStruc H) a b))) - (surj b)) - (surj a)) - (λ x inker → cong fst inker) - λ x → pRec propTruncIsProp (λ inimx → ∣ (ϕ (inimx .fst)) , ΣProp≡ (λ _ → propTruncIsProp) (inimx .snd) ∣₋₁) (snd x)) -{- -H¹-S¹≃Int : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) -H¹-S¹≃Int = - Iso''→Iso _ _ - (iso'' ((λ x → ∣ theFuns x ∣₀) , λ a b → cong ∣_∣₀ (funExt λ x → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 1) _) ∙ sym (cong (ΩKn+1→Kn) (theFunsId2 a b x)))) - (λ x inker → pRec (isSetInt _ _) (inj x) (Iso.fun PathIdTrunc₀Iso inker)) - finIm) where - d : _ - d = MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 - - i : _ - i = MV.i _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1 - - Δ : _ - Δ = MV.Δ _ _ (S₊ 1) (λ _ → tt) (λ _ → tt) 0 - - - d-surj : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) - → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x - d-surj = sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) - λ x → MV.Ker-i⊂Im-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ - (sym (isContrHelper .snd _)) - where - isContrHelper : isContr (Group.type (×coHomGr 1 Unit Unit)) - isContrHelper = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) - , λ y → prodId (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₁ y)) - (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₂ y)) - - theFuns : (a : Int) → Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 - theFuns a (inl x) = 0ₖ - theFuns a (inr x) = 0ₖ - theFuns a (push north i) = Kn→ΩKn+1 0 a i - theFuns a (push south i) = 0ₖ - - - theFunsId2 : (a b : Int) (x : Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) - → Kn→ΩKn+1 1 (theFuns a x) ∙ Kn→ΩKn+1 1 (theFuns b x) ≡ Kn→ΩKn+1 1 (theFuns (a +ℤ b) x) - theFunsId2 a b (inl x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) - theFunsId2 a b (inr x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) - theFunsId2 a b (push north i) j = - hcomp (λ k → λ {(i = i0) → ((λ i₁ → - (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) - ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) - j - ; (i = i1) → ((λ i₁ → - (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) - ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) - j - ; (j = i0) → cong₂Funct2 (λ p q → Kn→ΩKn+1 1 p ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b) (~ k) i - ; (j = i1) → (helper2 a b) k i }) - (hcomp (λ k → λ { (j = i0) → compPath-filler (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) k i - ; (j = i1) → compPath-filler (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) k i - ; (i = i1) → RHS-filler b j k - ; (i = i0) → ((λ i₁ → - (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) - ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) - j}) - (bottom-filler a j i)) + + + helper : (x y : hLevelTrunc 3 (S₊ 1)) (f g : S₊ 1 → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1) + → (f (S¹→S1 base)) ≡ x + → (g (S¹→S1 base)) ≡ y + → (basechange2⁻ (S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)) ∙ Kn→ΩKn+1 1 (g (S¹→S1 base))))))) + (λ i → S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 (loop i))) ∙ Kn→ΩKn+1 1 (g (S¹→S1 (loop i))))))) + ≡ (basechange2⁻ (S¹map (trMap S1→S¹ ((f (S¹→S1 base)))))) + (λ i → S¹map (trMap S1→S¹ (f (S¹→S1 (loop i))))) + ∙ (basechange2⁻ (S¹map (trMap S1→S¹ ((g (S¹→S1 base))))) + (λ i → S¹map (trMap S1→S¹ ((g (S¹→S1 (loop i))))))) + helper = elim2 + (λ _ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 + λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 + λ _ → isOfHLevelPath 3 (isOfHLevelSuc 3 (isGroupoidS¹) base base) _ _) + (toProdElimSuspElim2 {A = S₊ 0} north + (λ _ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → (isGroupoidS¹) _ _ _ _ ) + λ f g reflf reflg → + (basechange-lemma + base + base + (λ x → S¹map (trMap S1→S¹ (ΩKn+1→Kn x))) + (λ x → Kn→ΩKn+1 1 (f (S¹→S1 x))) ((λ x → Kn→ΩKn+1 1 (g (S¹→S1 x)))) + (cong (Kn→ΩKn+1 1) reflf ∙ Kn→ΩKn+10ₖ 1) + (cong (Kn→ΩKn+1 1) reflg ∙ Kn→ΩKn+10ₖ 1)) + ∙ λ j → basechange2⁻ (S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 base)) j))) + (λ i → S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 (loop i))) j))) + ∙ basechange2⁻ (S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 base)) j))) + (λ i → S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 (loop i))) j)))) +-} - where +{- +Goal: basechange2⁻ + (S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)))))) + (cong + (λ x → S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 x)))))) + loop) + ∙ + basechange2⁻ + (S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (g (S¹→S1 base)))))) + (cong + (λ x → S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (g (S¹→S1 x)))))) + loop) + ≡ + basechange2⁻ (S¹map (trMap S1→S¹ (f (S¹→S1 base)))) + (λ i → S¹map (trMap S1→S¹ (f (S¹→S1 (loop i))))) + ∙ + basechange2⁻ (S¹map (trMap S1→S¹ (g (S¹→S1 base)))) + (λ i → S¹map (trMap S1→S¹ (g (S¹→S1 (loop i))))) +-} +-- (f g : S₊ 1 → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1) → (f (S¹→S1 base)) ≡ ∣ x ∣ → (g (S¹→S1 base)) ≡ ∣ y ∣ → ? - bottom-filler : (a : Int) → - PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) - (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) - ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) - j ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) - (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) - ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) - j) (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) - bottom-filler a j i = - hcomp (λ k → λ {(j = i0) → helper2 (~ k) i ; - (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 a) i}) - ((λ j₂ → ∣ rCancel (merid north) j j₂ ∣) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) - - where - helper2 : Path (Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣ ≡ Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣) - (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i) ∙ Kn→ΩKn+1 1 ∣ north ∣) - (λ i → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) - helper2 = cong₂Sym1 (Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) (~ i) j ∣) (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) - - RHS-filler : (b : Int) → - PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j - ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j) - (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) - (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) - RHS-filler b j i = - hcomp (λ k → λ {(j = i0) → cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i ; - (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 b) i}) - (cong (λ q → (λ i → ∣ rCancel (merid north) j i ∣) ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i) - - helper2 : (a b : Int) → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b) ≡ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ℤ b)) - helper2 a b = - sym (congFunct (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b)) - ∙ (λ i → cong (Kn→ΩKn+1 1) (Iso.rightInv (Iso3-Kn-ΩKn+1 0) (Kn→ΩKn+1 0 a ∙ Kn→ΩKn+1 0 b) (~ i))) - ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ₖ b)) ) - ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (addLemma a b i))) - - theFunsId2 a b (push south i) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) - ∙ sym (lUnit _) - - inj : (a : Int) → theFuns a ≡ (λ _ → ∣ north ∣) → a ≡ pos 0 - inj a id = - pRec (isSetInt _ _) - (sigmaElim (λ _ → isOfHLevelPath 2 isSetInt _ _) - (λ a p → pRec (isSetInt _ _) - (λ id2 → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 0) _) - ∙ cong (ΩKn+1→Kn) - (PathP→compPathR - (cong (λ f → cong f (push north)) id) - ∙ test)) - (Iso.fun PathIdTrunc₀Iso p))) (d-surj ∣ theFuns a ∣₀) - where - test : (λ i → id i (inl tt)) ∙ (λ i → ∣ north ∣) ∙ sym (λ i → id i (inr tt)) ≡ refl - test = (λ i → cong (λ f → f (inl tt)) id - ∙ lUnit (sym (cong (λ f → f (inr tt)) id)) (~ i)) - ∙ (λ i → cong (λ f → f (push south i)) id - ∙ sym (cong (λ f → f (inr tt)) id)) - ∙ rCancel (cong (λ f → f (inr tt)) id) - - - consMember : (a : Int) → coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) - consMember a = d ∣ (λ _ → a) ∣₀ - - consMember≡0 : (a : Int) → consMember a ≡ 0ₕ - consMember≡0 a = - (MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ (λ _ → a) ∣₀ ∣ - (∣ (λ _ → a) ∣₀ , ∣ (λ _ → 0) ∣₀) - , cong ∣_∣₀ (λ i x → (rUnitₖ a i)) ∣₋₁) - - f+consMember' : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int × Int ] (f +ₕ (-ₕ (consMember (proj₁ x))) ≡ ∣ theFuns (proj₂ x) ∣₀) - f+consMember' = - sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) - λ f → pRec propTruncIsProp - (sigmaElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) - (λ g id → ∣ ((g south) , ((g north) +ₖ (-ₖ g south))) - , (pRec (setTruncIsSet _ _) - (λ id → (λ i → ∣ id (~ i) ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀) ∙ funId1 g) - (Iso.fun PathIdTrunc₀Iso id)) ∣₋₁)) - (d-surj ∣ f ∣₀) - where - funId1 : (g : S₊ 0 → Int) - → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀ ≡ - ∣ theFuns ((g north) +ₖ (-ₖ (g south))) ∣₀ - funId1 g = (λ i → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ - +ₕ (morphMinus (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) d - (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ∣ (λ _ → g south) ∣₀ (~ i))) - ∙ sym (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ g ∣₀ (-ₕ ∣ (λ _ → g south) ∣₀)) - ∙ (cong (λ x → d ∣ x ∣₀) g'Id) - ∙ cong ∣_∣₀ helper - where - g' : S₊ 0 → Int - g' north = (g north) +ₖ (-ₖ (g south)) - g' south = 0 - - g'Id : (λ x → g x +ₖ (-ₖ (g south))) ≡ g' - g'Id = funExt (λ {north → refl - ; south → rCancelₖ (g south)}) - - helper : MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g' ≡ theFuns (g north +ₖ (-ₖ g south)) - helper = funExt λ {(inl tt) → refl - ; (inr tt) → refl - ; (push north i) → refl - ; (push south i) → refl} - finIm : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int ] (∣ theFuns x ∣₀ ≡ f) - finIm f = - pRec propTruncIsProp - (λ {((a , b) , id) → ∣ b , (sym id ∙ cong (λ x → f +ₕ x) ((λ i → (-ₕ (consMember≡0 a i))) ∙ sym (lUnitₕ (-ₕ 0ₕ)) ∙ rCancelₕ 0ₕ) ∙ (rUnitₕ f)) ∣₋₁}) - (f+consMember' f) --} -Hⁿ-Sⁿ≃Int : (n : ℕ) → grIso intGroup (coHomGr (suc n) (S₊ (suc n))) -Hⁿ-Sⁿ≃Int zero = - compGrIso {F = intGroup} {G = {!!}} {H = {!coHomGr 1 (S₊ 1)!}} - (Iso''→Iso _ _ - (iso'' ((λ x → ∣ theFuns x ∣₀) , λ a b → cong ∣_∣₀ (funExt λ x → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 1) _) ∙ sym (cong (ΩKn+1→Kn) (theFunsId2 a b x)))) - (λ x inker → pRec (isSetInt _ _) (inj x) (Iso.fun PathIdTrunc₀Iso inker)) - finIm)) - {!invGrIso _ _ (coHomPushout≡coHomSn 0 1)!} +{- +coHom1S1≃ℤ : grIso (coHomGr 1 (S₊ 1)) intGroup +coHom1S1≃ℤ = + Iso'→Iso + (iso' + (iso + (sRec isSetInt λ f → proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f))) + {!winding (basechange2⁻ (S¹map (f base)) λ i → S¹map (f (loop i)))!} + {!!} + {!!}) + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) + λ f g → (λ i → winding (basechange2⁻ (S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)) ∙ Kn→ΩKn+1 1 (g (S¹→S1 base)))))) + (λ i → S¹map (trMap S1→S¹ ((ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 (loop i))) ∙ Kn→ΩKn+1 1 (g (S¹→S1 (loop i)))))))))) + ∙ (λ j → winding (basechange2⁻ (S¹map (trMap S1→S¹ (ΩKn+1→Kn {!basechange2⁻!}))) + (λ i → S¹map (trMap S1→S¹ (ΩKn+1→Kn {!!}))))) + ∙ {!!} + ∙ {!λ i → winding !} + ∙ winding-hom (basechange2⁻ (S¹map (trMap S1→S¹ (f (S¹→S1 base)))) + (λ i → S¹map (trMap S1→S¹ (f (S¹→S1 (loop i)))))) + (basechange2⁻ (S¹map (trMap S1→S¹ (g (S¹→S1 base)))) + (λ i → S¹map (trMap S1→S¹ (g (S¹→S1 (loop i)))))) )) -- sym (winding-hom _ _))) where - d : _ - d = MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 - - i : _ - i = MV.i _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1 - - Δ : _ - Δ = MV.Δ _ _ (S₊ 1) (λ _ → tt) (λ _ → tt) 0 - - - d-surj : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) - → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x - d-surj = sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) - λ x → MV.Ker-i⊂Im-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ - (sym (isContrHelper .snd _)) - where - isContrHelper : isContr (Group.type (×coHomGr 1 Unit Unit)) - isContrHelper = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) - , λ y → prodId (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₁ y)) - (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₂ y)) - - theFuns : (a : Int) → Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 - theFuns a (inl x) = 0ₖ - theFuns a (inr x) = 0ₖ - theFuns a (push north i) = Kn→ΩKn+1 0 a i - theFuns a (push south i) = 0ₖ - - - theFunsId2 : (a b : Int) (x : Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) - → Kn→ΩKn+1 1 (theFuns a x) ∙ Kn→ΩKn+1 1 (theFuns b x) ≡ Kn→ΩKn+1 1 (theFuns (a +ℤ b) x) - theFunsId2 a b (inl x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) - theFunsId2 a b (inr x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) - theFunsId2 a b (push north i) j = - hcomp (λ k → λ {(i = i0) → ((λ i₁ → - (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) - ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) - j - ; (i = i1) → ((λ i₁ → - (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) - ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) - j - ; (j = i0) → cong₂Funct2 (λ p q → Kn→ΩKn+1 1 p ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b) (~ k) i - ; (j = i1) → (helper2 a b) k i }) - (hcomp (λ k → λ { (j = i0) → compPath-filler (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) k i - ; (j = i1) → compPath-filler (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) k i - ; (i = i1) → RHS-filler b j k - ; (i = i0) → ((λ i₁ → - (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) - ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) - j}) - (bottom-filler a j i)) - - where + helper : {!!} + helper = {!!} -} + + + +-- indIntGroup : ∀ {ℓ} {G : Group ℓ} → (ϕ : Int → Group.type G) +-- → ϕ 0 ≡ isGroup.id (Group.groupStruc G) +-- → ((a : Int) → ϕ (a +ℤ 1) ≡ isGroup.comp (Group.groupStruc G) (ϕ a) (ϕ 1)) +-- → ((n : Int) → ϕ (predInt n) ≡ isGroup.comp (Group.groupStruc G) (ϕ n) (ϕ (negsuc zero))) +-- → isMorph intGroup G ϕ +-- indIntGroup {G = group G gSet (group-struct _ _ _+G_ _ rUnit₁ _ _ _)} ϕ zeroId _ _ n (pos zero) = +-- sym (rUnit₁ (ϕ n)) ∙ cong (λ x → ϕ n +G x) (sym zeroId) +-- indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId oneId minOneId n (pos (suc m)) = +-- (λ i → ϕ ((n +pos m) +ℤ 1)) +-- ∙ oneId (n +pos m) +-- ∙ cong (λ x → x +G ϕ (pos 1)) +-- (indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} +-- ϕ zeroId oneId minOneId n (pos m)) +-- ∙ assoc₁ (ϕ n) (ϕ (pos m)) (ϕ (pos 1)) +-- ∙ sym (cong (λ x → ϕ n +G x) (oneId (pos m))) +-- indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId _ minOneId n (negsuc zero) = minOneId n +-- indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId a minOneId n (negsuc (suc m)) = +-- (λ i → ϕ ((n +negsuc m) +ℤ (negsuc zero))) +-- ∙ minOneId (n +negsuc m) +-- ∙ cong (λ x → x +G ϕ (negsuc zero)) (indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId a minOneId n (negsuc m)) +-- ∙ assoc₁ (ϕ n) (ϕ (negsuc m)) (ϕ (negsuc zero)) +-- ∙ cong (λ x → ϕ n +G x) (sym (minOneId (negsuc m))) + +-- pushoutSn : (n : ℕ) → Iso (S₊ (suc n)) (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt) +-- Iso.fun (pushoutSn n) north = inl tt +-- Iso.fun (pushoutSn n) south = inr tt +-- Iso.fun (pushoutSn n) (merid a i) = push a i +-- Iso.inv (pushoutSn n) (inl x) = north +-- Iso.inv (pushoutSn n) (inr x) = south +-- Iso.inv (pushoutSn n) (push a i) = merid a i +-- Iso.rightInv (pushoutSn n) (inl x) = refl +-- Iso.rightInv (pushoutSn n) (inr x) = refl +-- Iso.rightInv (pushoutSn n) (push a i) = refl +-- Iso.leftInv (pushoutSn n) north = refl +-- Iso.leftInv (pushoutSn n) south = refl +-- Iso.leftInv (pushoutSn n) (merid a i) = refl + +-- Sn≡Pushout : (n : ℕ) → (S₊ (suc n)) ≡ (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt) +-- Sn≡Pushout n = isoToPath (pushoutSn n) + +-- coHomPushout≡coHomSn' : (n m : ℕ) → grIso (coHomGr m (S₊ (suc n))) (coHomGr m (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt)) +-- morph.fun (grIso.fun (coHomPushout≡coHomSn' n m)) = +-- sRec setTruncIsSet +-- λ f → ∣ (λ {(inl x) → f north ; (inr x) → f south ; (push a i) → f (merid a i)}) ∣₀ +-- morph.ismorph (grIso.fun (coHomPushout≡coHomSn' n m)) = +-- sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- λ a b → cong ∣_∣₀ (funExt λ {(inl x) → refl ; (inr x) → refl ; (push a i) → refl }) +-- morph.fun (grIso.inv (coHomPushout≡coHomSn' n m)) = sRec setTruncIsSet (λ f → ∣ (λ {north → f (inl tt) ; south → f (inr tt) ; (merid a i) → f (push a i)}) ∣₀) +-- morph.ismorph (grIso.inv (coHomPushout≡coHomSn' n m)) = +-- sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- λ a b → cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → refl }) +-- grIso.rightInv (coHomPushout≡coHomSn' n m) = +-- sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- λ f → cong ∣_∣₀ (funExt λ {(inl x) → refl ; (inr x) → refl ; (push a i) → refl }) +-- grIso.leftInv (coHomPushout≡coHomSn' n m) = +-- sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- λ f → cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → refl }) + + +-- isContr→≡Unit : {A : Type₀} → isContr A → A ≡ Unit +-- isContr→≡Unit contr = isoToPath (iso (λ _ → tt) (λ _ → fst contr) (λ _ → refl) λ _ → snd contr _) + +-- isContr→isContrTrunc : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → isContr A → isContr (hLevelTrunc n A) +-- isContr→isContrTrunc n contr = ∣ fst contr ∣ , (trElim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → cong ∣_∣ (snd contr a)) +-- isContr→isContrSetTrunc : ∀ {ℓ} {A : Type ℓ} → isContr A → isContr (∥ A ∥₀) +-- isContr→isContrSetTrunc contr = ∣ fst contr ∣₀ , sElim (λ _ → isOfHLevelPath 2 (setTruncIsSet) _ _) λ a → cong ∣_∣₀ (snd contr a) + +-- coHomGrUnit0 : grIso (coHomGr 0 Unit) intGroup +-- grIso.fun coHomGrUnit0 = mph (sRec isSetInt (λ f → f tt)) +-- (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) +-- (λ a b → addLemma (a tt) (b tt))) +-- grIso.inv coHomGrUnit0 = mph (λ a → ∣ (λ _ → a) ∣₀) (λ a b i → ∣ (λ _ → addLemma a b (~ i)) ∣₀) +-- grIso.rightInv coHomGrUnit0 a = refl +-- grIso.leftInv coHomGrUnit0 = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ a → refl + +-- isContrCohomUnit : (n : ℕ) → isContr (coHom (suc n) Unit) +-- isContrCohomUnit n = subst isContr (λ i → ∥ UnitToTypeId (coHomK (suc n)) (~ i) ∥₀) (helper n) +-- where +-- helper : (n : ℕ) → isContr (∥ coHomK (suc n) ∥₀) +-- helper n = subst isContr ((isoToPath (truncOfTruncIso {A = S₊ (1 + n)} 2 (1 + n))) ∙ sym propTrunc≡Trunc0 ∙ λ i → ∥ hLevelTrunc (suc (+-comm n 2 i)) (S₊ (1 + n)) ∥₀) +-- (isConnectedSubtr 2 (helper2 n .fst) (subst (λ x → isHLevelConnected x (S₊ (suc n))) (sym (helper2 n .snd)) (sphereConnected (suc n))) ) +-- where +-- helper2 : (n : ℕ) → Σ[ m ∈ ℕ ] m + 2 ≡ 2 + n +-- helper2 zero = 0 , refl +-- helper2 (suc n) = (suc n) , λ i → suc (+-comm n 2 i) + +-- coHomGrUnit≥1 : (n : ℕ) → grIso (coHomGr (suc n) Unit) trivialGroup +-- grIso.fun (coHomGrUnit≥1 n) = mph (λ _ → tt) (λ _ _ → refl) +-- grIso.inv (coHomGrUnit≥1 n) = mph (λ _ → ∣ (λ _ → ∣ north ∣) ∣₀) (λ _ _ → sym (rUnitₕ 0ₕ)) +-- grIso.rightInv (coHomGrUnit≥1 n) a = refl +-- grIso.leftInv (coHomGrUnit≥1 n) a = sym (isContrCohomUnit n .snd 0ₕ) ∙ isContrCohomUnit n .snd a + +-- S0→Int : (a : Int × Int) → S₊ 0 → Int +-- S0→Int a north = proj₁ a +-- S0→Int a south = proj₂ a + +-- coHom0-S0 : grIso (coHomGr 0 (S₊ 0)) (dirProd intGroup intGroup) +-- coHom0-S0 = +-- Iso'→Iso +-- (iso' +-- (iso (sRec (isOfHLevelProd 2 isSetInt isSetInt) +-- λ f → (f north) , (f south)) +-- (λ a → ∣ S0→Int a ∣₀) +-- (λ { (a , b) → refl}) +-- (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ f → cong ∣_∣₀ (funExt (λ {north → refl ; south → refl})))) +-- (sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 isSetInt isSetInt) _ _) +-- λ a b i → addLemma (a north) (b north) i , addLemma (a south) (b south) i)) + +-- ×morph : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group ℓ} {B : Group ℓ'} {C : Group ℓ''} {D : Group ℓ'''} → morph A B → morph C D → morph (dirProd A C) (dirProd B D) +-- morph.fun (×morph mf1 mf2) = +-- (λ {(a , b) → (morph.fun mf1 a) , morph.fun mf2 b}) +-- morph.ismorph (×morph mf1 mf2) = +-- (λ {(a , b) (c , d) i → morph.ismorph mf1 a c i , morph.ismorph mf2 b d i}) + + + + + +-- coHom1S1 : grIso intGroup (coHomGr 1 (S₊ 1)) +-- coHom1S1 = +-- compGrIso +-- (diagonalIso1 +-- _ +-- (coHomGr 0 (S₊ 0)) +-- _ +-- (invGrIso coHom0-S0) +-- (d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) +-- (λ x → MV.Ker-i⊂Im-d _ _(S₊ 0) (λ _ → tt) (λ _ → tt) 0 x +-- (×≡ (isOfHLevelSuc 0 (isContrCohomUnit 0) _ _) +-- (isOfHLevelSuc 0 (isContrCohomUnit 0) _ _))) +-- (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- (λ x inker +-- → pRec propTruncIsProp +-- (λ {((f , g) , id') → helper x f g id' inker}) +-- ((MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ inker)))) +-- (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- λ F surj +-- → pRec (setTruncIsSet _ _) (λ { (x , id) → MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ F ∣₀ ∣ (∣ (λ _ → x) ∣₀ , ∣ (λ _ → 0) ∣₀) , +-- (cong ∣_∣₀ (funExt (surjHelper x))) ∙ sym id ∣₋₁ }) surj) ) +-- (invGrIso (coHomPushout≡coHomSn' 0 1)) + +-- where +-- surjHelper : (x : Int) (x₁ : S₊ 0) → x +ₖ (-ₖ pos 0) ≡ S0→Int (x , x) x₁ +-- surjHelper x north = cong (x +ₖ_) (-0ₖ) ∙ rUnitₖ x +-- surjHelper x south = cong (x +ₖ_) (-0ₖ) ∙ rUnitₖ x + +-- helper : (F : S₊ 0 → Int) (f g : ∥ (Unit → Int) ∥₀) (id : MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (f , g) ≡ ∣ F ∣₀) +-- → isInKer (coHomGr 0 (S₊ 0)) +-- (coHomGr 1 (Pushout (λ _ → tt) (λ _ → tt))) +-- (morph.fun (d-morph Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) +-- ∣ F ∣₀ +-- → ∃[ x ∈ Int ] ∣ F ∣₀ ≡ morph.fun (grIso.inv coHom0-S0) (x , x) +-- helper F = +-- sElim2 (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- λ f g id inker +-- → pRec propTruncIsProp +-- (λ {((a , b) , id2) +-- → sElim2 {B = λ f g → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (f , g) ≡ ∣ F ∣₀ → _ } +-- (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- (λ f g id → ∣ (helper2 f g .fst) , (sym id ∙ sym (helper2 f g .snd)) ∣₋₁) +-- a b id2}) +-- (MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ F ∣₀ inker) +-- where +-- helper2 : (f g : Unit → Int) +-- → Σ[ x ∈ Int ] morph.fun (grIso.inv coHom0-S0) (x , x) +-- ≡ MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (∣ f ∣₀ , ∣ g ∣₀) +-- helper2 f g = (f _ +ₖ (-ₖ g _) ) , cong ∣_∣₀ (funExt (λ {north → refl ; south → refl})) + + +-- coHom-n-Sn : (n : ℕ) → grIso intGroup (coHomGr (suc n) (S₊ (suc n))) +-- coHom-n-Sn zero = coHom1S1 +-- coHom-n-Sn (suc n) = +-- compGrIso +-- (compGrIso +-- (coHom-n-Sn n) +-- theIso) +-- (invGrIso (coHomPushout≡coHomSn' (suc n) (suc (suc n)))) +-- where +-- theIso : grIso (coHomGr (suc n) (S₊ (suc n))) (coHomGr (suc (suc n)) +-- (Pushout {A = S₊ (suc n)} (λ _ → tt) (λ _ → tt))) +-- theIso = +-- SES→Iso +-- (×coHomGr (suc n) Unit Unit) +-- (×coHomGr (suc (suc n)) Unit Unit) +-- (ses (λ p q → ×≡ (isOfHLevelSuc 0 (isContrCohomUnit n) (proj₁ p) (proj₁ q)) +-- (isOfHLevelSuc 0 (isContrCohomUnit n) (proj₂ p) (proj₂ q))) +-- (λ p q → ×≡ (isOfHLevelSuc 0 (isContrCohomUnit (suc n)) (proj₁ p) (proj₁ q)) +-- (isOfHLevelSuc 0 (isContrCohomUnit (suc n)) (proj₂ p) (proj₂ q))) +-- (Δ-morph _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) +-- (i-morph _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (2 + n)) +-- (d-morph _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) +-- (MV.Ker-d⊂Im-Δ _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) +-- (MV.Ker-i⊂Im-d _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n))) + + + + + + + + + +-- -- basechange* : (x y : S¹) → x ≡ y → x ≡ y → ΩS¹ +-- -- basechange* x y = J (λ y p → (x ≡ y) → ΩS¹) (basechange x) + + +-- -- test1 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S₊ 1 × Int) +-- -- Iso.fun test1 f = (trRec isGroupoidS1 (λ a → a) (f north)) +-- -- , winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (f (loop* i))))) +-- -- Iso.inv test1 (north , b) x = ∣ x ∣ +-- -- Iso.inv test1 (south , b) x = ∣ x ∣ +-- -- Iso.inv test1 (merid a i , b) x = {!!} +-- -- Iso.rightInv test1 = {!!} +-- -- Iso.leftInv test1 = {!!} + +-- -- funRec : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) {C : (A → hLevelTrunc n B) → Type ℓ''} +-- -- → isOfHLevel n B +-- -- → ((f : A → B) → C (λ a → ∣ f a ∣)) +-- -- → (f : A → hLevelTrunc n B) → C f +-- -- funRec {A = A} {B = B} n {C = C} hLev ind f = subst C (helper f) (ind (λ a → trRec hLev (λ x → x) (f a))) +-- -- where +-- -- helper : retract {A = A → hLevelTrunc n B} {B = A → B} (λ f₁ a → trRec hLev (λ x → x) (f₁ a)) (λ f₁ a → ∣ f₁ a ∣) +-- -- helper f = funExt λ a → helper2 (f a) +-- -- where +-- -- helper2 : (x : hLevelTrunc n B) → ∣ trRec hLev (λ x → x) x ∣ ≡ x +-- -- helper2 = trElim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → refl + +-- -- invMapSurj : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (ϕ : morph G H) → ((x : Group.type H) → isInIm G H (fst ϕ) x) +-- -- → morph H G +-- -- fst (invMapSurj G H (ϕ , pf) surj) a = {!pRec!} +-- -- snd (invMapSurj G H (ϕ , pf) surj) = {!!} + +-- {- +-- ImIso : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (ϕ : morph G H) → ((x : Group.type H) → isInIm G H (fst ϕ) x) +-- → grIso H (imGroup G H ϕ) +-- ImIso G H (ϕ , mf) surj = +-- let idH = isGroup.id (Group.groupStruc H) +-- idG = isGroup.id (Group.groupStruc G) +-- _+G_ = isGroup.comp (Group.groupStruc G) +-- _+H_ = isGroup.comp (Group.groupStruc H) +-- _+Im_ = isGroup.comp (Group.groupStruc (imGroup G H (ϕ , mf))) +-- invG = isGroup.inv (Group.groupStruc G) +-- invH = isGroup.inv (Group.groupStruc H) +-- lUnit = isGroup.lUnit (Group.groupStruc H) +-- lCancel = isGroup.rCancel (Group.groupStruc H) +-- in +-- Iso''→Iso _ _ +-- (iso'' ((λ x → x , pRec propTruncIsProp (λ (a , b) → ∣ a , b ∣₋₁) (surj x)) +-- , λ a b → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) +-- (λ surja → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) +-- (λ surjb → +-- pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) +-- (λ surja+b → +-- (λ i → (a +H b) , (pRec (propTruncIsProp) +-- (λ (a , b) → ∣ a , b ∣₋₁) +-- (propTruncIsProp (surj (isGroup.comp (Group.groupStruc H) a b)) ∣ surja+b ∣₋₁ i))) ∙ +-- (λ i → (a +H b) , ∣ (fst surja+b) , (snd surja+b) ∣₋₁) ∙ +-- ΣProp≡ (λ _ → propTruncIsProp) refl ∙ +-- λ i → (a +H b) , pRec (propTruncIsProp) +-- (λ p1 → +-- pRec (λ x y → squash x y) +-- (λ p2 → +-- ∣ +-- isGroup.comp (Group.groupStruc G) (fst p1) (fst p2) , +-- mf (fst p1) (fst p2) ∙ +-- cong₂ (isGroup.comp (Group.groupStruc H)) (snd p1) (snd p2) +-- ∣₋₁) +-- (pRec (propTruncIsProp) +-- ∣_∣₋₁ (propTruncIsProp ∣ surjb ∣₋₁ (surj b) i))) +-- (pRec (propTruncIsProp) +-- ∣_∣₋₁ (propTruncIsProp ∣ surja ∣₋₁ (surj a) i ))) +-- (surj (isGroup.comp (Group.groupStruc H) a b))) +-- (surj b)) +-- (surj a)) +-- (λ x inker → cong fst inker) +-- λ x → pRec propTruncIsProp (λ inimx → ∣ (ϕ (inimx .fst)) , ΣProp≡ (λ _ → propTruncIsProp) (inimx .snd) ∣₋₁) (snd x)) +-- -} + + +-- {- +-- H¹-S¹≃Int : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- H¹-S¹≃Int = +-- Iso''→Iso _ _ +-- (iso'' ((λ x → ∣ theFuns x ∣₀) , λ a b → cong ∣_∣₀ (funExt λ x → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 1) _) ∙ sym (cong (ΩKn+1→Kn) (theFunsId2 a b x)))) +-- (λ x inker → pRec (isSetInt _ _) (inj x) (Iso.fun PathIdTrunc₀Iso inker)) +-- finIm) +-- where +-- d : _ +-- d = MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 + +-- i : _ +-- i = MV.i _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1 + +-- Δ : _ +-- Δ = MV.Δ _ _ (S₊ 1) (λ _ → tt) (λ _ → tt) 0 + + +-- d-surj : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x +-- d-surj = sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- λ x → MV.Ker-i⊂Im-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ +-- (sym (isContrHelper .snd _)) +-- where +-- isContrHelper : isContr (Group.type (×coHomGr 1 Unit Unit)) +-- isContrHelper = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) +-- , λ y → prodId (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₁ y)) +-- (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₂ y)) + +-- theFuns : (a : Int) → Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 +-- theFuns a (inl x) = 0ₖ +-- theFuns a (inr x) = 0ₖ +-- theFuns a (push north i) = Kn→ΩKn+1 0 a i +-- theFuns a (push south i) = 0ₖ + + +-- theFunsId2 : (a b : Int) (x : Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) +-- → Kn→ΩKn+1 1 (theFuns a x) ∙ Kn→ΩKn+1 1 (theFuns b x) ≡ Kn→ΩKn+1 1 (theFuns (a +ℤ b) x) +-- theFunsId2 a b (inl x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) +-- theFunsId2 a b (inr x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) +-- theFunsId2 a b (push north i) j = +-- hcomp (λ k → λ {(i = i0) → ((λ i₁ → +-- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) +-- j +-- ; (i = i1) → ((λ i₁ → +-- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) +-- j +-- ; (j = i0) → cong₂Funct2 (λ p q → Kn→ΩKn+1 1 p ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b) (~ k) i +-- ; (j = i1) → (helper2 a b) k i }) +-- (hcomp (λ k → λ { (j = i0) → compPath-filler (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) k i +-- ; (j = i1) → compPath-filler (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) k i +-- ; (i = i1) → RHS-filler b j k +-- ; (i = i0) → ((λ i₁ → +-- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) +-- j}) +-- (bottom-filler a j i)) - bottom-filler : (a : Int) → - PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) - (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) - ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) - j ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) - (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) - ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) - j) (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) - bottom-filler a j i = - hcomp (λ k → λ {(j = i0) → helper2 (~ k) i ; - (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 a) i}) - ((λ j₂ → ∣ rCancel (merid north) j j₂ ∣) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) - - where - helper2 : Path (Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣ ≡ Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣) - (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i) ∙ Kn→ΩKn+1 1 ∣ north ∣) - (λ i → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) - helper2 = cong₂Sym1 (Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) (~ i) j ∣) (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) - - RHS-filler : (b : Int) → - PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j - ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j) - (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) - (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) - RHS-filler b j i = - hcomp (λ k → λ {(j = i0) → cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i ; - (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 b) i}) - (cong (λ q → (λ i → ∣ rCancel (merid north) j i ∣) ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i) - - helper2 : (a b : Int) → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b) ≡ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ℤ b)) - helper2 a b = - sym (congFunct (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b)) - ∙ (λ i → cong (Kn→ΩKn+1 1) (Iso.rightInv (Iso3-Kn-ΩKn+1 0) (Kn→ΩKn+1 0 a ∙ Kn→ΩKn+1 0 b) (~ i))) - ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ₖ b)) ) - ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (addLemma a b i))) - - theFunsId2 a b (push south i) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) - ∙ sym (lUnit _) - - inj : (a : Int) → theFuns a ≡ (λ _ → ∣ north ∣) → a ≡ pos 0 - inj a id = - pRec (isSetInt _ _) - (sigmaElim (λ _ → isOfHLevelPath 2 isSetInt _ _) - (λ a p → pRec (isSetInt _ _) - (λ id2 → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 0) _) - ∙ cong (ΩKn+1→Kn) - (PathP→compPathR - (cong (λ f → cong f (push north)) id) - ∙ test)) - (Iso.fun PathIdTrunc₀Iso p))) (d-surj ∣ theFuns a ∣₀) - where +-- where - test : (λ i → id i (inl tt)) ∙ (λ i → ∣ north ∣) ∙ sym (λ i → id i (inr tt)) ≡ refl - test = (λ i → cong (λ f → f (inl tt)) id - ∙ lUnit (sym (cong (λ f → f (inr tt)) id)) (~ i)) - ∙ (λ i → cong (λ f → f (push south i)) id - ∙ sym (cong (λ f → f (inr tt)) id)) - ∙ rCancel (cong (λ f → f (inr tt)) id) - - - consMember : (a : Int) → coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) - consMember a = d ∣ (λ _ → a) ∣₀ - - consMember≡0 : (a : Int) → consMember a ≡ 0ₕ - consMember≡0 a = - (MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ (λ _ → a) ∣₀ ∣ - (∣ (λ _ → a) ∣₀ , ∣ (λ _ → 0) ∣₀) - , cong ∣_∣₀ (λ i x → (rUnitₖ a i)) ∣₋₁) - - f+consMember' : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int × Int ] (f +ₕ (-ₕ (consMember (proj₁ x))) ≡ ∣ theFuns (proj₂ x) ∣₀) - f+consMember' = - sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) - λ f → pRec propTruncIsProp - (sigmaElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) - (λ g id → ∣ ((g south) , ((g north) +ₖ (-ₖ g south))) - , (pRec (setTruncIsSet _ _) - (λ id → (λ i → ∣ id (~ i) ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀) ∙ funId1 g) - (Iso.fun PathIdTrunc₀Iso id)) ∣₋₁)) - (d-surj ∣ f ∣₀) - where - funId1 : (g : S₊ 0 → Int) - → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀ ≡ - ∣ theFuns ((g north) +ₖ (-ₖ (g south))) ∣₀ - funId1 g = (λ i → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ - +ₕ (morphMinus (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) d - (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ∣ (λ _ → g south) ∣₀ (~ i))) - ∙ sym (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ g ∣₀ (-ₕ ∣ (λ _ → g south) ∣₀)) - ∙ (cong (λ x → d ∣ x ∣₀) g'Id) - ∙ cong ∣_∣₀ helper - where - g' : S₊ 0 → Int - g' north = (g north) +ₖ (-ₖ (g south)) - g' south = 0 - - g'Id : (λ x → g x +ₖ (-ₖ (g south))) ≡ g' - g'Id = funExt (λ {north → refl - ; south → rCancelₖ (g south)}) - - helper : MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g' ≡ theFuns (g north +ₖ (-ₖ g south)) - helper = funExt λ {(inl tt) → refl - ; (inr tt) → refl - ; (push north i) → refl - ; (push south i) → refl} - finIm : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int ] (∣ theFuns x ∣₀ ≡ f) - finIm f = - pRec propTruncIsProp - (λ {((a , b) , id) → ∣ b , (sym id ∙ cong (λ x → f +ₕ x) ((λ i → (-ₕ (consMember≡0 a i))) ∙ sym (lUnitₕ (-ₕ 0ₕ)) ∙ rCancelₕ 0ₕ) ∙ (rUnitₕ f)) ∣₋₁}) - (f+consMember' f) -Hⁿ-Sⁿ≃Int (suc n) = - compGrIso (Hⁿ-Sⁿ≃Int n) - (transport (λ i → grIso {!!} {!coHomGr (suc (suc n)) (Pushout (λ _ → tt) (λ _ → tt))!}) {!!}) +-- bottom-filler : (a : Int) → +-- PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) +-- ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) +-- j ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) +-- ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) +-- j) (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) +-- bottom-filler a j i = +-- hcomp (λ k → λ {(j = i0) → helper2 (~ k) i ; +-- (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 a) i}) +-- ((λ j₂ → ∣ rCancel (merid north) j j₂ ∣) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + +-- where +-- helper2 : Path (Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣ ≡ Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- (λ i → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) +-- helper2 = cong₂Sym1 (Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) (~ i) j ∣) (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + +-- RHS-filler : (b : Int) → +-- PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j +-- ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j) +-- (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) +-- (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) +-- RHS-filler b j i = +-- hcomp (λ k → λ {(j = i0) → cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i ; +-- (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 b) i}) +-- (cong (λ q → (λ i → ∣ rCancel (merid north) j i ∣) ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i) + +-- helper2 : (a b : Int) → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b) ≡ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ℤ b)) +-- helper2 a b = +-- sym (congFunct (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b)) +-- ∙ (λ i → cong (Kn→ΩKn+1 1) (Iso.rightInv (Iso3-Kn-ΩKn+1 0) (Kn→ΩKn+1 0 a ∙ Kn→ΩKn+1 0 b) (~ i))) +-- ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ₖ b)) ) +-- ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (addLemma a b i))) + +-- theFunsId2 a b (push south i) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- ∙ sym (lUnit _) + +-- inj : (a : Int) → theFuns a ≡ (λ _ → ∣ north ∣) → a ≡ pos 0 +-- inj a id = +-- pRec (isSetInt _ _) +-- (sigmaElim (λ _ → isOfHLevelPath 2 isSetInt _ _) +-- (λ a p → pRec (isSetInt _ _) +-- (λ id2 → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 0) _) +-- ∙ cong (ΩKn+1→Kn) +-- (PathP→compPathR +-- (cong (λ f → cong f (push north)) id) +-- ∙ test)) +-- (Iso.fun PathIdTrunc₀Iso p))) (d-surj ∣ theFuns a ∣₀) +-- where +-- test : (λ i → id i (inl tt)) ∙ (λ i → ∣ north ∣) ∙ sym (λ i → id i (inr tt)) ≡ refl +-- test = (λ i → cong (λ f → f (inl tt)) id +-- ∙ lUnit (sym (cong (λ f → f (inr tt)) id)) (~ i)) +-- ∙ (λ i → cong (λ f → f (push south i)) id +-- ∙ sym (cong (λ f → f (inr tt)) id)) +-- ∙ rCancel (cong (λ f → f (inr tt)) id) + + +-- consMember : (a : Int) → coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) +-- consMember a = d ∣ (λ _ → a) ∣₀ + +-- consMember≡0 : (a : Int) → consMember a ≡ 0ₕ +-- consMember≡0 a = +-- (MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ (λ _ → a) ∣₀ ∣ +-- (∣ (λ _ → a) ∣₀ , ∣ (λ _ → 0) ∣₀) +-- , cong ∣_∣₀ (λ i x → (rUnitₖ a i)) ∣₋₁) + +-- f+consMember' : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int × Int ] (f +ₕ (-ₕ (consMember (proj₁ x))) ≡ ∣ theFuns (proj₂ x) ∣₀) +-- f+consMember' = +-- sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- λ f → pRec propTruncIsProp +-- (sigmaElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- (λ g id → ∣ ((g south) , ((g north) +ₖ (-ₖ g south))) +-- , (pRec (setTruncIsSet _ _) +-- (λ id → (λ i → ∣ id (~ i) ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀) ∙ funId1 g) +-- (Iso.fun PathIdTrunc₀Iso id)) ∣₋₁)) +-- (d-surj ∣ f ∣₀) +-- where +-- funId1 : (g : S₊ 0 → Int) +-- → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀ ≡ +-- ∣ theFuns ((g north) +ₖ (-ₖ (g south))) ∣₀ +-- funId1 g = (λ i → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +-- +ₕ (morphMinus (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) d +-- (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ∣ (λ _ → g south) ∣₀ (~ i))) +-- ∙ sym (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ g ∣₀ (-ₕ ∣ (λ _ → g south) ∣₀)) +-- ∙ (cong (λ x → d ∣ x ∣₀) g'Id) +-- ∙ cong ∣_∣₀ helper +-- where +-- g' : S₊ 0 → Int +-- g' north = (g north) +ₖ (-ₖ (g south)) +-- g' south = 0 + +-- g'Id : (λ x → g x +ₖ (-ₖ (g south))) ≡ g' +-- g'Id = funExt (λ {north → refl +-- ; south → rCancelₖ (g south)}) + +-- helper : MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g' ≡ theFuns (g north +ₖ (-ₖ g south)) +-- helper = funExt λ {(inl tt) → refl +-- ; (inr tt) → refl +-- ; (push north i) → refl +-- ; (push south i) → refl} +-- finIm : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int ] (∣ theFuns x ∣₀ ≡ f) +-- finIm f = +-- pRec propTruncIsProp +-- (λ {((a , b) , id) → ∣ b , (sym id ∙ cong (λ x → f +ₕ x) ((λ i → (-ₕ (consMember≡0 a i))) ∙ sym (lUnitₕ (-ₕ 0ₕ)) ∙ rCancelₕ 0ₕ) ∙ (rUnitₕ f)) ∣₋₁}) +-- (f+consMember' f) +-- -} + + + +-- -- Hⁿ-Sⁿ≃Int : (n : ℕ) → grIso intGroup (coHomGr (suc n) (S₊ (suc n))) +-- -- Hⁿ-Sⁿ≃Int zero = +-- -- compGrIso {F = intGroup} {G = {!!}} {H = {!coHomGr 1 (S₊ 1)!}} +-- -- (Iso''→Iso +-- -- (iso'' ((λ x → ∣ theFuns x ∣₀) , λ a b → cong ∣_∣₀ (funExt λ x → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 1) _) ∙ sym (cong (ΩKn+1→Kn) (theFunsId2 a b x)))) +-- -- (λ x inker → pRec (isSetInt _ _) (inj x) (Iso.fun PathIdTrunc₀Iso inker)) +-- -- finIm)) +-- -- {!invGrIso _ _ (coHomPushout≡coHomSn 0 1)!} +-- -- where +-- -- d : _ +-- -- d = MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 + +-- -- i : _ +-- -- i = MV.i _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1 + +-- -- Δ : _ +-- -- Δ = MV.Δ _ _ (S₊ 1) (λ _ → tt) (λ _ → tt) 0 + + +-- -- d-surj : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x +-- -- d-surj = sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- -- λ x → MV.Ker-i⊂Im-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ +-- -- (sym (isContrHelper .snd _)) +-- -- where +-- -- isContrHelper : isContr (Group.type (×coHomGr 1 Unit Unit)) +-- -- isContrHelper = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) +-- -- , λ y → prodId (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₁ y)) +-- -- (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₂ y)) + +-- -- theFuns : (a : Int) → Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 +-- -- theFuns a (inl x) = 0ₖ +-- -- theFuns a (inr x) = 0ₖ +-- -- theFuns a (push north i) = Kn→ΩKn+1 0 a i +-- -- theFuns a (push south i) = 0ₖ + + +-- -- theFunsId2 : (a b : Int) (x : Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) +-- -- → Kn→ΩKn+1 1 (theFuns a x) ∙ Kn→ΩKn+1 1 (theFuns b x) ≡ Kn→ΩKn+1 1 (theFuns (a +ℤ b) x) +-- -- theFunsId2 a b (inl x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) +-- -- theFunsId2 a b (inr x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) +-- -- theFunsId2 a b (push north i) j = +-- -- hcomp (λ k → λ {(i = i0) → ((λ i₁ → +-- -- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- -- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) +-- -- j +-- -- ; (i = i1) → ((λ i₁ → +-- -- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- -- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) +-- -- j +-- -- ; (j = i0) → cong₂Funct2 (λ p q → Kn→ΩKn+1 1 p ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b) (~ k) i +-- -- ; (j = i1) → (helper2 a b) k i }) +-- -- (hcomp (λ k → λ { (j = i0) → compPath-filler (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) k i +-- -- ; (j = i1) → compPath-filler (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) k i +-- -- ; (i = i1) → RHS-filler b j k +-- -- ; (i = i0) → ((λ i₁ → +-- -- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- -- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) +-- -- j}) +-- -- (bottom-filler a j i)) + +-- -- where + +-- -- bottom-filler : (a : Int) → +-- -- PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- -- (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) +-- -- ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) +-- -- j ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- -- (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) +-- -- ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) +-- -- j) (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) +-- -- bottom-filler a j i = +-- -- hcomp (λ k → λ {(j = i0) → helper2 (~ k) i ; +-- -- (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 a) i}) +-- -- ((λ j₂ → ∣ rCancel (merid north) j j₂ ∣) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + +-- -- where +-- -- helper2 : Path (Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣ ≡ Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- -- (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- -- (λ i → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) +-- -- helper2 = cong₂Sym1 (Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) (~ i) j ∣) (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + +-- -- RHS-filler : (b : Int) → +-- -- PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j +-- -- ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j) +-- -- (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) +-- -- (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) +-- -- RHS-filler b j i = +-- -- hcomp (λ k → λ {(j = i0) → cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i ; +-- -- (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 b) i}) +-- -- (cong (λ q → (λ i → ∣ rCancel (merid north) j i ∣) ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i) + +-- -- helper2 : (a b : Int) → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b) ≡ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ℤ b)) +-- -- helper2 a b = +-- -- sym (congFunct (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b)) +-- -- ∙ (λ i → cong (Kn→ΩKn+1 1) (Iso.rightInv (Iso3-Kn-ΩKn+1 0) (Kn→ΩKn+1 0 a ∙ Kn→ΩKn+1 0 b) (~ i))) +-- -- ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ₖ b)) ) +-- -- ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (addLemma a b i))) + +-- -- theFunsId2 a b (push south i) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- -- ∙ sym (lUnit _) + +-- -- inj : (a : Int) → theFuns a ≡ (λ _ → ∣ north ∣) → a ≡ pos 0 +-- -- inj a id = +-- -- pRec (isSetInt _ _) +-- -- (sigmaElim (λ _ → isOfHLevelPath 2 isSetInt _ _) +-- -- (λ a p → pRec (isSetInt _ _) +-- -- (λ id2 → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 0) _) +-- -- ∙ cong (ΩKn+1→Kn) +-- -- (PathP→compPathR +-- -- (cong (λ f → cong f (push north)) id) +-- -- ∙ test)) +-- -- (Iso.fun PathIdTrunc₀Iso p))) (d-surj ∣ theFuns a ∣₀) +-- -- where + +-- -- test : (λ i → id i (inl tt)) ∙ (λ i → ∣ north ∣) ∙ sym (λ i → id i (inr tt)) ≡ refl +-- -- test = (λ i → cong (λ f → f (inl tt)) id +-- -- ∙ lUnit (sym (cong (λ f → f (inr tt)) id)) (~ i)) +-- -- ∙ (λ i → cong (λ f → f (push south i)) id +-- -- ∙ sym (cong (λ f → f (inr tt)) id)) +-- -- ∙ rCancel (cong (λ f → f (inr tt)) id) + + +-- -- consMember : (a : Int) → coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) +-- -- consMember a = d ∣ (λ _ → a) ∣₀ + +-- -- consMember≡0 : (a : Int) → consMember a ≡ 0ₕ +-- -- consMember≡0 a = +-- -- (MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ (λ _ → a) ∣₀ ∣ +-- -- (∣ (λ _ → a) ∣₀ , ∣ (λ _ → 0) ∣₀) +-- -- , cong ∣_∣₀ (λ i x → (rUnitₖ a i)) ∣₋₁) + +-- -- f+consMember' : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int × Int ] (f +ₕ (-ₕ (consMember (proj₁ x))) ≡ ∣ theFuns (proj₂ x) ∣₀) +-- -- f+consMember' = +-- -- sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- -- λ f → pRec propTruncIsProp +-- -- (sigmaElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- -- (λ g id → ∣ ((g south) , ((g north) +ₖ (-ₖ g south))) +-- -- , (pRec (setTruncIsSet _ _) +-- -- (λ id → (λ i → ∣ id (~ i) ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀) ∙ funId1 g) +-- -- (Iso.fun PathIdTrunc₀Iso id)) ∣₋₁)) +-- -- (d-surj ∣ f ∣₀) +-- -- where +-- -- funId1 : (g : S₊ 0 → Int) +-- -- → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀ ≡ +-- -- ∣ theFuns ((g north) +ₖ (-ₖ (g south))) ∣₀ +-- -- funId1 g = (λ i → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +-- -- +ₕ (morphMinus (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) d +-- -- (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ∣ (λ _ → g south) ∣₀ (~ i))) +-- -- ∙ sym (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ g ∣₀ (-ₕ ∣ (λ _ → g south) ∣₀)) +-- -- ∙ (cong (λ x → d ∣ x ∣₀) g'Id) +-- -- ∙ cong ∣_∣₀ helper +-- -- where +-- -- g' : S₊ 0 → Int +-- -- g' north = (g north) +ₖ (-ₖ (g south)) +-- -- g' south = 0 + +-- -- g'Id : (λ x → g x +ₖ (-ₖ (g south))) ≡ g' +-- -- g'Id = funExt (λ {north → refl +-- -- ; south → rCancelₖ (g south)}) + +-- -- helper : MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g' ≡ theFuns (g north +ₖ (-ₖ g south)) +-- -- helper = funExt λ {(inl tt) → refl +-- -- ; (inr tt) → refl +-- -- ; (push north i) → refl +-- -- ; (push south i) → refl} +-- -- finIm : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int ] (∣ theFuns x ∣₀ ≡ f) +-- -- finIm f = +-- -- pRec propTruncIsProp +-- -- (λ {((a , b) , id) → ∣ b , (sym id ∙ cong (λ x → f +ₕ x) ((λ i → (-ₕ (consMember≡0 a i))) ∙ sym (lUnitₕ (-ₕ 0ₕ)) ∙ rCancelₕ 0ₕ) ∙ (rUnitₕ f)) ∣₋₁}) +-- -- (f+consMember' f) +-- -- Hⁿ-Sⁿ≃Int (suc n) = +-- -- compGrIso (Hⁿ-Sⁿ≃Int n) +-- -- (transport (λ i → grIso {!!} {!coHomGr (suc (suc n)) (Pushout (λ _ → tt) (λ _ → tt))!}) {!!}) + + +-- -- {- +-- -- coHom1S1≃ℤ : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- grIso.fun coHom1S1≃ℤ = grIso.fun {!compGrIso coHom1Iso (invGrIso _ _ (ImIso _ _ (d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ?))!} +-- -- grIso.inv coHom1S1≃ℤ = {!!} +-- -- grIso.rightInv coHom1S1≃ℤ = {!!} +-- -- grIso.leftInv coHom1S1≃ℤ = {!!} +-- -- -} + +-- -- -- compGrIso coHom1Iso (invGrIso _ _ (ImIso _ _ {!d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0!} {!!})) + + +-- -- -- coHomGrIm : grIso (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- (imGroup (coHomGr 0 (S₊ 0)) +-- -- -- (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 +-- -- -- , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) +-- -- -- coHomGrIm = +-- -- -- ImIso _ +-- -- -- _ +-- -- -- _ +-- -- -- {!!} + + +-- -- -- -- coHom1RedS1 : Iso (coHom 1 (S₊ 1)) (coHomRed 1 (S₊ 1 , north)) +-- -- -- -- Iso.fun coHom1RedS1 = sRec setTruncIsSet λ f → ∣ f , (pRec {!!} {!!} ((transport (λ i → (b : truncIdempotent 3 {!S₊ 1!} (~ i)) → ∥ (transp (λ j → truncIdempotent {!3!} {!!} (~ i ∨ (~ j))) (~ i) north) ≡ b ∥₋₁) isConnectedS1) (f north)) ) ∣₀ +-- -- -- -- Iso.inv coHom1RedS1 = {!!} +-- -- -- -- Iso.rightInv coHom1RedS1 = {!setTruncIdempotent!} +-- -- -- -- Iso.leftInv coHom1RedS1 = {!!} + +-- -- -- -- coHom1Red : ∀ {ℓ} (A : Pointed ℓ) → Iso (coHom 1 (typ A)) (coHomRed 1 A) +-- -- -- -- Iso.fun (coHom1Red A) = sRec setTruncIsSet λ f → ∣ f , {!!} ∣₀ +-- -- -- -- Iso.inv (coHom1Red A) = {!!} +-- -- -- -- Iso.rightInv (coHom1Red A) = {!!} +-- -- -- -- Iso.leftInv (coHom1Red A) = {!!} + +-- -- -- -- -- morphtest : morph (coHomGr 1 (S₊ 1)) intGroup +-- -- -- -- -- fst morphtest = sRec isSetInt λ f → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (f (loop* i))))) +-- -- -- -- -- snd morphtest = sElim2 {!!} +-- -- -- -- -- (funRec 3 isGroupoidS1 +-- -- -- -- -- λ f → funRec 3 isGroupoidS1 +-- -- -- -- -- λ g → pRec (isSetInt _ _) +-- -- -- -- -- (λ n=fn → +-- -- -- -- -- pRec (isSetInt _ _) +-- -- -- -- -- (λ n=gn → (λ j → winding (basechange (SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (∣ f (north) ∣ +ₖ ∣ n=gn (~ j) ∣)))) (λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (∣ f (loop* i) ∣ +ₖ ∣ transp (λ i → n=gn ((~ i) ∨ (~ j)) ≡ n=gn ((~ i) ∨ (~ j))) (~ j) (λ i → g (loop* i)) i ∣)))))) +-- -- -- -- -- ∙ {!!} +-- -- -- -- -- ∙ {!!}) +-- -- -- -- -- (isConnectedS1 (g north))) +-- -- -- -- -- (isConnectedS1 (f north))) + + +-- -- -- -- -- {- (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (∣ f (loop* i) ∣ +ₖ ∣ g (loop* i) ∣))))) +-- -- -- -- -- ∙ (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (Kn→ΩKn+1 1 ∣ f (loop* i) ∣ ∙ Kn→ΩKn+1 1 ∣ g (loop* i) ∣)))))) +-- -- -- -- -- ∙ (λ j → winding (basechange _ (cong (λ x → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (Kn→ΩKn+1 1 ∣ f x ∣ ∙ Kn→ΩKn+1 1 ∣ g x ∣))))) loop*)) ) +-- -- -- -- -- ∙ (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn ((cong ∣_∣ (merid (f (loop* i)) ∙ sym (merid north)) ∙ cong ∣_∣ (merid (g (loop* i)) ∙ sym (merid north))))))))) +-- -- -- -- -- ∙ (λ j → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (congFunct ∣_∣ (merid (f (loop* i)) ∙ sym (merid north)) (merid (g (loop* i)) ∙ sym (merid north)) (~ j))))))) +-- -- -- -- -- ∙ (λ j → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (cong ∣_∣ (({!!} ∙ {!!}) ∙ {!!}))))))) +-- -- -- -- -- ∙ {!!} +-- -- -- -- -- ∙ {!!} +-- -- -- -- -- ∙ {!!}) -} + +-- -- -- -- -- where +-- -- -- -- -- helper : ∀ {ℓ} {A : Type ℓ} (a : A) (f : A → S¹) (p q : a ≡ a) → winding (basechange (f a) (cong f (p ∙ q))) ≡ winding (basechange (f a) (cong f p ∙ cong f q)) +-- -- -- -- -- helper a f p q i = winding (basechange (f a) (congFunct f p q i)) +-- -- -- -- -- helper2 : (x : S¹) (p q : x ≡ x) → basechange x (p ∙ q) ≡ basechange x p ∙ basechange x q +-- -- -- -- -- helper2 base p q = refl +-- -- -- -- -- helper2 (loop i) p q = {!!} +-- -- -- -- -- helper4 : (x y z : S¹) (p : x ≡ z) (q r : x ≡ y) (s : y ≡ z) → basechange* x z p (q ∙ s) ≡ basechange* x y {!!} q ∙ {!!} +-- -- -- -- -- helper4 = {!!} +-- -- -- -- -- helper3 : (p q : ΩS¹) → winding (p ∙ q) ≡ (winding p +ℤ winding q) +-- -- -- -- -- helper3 = {!!} + + +-- -- -- -- -- -- fstmap : morph (dirProd intGroup intGroup) (coHomGr 0 (S₊ 0)) +-- -- -- -- -- -- fstmap = compMorph {F = dirProd intGroup intGroup} {G = ×coHomGr 0 Unit Unit} {H = coHomGr 0 (S₊ 0)} +-- -- -- -- -- -- (×morph (grIso.inv coHomGrUnit0) (grIso.inv coHomGrUnit0)) +-- -- -- -- -- -- (((MV.Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) zero)) , +-- -- -- -- -- -- {!MV.ΔIsHom _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) zero!}) + +-- -- -- -- -- -- fstMapId : (a : Int × Int) → fstmap .fst a ≡ ∣ (λ _ → proj₁ a +ℤ (0 - proj₂ a)) ∣₀ +-- -- -- -- -- -- fstMapId (a , b) = (λ i → ∣ (λ _ → a +ₖ (-ₖ b)) ∣₀) ∙ {!addLemma!} ∙ {!!} ∙ {!!} + +-- -- -- -- -- -- isoAgain : grIso intGroup (coHomGr 1 (S₊ 1)) +-- -- -- -- -- -- isoAgain = +-- -- -- -- -- -- Iso''→Iso _ _ +-- -- -- -- -- -- (iso'' ((λ a → ∣ (λ {north → 0ₖ ; south → 0ₖ ; (merid north i) → {!a!} ; (merid south i) → {!!}}) ∣₀) , {!!}) +-- -- -- -- -- -- {!!} +-- -- -- -- -- -- {!!}) + +-- -- -- -- -- -- -- -- test2 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → S¹) +-- -- -- -- -- -- -- -- Iso.fun test2 f = {!!} +-- -- -- -- -- -- -- -- Iso.inv test2 f north = ∣ transport (sym S¹≡S1) (f base) ∣ +-- -- -- -- -- -- -- -- Iso.inv test2 f south = ∣ transport (sym S¹≡S1) (f base) ∣ +-- -- -- -- -- -- -- -- Iso.inv test2 f (merid a i) = cong ∣_∣ {!transport (sym S¹≡S1) (f base)!} i +-- -- -- -- -- -- -- -- Iso.rightInv test2 = {!!} +-- -- -- -- -- -- -- -- Iso.leftInv test2 = {!!} + +-- -- -- -- -- -- -- -- F : winding (basechange base loop) ≡ 1 +-- -- -- -- -- -- -- -- F = refl + +-- -- -- -- -- -- -- -- another : (f g : Int) → winding (basechange {!!} {!!}) ≡ {!!} +-- -- -- -- -- -- -- -- another = {!!} + +-- -- -- -- -- -- -- -- test : Iso (S¹ → S¹) (S¹ × Int) +-- -- -- -- -- -- -- -- Iso.fun test f = f base , winding (basechange (f base) (cong f loop)) +-- -- -- -- -- -- -- -- Iso.inv test (x , int) base = x +-- -- -- -- -- -- -- -- Iso.inv test (x , int) (loop i) = {!!} +-- -- -- -- -- -- -- -- Iso.rightInv test = {!!} +-- -- -- -- -- -- -- -- Iso.leftInv test = {!!} + +-- -- -- -- -- -- -- -- -- test13 : Iso ∥ (S¹ → S¹) ∥₀ Int +-- -- -- -- -- -- -- -- -- Iso.fun test13 = sRec isSetInt λ f → winding (basechange (f base) (λ i → f (loop i))) +-- -- -- -- -- -- -- -- -- Iso.inv test13 a = ∣ (λ {base → {!!} ; (loop i) → {!!}}) ∣₀ +-- -- -- -- -- -- -- -- -- Iso.rightInv test13 = {!!} +-- -- -- -- -- -- -- -- -- Iso.leftInv test13 = {!!} + +-- -- -- -- -- -- -- -- -- test : Iso (S¹ → S¹) (S¹ × Int) +-- -- -- -- -- -- -- -- -- Iso.fun test f = (f base) , transport (basedΩS¹≡Int (f base)) λ i → f (loop i) +-- -- -- -- -- -- -- -- -- Iso.inv test (x , int) base = x +-- -- -- -- -- -- -- -- -- Iso.inv test (x , int) (loop i) = transport (sym (basedΩS¹≡Int x)) int i +-- -- -- -- -- -- -- -- -- Iso.rightInv test (x , int) i = (x , transportTransport⁻ (basedΩS¹≡Int x) int i) +-- -- -- -- -- -- -- -- -- Iso.leftInv test f = +-- -- -- -- -- -- -- -- -- funExt λ { base → refl +-- -- -- -- -- -- -- -- -- ; (loop i) j → transport⁻Transport (basedΩS¹≡Int (f base)) (λ i → f (loop i)) j i} + + +-- -- -- -- -- -- -- -- -- lem : S¹ ≡ hLevelTrunc 3 (S₊ 1) +-- -- -- -- -- -- -- -- -- lem = sym (truncIdempotent 3 isGroupoidS¹) ∙ λ i → hLevelTrunc 3 (S¹≡S1 (~ i)) + +-- -- -- -- -- -- -- -- -- prodId : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (a b : A × B) → proj₁ a ≡ proj₁ b → proj₂ a ≡ proj₂ b → a ≡ b +-- -- -- -- -- -- -- -- -- prodId (_ , _) (_ , _) id1 id2 i = (id1 i) , (id2 i) + +-- -- -- -- -- -- -- -- -- test22 : Iso (S₊ 1 → coHomK 1) (S₊ 1 × Int) +-- -- -- -- -- -- -- -- -- Iso.fun test22 f = {!f north!} , {!!} +-- -- -- -- -- -- -- -- -- Iso.inv test22 = {!!} +-- -- -- -- -- -- -- -- -- Iso.rightInv test22 = {!!} +-- -- -- -- -- -- -- -- -- Iso.leftInv test22 = {!!} + + + + + + +-- -- -- -- -- -- -- -- -- coHom1≃∥S1×ℤ∥₀ : Iso (coHom 1 (S₊ 1)) ∥ S₊ 1 × Int ∥₀ +-- -- -- -- -- -- -- -- -- coHom1≃∥S1×ℤ∥₀ = setTruncIso test2 +-- -- -- -- -- -- -- -- -- where +-- -- -- -- -- -- -- -- -- test2 : Iso (S₊ 1 → coHomK 1) (S₊ 1 × Int) +-- -- -- -- -- -- -- -- -- Iso.fun test2 f = transport (λ i → S¹≡S1 (~ i) × Int) (Iso.fun test (transport (λ i → (S¹≡S1 i → lem (~ i))) f)) +-- -- -- -- -- -- -- -- -- Iso.inv test2 x = transport (λ i → (S¹≡S1 (~ i) → lem i)) (Iso.inv test (transport (λ i → S¹≡S1 i × Int) x)) +-- -- -- -- -- -- -- -- -- Iso.rightInv test2 (s , int) = prodId _ _ {!!} {!!} +-- -- -- -- -- -- -- -- -- Iso.leftInv test2 f = {!!} ∙ {!!} ∙ {!!} + +-- -- -- -- -- -- -- -- -- test2Id : (a b : (S₊ 1 → coHomK 1)) → proj₂ (Iso.fun test2 (λ x → a x +ₖ b x)) ≡ (proj₂ (Iso.fun test2 a) +ₖ proj₂ (Iso.fun test2 a)) +-- -- -- -- -- -- -- -- -- test2Id a b = {! +-- -- -- -- -- -- -- -- -- transport +-- -- -- -- -- -- -- -- -- (basedΩS¹≡Int +-- -- -- -- -- -- -- -- -- (transport (λ i → S¹≡S1 i → lem (~ i)) (λ x → a x +ₖ b x) base)) +-- -- -- -- -- -- -- -- -- (λ i → +-- -- -- -- -- -- -- -- -- transport (λ i₁ → S¹≡S1 i₁ → lem (~ i₁)) (λ x → a x +ₖ b x) +-- -- -- -- -- -- -- -- -- (loop i))!} ∙ {!transport (λ i → S¹≡S1 i → lem (~ i)) (λ x → a x +ₖ b x) base!} + + +-- -- -- -- -- -- -- -- -- main : grIso intGroup (coHomGr 1 (S₊ 1)) +-- -- -- -- -- -- -- -- -- main = Iso'→Iso +-- -- -- -- -- -- -- -- -- (iso' {!!} +-- -- -- -- -- -- -- -- -- {!!}) -{- -coHom1S1≃ℤ : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) -grIso.fun coHom1S1≃ℤ = grIso.fun {!compGrIso coHom1Iso (invGrIso _ _ (ImIso _ _ (d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ?))!} -grIso.inv coHom1S1≃ℤ = {!!} -grIso.rightInv coHom1S1≃ℤ = {!!} -grIso.leftInv coHom1S1≃ℤ = {!!} --} --- compGrIso coHom1Iso (invGrIso _ _ (ImIso _ _ {!d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0!} {!!})) - - --- coHomGrIm : grIso (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- (imGroup (coHomGr 0 (S₊ 0)) --- (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 --- , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) --- coHomGrIm = --- ImIso _ --- _ --- _ --- {!!} - - --- -- coHom1RedS1 : Iso (coHom 1 (S₊ 1)) (coHomRed 1 (S₊ 1 , north)) --- -- Iso.fun coHom1RedS1 = sRec setTruncIsSet λ f → ∣ f , (pRec {!!} {!!} ((transport (λ i → (b : truncIdempotent 3 {!S₊ 1!} (~ i)) → ∥ (transp (λ j → truncIdempotent {!3!} {!!} (~ i ∨ (~ j))) (~ i) north) ≡ b ∥₋₁) isConnectedS1) (f north)) ) ∣₀ --- -- Iso.inv coHom1RedS1 = {!!} --- -- Iso.rightInv coHom1RedS1 = {!setTruncIdempotent!} --- -- Iso.leftInv coHom1RedS1 = {!!} - --- -- coHom1Red : ∀ {ℓ} (A : Pointed ℓ) → Iso (coHom 1 (typ A)) (coHomRed 1 A) --- -- Iso.fun (coHom1Red A) = sRec setTruncIsSet λ f → ∣ f , {!!} ∣₀ --- -- Iso.inv (coHom1Red A) = {!!} --- -- Iso.rightInv (coHom1Red A) = {!!} --- -- Iso.leftInv (coHom1Red A) = {!!} - --- -- -- morphtest : morph (coHomGr 1 (S₊ 1)) intGroup --- -- -- fst morphtest = sRec isSetInt λ f → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (f (loop* i))))) --- -- -- snd morphtest = sElim2 {!!} --- -- -- (funRec 3 isGroupoidS1 --- -- -- λ f → funRec 3 isGroupoidS1 --- -- -- λ g → pRec (isSetInt _ _) --- -- -- (λ n=fn → --- -- -- pRec (isSetInt _ _) --- -- -- (λ n=gn → (λ j → winding (basechange (SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (∣ f (north) ∣ +ₖ ∣ n=gn (~ j) ∣)))) (λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (∣ f (loop* i) ∣ +ₖ ∣ transp (λ i → n=gn ((~ i) ∨ (~ j)) ≡ n=gn ((~ i) ∨ (~ j))) (~ j) (λ i → g (loop* i)) i ∣)))))) --- -- -- ∙ {!!} --- -- -- ∙ {!!}) --- -- -- (isConnectedS1 (g north))) --- -- -- (isConnectedS1 (f north))) - - --- -- -- {- (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (∣ f (loop* i) ∣ +ₖ ∣ g (loop* i) ∣))))) --- -- -- ∙ (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (Kn→ΩKn+1 1 ∣ f (loop* i) ∣ ∙ Kn→ΩKn+1 1 ∣ g (loop* i) ∣)))))) --- -- -- ∙ (λ j → winding (basechange _ (cong (λ x → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (Kn→ΩKn+1 1 ∣ f x ∣ ∙ Kn→ΩKn+1 1 ∣ g x ∣))))) loop*)) ) --- -- -- ∙ (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn ((cong ∣_∣ (merid (f (loop* i)) ∙ sym (merid north)) ∙ cong ∣_∣ (merid (g (loop* i)) ∙ sym (merid north))))))))) --- -- -- ∙ (λ j → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (congFunct ∣_∣ (merid (f (loop* i)) ∙ sym (merid north)) (merid (g (loop* i)) ∙ sym (merid north)) (~ j))))))) --- -- -- ∙ (λ j → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (cong ∣_∣ (({!!} ∙ {!!}) ∙ {!!}))))))) --- -- -- ∙ {!!} --- -- -- ∙ {!!} --- -- -- ∙ {!!}) -} - --- -- -- where --- -- -- helper : ∀ {ℓ} {A : Type ℓ} (a : A) (f : A → S¹) (p q : a ≡ a) → winding (basechange (f a) (cong f (p ∙ q))) ≡ winding (basechange (f a) (cong f p ∙ cong f q)) --- -- -- helper a f p q i = winding (basechange (f a) (congFunct f p q i)) --- -- -- helper2 : (x : S¹) (p q : x ≡ x) → basechange x (p ∙ q) ≡ basechange x p ∙ basechange x q --- -- -- helper2 base p q = refl --- -- -- helper2 (loop i) p q = {!!} --- -- -- helper4 : (x y z : S¹) (p : x ≡ z) (q r : x ≡ y) (s : y ≡ z) → basechange* x z p (q ∙ s) ≡ basechange* x y {!!} q ∙ {!!} --- -- -- helper4 = {!!} --- -- -- helper3 : (p q : ΩS¹) → winding (p ∙ q) ≡ (winding p +ℤ winding q) --- -- -- helper3 = {!!} - - --- -- -- -- fstmap : morph (dirProd intGroup intGroup) (coHomGr 0 (S₊ 0)) --- -- -- -- fstmap = compMorph {F = dirProd intGroup intGroup} {G = ×coHomGr 0 Unit Unit} {H = coHomGr 0 (S₊ 0)} --- -- -- -- (×morph (grIso.inv coHomGrUnit0) (grIso.inv coHomGrUnit0)) --- -- -- -- (((MV.Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) zero)) , --- -- -- -- {!MV.ΔIsHom _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) zero!}) - --- -- -- -- fstMapId : (a : Int × Int) → fstmap .fst a ≡ ∣ (λ _ → proj₁ a +ℤ (0 - proj₂ a)) ∣₀ --- -- -- -- fstMapId (a , b) = (λ i → ∣ (λ _ → a +ₖ (-ₖ b)) ∣₀) ∙ {!addLemma!} ∙ {!!} ∙ {!!} - --- -- -- -- isoAgain : grIso intGroup (coHomGr 1 (S₊ 1)) --- -- -- -- isoAgain = --- -- -- -- Iso''→Iso _ _ --- -- -- -- (iso'' ((λ a → ∣ (λ {north → 0ₖ ; south → 0ₖ ; (merid north i) → {!a!} ; (merid south i) → {!!}}) ∣₀) , {!!}) --- -- -- -- {!!} --- -- -- -- {!!}) - --- -- -- -- -- -- test2 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → S¹) --- -- -- -- -- -- Iso.fun test2 f = {!!} --- -- -- -- -- -- Iso.inv test2 f north = ∣ transport (sym S¹≡S1) (f base) ∣ --- -- -- -- -- -- Iso.inv test2 f south = ∣ transport (sym S¹≡S1) (f base) ∣ --- -- -- -- -- -- Iso.inv test2 f (merid a i) = cong ∣_∣ {!transport (sym S¹≡S1) (f base)!} i --- -- -- -- -- -- Iso.rightInv test2 = {!!} --- -- -- -- -- -- Iso.leftInv test2 = {!!} - --- -- -- -- -- -- F : winding (basechange base loop) ≡ 1 --- -- -- -- -- -- F = refl - --- -- -- -- -- -- another : (f g : Int) → winding (basechange {!!} {!!}) ≡ {!!} --- -- -- -- -- -- another = {!!} - --- -- -- -- -- -- test : Iso (S¹ → S¹) (S¹ × Int) --- -- -- -- -- -- Iso.fun test f = f base , winding (basechange (f base) (cong f loop)) --- -- -- -- -- -- Iso.inv test (x , int) base = x --- -- -- -- -- -- Iso.inv test (x , int) (loop i) = {!!} --- -- -- -- -- -- Iso.rightInv test = {!!} --- -- -- -- -- -- Iso.leftInv test = {!!} - --- -- -- -- -- -- -- test13 : Iso ∥ (S¹ → S¹) ∥₀ Int --- -- -- -- -- -- -- Iso.fun test13 = sRec isSetInt λ f → winding (basechange (f base) (λ i → f (loop i))) --- -- -- -- -- -- -- Iso.inv test13 a = ∣ (λ {base → {!!} ; (loop i) → {!!}}) ∣₀ --- -- -- -- -- -- -- Iso.rightInv test13 = {!!} --- -- -- -- -- -- -- Iso.leftInv test13 = {!!} - --- -- -- -- -- -- -- test : Iso (S¹ → S¹) (S¹ × Int) --- -- -- -- -- -- -- Iso.fun test f = (f base) , transport (basedΩS¹≡Int (f base)) λ i → f (loop i) --- -- -- -- -- -- -- Iso.inv test (x , int) base = x --- -- -- -- -- -- -- Iso.inv test (x , int) (loop i) = transport (sym (basedΩS¹≡Int x)) int i --- -- -- -- -- -- -- Iso.rightInv test (x , int) i = (x , transportTransport⁻ (basedΩS¹≡Int x) int i) --- -- -- -- -- -- -- Iso.leftInv test f = --- -- -- -- -- -- -- funExt λ { base → refl --- -- -- -- -- -- -- ; (loop i) j → transport⁻Transport (basedΩS¹≡Int (f base)) (λ i → f (loop i)) j i} - - --- -- -- -- -- -- -- lem : S¹ ≡ hLevelTrunc 3 (S₊ 1) --- -- -- -- -- -- -- lem = sym (truncIdempotent 3 isGroupoidS¹) ∙ λ i → hLevelTrunc 3 (S¹≡S1 (~ i)) - --- -- -- -- -- -- -- prodId : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (a b : A × B) → proj₁ a ≡ proj₁ b → proj₂ a ≡ proj₂ b → a ≡ b --- -- -- -- -- -- -- prodId (_ , _) (_ , _) id1 id2 i = (id1 i) , (id2 i) - --- -- -- -- -- -- -- test22 : Iso (S₊ 1 → coHomK 1) (S₊ 1 × Int) --- -- -- -- -- -- -- Iso.fun test22 f = {!f north!} , {!!} --- -- -- -- -- -- -- Iso.inv test22 = {!!} --- -- -- -- -- -- -- Iso.rightInv test22 = {!!} --- -- -- -- -- -- -- Iso.leftInv test22 = {!!} - - - - - - --- -- -- -- -- -- -- coHom1≃∥S1×ℤ∥₀ : Iso (coHom 1 (S₊ 1)) ∥ S₊ 1 × Int ∥₀ --- -- -- -- -- -- -- coHom1≃∥S1×ℤ∥₀ = setTruncIso test2 --- -- -- -- -- -- -- where --- -- -- -- -- -- -- test2 : Iso (S₊ 1 → coHomK 1) (S₊ 1 × Int) --- -- -- -- -- -- -- Iso.fun test2 f = transport (λ i → S¹≡S1 (~ i) × Int) (Iso.fun test (transport (λ i → (S¹≡S1 i → lem (~ i))) f)) --- -- -- -- -- -- -- Iso.inv test2 x = transport (λ i → (S¹≡S1 (~ i) → lem i)) (Iso.inv test (transport (λ i → S¹≡S1 i × Int) x)) --- -- -- -- -- -- -- Iso.rightInv test2 (s , int) = prodId _ _ {!!} {!!} --- -- -- -- -- -- -- Iso.leftInv test2 f = {!!} ∙ {!!} ∙ {!!} - --- -- -- -- -- -- -- test2Id : (a b : (S₊ 1 → coHomK 1)) → proj₂ (Iso.fun test2 (λ x → a x +ₖ b x)) ≡ (proj₂ (Iso.fun test2 a) +ₖ proj₂ (Iso.fun test2 a)) --- -- -- -- -- -- -- test2Id a b = {! --- -- -- -- -- -- -- transport --- -- -- -- -- -- -- (basedΩS¹≡Int --- -- -- -- -- -- -- (transport (λ i → S¹≡S1 i → lem (~ i)) (λ x → a x +ₖ b x) base)) --- -- -- -- -- -- -- (λ i → --- -- -- -- -- -- -- transport (λ i₁ → S¹≡S1 i₁ → lem (~ i₁)) (λ x → a x +ₖ b x) --- -- -- -- -- -- -- (loop i))!} ∙ {!transport (λ i → S¹≡S1 i → lem (~ i)) (λ x → a x +ₖ b x) base!} - - --- -- -- -- -- -- -- main : grIso intGroup (coHomGr 1 (S₊ 1)) --- -- -- -- -- -- -- main = Iso'→Iso --- -- -- -- -- -- -- (iso' {!!} --- -- -- -- -- -- -- {!!}) - - --- -- -- -- -- coHom1 : grIso (coHomGr 1 (S₊ 1)) intGroup --- -- -- -- -- coHom1 = --- -- -- -- -- Iso'→Iso --- -- -- -- -- (iso' (iso ({!!} ∘ {!!} ∘ {!!} ∘ {!!}) --- -- -- -- -- {!!} --- -- -- -- -- {!!} --- -- -- -- -- {!!}) --- -- -- -- -- {!!}) - - - --- -- -- -- -- schonf : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : (A × B) → Type ℓ''} --- -- -- -- -- → ((a : A) (b : B) → C (a , b)) --- -- -- -- -- → (x : A × B) → C x --- -- -- -- -- schonf f (a , b) = f a b - --- -- -- -- -- -- -- setTruncProdIso : ∀ {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') → Iso ∥ (A × B) ∥₀ (∥ A ∥₀ × ∥ B ∥₀) --- -- -- -- -- -- -- Iso.fun (setTruncProdIso A B) = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) λ {(a , b) → ∣ a ∣₀ , ∣ b ∣₀} --- -- -- -- -- -- -- Iso.inv (setTruncProdIso A B) (a , b) = sRec setTruncIsSet (λ a → sRec setTruncIsSet (λ b → ∣ a , b ∣₀) b) a --- -- -- -- -- -- -- Iso.rightInv (setTruncProdIso A B) = --- -- -- -- -- -- -- schonf (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) --- -- -- -- -- -- -- λ _ → sElim (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) --- -- -- -- -- -- -- λ _ → refl) --- -- -- -- -- -- -- Iso.leftInv (setTruncProdIso A B) = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ {(a , b) → refl} - --- -- -- -- -- -- -- setTruncProdLemma : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (a1 a2 : A) (b : B) → isHLevelConnected 2 A --- -- -- -- -- -- -- → Path ∥ A × B ∥₀ ∣ a1 , b ∣₀ ∣ a2 , b ∣₀ --- -- -- -- -- -- -- setTruncProdLemma {A = A} {B = B} a1 a2 b conA i = Iso.inv (setTruncProdIso A B) (Iso.inv setTruncTrunc0Iso ((sym (conA .snd ∣ a1 ∣) ∙ (conA .snd ∣ a2 ∣)) i) , ∣ b ∣₀) - --- -- -- -- -- -- -- test3 : Iso ∥ S₊ 1 × Int ∥₀ Int --- -- -- -- -- -- -- Iso.fun test3 = sRec isSetInt proj₂ --- -- -- -- -- -- -- Iso.inv test3 a = ∣ north , a ∣₀ --- -- -- -- -- -- -- Iso.rightInv test3 a = refl --- -- -- -- -- -- -- Iso.leftInv test3 = --- -- -- -- -- -- -- sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) --- -- -- -- -- -- -- λ {(s , int) → setTruncProdLemma north s int (sphereConnected 1)} - --- -- -- -- -- -- -- coHomGr0-S1 : grIso intGroup (coHomGr 1 (S₊ 1)) --- -- -- -- -- -- -- coHomGr0-S1 = +-- -- -- -- -- -- -- coHom1 : grIso (coHomGr 1 (S₊ 1)) intGroup +-- -- -- -- -- -- -- coHom1 = -- -- -- -- -- -- -- Iso'→Iso --- -- -- -- -- -- -- (iso' (compIso (symIso test3) (symIso coHom1≃∥S1×ℤ∥₀)) --- -- -- -- -- -- -- (indIntGroup {G = coHomGr 1 (S₊ 1)} --- -- -- -- -- -- -- (Iso.fun (compIso (symIso test3) (symIso coHom1≃∥S1×ℤ∥₀))) --- -- -- -- -- -- -- ((λ i → ∣ transport (λ i → (S¹≡S1 (~ i) → lem i)) (Iso.inv test (base , 0)) ∣₀) --- -- -- -- -- -- -- ∙ (λ i → ∣ transport (λ i → (S¹≡S1 (~ i) → lem i)) (helper2 i) ∣₀) --- -- -- -- -- -- -- ∙ cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → {!!}})) --- -- -- -- -- -- -- {!!} --- -- -- -- -- -- -- {!!})) +-- -- -- -- -- -- -- (iso' (iso ({!!} ∘ {!!} ∘ {!!} ∘ {!!}) +-- -- -- -- -- -- -- {!!} +-- -- -- -- -- -- -- {!!} +-- -- -- -- -- -- -- {!!}) +-- -- -- -- -- -- -- {!!}) + + + +-- -- -- -- -- -- -- schonf : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : (A × B) → Type ℓ''} +-- -- -- -- -- -- -- → ((a : A) (b : B) → C (a , b)) +-- -- -- -- -- -- -- → (x : A × B) → C x +-- -- -- -- -- -- -- schonf f (a , b) = f a b + +-- -- -- -- -- -- -- -- -- setTruncProdIso : ∀ {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') → Iso ∥ (A × B) ∥₀ (∥ A ∥₀ × ∥ B ∥₀) +-- -- -- -- -- -- -- -- -- Iso.fun (setTruncProdIso A B) = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) λ {(a , b) → ∣ a ∣₀ , ∣ b ∣₀} +-- -- -- -- -- -- -- -- -- Iso.inv (setTruncProdIso A B) (a , b) = sRec setTruncIsSet (λ a → sRec setTruncIsSet (λ b → ∣ a , b ∣₀) b) a +-- -- -- -- -- -- -- -- -- Iso.rightInv (setTruncProdIso A B) = +-- -- -- -- -- -- -- -- -- schonf (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- -- -- -- -- -- -- -- -- λ _ → sElim (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- -- -- -- -- -- -- -- -- λ _ → refl) +-- -- -- -- -- -- -- -- -- Iso.leftInv (setTruncProdIso A B) = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ {(a , b) → refl} + +-- -- -- -- -- -- -- -- -- setTruncProdLemma : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (a1 a2 : A) (b : B) → isHLevelConnected 2 A +-- -- -- -- -- -- -- -- -- → Path ∥ A × B ∥₀ ∣ a1 , b ∣₀ ∣ a2 , b ∣₀ +-- -- -- -- -- -- -- -- -- setTruncProdLemma {A = A} {B = B} a1 a2 b conA i = Iso.inv (setTruncProdIso A B) (Iso.inv setTruncTrunc0Iso ((sym (conA .snd ∣ a1 ∣) ∙ (conA .snd ∣ a2 ∣)) i) , ∣ b ∣₀) + +-- -- -- -- -- -- -- -- -- test3 : Iso ∥ S₊ 1 × Int ∥₀ Int +-- -- -- -- -- -- -- -- -- Iso.fun test3 = sRec isSetInt proj₂ +-- -- -- -- -- -- -- -- -- Iso.inv test3 a = ∣ north , a ∣₀ +-- -- -- -- -- -- -- -- -- Iso.rightInv test3 a = refl +-- -- -- -- -- -- -- -- -- Iso.leftInv test3 = +-- -- -- -- -- -- -- -- -- sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- -- -- -- -- -- -- -- -- λ {(s , int) → setTruncProdLemma north s int (sphereConnected 1)} + +-- -- -- -- -- -- -- -- -- coHomGr0-S1 : grIso intGroup (coHomGr 1 (S₊ 1)) +-- -- -- -- -- -- -- -- -- coHomGr0-S1 = +-- -- -- -- -- -- -- -- -- Iso'→Iso +-- -- -- -- -- -- -- -- -- (iso' (compIso (symIso test3) (symIso coHom1≃∥S1×ℤ∥₀)) +-- -- -- -- -- -- -- -- -- (indIntGroup {G = coHomGr 1 (S₊ 1)} +-- -- -- -- -- -- -- -- -- (Iso.fun (compIso (symIso test3) (symIso coHom1≃∥S1×ℤ∥₀))) +-- -- -- -- -- -- -- -- -- ((λ i → ∣ transport (λ i → (S¹≡S1 (~ i) → lem i)) (Iso.inv test (base , 0)) ∣₀) +-- -- -- -- -- -- -- -- -- ∙ (λ i → ∣ transport (λ i → (S¹≡S1 (~ i) → lem i)) (helper2 i) ∣₀) +-- -- -- -- -- -- -- -- -- ∙ cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → {!!}})) +-- -- -- -- -- -- -- -- -- {!!} +-- -- -- -- -- -- -- -- -- {!!})) +-- -- -- -- -- -- -- -- -- where +-- -- -- -- -- -- -- -- -- helper : basedΩS¹≡ΩS¹ base ≡ {!basechange!} +-- -- -- -- -- -- -- -- -- helper = {!substComposite!} + +-- -- -- -- -- -- -- -- -- substComposite2 : ∀ {ℓ} {A B C : Type ℓ} +-- -- -- -- -- -- -- -- -- (P : A ≡ B) (Q : B ≡ C) (a : A) +-- -- -- -- -- -- -- -- -- → transport (P ∙ Q) a ≡ transport Q (transport P a) +-- -- -- -- -- -- -- -- -- substComposite2 = {!!} + +-- -- -- -- -- -- -- -- -- helper1 : transport (λ i → S¹≡S1 i × Int) (north , 0) ≡ (base , 0) +-- -- -- -- -- -- -- -- -- helper1 = refl +-- -- -- -- -- -- -- -- -- helper3 : transport (sym (basedΩS¹≡Int base)) 0 ≡ refl +-- -- -- -- -- -- -- -- -- helper3 = (λ i → transport (symDistr (basedΩS¹≡ΩS¹ base) ΩS¹≡Int i) 0) +-- -- -- -- -- -- -- -- -- ∙ substComposite2 (sym ΩS¹≡Int) (sym (basedΩS¹≡ΩS¹ base)) 0 +-- -- -- -- -- -- -- -- -- ∙ (λ i → transport (λ i → basedΩS¹≡ΩS¹ base (~ i)) refl) -- +-- -- -- -- -- -- -- -- -- ∙ transportRefl ((equiv-proof (basechange-isequiv base) refl) .fst .fst) +-- -- -- -- -- -- -- -- -- ∙ (λ i → equiv-proof (transport (λ j → isEquiv (refl-conjugation j)) (basedΩS¹→ΩS¹-isequiv i0)) refl .fst .fst) +-- -- -- -- -- -- -- -- -- ∙ (λ i → {!equiv-proof (transport (λ j → isEquiv (refl-conjugation j)) (basedΩS¹→ΩS¹-isequiv i0)) refl .fst!}) +-- -- -- -- -- -- -- -- -- ∙ {!basedΩS¹→ΩS¹!} +-- -- -- -- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- -- -- -- helper4 : (x : S¹) → Iso.inv test (base , 0) x ≡ x +-- -- -- -- -- -- -- -- -- helper4 = {!!} +-- -- -- -- -- -- -- -- -- helper2 : Iso.inv test (transport (λ i → S¹≡S1 i × Int) (north , 0)) ≡ λ x → x +-- -- -- -- -- -- -- -- -- helper2 = (λ i → Iso.inv test (base , 0)) ∙ {!!} ∙ {!!} + +-- -- -- -- -- -- -- prodId : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A × B} → proj₁ x ≡ proj₁ y → proj₂ x ≡ proj₂ y → x ≡ y +-- -- -- -- -- -- -- prodId {x = (a , b)} {y = (c , d)} id1 id2 i = (id1 i) , (id2 i) + +-- -- -- -- -- -- -- fstFunHelper : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- -- -- → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 _) (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x +-- -- -- -- -- -- -- fstFunHelper a = MV.Ker-i⊂Im-d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a +-- -- -- -- -- -- -- (sym (isContrH1Unit×H1Unit .snd _) ∙ (isContrH1Unit×H1Unit .snd _)) +-- -- -- -- -- -- -- where +-- -- -- -- -- -- -- isContrH1Unit×H1Unit : isContr (Group.type (×coHomGr 1 Unit Unit)) +-- -- -- -- -- -- -- isContrH1Unit×H1Unit = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) +-- -- -- -- -- -- -- , λ {(a , b) → sigmaProdElim {D = λ (x : Σ[ x ∈ Group.type (×coHomGr 1 Unit Unit)] Unit) → (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) ≡ fst x} +-- -- -- -- -- -- -- (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- -- -- -- -- -- -- (λ a b _ → prodId (grIso.leftInv (coHomGrUnit≥1 0) ∣ a ∣₀) (grIso.leftInv (coHomGrUnit≥1 0) ∣ b ∣₀)) ((a , b) , tt)} + + + +-- -- -- -- -- -- -- helperMorph : isMorph intGroup (dirProd intGroup intGroup) λ a → a , (0 - a) +-- -- -- -- -- -- -- helperMorph = +-- -- -- -- -- -- -- indIntGroup {G = dirProd intGroup intGroup} +-- -- -- -- -- -- -- (λ a → a , (0 - a)) +-- -- -- -- -- -- -- refl +-- -- -- -- -- -- -- (λ a → prodId refl (helper2 a)) +-- -- -- -- -- -- -- λ a → prodId refl (helper3 a) +-- -- -- -- -- -- -- where +-- -- -- -- -- -- -- helper1 : (a : Int) → predInt (sucInt a) ≡ a +-- -- -- -- -- -- -- helper1 (pos zero) = refl +-- -- -- -- -- -- -- helper1 (pos (suc n)) = refl +-- -- -- -- -- -- -- helper1 (negsuc zero) = refl +-- -- -- -- -- -- -- helper1 (negsuc (suc n)) = refl + +-- -- -- -- -- -- -- helper4 : (a : Int) → sucInt (predInt a) ≡ a +-- -- -- -- -- -- -- helper4 (pos zero) = refl +-- -- -- -- -- -- -- helper4 (pos (suc n)) = refl +-- -- -- -- -- -- -- helper4 (negsuc zero) = refl +-- -- -- -- -- -- -- helper4 (negsuc (suc n)) = refl + +-- -- -- -- -- -- -- helper2 : (a : Int) → (pos 0 - sucInt a) ≡ predInt (pos 0 - a) +-- -- -- -- -- -- -- helper2 (pos zero) = refl +-- -- -- -- -- -- -- helper2 (pos (suc n)) = refl +-- -- -- -- -- -- -- helper2 (negsuc zero) = refl +-- -- -- -- -- -- -- helper2 (negsuc (suc n)) = sym (helper1 _) + +-- -- -- -- -- -- -- helper3 : (a : Int) → (pos 0 - predInt a) ≡ sucInt (pos 0 - a) +-- -- -- -- -- -- -- helper3 (pos zero) = refl +-- -- -- -- -- -- -- helper3 (pos (suc zero)) = refl +-- -- -- -- -- -- -- helper3 (pos (suc (suc n))) = sym (helper4 _) +-- -- -- -- -- -- -- helper3 (negsuc zero) = refl +-- -- -- -- -- -- -- helper3 (negsuc (suc n)) = refl + +-- -- -- -- -- -- -- helper : (a b : Int) → (pos 0 - (a +ℤ b)) ≡ ((pos 0 - a) +ℤ (pos 0 - b)) +-- -- -- -- -- -- -- helper a (pos zero) = refl +-- -- -- -- -- -- -- helper (pos zero) (pos (suc n)) = +-- -- -- -- -- -- -- cong (λ x → pos 0 - sucInt x) (+ℤ-comm (pos zero) (pos n)) +-- -- -- -- -- -- -- ∙ +ℤ-comm (pos 0 +negsuc n) (pos zero) +-- -- -- -- -- -- -- helper (pos (suc n₁)) (pos (suc n)) = +-- -- -- -- -- -- -- {!!} +-- -- -- -- -- -- -- helper (negsuc n₁) (pos (suc n)) = {!!} +-- -- -- -- -- -- -- helper a (negsuc n) = {!!} + +-- -- -- -- -- -- -- fun : morph intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- -- -- fst fun = MV.d _ _ _ (λ _ → tt) (λ _ → tt) 0 ∘ grIso.inv coHom0-S0 .fst ∘ λ a → a , (0 - a) +-- -- -- -- -- -- -- snd fun = {!!} +-- -- -- -- -- -- -- {- compMorph {F = intGroup} {G = dirProd intGroup intGroup} {H = coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))} +-- -- -- -- -- -- -- ((λ a → a , a) , (λ a b → refl)) +-- -- -- -- -- -- -- (compMorph {F = dirProd intGroup intGroup} {G = coHomGr 0 (S₊ 0)} {H = coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))} (grIso.inv coHom0-S0) +-- -- -- -- -- -- -- (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 +-- -- -- -- -- -- -- , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) .snd -} +-- -- -- -- -- -- -- {- theIso : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- -- -- theIso = Iso''→Iso _ _ +-- -- -- -- -- -- -- (iso'' ((λ x → ∣ (λ {(inl tt) → 0ₖ ; (inr tt) → 0ₖ ; (push a i) → Kn→ΩKn+1 0 x i}) ∣₀) , {!!}) +-- -- -- -- -- -- -- {!!} +-- -- -- -- -- -- -- {!MV.d!}) +-- -- -- -- -- -- -- -} + + + +-- -- -- -- -- -- -- theIso : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- -- -- theIso = +-- -- -- -- -- -- -- Iso''→Iso _ _ +-- -- -- -- -- -- -- (iso'' fun +-- -- -- -- -- -- -- (λ x inker → {!MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 +-- -- -- -- -- -- -- (grIso.inv coHom0-S0 .fst (g , g))!}) +-- -- -- -- -- -- -- (sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- -- -- -- -- -- -- λ x → pRec propTruncIsProp +-- -- -- -- -- -- -- (λ {(a , b) → {!fun!} }) +-- -- -- -- -- -- -- (fstFunHelper (∣ x ∣₀)))) +-- -- -- -- -- -- -- where +-- -- -- -- -- -- -- whoKnows : (a : S₊ 0 → Int) (x : MV.D Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt)) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (λ _ → a north) x +-- -- -- -- -- -- -- ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 a x +-- -- -- -- -- -- -- whoKnows a (inl x) = refl +-- -- -- -- -- -- -- whoKnows a (inr x) = refl +-- -- -- -- -- -- -- whoKnows a (push north i) = refl +-- -- -- -- -- -- -- whoKnows a (push south i) j = {!!} + +-- -- -- -- -- -- -- helper : (a : Int) → (grIso.inv coHom0-S0 .fst (a , a)) ≡ ∣ S0→Int (a , a) ∣₀ +-- -- -- -- -- -- -- helper a = {!have : + +-- -- -- -- -- -- -- ∣ +-- -- -- -- -- -- -- MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 +-- -- -- -- -- -- -- (S0→Int (x , x)) +-- -- -- -- -- -- -- ∣₀ +-- -- -- -- -- -- -- ≡ ∣ (λ _ → ∣ north ∣) ∣₀!} + +-- -- -- -- -- -- -- helper2 : (a b : Int) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a)) ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (b , b)) +-- -- -- -- -- -- -- → a ≡ b +-- -- -- -- -- -- -- helper2 a b id = pRec (isSetInt a b) (λ {(pt , id) → {!!}}) (fstFunHelper ∣ (MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a))) ∣₀) + +-- -- -- -- -- -- -- idFun : (a : Int) → MV.D Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 +-- -- -- -- -- -- -- idFun a (inl x) = 0ₖ +-- -- -- -- -- -- -- idFun a (inr x) = 0ₖ +-- -- -- -- -- -- -- idFun a (push north i) = Kn→ΩKn+1 zero a i +-- -- -- -- -- -- -- idFun a (push south i) = Kn→ΩKn+1 zero a i + +-- -- -- -- -- -- -- helper3 : (a : Int) → (MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a))) ≡ idFun a +-- -- -- -- -- -- -- helper3 a = funExt λ {(inl x) → refl ; (inr x) → refl ; (push north i) → refl ; (push south i) → refl } + +-- -- -- -- -- -- -- helper4 : (a b : Int) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a)) ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (b , b)) +-- -- -- -- -- -- -- → a ≡ b +-- -- -- -- -- -- -- helper4 a b id = +-- -- -- -- -- -- -- {!!} +-- -- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- -- ∙ {!!} -- -- -- -- -- -- -- where --- -- -- -- -- -- -- helper : basedΩS¹≡ΩS¹ base ≡ {!basechange!} --- -- -- -- -- -- -- helper = {!substComposite!} - --- -- -- -- -- -- -- substComposite2 : ∀ {ℓ} {A B C : Type ℓ} --- -- -- -- -- -- -- (P : A ≡ B) (Q : B ≡ C) (a : A) --- -- -- -- -- -- -- → transport (P ∙ Q) a ≡ transport Q (transport P a) --- -- -- -- -- -- -- substComposite2 = {!!} - --- -- -- -- -- -- -- helper1 : transport (λ i → S¹≡S1 i × Int) (north , 0) ≡ (base , 0) --- -- -- -- -- -- -- helper1 = refl --- -- -- -- -- -- -- helper3 : transport (sym (basedΩS¹≡Int base)) 0 ≡ refl --- -- -- -- -- -- -- helper3 = (λ i → transport (symDistr (basedΩS¹≡ΩS¹ base) ΩS¹≡Int i) 0) --- -- -- -- -- -- -- ∙ substComposite2 (sym ΩS¹≡Int) (sym (basedΩS¹≡ΩS¹ base)) 0 --- -- -- -- -- -- -- ∙ (λ i → transport (λ i → basedΩS¹≡ΩS¹ base (~ i)) refl) -- --- -- -- -- -- -- -- ∙ transportRefl ((equiv-proof (basechange-isequiv base) refl) .fst .fst) --- -- -- -- -- -- -- ∙ (λ i → equiv-proof (transport (λ j → isEquiv (refl-conjugation j)) (basedΩS¹→ΩS¹-isequiv i0)) refl .fst .fst) --- -- -- -- -- -- -- ∙ (λ i → {!equiv-proof (transport (λ j → isEquiv (refl-conjugation j)) (basedΩS¹→ΩS¹-isequiv i0)) refl .fst!}) --- -- -- -- -- -- -- ∙ {!basedΩS¹→ΩS¹!} --- -- -- -- -- -- -- ∙ {!!} --- -- -- -- -- -- -- ∙ {!!} --- -- -- -- -- -- -- helper4 : (x : S¹) → Iso.inv test (base , 0) x ≡ x --- -- -- -- -- -- -- helper4 = {!!} --- -- -- -- -- -- -- helper2 : Iso.inv test (transport (λ i → S¹≡S1 i × Int) (north , 0)) ≡ λ x → x --- -- -- -- -- -- -- helper2 = (λ i → Iso.inv test (base , 0)) ∙ {!!} ∙ {!!} - --- -- -- -- -- prodId : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A × B} → proj₁ x ≡ proj₁ y → proj₂ x ≡ proj₂ y → x ≡ y --- -- -- -- -- prodId {x = (a , b)} {y = (c , d)} id1 id2 i = (id1 i) , (id2 i) - --- -- -- -- -- fstFunHelper : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- -- -- -- → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 _) (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x --- -- -- -- -- fstFunHelper a = MV.Ker-i⊂Im-d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a --- -- -- -- -- (sym (isContrH1Unit×H1Unit .snd _) ∙ (isContrH1Unit×H1Unit .snd _)) --- -- -- -- -- where --- -- -- -- -- isContrH1Unit×H1Unit : isContr (Group.type (×coHomGr 1 Unit Unit)) --- -- -- -- -- isContrH1Unit×H1Unit = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) --- -- -- -- -- , λ {(a , b) → sigmaProdElim {D = λ (x : Σ[ x ∈ Group.type (×coHomGr 1 Unit Unit)] Unit) → (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) ≡ fst x} --- -- -- -- -- (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) --- -- -- -- -- (λ a b _ → prodId (grIso.leftInv (coHomGrUnit≥1 0) ∣ a ∣₀) (grIso.leftInv (coHomGrUnit≥1 0) ∣ b ∣₀)) ((a , b) , tt)} - - - --- -- -- -- -- helperMorph : isMorph intGroup (dirProd intGroup intGroup) λ a → a , (0 - a) --- -- -- -- -- helperMorph = --- -- -- -- -- indIntGroup {G = dirProd intGroup intGroup} --- -- -- -- -- (λ a → a , (0 - a)) --- -- -- -- -- refl --- -- -- -- -- (λ a → prodId refl (helper2 a)) --- -- -- -- -- λ a → prodId refl (helper3 a) --- -- -- -- -- where --- -- -- -- -- helper1 : (a : Int) → predInt (sucInt a) ≡ a --- -- -- -- -- helper1 (pos zero) = refl --- -- -- -- -- helper1 (pos (suc n)) = refl --- -- -- -- -- helper1 (negsuc zero) = refl --- -- -- -- -- helper1 (negsuc (suc n)) = refl - --- -- -- -- -- helper4 : (a : Int) → sucInt (predInt a) ≡ a --- -- -- -- -- helper4 (pos zero) = refl --- -- -- -- -- helper4 (pos (suc n)) = refl --- -- -- -- -- helper4 (negsuc zero) = refl --- -- -- -- -- helper4 (negsuc (suc n)) = refl - --- -- -- -- -- helper2 : (a : Int) → (pos 0 - sucInt a) ≡ predInt (pos 0 - a) --- -- -- -- -- helper2 (pos zero) = refl --- -- -- -- -- helper2 (pos (suc n)) = refl --- -- -- -- -- helper2 (negsuc zero) = refl --- -- -- -- -- helper2 (negsuc (suc n)) = sym (helper1 _) - --- -- -- -- -- helper3 : (a : Int) → (pos 0 - predInt a) ≡ sucInt (pos 0 - a) --- -- -- -- -- helper3 (pos zero) = refl --- -- -- -- -- helper3 (pos (suc zero)) = refl --- -- -- -- -- helper3 (pos (suc (suc n))) = sym (helper4 _) --- -- -- -- -- helper3 (negsuc zero) = refl --- -- -- -- -- helper3 (negsuc (suc n)) = refl - --- -- -- -- -- helper : (a b : Int) → (pos 0 - (a +ℤ b)) ≡ ((pos 0 - a) +ℤ (pos 0 - b)) --- -- -- -- -- helper a (pos zero) = refl --- -- -- -- -- helper (pos zero) (pos (suc n)) = --- -- -- -- -- cong (λ x → pos 0 - sucInt x) (+ℤ-comm (pos zero) (pos n)) --- -- -- -- -- ∙ +ℤ-comm (pos 0 +negsuc n) (pos zero) --- -- -- -- -- helper (pos (suc n₁)) (pos (suc n)) = --- -- -- -- -- {!!} --- -- -- -- -- helper (negsuc n₁) (pos (suc n)) = {!!} --- -- -- -- -- helper a (negsuc n) = {!!} - --- -- -- -- -- fun : morph intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- -- -- -- fst fun = MV.d _ _ _ (λ _ → tt) (λ _ → tt) 0 ∘ grIso.inv coHom0-S0 .fst ∘ λ a → a , (0 - a) --- -- -- -- -- snd fun = {!!} --- -- -- -- -- {- compMorph {F = intGroup} {G = dirProd intGroup intGroup} {H = coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))} --- -- -- -- -- ((λ a → a , a) , (λ a b → refl)) --- -- -- -- -- (compMorph {F = dirProd intGroup intGroup} {G = coHomGr 0 (S₊ 0)} {H = coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))} (grIso.inv coHom0-S0) --- -- -- -- -- (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 --- -- -- -- -- , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) .snd -} --- -- -- -- -- {- theIso : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- -- -- -- theIso = Iso''→Iso _ _ --- -- -- -- -- (iso'' ((λ x → ∣ (λ {(inl tt) → 0ₖ ; (inr tt) → 0ₖ ; (push a i) → Kn→ΩKn+1 0 x i}) ∣₀) , {!!}) --- -- -- -- -- {!!} --- -- -- -- -- {!MV.d!}) --- -- -- -- -- -} - - - --- -- -- -- -- theIso : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- -- -- -- theIso = --- -- -- -- -- Iso''→Iso _ _ --- -- -- -- -- (iso'' fun --- -- -- -- -- (λ x inker → {!MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 --- -- -- -- -- (grIso.inv coHom0-S0 .fst (g , g))!}) --- -- -- -- -- (sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) --- -- -- -- -- λ x → pRec propTruncIsProp --- -- -- -- -- (λ {(a , b) → {!fun!} }) --- -- -- -- -- (fstFunHelper (∣ x ∣₀)))) --- -- -- -- -- where --- -- -- -- -- whoKnows : (a : S₊ 0 → Int) (x : MV.D Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt)) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (λ _ → a north) x --- -- -- -- -- ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 a x --- -- -- -- -- whoKnows a (inl x) = refl --- -- -- -- -- whoKnows a (inr x) = refl --- -- -- -- -- whoKnows a (push north i) = refl --- -- -- -- -- whoKnows a (push south i) j = {!!} - --- -- -- -- -- helper : (a : Int) → (grIso.inv coHom0-S0 .fst (a , a)) ≡ ∣ S0→Int (a , a) ∣₀ --- -- -- -- -- helper a = {!have : - --- -- -- -- -- ∣ --- -- -- -- -- MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 --- -- -- -- -- (S0→Int (x , x)) --- -- -- -- -- ∣₀ --- -- -- -- -- ≡ ∣ (λ _ → ∣ north ∣) ∣₀!} - --- -- -- -- -- helper2 : (a b : Int) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a)) ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (b , b)) --- -- -- -- -- → a ≡ b --- -- -- -- -- helper2 a b id = pRec (isSetInt a b) (λ {(pt , id) → {!!}}) (fstFunHelper ∣ (MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a))) ∣₀) - --- -- -- -- -- idFun : (a : Int) → MV.D Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 --- -- -- -- -- idFun a (inl x) = 0ₖ --- -- -- -- -- idFun a (inr x) = 0ₖ --- -- -- -- -- idFun a (push north i) = Kn→ΩKn+1 zero a i --- -- -- -- -- idFun a (push south i) = Kn→ΩKn+1 zero a i - --- -- -- -- -- helper3 : (a : Int) → (MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a))) ≡ idFun a --- -- -- -- -- helper3 a = funExt λ {(inl x) → refl ; (inr x) → refl ; (push north i) → refl ; (push south i) → refl } - --- -- -- -- -- helper4 : (a b : Int) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a)) ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (b , b)) --- -- -- -- -- → a ≡ b --- -- -- -- -- helper4 a b id = --- -- -- -- -- {!!} --- -- -- -- -- ∙ {!!} --- -- -- -- -- ∙ {!!} --- -- -- -- -- where --- -- -- -- -- helper5 : {!!} --PathP (λ k → id k (inl tt) ≡ id k (inr tt)) (Kn→ΩKn+1 zero a) (Kn→ΩKn+1 zero a) --- -- -- -- -- helper5 i j = {!id i!} - --- -- -- -- -- -- fun : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) → coHom 0 (S₊ 0) --- -- -- -- -- -- fun a = (pRec {P = Σ[ x ∈ coHom 0 (S₊ 0)] (MV.d _ _ _ (λ _ → tt) (λ _ → tt) 0 x ≡ a) × isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) (MV.Δ _ _ _ (λ _ → tt) (λ _ → tt) 0) x} --- -- -- -- -- -- (λ {(a1 , b) (c , d) → ΣProp≡ (λ x → isOfHLevelProd 1 (setTruncIsSet _ _) propTruncIsProp) --- -- -- -- -- -- (sElim2 {B = λ a1 c → (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a1 ≡ a) --- -- -- -- -- -- → MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 c ≡ a --- -- -- -- -- -- → isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) --- -- -- -- -- -- (λ z → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 z) a1 --- -- -- -- -- -- → isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) --- -- -- -- -- -- (λ z → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 z) c → a1 ≡ c} --- -- -- -- -- -- (λ _ _ → {!!}) --- -- -- -- -- -- (λ a c b1 d1 → pElim (λ _ → isOfHLevelΠ 1 λ _ → setTruncIsSet _ _) --- -- -- -- -- -- λ b2 → pElim (λ _ → setTruncIsSet _ _) --- -- -- -- -- -- λ d2 → {!d2!}) --- -- -- -- -- -- a1 c (proj₁ b) (proj₁ d) (proj₂ b) (proj₂ d))}) --- -- -- -- -- -- (λ {(a , b) → a , b , MV.Ker-d⊂Im-Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a {!!}}) --- -- -- -- -- -- (fstFunHelper a)) .fst -- pRec {!!} {!!} (fstFunHelper a) +-- -- -- -- -- -- -- helper5 : {!!} --PathP (λ k → id k (inl tt) ≡ id k (inr tt)) (Kn→ΩKn+1 zero a) (Kn→ΩKn+1 zero a) +-- -- -- -- -- -- -- helper5 i j = {!id i!} + +-- -- -- -- -- -- -- -- fun : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) → coHom 0 (S₊ 0) +-- -- -- -- -- -- -- -- fun a = (pRec {P = Σ[ x ∈ coHom 0 (S₊ 0)] (MV.d _ _ _ (λ _ → tt) (λ _ → tt) 0 x ≡ a) × isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) (MV.Δ _ _ _ (λ _ → tt) (λ _ → tt) 0) x} +-- -- -- -- -- -- -- -- (λ {(a1 , b) (c , d) → ΣProp≡ (λ x → isOfHLevelProd 1 (setTruncIsSet _ _) propTruncIsProp) +-- -- -- -- -- -- -- -- (sElim2 {B = λ a1 c → (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a1 ≡ a) +-- -- -- -- -- -- -- -- → MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 c ≡ a +-- -- -- -- -- -- -- -- → isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) +-- -- -- -- -- -- -- -- (λ z → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 z) a1 +-- -- -- -- -- -- -- -- → isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) +-- -- -- -- -- -- -- -- (λ z → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 z) c → a1 ≡ c} +-- -- -- -- -- -- -- -- (λ _ _ → {!!}) +-- -- -- -- -- -- -- -- (λ a c b1 d1 → pElim (λ _ → isOfHLevelΠ 1 λ _ → setTruncIsSet _ _) +-- -- -- -- -- -- -- -- λ b2 → pElim (λ _ → setTruncIsSet _ _) +-- -- -- -- -- -- -- -- λ d2 → {!d2!}) +-- -- -- -- -- -- -- -- a1 c (proj₁ b) (proj₁ d) (proj₂ b) (proj₂ d))}) +-- -- -- -- -- -- -- -- (λ {(a , b) → a , b , MV.Ker-d⊂Im-Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a {!!}}) +-- -- -- -- -- -- -- -- (fstFunHelper a)) .fst -- pRec {!!} {!!} (fstFunHelper a) From 0d8db9659b780b2cf325336e3bdaf4a5966b6828 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Sun, 7 Jun 2020 03:42:02 +0200 Subject: [PATCH 40/89] cohom1Torus-maps+rinv --- Cubical/ZCohomology/Sn/Sn.agda | 2373 +++++++++++++++++--------------- 1 file changed, 1298 insertions(+), 1075 deletions(-) diff --git a/Cubical/ZCohomology/Sn/Sn.agda b/Cubical/ZCohomology/Sn/Sn.agda index a190b3300..bb093ed34 100644 --- a/Cubical/ZCohomology/Sn/Sn.agda +++ b/Cubical/ZCohomology/Sn/Sn.agda @@ -48,6 +48,19 @@ open import Cubical.Data.HomotopyGroup open import Cubical.Data.Bool hiding (_⊕_) + +funIso : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} → Iso (A → B × C) ((A → B) × (A → C)) +Iso.fun funIso = λ f → (λ a → proj₁ (f a)) , (λ a → proj₂ (f a)) +Iso.inv funIso (f , g) = λ a → (f a) , (g a) +Iso.rightInv funIso (f , g) = refl +Iso.leftInv funIso b = funExt λ a → sym (×-η _) + +schonfIso : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} → Iso (A × B → C) (A → B → C) +Iso.fun schonfIso f a b = f (a , b) +Iso.inv schonfIso f (a , b) = f a b +Iso.rightInv schonfIso a = refl +Iso.leftInv schonfIso f = funExt λ {(a , b) → refl} + toProdElimSuspRec : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Susp A → Type ℓ'} (a : A) → ((x : Susp A) → isProp (B x)) → B north @@ -107,716 +120,751 @@ prodId : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A × B} → proj prodId {x = (a , b)} {y = (c , d)} id1 id2 i = (id1 i) , (id2 i) -basechange2⁻Gen : (x y : S¹) → x ≡ y → x ≡ y → ΩS¹ -basechange2⁻Gen x y = J (λ y p → x ≡ y → ΩS¹) (basechange2⁻ x) - -basechange2⁻Gen-morphIsh : (x y z : S¹) → (P : x ≡ y) (Q : y ≡ z) (p : x ≡ y) (q : y ≡ z) → basechange2⁻Gen x z (P ∙ Q) (p ∙ q) ≡ basechange2⁻Gen x y P p ∙ basechange2⁻Gen y z Q q -basechange2⁻Gen-morphIsh x y z = - J (λ y P → (Q : y ≡ z) (p : x ≡ y) (q : y ≡ z) → basechange2⁻Gen x z (P ∙ Q) (p ∙ q) ≡ basechange2⁻Gen x y P p ∙ basechange2⁻Gen y z Q q) - (J (λ z Q → (p : x ≡ x) (q : x ≡ z) → basechange2⁻Gen x z (refl ∙ Q) (p ∙ q) ≡ basechange2⁻Gen x x refl p ∙ basechange2⁻Gen x z Q q) - λ p q → cong (λ y → basechange2⁻Gen x x y (p ∙ q)) (sym (lUnit refl)) - ∙ (λ i → JRefl (λ y p → x ≡ y → ΩS¹) (basechange2⁻ x) i (p ∙ q)) - ∙ basechange2⁻-morph x p q - ∙ λ i → JRefl (λ y p → x ≡ y → ΩS¹) (basechange2⁻ x) (~ i) p ∙ JRefl (λ y p → x ≡ y → ΩS¹) (basechange2⁻ x) (~ i) q) - -S1→S¹ : S₊ 1 → S¹ -S1→S¹ x = SuspBool→S¹ (S1→SuspBool x) - -S¹→S1 : S¹ → S₊ 1 -S¹→S1 x = SuspBool→S1 (S¹→SuspBool x) - -S1→S¹-sect : section S1→S¹ S¹→S1 -S1→S¹-sect x = - cong SuspBool→S¹ (SuspBool→S1-retr (S¹→SuspBool x)) - ∙ S¹→SuspBool→S¹ x - -S1→S¹-retr : retract S1→S¹ S¹→S1 -S1→S¹-retr x = - cong SuspBool→S1 (SuspBool→S¹→SuspBool (S1→SuspBool x)) - ∙ SuspBool→S1-sect x - - -prodElim : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : ∥ A ∥₀ × ∥ B ∥₀ → Type ℓ''} - → ((x : ∥ A ∥₀ × ∥ B ∥₀) → isOfHLevel 2 (C x)) - → ((a : A) (b : B) → C (∣ a ∣₀ , ∣ b ∣₀)) - → (x : ∥ A ∥₀ × ∥ B ∥₀) → C x -prodElim {A = A} {B = B} {C = C} hlevel ind (a , b) = schonf a b - where - schonf : (a : ∥ A ∥₀) (b : ∥ B ∥₀) → C (a , b) - schonf = sElim (λ x → isOfHLevelΠ 2 λ y → hlevel (_ , _)) λ a → sElim (λ x → hlevel (_ , _)) - λ b → ind a b - - -isGroupoidS1 : isGroupoid (S₊ 1) -isGroupoidS1 = transport (λ i → isGroupoid (S¹≡S1 (~ i))) isGroupoidS¹ -isConnectedS1 : (x : S₊ 1) → ∥ north ≡ x ∥₋₁ -isConnectedS1 x = pRec propTruncIsProp - (λ p → ∣ cong (transport (sym (S¹≡S1))) p ∙ transport⁻Transport (S¹≡S1) x ∣₋₁) - (isConnectedS¹ (transport S¹≡S1 x)) - -private - S¹map : hLevelTrunc 3 S¹ → S¹ - S¹map = trElim (λ _ → isGroupoidS¹) (idfun S¹) - - S¹map-id : (x : hLevelTrunc 3 S¹) → Path (hLevelTrunc 3 S¹) ∣ S¹map x ∣ x - S¹map-id = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) - λ a → refl - - S1map : hLevelTrunc 3 (S₊ 1) → (S₊ 1) - S1map = trElim (λ _ → isGroupoidS1) (idfun _) - -S¹→S¹ : Iso (S¹ → hLevelTrunc 3 S¹) (S¹ × Int) -Iso.fun S¹→S¹ f = S¹map (f base) - , winding (basechange2⁻ (S¹map (f base)) λ i → S¹map (f (loop i))) -Iso.inv S¹→S¹ (s , int) base = ∣ s ∣ -Iso.inv S¹→S¹ (s , int) (loop i) = ∣ basechange2 s (intLoop int) i ∣ -Iso.rightInv S¹→S¹ (s , int) = ×≡ refl ((λ i → winding (basechange2-retr s (λ i → intLoop int i) i)) - ∙ windingIntLoop int) -Iso.leftInv S¹→S¹ f = funExt λ { base → S¹map-id (f base) - ; (loop i) j → helper2 j i} - where - helper2 : PathP (λ i → S¹map-id (f base) i ≡ S¹map-id (f base) i) - (λ i → ∣ basechange2 (S¹map (f base)) (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁)))))) i ∣) - (cong f loop) - helper2 i j = - hcomp (λ k → λ { (i = i0) → cong ∣_∣ (basechange2 (S¹map (f base)) (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁))))))) j - ; (i = i1) → S¹map-id (f (loop j)) k - ; (j = i0) → S¹map-id (f base) (i ∧ k) - ; (j = i1) → S¹map-id (f base) (i ∧ k)}) - (helper4 i j) - where - helper4 : Path (Path (hLevelTrunc 3 _) _ _) - (cong ∣_∣ (basechange2 (S¹map (f base)) - (intLoop - (winding - (basechange2⁻ (S¹map (f base)) - (λ i₁ → S¹map (f (loop i₁)))))))) - λ i → ∣ S¹map (f (loop i)) ∣ - helper4 i = - cong ∣_∣ - ((cong (basechange2 (S¹map (f base))) - (decodeEncode base (basechange2⁻ (S¹map (f base)) - (λ i₁ → S¹map (f (loop i₁))))) - ∙ basechange2-sect (S¹map (f base)) - (λ i → S¹map (f (loop i)))) i) - - - -S¹→S¹≡S1→S1 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → hLevelTrunc 3 S¹) -Iso.fun S¹→S¹≡S1→S1 f x = trMap S1→S¹ (f (S¹→S1 x)) -Iso.inv S¹→S¹≡S1→S1 f x = trMap S¹→S1 (f (S1→S¹ x)) -Iso.rightInv S¹→S¹≡S1→S1 F = funExt λ x i → helper (F (S1→S¹-sect x i)) i - where - helper : (x : hLevelTrunc 3 S¹) → trMap S1→S¹ (trMap S¹→S1 x) ≡ x - helper = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) - λ a → cong ∣_∣ (S1→S¹-sect a) -Iso.leftInv S¹→S¹≡S1→S1 F = funExt λ x i → helper (F (S1→S¹-retr x i)) i - where - helper : (x : hLevelTrunc 3 (S₊ 1)) → trMap S¹→S1 (trMap S1→S¹ x) ≡ x - helper = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) - λ a → cong ∣_∣ (S1→S¹-retr a) - - - - -basechange-lemma : ∀ {ℓ} {A : Type ℓ} {a : A} (x y : S¹) (F : a ≡ a → S¹) (f : S¹ → a ≡ a) (g : S¹ → a ≡ a) - → (f base ≡ refl) - → (g base ≡ refl) - → basechange2⁻ (F (f base ∙ g base)) (cong₂ {A = S¹} {B = λ x → S¹} (λ x y → F (f x ∙ g y)) loop loop) - ≡ basechange2⁻ (F (f base)) (cong (λ x → F (f x)) loop) ∙ basechange2⁻ (F (g base)) (cong (λ x → F (g x)) loop) -basechange-lemma x y F f g frefl grefl = - (λ i → basechange2⁻ (F (f base ∙ g base)) (cong₂Funct2 (λ x y → F (f x ∙ g y)) loop loop i)) - ∙ (λ i → basechange2⁻ (F (f base ∙ g base)) (cong (λ x₁ → F (f x₁ ∙ g base)) loop ∙ cong (λ y₁ → F (f base ∙ g y₁)) loop)) - ∙ basechange2⁻-morph (F (f base ∙ g base)) _ _ - ∙ (λ j → basechange2⁻ (F (f base ∙ grefl j)) - (λ i → F (f (loop i) ∙ grefl j)) - ∙ basechange2⁻ (F (frefl j ∙ g base)) - (λ i → F (frefl j ∙ g (loop i)))) - ∙ ((λ j → basechange2⁻ (F (rUnit (f base) (~ j))) - (λ i → F (rUnit (f (loop i)) (~ j))) - ∙ basechange2⁻ (F (lUnit (g base) (~ j))) - (λ i → F (lUnit (g (loop i)) (~ j))))) - - -coHom1S1≃ℤ : grIso (coHomGr 1 (S₊ 1)) intGroup -coHom1S1≃ℤ = - Iso'→Iso - (iso' - (iso - (sRec isSetInt (λ f → proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) - (λ a → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , a)) ∣₀) - (λ a → (λ i → proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 (Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , a)))))) - ∙ (λ i → proj₂ (Iso.fun S¹→S¹ (Iso.rightInv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , a)) i))) - ∙ λ i → proj₂ (Iso.rightInv S¹→S¹ (base , a) i)) - (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) - λ f → (λ i → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) ∣₀) - ∙ (λ i → ∣ Iso.inv S¹→S¹≡S1→S1 {!!} ∣₀) - ∙ {!!} - ∙ {!!} - ∙ {!!})) - {!!}) + + + +test13 : Iso (S₊ 1 → hLevelTrunc 4 (S₊ 2)) (hLevelTrunc 4 (S₊ 2) × hLevelTrunc 3 (S₊ 1)) +Iso.fun test13 f = f north , ΩKn+1→Kn (sym (rCancelₖ (f north)) + ∙ (λ i → f (merid south i) +ₖ (-ₖ f (merid north i))) + ∙ rCancelₖ (f south)) +Iso.inv test13 (a , b) north = a +ₖ 0ₖ +Iso.inv test13 (a , b) south = a +ₖ 0ₖ +Iso.inv test13 (a , b) (merid south i) = a +ₖ (Kn→ΩKn+1 1 b i) +Iso.inv test13 (a , b) (merid north i) = a +ₖ 0ₖ +Iso.rightInv test13 (a , b) = + ×≡ (rUnitₖ a) + ((cong ΩKn+1→Kn (congHelper++ (Kn→ΩKn+1 1 b) (λ x → (a +ₖ x) +ₖ (-ₖ (a +ₖ 0ₖ))) (funExt (λ x → sym (cancelHelper a x))) (rCancelₖ (a +ₖ 0ₖ)))) + ∙ Iso.leftInv (Iso3-Kn-ΩKn+1 1) b) where - helper : (f : _) (x : S¹) → ∣ (Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f))))) ∣₀ ≡ ∣ (Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) ∣₀ - helper f x = sym ({!!} - {!!} ∙ - {!!} ∙ - {!!} ∙ - {!!} ∙ - {!!}) - -{- - (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) - λ f g → (λ i → winding (basechange2⁻ (S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)) ∙ Kn→ΩKn+1 1 (g (S¹→S1 base)))))) - (λ i → S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 (loop i))) ∙ Kn→ΩKn+1 1 (g (S¹→S1 (loop i))))))))) - ∙ cong winding (helper (f (S¹→S1 base)) (g (S¹→S1 base)) f g refl refl) - ∙ winding-hom ((basechange2⁻ (S¹map (trMap S1→S¹ (f north))) - (λ i → S¹map (trMap S1→S¹ (f (S¹→S1 (loop i))))))) - ((basechange2⁻ (S¹map (trMap S1→S¹ (g north))) - (λ i → S¹map (trMap S1→S¹ (g (S¹→S1 (loop i))))))))) - - - where - - - helper : (x y : hLevelTrunc 3 (S₊ 1)) (f g : S₊ 1 → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1) - → (f (S¹→S1 base)) ≡ x - → (g (S¹→S1 base)) ≡ y - → (basechange2⁻ (S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)) ∙ Kn→ΩKn+1 1 (g (S¹→S1 base))))))) - (λ i → S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 (loop i))) ∙ Kn→ΩKn+1 1 (g (S¹→S1 (loop i))))))) - ≡ (basechange2⁻ (S¹map (trMap S1→S¹ ((f (S¹→S1 base)))))) - (λ i → S¹map (trMap S1→S¹ (f (S¹→S1 (loop i))))) - ∙ (basechange2⁻ (S¹map (trMap S1→S¹ ((g (S¹→S1 base))))) - (λ i → S¹map (trMap S1→S¹ ((g (S¹→S1 (loop i))))))) - helper = elim2 - (λ _ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 - λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 - λ _ → isOfHLevelPath 3 (isOfHLevelSuc 3 (isGroupoidS¹) base base) _ _) - (toProdElimSuspElim2 {A = S₊ 0} north - (λ _ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → (isGroupoidS¹) _ _ _ _ ) - λ f g reflf reflg → - (basechange-lemma - base - base - (λ x → S¹map (trMap S1→S¹ (ΩKn+1→Kn x))) - (λ x → Kn→ΩKn+1 1 (f (S¹→S1 x))) ((λ x → Kn→ΩKn+1 1 (g (S¹→S1 x)))) - (cong (Kn→ΩKn+1 1) reflf ∙ Kn→ΩKn+10ₖ 1) - (cong (Kn→ΩKn+1 1) reflg ∙ Kn→ΩKn+10ₖ 1)) - ∙ λ j → basechange2⁻ (S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 base)) j))) - (λ i → S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 (loop i))) j))) - ∙ basechange2⁻ (S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 base)) j))) - (λ i → S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 (loop i))) j)))) --} - -{- -Goal: basechange2⁻ - (S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)))))) - (cong - (λ x → S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 x)))))) - loop) - ∙ - basechange2⁻ - (S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (g (S¹→S1 base)))))) - (cong - (λ x → S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (g (S¹→S1 x)))))) - loop) - ≡ - basechange2⁻ (S¹map (trMap S1→S¹ (f (S¹→S1 base)))) - (λ i → S¹map (trMap S1→S¹ (f (S¹→S1 (loop i))))) - ∙ - basechange2⁻ (S¹map (trMap S1→S¹ (g (S¹→S1 base)))) - (λ i → S¹map (trMap S1→S¹ (g (S¹→S1 (loop i))))) --} --- (f g : S₊ 1 → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1) → (f (S¹→S1 base)) ≡ ∣ x ∣ → (g (S¹→S1 base)) ≡ ∣ y ∣ → ? - - -{- -coHom1S1≃ℤ : grIso (coHomGr 1 (S₊ 1)) intGroup -coHom1S1≃ℤ = - Iso'→Iso - (iso' - (iso - (sRec isSetInt λ f → proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f))) - {!winding (basechange2⁻ (S¹map (f base)) λ i → S¹map (f (loop i)))!} - {!!} - {!!}) - (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) - λ f g → (λ i → winding (basechange2⁻ (S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)) ∙ Kn→ΩKn+1 1 (g (S¹→S1 base)))))) - (λ i → S¹map (trMap S1→S¹ ((ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 (loop i))) ∙ Kn→ΩKn+1 1 (g (S¹→S1 (loop i)))))))))) - ∙ (λ j → winding (basechange2⁻ (S¹map (trMap S1→S¹ (ΩKn+1→Kn {!basechange2⁻!}))) - (λ i → S¹map (trMap S1→S¹ (ΩKn+1→Kn {!!}))))) - ∙ {!!} - ∙ {!λ i → winding !} - ∙ winding-hom (basechange2⁻ (S¹map (trMap S1→S¹ (f (S¹→S1 base)))) - (λ i → S¹map (trMap S1→S¹ (f (S¹→S1 (loop i)))))) - (basechange2⁻ (S¹map (trMap S1→S¹ (g (S¹→S1 base)))) - (λ i → S¹map (trMap S1→S¹ (g (S¹→S1 (loop i)))))) )) -- sym (winding-hom _ _))) - where - helper : {!!} - helper = {!!} -} - - - --- indIntGroup : ∀ {ℓ} {G : Group ℓ} → (ϕ : Int → Group.type G) --- → ϕ 0 ≡ isGroup.id (Group.groupStruc G) --- → ((a : Int) → ϕ (a +ℤ 1) ≡ isGroup.comp (Group.groupStruc G) (ϕ a) (ϕ 1)) --- → ((n : Int) → ϕ (predInt n) ≡ isGroup.comp (Group.groupStruc G) (ϕ n) (ϕ (negsuc zero))) --- → isMorph intGroup G ϕ --- indIntGroup {G = group G gSet (group-struct _ _ _+G_ _ rUnit₁ _ _ _)} ϕ zeroId _ _ n (pos zero) = --- sym (rUnit₁ (ϕ n)) ∙ cong (λ x → ϕ n +G x) (sym zeroId) --- indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId oneId minOneId n (pos (suc m)) = --- (λ i → ϕ ((n +pos m) +ℤ 1)) --- ∙ oneId (n +pos m) --- ∙ cong (λ x → x +G ϕ (pos 1)) --- (indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} --- ϕ zeroId oneId minOneId n (pos m)) --- ∙ assoc₁ (ϕ n) (ϕ (pos m)) (ϕ (pos 1)) --- ∙ sym (cong (λ x → ϕ n +G x) (oneId (pos m))) --- indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId _ minOneId n (negsuc zero) = minOneId n --- indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId a minOneId n (negsuc (suc m)) = --- (λ i → ϕ ((n +negsuc m) +ℤ (negsuc zero))) --- ∙ minOneId (n +negsuc m) --- ∙ cong (λ x → x +G ϕ (negsuc zero)) (indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId a minOneId n (negsuc m)) --- ∙ assoc₁ (ϕ n) (ϕ (negsuc m)) (ϕ (negsuc zero)) --- ∙ cong (λ x → ϕ n +G x) (sym (minOneId (negsuc m))) - --- pushoutSn : (n : ℕ) → Iso (S₊ (suc n)) (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt) --- Iso.fun (pushoutSn n) north = inl tt --- Iso.fun (pushoutSn n) south = inr tt --- Iso.fun (pushoutSn n) (merid a i) = push a i --- Iso.inv (pushoutSn n) (inl x) = north --- Iso.inv (pushoutSn n) (inr x) = south --- Iso.inv (pushoutSn n) (push a i) = merid a i --- Iso.rightInv (pushoutSn n) (inl x) = refl --- Iso.rightInv (pushoutSn n) (inr x) = refl --- Iso.rightInv (pushoutSn n) (push a i) = refl --- Iso.leftInv (pushoutSn n) north = refl --- Iso.leftInv (pushoutSn n) south = refl --- Iso.leftInv (pushoutSn n) (merid a i) = refl - --- Sn≡Pushout : (n : ℕ) → (S₊ (suc n)) ≡ (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt) --- Sn≡Pushout n = isoToPath (pushoutSn n) - --- coHomPushout≡coHomSn' : (n m : ℕ) → grIso (coHomGr m (S₊ (suc n))) (coHomGr m (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt)) --- morph.fun (grIso.fun (coHomPushout≡coHomSn' n m)) = --- sRec setTruncIsSet --- λ f → ∣ (λ {(inl x) → f north ; (inr x) → f south ; (push a i) → f (merid a i)}) ∣₀ --- morph.ismorph (grIso.fun (coHomPushout≡coHomSn' n m)) = --- sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) --- λ a b → cong ∣_∣₀ (funExt λ {(inl x) → refl ; (inr x) → refl ; (push a i) → refl }) --- morph.fun (grIso.inv (coHomPushout≡coHomSn' n m)) = sRec setTruncIsSet (λ f → ∣ (λ {north → f (inl tt) ; south → f (inr tt) ; (merid a i) → f (push a i)}) ∣₀) --- morph.ismorph (grIso.inv (coHomPushout≡coHomSn' n m)) = --- sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) --- λ a b → cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → refl }) --- grIso.rightInv (coHomPushout≡coHomSn' n m) = --- sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) --- λ f → cong ∣_∣₀ (funExt λ {(inl x) → refl ; (inr x) → refl ; (push a i) → refl }) --- grIso.leftInv (coHomPushout≡coHomSn' n m) = --- sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) --- λ f → cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → refl }) - - --- isContr→≡Unit : {A : Type₀} → isContr A → A ≡ Unit --- isContr→≡Unit contr = isoToPath (iso (λ _ → tt) (λ _ → fst contr) (λ _ → refl) λ _ → snd contr _) - --- isContr→isContrTrunc : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → isContr A → isContr (hLevelTrunc n A) --- isContr→isContrTrunc n contr = ∣ fst contr ∣ , (trElim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → cong ∣_∣ (snd contr a)) --- isContr→isContrSetTrunc : ∀ {ℓ} {A : Type ℓ} → isContr A → isContr (∥ A ∥₀) --- isContr→isContrSetTrunc contr = ∣ fst contr ∣₀ , sElim (λ _ → isOfHLevelPath 2 (setTruncIsSet) _ _) λ a → cong ∣_∣₀ (snd contr a) - --- coHomGrUnit0 : grIso (coHomGr 0 Unit) intGroup --- grIso.fun coHomGrUnit0 = mph (sRec isSetInt (λ f → f tt)) --- (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) --- (λ a b → addLemma (a tt) (b tt))) --- grIso.inv coHomGrUnit0 = mph (λ a → ∣ (λ _ → a) ∣₀) (λ a b i → ∣ (λ _ → addLemma a b (~ i)) ∣₀) --- grIso.rightInv coHomGrUnit0 a = refl --- grIso.leftInv coHomGrUnit0 = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ a → refl - --- isContrCohomUnit : (n : ℕ) → isContr (coHom (suc n) Unit) --- isContrCohomUnit n = subst isContr (λ i → ∥ UnitToTypeId (coHomK (suc n)) (~ i) ∥₀) (helper n) + cancelHelper : (a b : hLevelTrunc 4 (S₊ 2)) → (a +ₖ b) +ₖ (-ₖ (a +ₖ 0ₖ)) ≡ b + cancelHelper a b = + (a +ₖ b) +ₖ (-ₖ (a +ₖ 0ₖ)) ≡⟨ (λ i → (a +ₖ b) +ₖ (-ₖ (rUnitₖ a i))) ⟩ + (a +ₖ b) +ₖ (-ₖ a) ≡⟨ cong (λ x → x +ₖ (-ₖ a)) (commₖ a b) ⟩ + (b +ₖ a) +ₖ (-ₖ a) ≡⟨ assocₖ b a (-ₖ a) ⟩ + b +ₖ a +ₖ (-ₖ a) ≡⟨ cong (λ x → b +ₖ x) (rCancelₖ a) ⟩ + b +ₖ 0ₖ ≡⟨ rUnitₖ b ⟩ + b ∎ + + abstract + commHelper : (p q : Path (hLevelTrunc 4 (S₊ 2)) 0ₖ 0ₖ) → p ∙ q ≡ q ∙ p + commHelper p q = + cong₂ _∙_ (sym (Iso.rightInv (Iso3-Kn-ΩKn+1 1) p)) + (sym (Iso.rightInv (Iso3-Kn-ΩKn+1 1) q)) + ∙ sym (Iso.rightInv (Iso3-Kn-ΩKn+1 1) (Kn→ΩKn+1 1 (ΩKn+1→Kn p) ∙ Kn→ΩKn+1 1 (ΩKn+1→Kn q))) + ∙ cong (Kn→ΩKn+1 1) (commₖ (ΩKn+1→Kn p) (ΩKn+1→Kn q)) + ∙ Iso.rightInv (Iso3-Kn-ΩKn+1 1) (Kn→ΩKn+1 1 (ΩKn+1→Kn q) ∙ Kn→ΩKn+1 1 (ΩKn+1→Kn p)) + ∙ sym (cong₂ _∙_ (sym (Iso.rightInv (Iso3-Kn-ΩKn+1 1) q)) + (sym (Iso.rightInv (Iso3-Kn-ΩKn+1 1) p))) + + moveabout : {x : hLevelTrunc 4 (S₊ 2)} (p q : x ≡ 0ₖ) (mid : 0ₖ ≡ 0ₖ) + → sym q ∙ (p ∙ mid ∙ sym p) ∙ q ≡ mid + moveabout p q mid = assoc (sym q) _ _ + ∙ cong (_∙ q) (assoc (sym q) p (mid ∙ sym p)) + ∙ sym (assoc (sym q ∙ p) (mid ∙ sym p) q) + ∙ cong ((sym q ∙ p) ∙_) (sym (assoc mid (sym p) q)) + ∙ cong (λ x → (sym q ∙ p) ∙ (mid ∙ x)) (sym (symDistr (sym q) p)) + ∙ cong ((sym q ∙ p)∙_) (commHelper mid _) + ∙ assoc _ _ _ + ∙ cong (_∙ mid) (rCancel (sym q ∙ p)) + ∙ sym (lUnit mid) + + + + +-- (cong (λ x → (a +ₖ 0ₖ) +ₖ x) (-distrₖ a 0ₖ) + congHelper : ∀ {ℓ} {A : Type ℓ} {a1 : A} (p : a1 ≡ a1) (f : A → A) (id : (λ x → x) ≡ f) + → cong f p ≡ sym (funExt⁻ id a1) ∙ p ∙ funExt⁻ id a1 + congHelper {a1 = a1} p f id = + (λ i → lUnit (rUnit (cong f p) i) i) + ∙ (λ i → (λ j → id ((~ i) ∨ (~ j)) a1) ∙ cong (id (~ i)) p ∙ λ j → id (~ i ∨ j) a1) + + + congHelper++ : (p : 0ₖ ≡ 0ₖ) (f : hLevelTrunc 4 (S₊ 2) → hLevelTrunc 4 (S₊ 2)) (id : (λ x → x) ≡ f) + → (q : (f 0ₖ) ≡ 0ₖ) + → (sym q) ∙ cong f p ∙ q ≡ p + congHelper++ p f id q = + cong (λ x → sym q ∙ x ∙ q) (congHelper p f id) ∙ + moveabout (sym (funExt⁻ id ∣ north ∣)) q p + +Iso.leftInv test13 a = + funExt λ {north → rUnitₖ (a north) + ; south → rUnitₖ (a north) ∙ cong a (merid north) + ; (merid south i) j → {!!} + ; (merid north i) → {!!}} + + + + +-- S1→S¹ : S₊ 1 → S¹ +-- S1→S¹ x = SuspBool→S¹ (S1→SuspBool x) + +-- S¹→S1 : S¹ → S₊ 1 +-- S¹→S1 x = SuspBool→S1 (S¹→SuspBool x) + +-- S1→S¹-sect : section S1→S¹ S¹→S1 +-- S1→S¹-sect x = +-- cong SuspBool→S¹ (SuspBool→S1-retr (S¹→SuspBool x)) +-- ∙ S¹→SuspBool→S¹ x + +-- S1→S¹-retr : retract S1→S¹ S¹→S1 +-- S1→S¹-retr x = +-- cong SuspBool→S1 (SuspBool→S¹→SuspBool (S1→SuspBool x)) +-- ∙ SuspBool→S1-sect x + + +-- prodElim : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : ∥ A ∥₀ × ∥ B ∥₀ → Type ℓ''} +-- → ((x : ∥ A ∥₀ × ∥ B ∥₀) → isOfHLevel 2 (C x)) +-- → ((a : A) (b : B) → C (∣ a ∣₀ , ∣ b ∣₀)) +-- → (x : ∥ A ∥₀ × ∥ B ∥₀) → C x +-- prodElim {A = A} {B = B} {C = C} hlevel ind (a , b) = schonf a b -- where --- helper : (n : ℕ) → isContr (∥ coHomK (suc n) ∥₀) --- helper n = subst isContr ((isoToPath (truncOfTruncIso {A = S₊ (1 + n)} 2 (1 + n))) ∙ sym propTrunc≡Trunc0 ∙ λ i → ∥ hLevelTrunc (suc (+-comm n 2 i)) (S₊ (1 + n)) ∥₀) --- (isConnectedSubtr 2 (helper2 n .fst) (subst (λ x → isHLevelConnected x (S₊ (suc n))) (sym (helper2 n .snd)) (sphereConnected (suc n))) ) --- where --- helper2 : (n : ℕ) → Σ[ m ∈ ℕ ] m + 2 ≡ 2 + n --- helper2 zero = 0 , refl --- helper2 (suc n) = (suc n) , λ i → suc (+-comm n 2 i) - --- coHomGrUnit≥1 : (n : ℕ) → grIso (coHomGr (suc n) Unit) trivialGroup --- grIso.fun (coHomGrUnit≥1 n) = mph (λ _ → tt) (λ _ _ → refl) --- grIso.inv (coHomGrUnit≥1 n) = mph (λ _ → ∣ (λ _ → ∣ north ∣) ∣₀) (λ _ _ → sym (rUnitₕ 0ₕ)) --- grIso.rightInv (coHomGrUnit≥1 n) a = refl --- grIso.leftInv (coHomGrUnit≥1 n) a = sym (isContrCohomUnit n .snd 0ₕ) ∙ isContrCohomUnit n .snd a - --- S0→Int : (a : Int × Int) → S₊ 0 → Int --- S0→Int a north = proj₁ a --- S0→Int a south = proj₂ a - --- coHom0-S0 : grIso (coHomGr 0 (S₊ 0)) (dirProd intGroup intGroup) --- coHom0-S0 = +-- schonf : (a : ∥ A ∥₀) (b : ∥ B ∥₀) → C (a , b) +-- schonf = sElim (λ x → isOfHLevelΠ 2 λ y → hlevel (_ , _)) λ a → sElim (λ x → hlevel (_ , _)) +-- λ b → ind a b + +-- setTruncOfProdIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → Iso ∥ A × B ∥₀ (∥ A ∥₀ × ∥ B ∥₀) +-- Iso.fun setTruncOfProdIso = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) λ { (a , b) → ∣ a ∣₀ , ∣ b ∣₀ } +-- Iso.inv setTruncOfProdIso = prodElim (λ _ → setTruncIsSet) λ a b → ∣ a , b ∣₀ +-- Iso.rightInv setTruncOfProdIso = prodElim (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- λ _ _ → refl +-- Iso.leftInv setTruncOfProdIso = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- λ {(a , b) → refl} + + +-- isGroupoidS1 : isGroupoid (S₊ 1) +-- isGroupoidS1 = transport (λ i → isGroupoid (S¹≡S1 (~ i))) isGroupoidS¹ +-- isConnectedS1 : (x : S₊ 1) → ∥ north ≡ x ∥₋₁ +-- isConnectedS1 x = pRec propTruncIsProp +-- (λ p → ∣ cong (transport (sym (S¹≡S1))) p ∙ transport⁻Transport (S¹≡S1) x ∣₋₁) +-- (isConnectedS¹ (transport S¹≡S1 x)) + + +-- open import Cubical.HITs.S2 + +-- test : (S₊ 2) → S₊ 1 +-- test north = north +-- test south = south +-- test (merid a i) = merid north i + +-- test2 : (S₊ 1) → (S₊ 2) +-- test2 north = north +-- test2 south = south +-- test2 (merid a i) = merid (merid a i) i + +-- S1test : ∀ {ℓ} {A : S¹ → Type ℓ} → (Abase : A base) → subst A loop Abase ≡ Abase → (x : S¹) → A x +-- S1test = {!!} + +-- testS² : S² → S¹ +-- testS² base = base +-- testS² (surf i i₁) = base + +-- test4 : S₊ 1 → hLevelTrunc 4 (S₊ 2) +-- test4 north = ∣ north ∣ +-- test4 south = ∣ north ∣ +-- test4 (merid a i) = (Kn→ΩKn+1 1 ∣ south ∣) i + +-- test3 : hLevelTrunc 4 (S₊ 2) → S₊ 1 +-- test3 = +-- trElim (λ _ → {!!}) +-- λ {north → north ; south → north ; (merid a i) → loop* i} + +-- testIso2 : Iso ((S₊ 1) → hLevelTrunc 4 (S₊ 2)) ((S₊ 1) × hLevelTrunc 4 (S₊ 2)) +-- Iso.fun testIso2 f = (test3 (f north)) , trElim (λ _ → isOfHLevelTrunc 4) (λ s → ΩKn+1→Kn (cong ∣_∣ (merid s ∙ sym (merid north)))) (f north) +-- Iso.inv testIso2 (x , p) y = (test4 x) +ₖ (test4 y) +ₖ p +-- Iso.rightInv testIso2 (s , s2) = trElim {B = λ s2 → Iso.fun testIso2 (Iso.inv testIso2 (s , s2)) ≡ (s , s2)} +-- {!!} +-- (λ s3 → ×≡ {!!} {!!}) +-- s2 +-- Iso.leftInv testIso2 a = funExt λ x → {!!} + +-- testIso : Iso (S¹ → hLevelTrunc 4 S²) (S¹ × hLevelTrunc 4 S²) +-- Iso.fun testIso f = {!(f base)!} , {!!} -- trElim (λ _ → isOfHLevelSuc 3 (isGroupoidS1)) test (f north) , (f north) +-- Iso.inv testIso (s , tr) x = tr +-- Iso.rightInv testIso (x , tr) = {!!} +-- Iso.leftInv testIso a = funExt (S1test (cong a loop) {!!}) -- funExt (toPropElim {B = λ x → a base ≡ a x} {!a base!} refl) + + + + +-- coHom0Torus : grIso (coHomGr 0 (S₊ 1 × S₊ 1)) intGroup +-- coHom0Torus = -- Iso'→Iso -- (iso' --- (iso (sRec (isOfHLevelProd 2 isSetInt isSetInt) --- λ f → (f north) , (f south)) --- (λ a → ∣ S0→Int a ∣₀) --- (λ { (a , b) → refl}) --- (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ f → cong ∣_∣₀ (funExt (λ {north → refl ; south → refl})))) --- (sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 isSetInt isSetInt) _ _) --- λ a b i → addLemma (a north) (b north) i , addLemma (a south) (b south) i)) - --- ×morph : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group ℓ} {B : Group ℓ'} {C : Group ℓ''} {D : Group ℓ'''} → morph A B → morph C D → morph (dirProd A C) (dirProd B D) --- morph.fun (×morph mf1 mf2) = --- (λ {(a , b) → (morph.fun mf1 a) , morph.fun mf2 b}) --- morph.ismorph (×morph mf1 mf2) = --- (λ {(a , b) (c , d) i → morph.ismorph mf1 a c i , morph.ismorph mf2 b d i}) - - - - - --- coHom1S1 : grIso intGroup (coHomGr 1 (S₊ 1)) --- coHom1S1 = --- compGrIso --- (diagonalIso1 --- _ --- (coHomGr 0 (S₊ 0)) --- _ --- (invGrIso coHom0-S0) --- (d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) --- (λ x → MV.Ker-i⊂Im-d _ _(S₊ 0) (λ _ → tt) (λ _ → tt) 0 x --- (×≡ (isOfHLevelSuc 0 (isContrCohomUnit 0) _ _) --- (isOfHLevelSuc 0 (isContrCohomUnit 0) _ _))) --- (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) --- (λ x inker --- → pRec propTruncIsProp --- (λ {((f , g) , id') → helper x f g id' inker}) --- ((MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ inker)))) --- (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) --- λ F surj --- → pRec (setTruncIsSet _ _) (λ { (x , id) → MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ F ∣₀ ∣ (∣ (λ _ → x) ∣₀ , ∣ (λ _ → 0) ∣₀) , --- (cong ∣_∣₀ (funExt (surjHelper x))) ∙ sym id ∣₋₁ }) surj) ) --- (invGrIso (coHomPushout≡coHomSn' 0 1)) +-- (iso (sRec isSetInt (λ f → f (north , north))) +-- (λ a → ∣ (λ x → a) ∣₀) +-- (λ a → refl) +-- (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- λ f → cong ∣_∣₀ +-- (funExt λ {(x , y) → toProdElimSuspElim2 +-- {B = λ x y → f (north , north) ≡ f (x , y)} +-- north +-- (λ _ _ → isSetInt _ _) +-- refl +-- x y}))) +-- (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) λ a b → addLemma (a (north , north)) (b (north , north)))) + +-- -- private +-- -- S¹map : hLevelTrunc 3 S¹ → S¹ +-- -- S¹map = trElim (λ _ → isGroupoidS¹) (idfun S¹) + +-- -- S¹map-id : (x : hLevelTrunc 3 S¹) → Path (hLevelTrunc 3 S¹) ∣ S¹map x ∣ x +-- -- S¹map-id = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) +-- -- λ a → refl + +-- -- S1map : hLevelTrunc 3 (S₊ 1) → (S₊ 1) +-- -- S1map = trElim (λ _ → isGroupoidS1) (idfun _) + + +-- -- S¹→S¹ : Iso (S¹ → hLevelTrunc 3 S¹) (S¹ × Int) +-- -- Iso.fun S¹→S¹ f = S¹map (f base) +-- -- , winding (basechange2⁻ (S¹map (f base)) λ i → S¹map (f (loop i))) +-- -- Iso.inv S¹→S¹ (s , int) base = ∣ s ∣ +-- -- Iso.inv S¹→S¹ (s , int) (loop i) = ∣ basechange2 s (intLoop int) i ∣ +-- -- Iso.rightInv S¹→S¹ (s , int) = ×≡ refl ((λ i → winding (basechange2-retr s (λ i → intLoop int i) i)) +-- -- ∙ windingIntLoop int) +-- -- Iso.leftInv S¹→S¹ f = funExt λ { base → S¹map-id (f base) +-- -- ; (loop i) j → helper2 j i} +-- -- where +-- -- helper2 : PathP (λ i → S¹map-id (f base) i ≡ S¹map-id (f base) i) +-- -- (λ i → ∣ basechange2 (S¹map (f base)) (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁)))))) i ∣) +-- -- (cong f loop) +-- -- helper2 i j = +-- -- hcomp (λ k → λ { (i = i0) → cong ∣_∣ (basechange2 (S¹map (f base)) (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁))))))) j +-- -- ; (i = i1) → S¹map-id (f (loop j)) k +-- -- ; (j = i0) → S¹map-id (f base) (i ∧ k) +-- -- ; (j = i1) → S¹map-id (f base) (i ∧ k)}) +-- -- (helper4 i j) +-- -- where +-- -- helper4 : Path (Path (hLevelTrunc 3 _) _ _) +-- -- (cong ∣_∣ (basechange2 (S¹map (f base)) +-- -- (intLoop +-- -- (winding +-- -- (basechange2⁻ (S¹map (f base)) +-- -- (λ i₁ → S¹map (f (loop i₁)))))))) +-- -- λ i → ∣ S¹map (f (loop i)) ∣ +-- -- helper4 i = +-- -- cong ∣_∣ +-- -- ((cong (basechange2 (S¹map (f base))) +-- -- (decodeEncode base (basechange2⁻ (S¹map (f base)) +-- -- (λ i₁ → S¹map (f (loop i₁))))) +-- -- ∙ basechange2-sect (S¹map (f base)) +-- -- (λ i → S¹map (f (loop i)))) i) + +-- -- S1→S1→S1×Int : Iso ((S₊ 1) → hLevelTrunc 3 (S₊ 1)) ((hLevelTrunc 3 (S₊ 1)) × Int) +-- -- S1→S1→S1×Int = compIso helper2 (compIso S¹→S¹ helper) +-- -- where +-- -- helper : Iso (S¹ × Int) (hLevelTrunc 3 (S₊ 1) × Int) +-- -- Iso.fun helper (s , int) = ∣ S¹→S1 s ∣ , int +-- -- Iso.inv helper (s , int) = (S1→S¹ (S1map s)) , int +-- -- Iso.rightInv helper (s , int) = +-- -- trElim {B = λ s → (∣ S¹→S1 (S1→S¹ (S1map s)) ∣ , int) ≡ (s , int)} +-- -- (λ _ → isOfHLevelPath 3 (isOfHLevelProd 3 (isOfHLevelTrunc 3) (isOfHLevelSuc 2 isSetInt)) _ _) +-- -- (λ a → ×≡ (cong ∣_∣ (S1→S¹-retr a)) refl) +-- -- s +-- -- Iso.leftInv helper (s , int) = ×≡ (S1→S¹-sect s) refl + +-- -- helper2 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → hLevelTrunc 3 S¹) +-- -- Iso.fun helper2 f x = trMap S1→S¹ (f (S¹→S1 x)) +-- -- Iso.inv helper2 f x = trMap S¹→S1 (f (S1→S¹ x)) +-- -- Iso.rightInv helper2 f = funExt λ x i → helper3 (f (S1→S¹-sect x i)) i +-- -- where +-- -- helper3 : (x : hLevelTrunc 3 S¹) → trMap S1→S¹ (trMap S¹→S1 x) ≡ x +-- -- helper3 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) +-- -- λ a → cong ∣_∣ (S1→S¹-sect a) +-- -- Iso.leftInv helper2 f = funExt λ x i → helper3 (f (S1→S¹-retr x i)) i +-- -- where +-- -- helper3 : (x : hLevelTrunc 3 (S₊ 1)) → trMap S¹→S1 (trMap S1→S¹ x) ≡ x +-- -- helper3 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) +-- -- λ a → cong ∣_∣ (S1→S¹-retr a) + +-- -- S¹→S¹≡S1→S1 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → hLevelTrunc 3 S¹) +-- -- Iso.fun S¹→S¹≡S1→S1 f x = trMap S1→S¹ (f (S¹→S1 x)) +-- -- Iso.inv S¹→S¹≡S1→S1 f x = trMap S¹→S1 (f (S1→S¹ x)) +-- -- Iso.rightInv S¹→S¹≡S1→S1 F = funExt λ x i → helper2 (F (S1→S¹-sect x i)) i +-- -- where +-- -- helper2 : (x : hLevelTrunc 3 S¹) → trMap S1→S¹ (trMap S¹→S1 x) ≡ x +-- -- helper2 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) +-- -- λ a → cong ∣_∣ (S1→S¹-sect a) +-- -- Iso.leftInv S¹→S¹≡S1→S1 F = funExt λ x i → helper2 (F (S1→S¹-retr x i)) i +-- -- where +-- -- helper2 : (x : hLevelTrunc 3 (S₊ 1)) → trMap S¹→S1 (trMap S1→S¹ x) ≡ x +-- -- helper2 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) +-- -- λ a → cong ∣_∣ (S1→S¹-retr a) + + + + +-- -- basechange-lemma : ∀ {ℓ} {A : Type ℓ} {a : A} (x y : S¹) (F : a ≡ a → S¹) (f : S¹ → a ≡ a) (g : S¹ → a ≡ a) +-- -- → (f base ≡ refl) +-- -- → (g base ≡ refl) +-- -- → basechange2⁻ (F (f base ∙ g base)) (cong₂ {A = S¹} {B = λ x → S¹} (λ x y → F (f x ∙ g y)) loop loop) +-- -- ≡ basechange2⁻ (F (f base)) (cong (λ x → F (f x)) loop) ∙ basechange2⁻ (F (g base)) (cong (λ x → F (g x)) loop) +-- -- basechange-lemma x y F f g frefl grefl = +-- -- (λ i → basechange2⁻ (F (f base ∙ g base)) (cong₂Funct2 (λ x y → F (f x ∙ g y)) loop loop i)) +-- -- ∙ (λ i → basechange2⁻ (F (f base ∙ g base)) (cong (λ x₁ → F (f x₁ ∙ g base)) loop ∙ cong (λ y₁ → F (f base ∙ g y₁)) loop)) +-- -- ∙ basechange2⁻-morph (F (f base ∙ g base)) _ _ +-- -- ∙ (λ j → basechange2⁻ (F (f base ∙ grefl j)) +-- -- (λ i → F (f (loop i) ∙ grefl j)) +-- -- ∙ basechange2⁻ (F (frefl j ∙ g base)) +-- -- (λ i → F (frefl j ∙ g (loop i)))) +-- -- ∙ ((λ j → basechange2⁻ (F (rUnit (f base) (~ j))) +-- -- (λ i → F (rUnit (f (loop i)) (~ j))) +-- -- ∙ basechange2⁻ (F (lUnit (g base) (~ j))) +-- -- (λ i → F (lUnit (g (loop i)) (~ j))))) + + +-- -- basechange-lemma2 : (f g : S¹ → hLevelTrunc 3 (S₊ 1)) (F : hLevelTrunc 3 (S₊ 1) → S¹) +-- -- → ((basechange2⁻ (F (f base +ₖ g base)) λ i → F ((f (loop i)) +ₖ g (loop i))) +-- -- ≡ basechange2⁻ (F (f base)) (cong (F ∘ f) loop) +-- -- ∙ basechange2⁻ (F (g base)) (cong (F ∘ g) loop)) +-- -- basechange-lemma2 f g F = coInd f g F (f base) (g base) refl refl +-- -- where +-- -- coInd : (f g : S¹ → hLevelTrunc 3 (S₊ 1)) (F : hLevelTrunc 3 (S₊ 1) → S¹) (x y : hLevelTrunc 3 (S₊ 1)) +-- -- → f base ≡ x +-- -- → g base ≡ y +-- -- → ((basechange2⁻ (F (f base +ₖ g base)) λ i → F ((f (loop i)) +ₖ g (loop i))) +-- -- ≡ basechange2⁻ (F (f base)) (cong (F ∘ f) loop) +-- -- ∙ basechange2⁻ (F (g base)) (cong (F ∘ g) loop)) +-- -- coInd f g F = +-- -- elim2 (λ _ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 (isGroupoidS¹ base base)) _ _ ) +-- -- (toProdElimSuspElim2 +-- -- north +-- -- (λ _ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → isGroupoidS¹ _ _ _ _) +-- -- λ fid gid → +-- -- (λ i → basechange2⁻ (F (f base +ₖ g base)) (cong₂Funct2 (λ x y → F (f x +ₖ g y)) loop loop i)) -- (λ i → F (f (loop i) +ₖ g (loop i)))) +-- -- ∙ (basechange2⁻-morph (F (f base +ₖ g base)) +-- -- (cong (λ x → F (f x +ₖ g base)) loop) +-- -- (cong (λ x → F (f base +ₖ g x)) loop)) +-- -- ∙ (λ i → basechange2⁻ (F (f base +ₖ gid i)) (cong (λ x → F (f x +ₖ gid i)) loop) +-- -- ∙ basechange2⁻ (F (fid i +ₖ g base)) (cong (λ x → F (fid i +ₖ g x)) loop)) +-- -- ∙ (λ i → basechange2⁻ (F (rUnitₖ (f base) i)) (cong (λ x → F (rUnitₖ (f x) i)) loop) +-- -- ∙ basechange2⁻ (F (lUnitₖ (g base) i)) (cong (λ x → F (lUnitₖ (g x) i)) loop))) + + + +-- -- coHom1S1≃ℤ : grIso (coHomGr 1 (S₊ 1)) intGroup +-- -- coHom1S1≃ℤ = +-- -- Iso'→Iso +-- -- (iso' +-- -- (iso +-- -- (sRec isSetInt (λ f → proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) +-- -- (λ a → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , a)) ∣₀) +-- -- (λ a → (λ i → proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 (Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , a)))))) +-- -- ∙ (λ i → proj₂ (Iso.fun S¹→S¹ (Iso.rightInv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , a)) i))) +-- -- ∙ λ i → proj₂ (Iso.rightInv S¹→S¹ (base , a) i)) +-- -- (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- -- λ f → (λ i → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) ∣₀) +-- -- ∙ (λ i → sRec setTruncIsSet +-- -- (λ x → ∣ Iso.inv S¹→S¹≡S1→S1 x ∣₀) +-- -- (sRec setTruncIsSet +-- -- (λ x → ∣ Iso.inv S¹→S¹ (x , (proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) ∣₀) +-- -- ∣ base ∣₀)) +-- -- ∙ (λ i → sRec setTruncIsSet +-- -- (λ x → ∣ Iso.inv S¹→S¹≡S1→S1 x ∣₀) +-- -- (sRec setTruncIsSet +-- -- (λ x → ∣ Iso.inv S¹→S¹ (x , (proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) ∣₀) +-- -- (Iso.inv PathIdTrunc₀Iso (isConnectedS¹ (proj₁ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) i))) +-- -- ∙ (λ i → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.leftInv S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f) i) ∣₀) +-- -- ∙ (λ i → ∣ Iso.leftInv S¹→S¹≡S1→S1 f i ∣₀))) +-- -- (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) +-- -- λ f g → (λ i → winding (basechange2⁻ (S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)) ∙ Kn→ΩKn+1 1 (g (S¹→S1 base)))))) +-- -- (λ i → S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 (loop i))) ∙ Kn→ΩKn+1 1 (g (S¹→S1 (loop i))))))))) +-- -- ∙ cong winding (helper2 (f (S¹→S1 base)) (g (S¹→S1 base)) f g refl refl) +-- -- ∙ winding-hom ((basechange2⁻ (S¹map (trMap S1→S¹ (f north))) +-- -- (λ i → S¹map (trMap S1→S¹ (f (S¹→S1 (loop i))))))) +-- -- ((basechange2⁻ (S¹map (trMap S1→S¹ (g north))) +-- -- (λ i → S¹map (trMap S1→S¹ (g (S¹→S1 (loop i))))))))) + + +-- -- where +-- -- helper2 : (x y : hLevelTrunc 3 (S₊ 1)) (f g : S₊ 1 → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1) +-- -- → (f (S¹→S1 base)) ≡ x +-- -- → (g (S¹→S1 base)) ≡ y +-- -- → (basechange2⁻ (S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)) ∙ Kn→ΩKn+1 1 (g (S¹→S1 base))))))) +-- -- (λ i → S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 (loop i))) ∙ Kn→ΩKn+1 1 (g (S¹→S1 (loop i))))))) +-- -- ≡ (basechange2⁻ (S¹map (trMap S1→S¹ ((f (S¹→S1 base)))))) +-- -- (λ i → S¹map (trMap S1→S¹ (f (S¹→S1 (loop i))))) +-- -- ∙ (basechange2⁻ (S¹map (trMap S1→S¹ ((g (S¹→S1 base))))) +-- -- (λ i → S¹map (trMap S1→S¹ ((g (S¹→S1 (loop i))))))) +-- -- helper2 = elim2 +-- -- (λ _ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 +-- -- λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 +-- -- λ _ → isOfHLevelPath 3 (isOfHLevelSuc 3 (isGroupoidS¹) base base) _ _) +-- -- (toProdElimSuspElim2 {A = S₊ 0} north +-- -- (λ _ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → (isGroupoidS¹) _ _ _ _ ) +-- -- λ f g reflf reflg → +-- -- (basechange-lemma +-- -- base +-- -- base +-- -- (λ x → S¹map (trMap S1→S¹ (ΩKn+1→Kn x))) +-- -- (λ x → Kn→ΩKn+1 1 (f (S¹→S1 x))) ((λ x → Kn→ΩKn+1 1 (g (S¹→S1 x)))) +-- -- (cong (Kn→ΩKn+1 1) reflf ∙ Kn→ΩKn+10ₖ 1) +-- -- (cong (Kn→ΩKn+1 1) reflg ∙ Kn→ΩKn+10ₖ 1)) +-- -- ∙ λ j → basechange2⁻ (S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 base)) j))) +-- -- (λ i → S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 (loop i))) j))) +-- -- ∙ basechange2⁻ (S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 base)) j))) +-- -- (λ i → S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 (loop i))) j)))) + + + + + + + + +-- -- indIntGroup : ∀ {ℓ} {G : Group ℓ} → (ϕ : Int → Group.type G) +-- -- → ϕ 0 ≡ isGroup.id (Group.groupStruc G) +-- -- → ((a : Int) → ϕ (a +ℤ 1) ≡ isGroup.comp (Group.groupStruc G) (ϕ a) (ϕ 1)) +-- -- → ((n : Int) → ϕ (predInt n) ≡ isGroup.comp (Group.groupStruc G) (ϕ n) (ϕ (negsuc zero))) +-- -- → isMorph intGroup G ϕ +-- -- indIntGroup {G = group G gSet (group-struct _ _ _+G_ _ rUnit₁ _ _ _)} ϕ zeroId _ _ n (pos zero) = +-- -- sym (rUnit₁ (ϕ n)) ∙ cong (λ x → ϕ n +G x) (sym zeroId) +-- -- indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId oneId minOneId n (pos (suc m)) = +-- -- (λ i → ϕ ((n +pos m) +ℤ 1)) +-- -- ∙ oneId (n +pos m) +-- -- ∙ cong (λ x → x +G ϕ (pos 1)) +-- -- (indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} +-- -- ϕ zeroId oneId minOneId n (pos m)) +-- -- ∙ assoc₁ (ϕ n) (ϕ (pos m)) (ϕ (pos 1)) +-- -- ∙ sym (cong (λ x → ϕ n +G x) (oneId (pos m))) +-- -- indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId _ minOneId n (negsuc zero) = minOneId n +-- -- indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId a minOneId n (negsuc (suc m)) = +-- -- (λ i → ϕ ((n +negsuc m) +ℤ (negsuc zero))) +-- -- ∙ minOneId (n +negsuc m) +-- -- ∙ cong (λ x → x +G ϕ (negsuc zero)) (indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId a minOneId n (negsuc m)) +-- -- ∙ assoc₁ (ϕ n) (ϕ (negsuc m)) (ϕ (negsuc zero)) +-- -- ∙ cong (λ x → ϕ n +G x) (sym (minOneId (negsuc m))) + +-- -- pushoutSn : (n : ℕ) → Iso (S₊ (suc n)) (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt) +-- -- Iso.fun (pushoutSn n) north = inl tt +-- -- Iso.fun (pushoutSn n) south = inr tt +-- -- Iso.fun (pushoutSn n) (merid a i) = push a i +-- -- Iso.inv (pushoutSn n) (inl x) = north +-- -- Iso.inv (pushoutSn n) (inr x) = south +-- -- Iso.inv (pushoutSn n) (push a i) = merid a i +-- -- Iso.rightInv (pushoutSn n) (inl x) = refl +-- -- Iso.rightInv (pushoutSn n) (inr x) = refl +-- -- Iso.rightInv (pushoutSn n) (push a i) = refl +-- -- Iso.leftInv (pushoutSn n) north = refl +-- -- Iso.leftInv (pushoutSn n) south = refl +-- -- Iso.leftInv (pushoutSn n) (merid a i) = refl + +-- -- Sn≡Pushout : (n : ℕ) → (S₊ (suc n)) ≡ (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt) +-- -- Sn≡Pushout n = isoToPath (pushoutSn n) + +-- -- coHomPushout≡coHomSn' : (n m : ℕ) → grIso (coHomGr m (S₊ (suc n))) (coHomGr m (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt)) +-- -- morph.fun (grIso.fun (coHomPushout≡coHomSn' n m)) = +-- -- sRec setTruncIsSet +-- -- λ f → ∣ (λ {(inl x) → f north ; (inr x) → f south ; (push a i) → f (merid a i)}) ∣₀ +-- -- morph.ismorph (grIso.fun (coHomPushout≡coHomSn' n m)) = +-- -- sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- -- λ a b → cong ∣_∣₀ (funExt λ {(inl x) → refl ; (inr x) → refl ; (push a i) → refl }) +-- -- morph.fun (grIso.inv (coHomPushout≡coHomSn' n m)) = sRec setTruncIsSet (λ f → ∣ (λ {north → f (inl tt) ; south → f (inr tt) ; (merid a i) → f (push a i)}) ∣₀) +-- -- morph.ismorph (grIso.inv (coHomPushout≡coHomSn' n m)) = +-- -- sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- -- λ a b → cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → refl }) +-- -- grIso.rightInv (coHomPushout≡coHomSn' n m) = +-- -- sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- -- λ f → cong ∣_∣₀ (funExt λ {(inl x) → refl ; (inr x) → refl ; (push a i) → refl }) +-- -- grIso.leftInv (coHomPushout≡coHomSn' n m) = +-- -- sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- -- λ f → cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → refl }) + + +-- -- isContr→≡Unit : {A : Type₀} → isContr A → A ≡ Unit +-- -- isContr→≡Unit contr = isoToPath (iso (λ _ → tt) (λ _ → fst contr) (λ _ → refl) λ _ → snd contr _) + +-- -- isContr→isContrTrunc : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → isContr A → isContr (hLevelTrunc n A) +-- -- isContr→isContrTrunc n contr = ∣ fst contr ∣ , (trElim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → cong ∣_∣ (snd contr a)) +-- -- isContr→isContrSetTrunc : ∀ {ℓ} {A : Type ℓ} → isContr A → isContr (∥ A ∥₀) +-- -- isContr→isContrSetTrunc contr = ∣ fst contr ∣₀ , sElim (λ _ → isOfHLevelPath 2 (setTruncIsSet) _ _) λ a → cong ∣_∣₀ (snd contr a) + +-- -- coHomGrUnit0 : grIso (coHomGr 0 Unit) intGroup +-- -- grIso.fun coHomGrUnit0 = mph (sRec isSetInt (λ f → f tt)) +-- -- (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) +-- -- (λ a b → addLemma (a tt) (b tt))) +-- -- grIso.inv coHomGrUnit0 = mph (λ a → ∣ (λ _ → a) ∣₀) (λ a b i → ∣ (λ _ → addLemma a b (~ i)) ∣₀) +-- -- grIso.rightInv coHomGrUnit0 a = refl +-- -- grIso.leftInv coHomGrUnit0 = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ a → refl + +-- -- isContrCohomUnit : (n : ℕ) → isContr (coHom (suc n) Unit) +-- -- isContrCohomUnit n = subst isContr (λ i → ∥ UnitToTypeId (coHomK (suc n)) (~ i) ∥₀) (helper n) +-- -- where +-- -- helper : (n : ℕ) → isContr (∥ coHomK (suc n) ∥₀) +-- -- helper n = subst isContr ((isoToPath (truncOfTruncIso {A = S₊ (1 + n)} 2 (1 + n))) ∙ sym propTrunc≡Trunc0 ∙ λ i → ∥ hLevelTrunc (suc (+-comm n 2 i)) (S₊ (1 + n)) ∥₀) +-- -- (isConnectedSubtr 2 (helper2 n .fst) (subst (λ x → isHLevelConnected x (S₊ (suc n))) (sym (helper2 n .snd)) (sphereConnected (suc n))) ) +-- -- where +-- -- helper2 : (n : ℕ) → Σ[ m ∈ ℕ ] m + 2 ≡ 2 + n +-- -- helper2 zero = 0 , refl +-- -- helper2 (suc n) = (suc n) , λ i → suc (+-comm n 2 i) + +-- -- coHomGrUnit≥1 : (n : ℕ) → grIso (coHomGr (suc n) Unit) trivialGroup +-- -- grIso.fun (coHomGrUnit≥1 n) = mph (λ _ → tt) (λ _ _ → refl) +-- -- grIso.inv (coHomGrUnit≥1 n) = mph (λ _ → ∣ (λ _ → ∣ north ∣) ∣₀) (λ _ _ → sym (rUnitₕ 0ₕ)) +-- -- grIso.rightInv (coHomGrUnit≥1 n) a = refl +-- -- grIso.leftInv (coHomGrUnit≥1 n) a = sym (isContrCohomUnit n .snd 0ₕ) ∙ isContrCohomUnit n .snd a + +-- -- S0→Int : (a : Int × Int) → S₊ 0 → Int +-- -- S0→Int a north = proj₁ a +-- -- S0→Int a south = proj₂ a + +-- -- coHom0-S0 : grIso (coHomGr 0 (S₊ 0)) (dirProd intGroup intGroup) +-- -- coHom0-S0 = +-- -- Iso'→Iso +-- -- (iso' +-- -- (iso (sRec (isOfHLevelProd 2 isSetInt isSetInt) +-- -- λ f → (f north) , (f south)) +-- -- (λ a → ∣ S0→Int a ∣₀) +-- -- (λ { (a , b) → refl}) +-- -- (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ f → cong ∣_∣₀ (funExt (λ {north → refl ; south → refl})))) +-- -- (sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 isSetInt isSetInt) _ _) +-- -- λ a b i → addLemma (a north) (b north) i , addLemma (a south) (b south) i)) + +-- -- ×morph : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group ℓ} {B : Group ℓ'} {C : Group ℓ''} {D : Group ℓ'''} → morph A B → morph C D → morph (dirProd A C) (dirProd B D) +-- -- morph.fun (×morph mf1 mf2) = +-- -- (λ {(a , b) → (morph.fun mf1 a) , morph.fun mf2 b}) +-- -- morph.ismorph (×morph mf1 mf2) = +-- -- (λ {(a , b) (c , d) i → morph.ismorph mf1 a c i , morph.ismorph mf2 b d i}) + + +-- -- coHom0S1 : grIso intGroup (coHomGr 0 (S₊ 1)) +-- -- coHom0S1 = +-- -- Iso'→Iso +-- -- (iso' +-- -- (iso +-- -- (λ a → ∣ (λ x → a) ∣₀) +-- -- (sRec isSetInt (λ f → f north)) +-- -- (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- -- λ f → cong ∣_∣₀ (funExt (toProdElimSuspRec north (λ _ → isSetInt _ _) refl))) +-- -- (λ a → refl)) +-- -- λ a b → cong ∣_∣₀ (funExt λ x → sym (addLemma a b))) + +-- -- coHom1S1 : grIso intGroup (coHomGr 1 (S₊ 1)) +-- -- coHom1S1 = +-- -- compGrIso +-- -- (diagonalIso1 +-- -- _ +-- -- (coHomGr 0 (S₊ 0)) +-- -- _ +-- -- (invGrIso coHom0-S0) +-- -- (d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) +-- -- (λ x → MV.Ker-i⊂Im-d _ _(S₊ 0) (λ _ → tt) (λ _ → tt) 0 x +-- -- (×≡ (isOfHLevelSuc 0 (isContrCohomUnit 0) _ _) +-- -- (isOfHLevelSuc 0 (isContrCohomUnit 0) _ _))) +-- -- (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- -- (λ x inker +-- -- → pRec propTruncIsProp +-- -- (λ {((f , g) , id') → helper x f g id' inker}) +-- -- ((MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ inker)))) +-- -- (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- -- λ F surj +-- -- → pRec (setTruncIsSet _ _) (λ { (x , id) → MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ F ∣₀ ∣ (∣ (λ _ → x) ∣₀ , ∣ (λ _ → 0) ∣₀) , +-- -- (cong ∣_∣₀ (funExt (surjHelper x))) ∙ sym id ∣₋₁ }) surj) ) +-- -- (invGrIso (coHomPushout≡coHomSn' 0 1)) --- where --- surjHelper : (x : Int) (x₁ : S₊ 0) → x +ₖ (-ₖ pos 0) ≡ S0→Int (x , x) x₁ --- surjHelper x north = cong (x +ₖ_) (-0ₖ) ∙ rUnitₖ x --- surjHelper x south = cong (x +ₖ_) (-0ₖ) ∙ rUnitₖ x - --- helper : (F : S₊ 0 → Int) (f g : ∥ (Unit → Int) ∥₀) (id : MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (f , g) ≡ ∣ F ∣₀) --- → isInKer (coHomGr 0 (S₊ 0)) --- (coHomGr 1 (Pushout (λ _ → tt) (λ _ → tt))) --- (morph.fun (d-morph Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) --- ∣ F ∣₀ --- → ∃[ x ∈ Int ] ∣ F ∣₀ ≡ morph.fun (grIso.inv coHom0-S0) (x , x) --- helper F = --- sElim2 (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) --- λ f g id inker --- → pRec propTruncIsProp --- (λ {((a , b) , id2) --- → sElim2 {B = λ f g → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (f , g) ≡ ∣ F ∣₀ → _ } --- (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) --- (λ f g id → ∣ (helper2 f g .fst) , (sym id ∙ sym (helper2 f g .snd)) ∣₋₁) --- a b id2}) --- (MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ F ∣₀ inker) --- where --- helper2 : (f g : Unit → Int) --- → Σ[ x ∈ Int ] morph.fun (grIso.inv coHom0-S0) (x , x) --- ≡ MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (∣ f ∣₀ , ∣ g ∣₀) --- helper2 f g = (f _ +ₖ (-ₖ g _) ) , cong ∣_∣₀ (funExt (λ {north → refl ; south → refl})) - - --- coHom-n-Sn : (n : ℕ) → grIso intGroup (coHomGr (suc n) (S₊ (suc n))) --- coHom-n-Sn zero = coHom1S1 --- coHom-n-Sn (suc n) = --- compGrIso --- (compGrIso --- (coHom-n-Sn n) --- theIso) --- (invGrIso (coHomPushout≡coHomSn' (suc n) (suc (suc n)))) --- where --- theIso : grIso (coHomGr (suc n) (S₊ (suc n))) (coHomGr (suc (suc n)) --- (Pushout {A = S₊ (suc n)} (λ _ → tt) (λ _ → tt))) --- theIso = --- SES→Iso --- (×coHomGr (suc n) Unit Unit) --- (×coHomGr (suc (suc n)) Unit Unit) --- (ses (λ p q → ×≡ (isOfHLevelSuc 0 (isContrCohomUnit n) (proj₁ p) (proj₁ q)) --- (isOfHLevelSuc 0 (isContrCohomUnit n) (proj₂ p) (proj₂ q))) --- (λ p q → ×≡ (isOfHLevelSuc 0 (isContrCohomUnit (suc n)) (proj₁ p) (proj₁ q)) --- (isOfHLevelSuc 0 (isContrCohomUnit (suc n)) (proj₂ p) (proj₂ q))) --- (Δ-morph _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) --- (i-morph _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (2 + n)) --- (d-morph _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) --- (MV.Ker-d⊂Im-Δ _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) --- (MV.Ker-i⊂Im-d _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n))) +-- -- where +-- -- surjHelper : (x : Int) (x₁ : S₊ 0) → x +ₖ (-ₖ pos 0) ≡ S0→Int (x , x) x₁ +-- -- surjHelper x north = cong (x +ₖ_) (-0ₖ) ∙ rUnitₖ x +-- -- surjHelper x south = cong (x +ₖ_) (-0ₖ) ∙ rUnitₖ x + +-- -- helper : (F : S₊ 0 → Int) (f g : ∥ (Unit → Int) ∥₀) (id : MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (f , g) ≡ ∣ F ∣₀) +-- -- → isInKer (coHomGr 0 (S₊ 0)) +-- -- (coHomGr 1 (Pushout (λ _ → tt) (λ _ → tt))) +-- -- (morph.fun (d-morph Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) +-- -- ∣ F ∣₀ +-- -- → ∃[ x ∈ Int ] ∣ F ∣₀ ≡ morph.fun (grIso.inv coHom0-S0) (x , x) +-- -- helper F = +-- -- sElim2 (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- -- λ f g id inker +-- -- → pRec propTruncIsProp +-- -- (λ {((a , b) , id2) +-- -- → sElim2 {B = λ f g → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (f , g) ≡ ∣ F ∣₀ → _ } +-- -- (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- -- (λ f g id → ∣ (helper2 f g .fst) , (sym id ∙ sym (helper2 f g .snd)) ∣₋₁) +-- -- a b id2}) +-- -- (MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ F ∣₀ inker) +-- -- where +-- -- helper2 : (f g : Unit → Int) +-- -- → Σ[ x ∈ Int ] morph.fun (grIso.inv coHom0-S0) (x , x) +-- -- ≡ MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (∣ f ∣₀ , ∣ g ∣₀) +-- -- helper2 f g = (f _ +ₖ (-ₖ g _) ) , cong ∣_∣₀ (funExt (λ {north → refl ; south → refl})) + + +-- -- coHom-n-Sn : (n : ℕ) → grIso intGroup (coHomGr (suc n) (S₊ (suc n))) +-- -- coHom-n-Sn zero = coHom1S1 +-- -- coHom-n-Sn (suc n) = +-- -- compGrIso +-- -- (compGrIso +-- -- (coHom-n-Sn n) +-- -- theIso) +-- -- (invGrIso (coHomPushout≡coHomSn' (suc n) (suc (suc n)))) +-- -- where +-- -- theIso : grIso (coHomGr (suc n) (S₊ (suc n))) (coHomGr (suc (suc n)) +-- -- (Pushout {A = S₊ (suc n)} (λ _ → tt) (λ _ → tt))) +-- -- theIso = +-- -- SES→Iso +-- -- (×coHomGr (suc n) Unit Unit) +-- -- (×coHomGr (suc (suc n)) Unit Unit) +-- -- (ses (λ p q → ×≡ (isOfHLevelSuc 0 (isContrCohomUnit n) (proj₁ p) (proj₁ q)) +-- -- (isOfHLevelSuc 0 (isContrCohomUnit n) (proj₂ p) (proj₂ q))) +-- -- (λ p q → ×≡ (isOfHLevelSuc 0 (isContrCohomUnit (suc n)) (proj₁ p) (proj₁ q)) +-- -- (isOfHLevelSuc 0 (isContrCohomUnit (suc n)) (proj₂ p) (proj₂ q))) +-- -- (Δ-morph _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) +-- -- (i-morph _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (2 + n)) +-- -- (d-morph _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) +-- -- (MV.Ker-d⊂Im-Δ _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) +-- -- (MV.Ker-i⊂Im-d _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n))) + + +-- -- setTruncIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → Iso A B → Iso ∥ A ∥₀ ∥ B ∥₀ +-- -- Iso.fun (setTruncIso is) = sRec setTruncIsSet (λ x → ∣ Iso.fun is x ∣₀) +-- -- Iso.inv (setTruncIso is) = sRec setTruncIsSet (λ x → ∣ Iso.inv is x ∣₀) +-- -- Iso.rightInv (setTruncIso is) = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ a → cong ∣_∣₀ (Iso.rightInv is a) +-- -- Iso.leftInv (setTruncIso is) = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ a → cong ∣_∣₀ (Iso.leftInv is a) + +-- -- targetIso : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} → Iso B C → Iso (A → B) (A → C) +-- -- Iso.fun (targetIso is) f a = Iso.fun is (f a) +-- -- Iso.inv (targetIso is) f a = Iso.inv is (f a) +-- -- Iso.rightInv (targetIso is) f = funExt λ a → Iso.rightInv is (f a) +-- -- Iso.leftInv (targetIso is) f = funExt λ a → Iso.leftInv is (f a) + +-- -- coHom1Torus : grIso (coHomGr 1 ((S₊ 1) × (S₊ 1))) (dirProd intGroup intGroup) +-- -- coHom1Torus = +-- -- compGrIso +-- -- (Iso'→Iso +-- -- (iso' (compIso +-- -- (setTruncIso (compIso +-- -- schonfIso +-- -- (compIso +-- -- (targetIso S1→S1→S1×Int) +-- -- funIso))) +-- -- (setTruncOfProdIso)) +-- -- (sElim2 +-- -- (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- -- λ f g → ×≡ (cong ∣_∣₀ +-- -- (funExt (λ x → helper (f (x , S¹→S1 base) +ₖ g (x , S¹→S1 base)) +-- -- ∙ sym (cong₂ (λ x y → x +ₖ y) +-- -- (helper (f (x , S¹→S1 base))) +-- -- (helper (g (x , S¹→S1 base))))))) +-- -- (cong ∣_∣₀ +-- -- (funExt +-- -- (toProdElimSuspRec +-- -- north +-- -- (λ _ → isSetInt _ _) +-- -- (cong winding +-- -- (basechange-lemma2 +-- -- (λ x → f (north , S¹→S1 x)) +-- -- (λ x → g (north , S¹→S1 x)) +-- -- λ x → S¹map (trMap S1→S¹ x)) +-- -- ∙ winding-hom +-- -- (basechange2⁻ +-- -- (S¹map (trMap S1→S¹ (f (north , S¹→S1 base)))) +-- -- (λ i → S¹map (trMap S1→S¹ (f (north , S¹→S1 (loop i)))))) +-- -- (basechange2⁻ +-- -- (S¹map (trMap S1→S¹ (g (north , S¹→S1 base)))) +-- -- (λ i → S¹map (trMap S1→S¹ (g (north , S¹→S1 (loop i)))))) +-- -- ∙ sym (addLemma +-- -- (winding +-- -- (basechange2⁻ +-- -- (S¹map (trMap S1→S¹ (f (north , S¹→S1 base)))) +-- -- (λ i → S¹map (trMap S1→S¹ (f (north , S¹→S1 (loop i))))))) +-- -- (winding +-- -- (basechange2⁻ +-- -- (S¹map (trMap S1→S¹ (g (north , S¹→S1 base)))) +-- -- (λ i → S¹map (trMap S1→S¹ (g (north , S¹→S1 (loop i))))))))))))))) +-- -- (dirProdIso (invGrIso (coHom-n-Sn 0)) (invGrIso coHom0S1)) +-- -- where +-- -- helper : (x : hLevelTrunc 3 (S₊ 1)) → ∣ S¹→S1 (S¹map (trMap S1→S¹ x)) ∣ ≡ x +-- -- helper = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ a → cong ∣_∣ (S1→S¹-retr a) +-- -- -- basechange* : (x y : S¹) → x ≡ y → x ≡ y → ΩS¹ +-- -- -- basechange* x y = J (λ y p → (x ≡ y) → ΩS¹) (basechange x) +-- -- -- test1 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S₊ 1 × Int) +-- -- -- Iso.fun test1 f = (trRec isGroupoidS1 (λ a → a) (f north)) +-- -- -- , winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (f (loop* i))))) +-- -- -- Iso.inv test1 (north , b) x = ∣ x ∣ +-- -- -- Iso.inv test1 (south , b) x = ∣ x ∣ +-- -- -- Iso.inv test1 (merid a i , b) x = {!!} +-- -- -- Iso.rightInv test1 = {!!} +-- -- -- Iso.leftInv test1 = {!!} +-- -- -- funRec : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) {C : (A → hLevelTrunc n B) → Type ℓ''} +-- -- -- → isOfHLevel n B +-- -- -- → ((f : A → B) → C (λ a → ∣ f a ∣)) +-- -- -- → (f : A → hLevelTrunc n B) → C f +-- -- -- funRec {A = A} {B = B} n {C = C} hLev ind f = subst C (helper f) (ind (λ a → trRec hLev (λ x → x) (f a))) +-- -- -- where +-- -- -- helper : retract {A = A → hLevelTrunc n B} {B = A → B} (λ f₁ a → trRec hLev (λ x → x) (f₁ a)) (λ f₁ a → ∣ f₁ a ∣) +-- -- -- helper f = funExt λ a → helper2 (f a) +-- -- -- where +-- -- -- helper2 : (x : hLevelTrunc n B) → ∣ trRec hLev (λ x → x) x ∣ ≡ x +-- -- -- helper2 = trElim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → refl --- -- basechange* : (x y : S¹) → x ≡ y → x ≡ y → ΩS¹ --- -- basechange* x y = J (λ y p → (x ≡ y) → ΩS¹) (basechange x) +-- -- -- invMapSurj : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (ϕ : morph G H) → ((x : Group.type H) → isInIm G H (fst ϕ) x) +-- -- -- → morph H G +-- -- -- fst (invMapSurj G H (ϕ , pf) surj) a = {!pRec!} +-- -- -- snd (invMapSurj G H (ϕ , pf) surj) = {!!} +-- -- {- +-- -- ImIso : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (ϕ : morph G H) → ((x : Group.type H) → isInIm G H (fst ϕ) x) +-- -- → grIso H (imGroup G H ϕ) +-- -- ImIso G H (ϕ , mf) surj = +-- -- let idH = isGroup.id (Group.groupStruc H) +-- -- idG = isGroup.id (Group.groupStruc G) +-- -- _+G_ = isGroup.comp (Group.groupStruc G) +-- -- _+H_ = isGroup.comp (Group.groupStruc H) +-- -- _+Im_ = isGroup.comp (Group.groupStruc (imGroup G H (ϕ , mf))) +-- -- invG = isGroup.inv (Group.groupStruc G) +-- -- invH = isGroup.inv (Group.groupStruc H) +-- -- lUnit = isGroup.lUnit (Group.groupStruc H) +-- -- lCancel = isGroup.rCancel (Group.groupStruc H) +-- -- in +-- -- Iso''→Iso _ _ +-- -- (iso'' ((λ x → x , pRec propTruncIsProp (λ (a , b) → ∣ a , b ∣₋₁) (surj x)) +-- -- , λ a b → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) +-- -- (λ surja → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) +-- -- (λ surjb → +-- -- pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) +-- -- (λ surja+b → +-- -- (λ i → (a +H b) , (pRec (propTruncIsProp) +-- -- (λ (a , b) → ∣ a , b ∣₋₁) +-- -- (propTruncIsProp (surj (isGroup.comp (Group.groupStruc H) a b)) ∣ surja+b ∣₋₁ i))) ∙ +-- -- (λ i → (a +H b) , ∣ (fst surja+b) , (snd surja+b) ∣₋₁) ∙ +-- -- ΣProp≡ (λ _ → propTruncIsProp) refl ∙ +-- -- λ i → (a +H b) , pRec (propTruncIsProp) +-- -- (λ p1 → +-- -- pRec (λ x y → squash x y) +-- -- (λ p2 → +-- -- ∣ +-- -- isGroup.comp (Group.groupStruc G) (fst p1) (fst p2) , +-- -- mf (fst p1) (fst p2) ∙ +-- -- cong₂ (isGroup.comp (Group.groupStruc H)) (snd p1) (snd p2) +-- -- ∣₋₁) +-- -- (pRec (propTruncIsProp) +-- -- ∣_∣₋₁ (propTruncIsProp ∣ surjb ∣₋₁ (surj b) i))) +-- -- (pRec (propTruncIsProp) +-- -- ∣_∣₋₁ (propTruncIsProp ∣ surja ∣₋₁ (surj a) i ))) +-- -- (surj (isGroup.comp (Group.groupStruc H) a b))) +-- -- (surj b)) +-- -- (surj a)) +-- -- (λ x inker → cong fst inker) +-- -- λ x → pRec propTruncIsProp (λ inimx → ∣ (ϕ (inimx .fst)) , ΣProp≡ (λ _ → propTruncIsProp) (inimx .snd) ∣₋₁) (snd x)) +-- -- -} --- -- test1 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S₊ 1 × Int) --- -- Iso.fun test1 f = (trRec isGroupoidS1 (λ a → a) (f north)) --- -- , winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (f (loop* i))))) --- -- Iso.inv test1 (north , b) x = ∣ x ∣ --- -- Iso.inv test1 (south , b) x = ∣ x ∣ --- -- Iso.inv test1 (merid a i , b) x = {!!} --- -- Iso.rightInv test1 = {!!} --- -- Iso.leftInv test1 = {!!} --- -- funRec : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) {C : (A → hLevelTrunc n B) → Type ℓ''} --- -- → isOfHLevel n B --- -- → ((f : A → B) → C (λ a → ∣ f a ∣)) --- -- → (f : A → hLevelTrunc n B) → C f --- -- funRec {A = A} {B = B} n {C = C} hLev ind f = subst C (helper f) (ind (λ a → trRec hLev (λ x → x) (f a))) --- -- where --- -- helper : retract {A = A → hLevelTrunc n B} {B = A → B} (λ f₁ a → trRec hLev (λ x → x) (f₁ a)) (λ f₁ a → ∣ f₁ a ∣) --- -- helper f = funExt λ a → helper2 (f a) --- -- where --- -- helper2 : (x : hLevelTrunc n B) → ∣ trRec hLev (λ x → x) x ∣ ≡ x --- -- helper2 = trElim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → refl - --- -- invMapSurj : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (ϕ : morph G H) → ((x : Group.type H) → isInIm G H (fst ϕ) x) --- -- → morph H G --- -- fst (invMapSurj G H (ϕ , pf) surj) a = {!pRec!} --- -- snd (invMapSurj G H (ϕ , pf) surj) = {!!} - --- {- --- ImIso : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (ϕ : morph G H) → ((x : Group.type H) → isInIm G H (fst ϕ) x) --- → grIso H (imGroup G H ϕ) --- ImIso G H (ϕ , mf) surj = --- let idH = isGroup.id (Group.groupStruc H) --- idG = isGroup.id (Group.groupStruc G) --- _+G_ = isGroup.comp (Group.groupStruc G) --- _+H_ = isGroup.comp (Group.groupStruc H) --- _+Im_ = isGroup.comp (Group.groupStruc (imGroup G H (ϕ , mf))) --- invG = isGroup.inv (Group.groupStruc G) --- invH = isGroup.inv (Group.groupStruc H) --- lUnit = isGroup.lUnit (Group.groupStruc H) --- lCancel = isGroup.rCancel (Group.groupStruc H) --- in --- Iso''→Iso _ _ --- (iso'' ((λ x → x , pRec propTruncIsProp (λ (a , b) → ∣ a , b ∣₋₁) (surj x)) --- , λ a b → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) --- (λ surja → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) --- (λ surjb → --- pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) --- (λ surja+b → --- (λ i → (a +H b) , (pRec (propTruncIsProp) --- (λ (a , b) → ∣ a , b ∣₋₁) --- (propTruncIsProp (surj (isGroup.comp (Group.groupStruc H) a b)) ∣ surja+b ∣₋₁ i))) ∙ --- (λ i → (a +H b) , ∣ (fst surja+b) , (snd surja+b) ∣₋₁) ∙ --- ΣProp≡ (λ _ → propTruncIsProp) refl ∙ --- λ i → (a +H b) , pRec (propTruncIsProp) --- (λ p1 → --- pRec (λ x y → squash x y) --- (λ p2 → --- ∣ --- isGroup.comp (Group.groupStruc G) (fst p1) (fst p2) , --- mf (fst p1) (fst p2) ∙ --- cong₂ (isGroup.comp (Group.groupStruc H)) (snd p1) (snd p2) --- ∣₋₁) --- (pRec (propTruncIsProp) --- ∣_∣₋₁ (propTruncIsProp ∣ surjb ∣₋₁ (surj b) i))) --- (pRec (propTruncIsProp) --- ∣_∣₋₁ (propTruncIsProp ∣ surja ∣₋₁ (surj a) i ))) --- (surj (isGroup.comp (Group.groupStruc H) a b))) --- (surj b)) --- (surj a)) --- (λ x inker → cong fst inker) --- λ x → pRec propTruncIsProp (λ inimx → ∣ (ϕ (inimx .fst)) , ΣProp≡ (λ _ → propTruncIsProp) (inimx .snd) ∣₋₁) (snd x)) --- -} - - --- {- --- H¹-S¹≃Int : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- H¹-S¹≃Int = --- Iso''→Iso _ _ --- (iso'' ((λ x → ∣ theFuns x ∣₀) , λ a b → cong ∣_∣₀ (funExt λ x → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 1) _) ∙ sym (cong (ΩKn+1→Kn) (theFunsId2 a b x)))) --- (λ x inker → pRec (isSetInt _ _) (inj x) (Iso.fun PathIdTrunc₀Iso inker)) --- finIm) --- where --- d : _ --- d = MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 - --- i : _ --- i = MV.i _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1 - --- Δ : _ --- Δ = MV.Δ _ _ (S₊ 1) (λ _ → tt) (λ _ → tt) 0 - - --- d-surj : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x --- d-surj = sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) --- λ x → MV.Ker-i⊂Im-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ --- (sym (isContrHelper .snd _)) --- where --- isContrHelper : isContr (Group.type (×coHomGr 1 Unit Unit)) --- isContrHelper = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) --- , λ y → prodId (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₁ y)) --- (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₂ y)) - --- theFuns : (a : Int) → Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 --- theFuns a (inl x) = 0ₖ --- theFuns a (inr x) = 0ₖ --- theFuns a (push north i) = Kn→ΩKn+1 0 a i --- theFuns a (push south i) = 0ₖ - - --- theFunsId2 : (a b : Int) (x : Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) --- → Kn→ΩKn+1 1 (theFuns a x) ∙ Kn→ΩKn+1 1 (theFuns b x) ≡ Kn→ΩKn+1 1 (theFuns (a +ℤ b) x) --- theFunsId2 a b (inl x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) --- theFunsId2 a b (inr x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) --- theFunsId2 a b (push north i) j = --- hcomp (λ k → λ {(i = i0) → ((λ i₁ → --- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) --- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) --- j --- ; (i = i1) → ((λ i₁ → --- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) --- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) --- j --- ; (j = i0) → cong₂Funct2 (λ p q → Kn→ΩKn+1 1 p ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b) (~ k) i --- ; (j = i1) → (helper2 a b) k i }) --- (hcomp (λ k → λ { (j = i0) → compPath-filler (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) k i --- ; (j = i1) → compPath-filler (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) k i --- ; (i = i1) → RHS-filler b j k --- ; (i = i0) → ((λ i₁ → --- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) --- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) --- j}) --- (bottom-filler a j i)) - --- where - --- bottom-filler : (a : Int) → --- PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) --- (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) --- ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) --- j ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) --- (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) --- ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) --- j) (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) --- bottom-filler a j i = --- hcomp (λ k → λ {(j = i0) → helper2 (~ k) i ; --- (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 a) i}) --- ((λ j₂ → ∣ rCancel (merid north) j j₂ ∣) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) - --- where --- helper2 : Path (Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣ ≡ Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣) --- (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i) ∙ Kn→ΩKn+1 1 ∣ north ∣) --- (λ i → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) --- helper2 = cong₂Sym1 (Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) (~ i) j ∣) (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) - --- RHS-filler : (b : Int) → --- PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j --- ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j) --- (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) --- (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) --- RHS-filler b j i = --- hcomp (λ k → λ {(j = i0) → cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i ; --- (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 b) i}) --- (cong (λ q → (λ i → ∣ rCancel (merid north) j i ∣) ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i) - --- helper2 : (a b : Int) → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b) ≡ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ℤ b)) --- helper2 a b = --- sym (congFunct (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b)) --- ∙ (λ i → cong (Kn→ΩKn+1 1) (Iso.rightInv (Iso3-Kn-ΩKn+1 0) (Kn→ΩKn+1 0 a ∙ Kn→ΩKn+1 0 b) (~ i))) --- ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ₖ b)) ) --- ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (addLemma a b i))) - --- theFunsId2 a b (push south i) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) --- ∙ sym (lUnit _) - --- inj : (a : Int) → theFuns a ≡ (λ _ → ∣ north ∣) → a ≡ pos 0 --- inj a id = --- pRec (isSetInt _ _) --- (sigmaElim (λ _ → isOfHLevelPath 2 isSetInt _ _) --- (λ a p → pRec (isSetInt _ _) --- (λ id2 → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 0) _) --- ∙ cong (ΩKn+1→Kn) --- (PathP→compPathR --- (cong (λ f → cong f (push north)) id) --- ∙ test)) --- (Iso.fun PathIdTrunc₀Iso p))) (d-surj ∣ theFuns a ∣₀) --- where - --- test : (λ i → id i (inl tt)) ∙ (λ i → ∣ north ∣) ∙ sym (λ i → id i (inr tt)) ≡ refl --- test = (λ i → cong (λ f → f (inl tt)) id --- ∙ lUnit (sym (cong (λ f → f (inr tt)) id)) (~ i)) --- ∙ (λ i → cong (λ f → f (push south i)) id --- ∙ sym (cong (λ f → f (inr tt)) id)) --- ∙ rCancel (cong (λ f → f (inr tt)) id) - - --- consMember : (a : Int) → coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) --- consMember a = d ∣ (λ _ → a) ∣₀ - --- consMember≡0 : (a : Int) → consMember a ≡ 0ₕ --- consMember≡0 a = --- (MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ (λ _ → a) ∣₀ ∣ --- (∣ (λ _ → a) ∣₀ , ∣ (λ _ → 0) ∣₀) --- , cong ∣_∣₀ (λ i x → (rUnitₖ a i)) ∣₋₁) - --- f+consMember' : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int × Int ] (f +ₕ (-ₕ (consMember (proj₁ x))) ≡ ∣ theFuns (proj₂ x) ∣₀) --- f+consMember' = --- sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) --- λ f → pRec propTruncIsProp --- (sigmaElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) --- (λ g id → ∣ ((g south) , ((g north) +ₖ (-ₖ g south))) --- , (pRec (setTruncIsSet _ _) --- (λ id → (λ i → ∣ id (~ i) ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀) ∙ funId1 g) --- (Iso.fun PathIdTrunc₀Iso id)) ∣₋₁)) --- (d-surj ∣ f ∣₀) --- where --- funId1 : (g : S₊ 0 → Int) --- → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀ ≡ --- ∣ theFuns ((g north) +ₖ (-ₖ (g south))) ∣₀ --- funId1 g = (λ i → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ --- +ₕ (morphMinus (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) d --- (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ∣ (λ _ → g south) ∣₀ (~ i))) --- ∙ sym (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ g ∣₀ (-ₕ ∣ (λ _ → g south) ∣₀)) --- ∙ (cong (λ x → d ∣ x ∣₀) g'Id) --- ∙ cong ∣_∣₀ helper --- where --- g' : S₊ 0 → Int --- g' north = (g north) +ₖ (-ₖ (g south)) --- g' south = 0 - --- g'Id : (λ x → g x +ₖ (-ₖ (g south))) ≡ g' --- g'Id = funExt (λ {north → refl --- ; south → rCancelₖ (g south)}) - --- helper : MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g' ≡ theFuns (g north +ₖ (-ₖ g south)) --- helper = funExt λ {(inl tt) → refl --- ; (inr tt) → refl --- ; (push north i) → refl --- ; (push south i) → refl} --- finIm : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int ] (∣ theFuns x ∣₀ ≡ f) --- finIm f = --- pRec propTruncIsProp --- (λ {((a , b) , id) → ∣ b , (sym id ∙ cong (λ x → f +ₕ x) ((λ i → (-ₕ (consMember≡0 a i))) ∙ sym (lUnitₕ (-ₕ 0ₕ)) ∙ rCancelₕ 0ₕ) ∙ (rUnitₕ f)) ∣₋₁}) --- (f+consMember' f) --- -} - - - --- -- Hⁿ-Sⁿ≃Int : (n : ℕ) → grIso intGroup (coHomGr (suc n) (S₊ (suc n))) --- -- Hⁿ-Sⁿ≃Int zero = --- -- compGrIso {F = intGroup} {G = {!!}} {H = {!coHomGr 1 (S₊ 1)!}} --- -- (Iso''→Iso --- -- (iso'' ((λ x → ∣ theFuns x ∣₀) , λ a b → cong ∣_∣₀ (funExt λ x → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 1) _) ∙ sym (cong (ΩKn+1→Kn) (theFunsId2 a b x)))) --- -- (λ x inker → pRec (isSetInt _ _) (inj x) (Iso.fun PathIdTrunc₀Iso inker)) --- -- finIm)) --- -- {!invGrIso _ _ (coHomPushout≡coHomSn 0 1)!} +-- -- {- +-- -- H¹-S¹≃Int : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- H¹-S¹≃Int = +-- -- Iso''→Iso _ _ +-- -- (iso'' ((λ x → ∣ theFuns x ∣₀) , λ a b → cong ∣_∣₀ (funExt λ x → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 1) _) ∙ sym (cong (ΩKn+1→Kn) (theFunsId2 a b x)))) +-- -- (λ x inker → pRec (isSetInt _ _) (inj x) (Iso.fun PathIdTrunc₀Iso inker)) +-- -- finIm) -- -- where -- -- d : _ -- -- d = MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 @@ -980,391 +1028,566 @@ coHom1S1≃ℤ = -- -- pRec propTruncIsProp -- -- (λ {((a , b) , id) → ∣ b , (sym id ∙ cong (λ x → f +ₕ x) ((λ i → (-ₕ (consMember≡0 a i))) ∙ sym (lUnitₕ (-ₕ 0ₕ)) ∙ rCancelₕ 0ₕ) ∙ (rUnitₕ f)) ∣₋₁}) -- -- (f+consMember' f) --- -- Hⁿ-Sⁿ≃Int (suc n) = --- -- compGrIso (Hⁿ-Sⁿ≃Int n) --- -- (transport (λ i → grIso {!!} {!coHomGr (suc (suc n)) (Pushout (λ _ → tt) (λ _ → tt))!}) {!!}) +-- -- -} --- -- {- --- -- coHom1S1≃ℤ : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- grIso.fun coHom1S1≃ℤ = grIso.fun {!compGrIso coHom1Iso (invGrIso _ _ (ImIso _ _ (d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ?))!} --- -- grIso.inv coHom1S1≃ℤ = {!!} --- -- grIso.rightInv coHom1S1≃ℤ = {!!} --- -- grIso.leftInv coHom1S1≃ℤ = {!!} --- -- -} --- -- -- compGrIso coHom1Iso (invGrIso _ _ (ImIso _ _ {!d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0!} {!!})) - - --- -- -- coHomGrIm : grIso (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- -- (imGroup (coHomGr 0 (S₊ 0)) --- -- -- (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- -- (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 --- -- -- , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) --- -- -- coHomGrIm = --- -- -- ImIso _ --- -- -- _ --- -- -- _ --- -- -- {!!} - - --- -- -- -- coHom1RedS1 : Iso (coHom 1 (S₊ 1)) (coHomRed 1 (S₊ 1 , north)) --- -- -- -- Iso.fun coHom1RedS1 = sRec setTruncIsSet λ f → ∣ f , (pRec {!!} {!!} ((transport (λ i → (b : truncIdempotent 3 {!S₊ 1!} (~ i)) → ∥ (transp (λ j → truncIdempotent {!3!} {!!} (~ i ∨ (~ j))) (~ i) north) ≡ b ∥₋₁) isConnectedS1) (f north)) ) ∣₀ --- -- -- -- Iso.inv coHom1RedS1 = {!!} --- -- -- -- Iso.rightInv coHom1RedS1 = {!setTruncIdempotent!} --- -- -- -- Iso.leftInv coHom1RedS1 = {!!} - --- -- -- -- coHom1Red : ∀ {ℓ} (A : Pointed ℓ) → Iso (coHom 1 (typ A)) (coHomRed 1 A) --- -- -- -- Iso.fun (coHom1Red A) = sRec setTruncIsSet λ f → ∣ f , {!!} ∣₀ --- -- -- -- Iso.inv (coHom1Red A) = {!!} --- -- -- -- Iso.rightInv (coHom1Red A) = {!!} --- -- -- -- Iso.leftInv (coHom1Red A) = {!!} - --- -- -- -- -- morphtest : morph (coHomGr 1 (S₊ 1)) intGroup --- -- -- -- -- fst morphtest = sRec isSetInt λ f → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (f (loop* i))))) --- -- -- -- -- snd morphtest = sElim2 {!!} --- -- -- -- -- (funRec 3 isGroupoidS1 --- -- -- -- -- λ f → funRec 3 isGroupoidS1 --- -- -- -- -- λ g → pRec (isSetInt _ _) --- -- -- -- -- (λ n=fn → --- -- -- -- -- pRec (isSetInt _ _) --- -- -- -- -- (λ n=gn → (λ j → winding (basechange (SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (∣ f (north) ∣ +ₖ ∣ n=gn (~ j) ∣)))) (λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (∣ f (loop* i) ∣ +ₖ ∣ transp (λ i → n=gn ((~ i) ∨ (~ j)) ≡ n=gn ((~ i) ∨ (~ j))) (~ j) (λ i → g (loop* i)) i ∣)))))) --- -- -- -- -- ∙ {!!} --- -- -- -- -- ∙ {!!}) --- -- -- -- -- (isConnectedS1 (g north))) --- -- -- -- -- (isConnectedS1 (f north))) - - --- -- -- -- -- {- (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (∣ f (loop* i) ∣ +ₖ ∣ g (loop* i) ∣))))) --- -- -- -- -- ∙ (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (Kn→ΩKn+1 1 ∣ f (loop* i) ∣ ∙ Kn→ΩKn+1 1 ∣ g (loop* i) ∣)))))) --- -- -- -- -- ∙ (λ j → winding (basechange _ (cong (λ x → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (Kn→ΩKn+1 1 ∣ f x ∣ ∙ Kn→ΩKn+1 1 ∣ g x ∣))))) loop*)) ) --- -- -- -- -- ∙ (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn ((cong ∣_∣ (merid (f (loop* i)) ∙ sym (merid north)) ∙ cong ∣_∣ (merid (g (loop* i)) ∙ sym (merid north))))))))) --- -- -- -- -- ∙ (λ j → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (congFunct ∣_∣ (merid (f (loop* i)) ∙ sym (merid north)) (merid (g (loop* i)) ∙ sym (merid north)) (~ j))))))) --- -- -- -- -- ∙ (λ j → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (cong ∣_∣ (({!!} ∙ {!!}) ∙ {!!}))))))) --- -- -- -- -- ∙ {!!} --- -- -- -- -- ∙ {!!} --- -- -- -- -- ∙ {!!}) -} - --- -- -- -- -- where --- -- -- -- -- helper : ∀ {ℓ} {A : Type ℓ} (a : A) (f : A → S¹) (p q : a ≡ a) → winding (basechange (f a) (cong f (p ∙ q))) ≡ winding (basechange (f a) (cong f p ∙ cong f q)) --- -- -- -- -- helper a f p q i = winding (basechange (f a) (congFunct f p q i)) --- -- -- -- -- helper2 : (x : S¹) (p q : x ≡ x) → basechange x (p ∙ q) ≡ basechange x p ∙ basechange x q --- -- -- -- -- helper2 base p q = refl --- -- -- -- -- helper2 (loop i) p q = {!!} --- -- -- -- -- helper4 : (x y z : S¹) (p : x ≡ z) (q r : x ≡ y) (s : y ≡ z) → basechange* x z p (q ∙ s) ≡ basechange* x y {!!} q ∙ {!!} --- -- -- -- -- helper4 = {!!} --- -- -- -- -- helper3 : (p q : ΩS¹) → winding (p ∙ q) ≡ (winding p +ℤ winding q) --- -- -- -- -- helper3 = {!!} - - --- -- -- -- -- -- fstmap : morph (dirProd intGroup intGroup) (coHomGr 0 (S₊ 0)) --- -- -- -- -- -- fstmap = compMorph {F = dirProd intGroup intGroup} {G = ×coHomGr 0 Unit Unit} {H = coHomGr 0 (S₊ 0)} --- -- -- -- -- -- (×morph (grIso.inv coHomGrUnit0) (grIso.inv coHomGrUnit0)) --- -- -- -- -- -- (((MV.Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) zero)) , --- -- -- -- -- -- {!MV.ΔIsHom _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) zero!}) - --- -- -- -- -- -- fstMapId : (a : Int × Int) → fstmap .fst a ≡ ∣ (λ _ → proj₁ a +ℤ (0 - proj₂ a)) ∣₀ --- -- -- -- -- -- fstMapId (a , b) = (λ i → ∣ (λ _ → a +ₖ (-ₖ b)) ∣₀) ∙ {!addLemma!} ∙ {!!} ∙ {!!} - --- -- -- -- -- -- isoAgain : grIso intGroup (coHomGr 1 (S₊ 1)) --- -- -- -- -- -- isoAgain = --- -- -- -- -- -- Iso''→Iso _ _ --- -- -- -- -- -- (iso'' ((λ a → ∣ (λ {north → 0ₖ ; south → 0ₖ ; (merid north i) → {!a!} ; (merid south i) → {!!}}) ∣₀) , {!!}) --- -- -- -- -- -- {!!} --- -- -- -- -- -- {!!}) - --- -- -- -- -- -- -- -- test2 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → S¹) --- -- -- -- -- -- -- -- Iso.fun test2 f = {!!} --- -- -- -- -- -- -- -- Iso.inv test2 f north = ∣ transport (sym S¹≡S1) (f base) ∣ --- -- -- -- -- -- -- -- Iso.inv test2 f south = ∣ transport (sym S¹≡S1) (f base) ∣ --- -- -- -- -- -- -- -- Iso.inv test2 f (merid a i) = cong ∣_∣ {!transport (sym S¹≡S1) (f base)!} i --- -- -- -- -- -- -- -- Iso.rightInv test2 = {!!} --- -- -- -- -- -- -- -- Iso.leftInv test2 = {!!} - --- -- -- -- -- -- -- -- F : winding (basechange base loop) ≡ 1 --- -- -- -- -- -- -- -- F = refl - --- -- -- -- -- -- -- -- another : (f g : Int) → winding (basechange {!!} {!!}) ≡ {!!} --- -- -- -- -- -- -- -- another = {!!} - --- -- -- -- -- -- -- -- test : Iso (S¹ → S¹) (S¹ × Int) --- -- -- -- -- -- -- -- Iso.fun test f = f base , winding (basechange (f base) (cong f loop)) --- -- -- -- -- -- -- -- Iso.inv test (x , int) base = x --- -- -- -- -- -- -- -- Iso.inv test (x , int) (loop i) = {!!} --- -- -- -- -- -- -- -- Iso.rightInv test = {!!} --- -- -- -- -- -- -- -- Iso.leftInv test = {!!} - --- -- -- -- -- -- -- -- -- test13 : Iso ∥ (S¹ → S¹) ∥₀ Int --- -- -- -- -- -- -- -- -- Iso.fun test13 = sRec isSetInt λ f → winding (basechange (f base) (λ i → f (loop i))) --- -- -- -- -- -- -- -- -- Iso.inv test13 a = ∣ (λ {base → {!!} ; (loop i) → {!!}}) ∣₀ --- -- -- -- -- -- -- -- -- Iso.rightInv test13 = {!!} --- -- -- -- -- -- -- -- -- Iso.leftInv test13 = {!!} +-- -- -- Hⁿ-Sⁿ≃Int : (n : ℕ) → grIso intGroup (coHomGr (suc n) (S₊ (suc n))) +-- -- -- Hⁿ-Sⁿ≃Int zero = +-- -- -- compGrIso {F = intGroup} {G = {!!}} {H = {!coHomGr 1 (S₊ 1)!}} +-- -- -- (Iso''→Iso +-- -- -- (iso'' ((λ x → ∣ theFuns x ∣₀) , λ a b → cong ∣_∣₀ (funExt λ x → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 1) _) ∙ sym (cong (ΩKn+1→Kn) (theFunsId2 a b x)))) +-- -- -- (λ x inker → pRec (isSetInt _ _) (inj x) (Iso.fun PathIdTrunc₀Iso inker)) +-- -- -- finIm)) +-- -- -- {!invGrIso _ _ (coHomPushout≡coHomSn 0 1)!} +-- -- -- where +-- -- -- d : _ +-- -- -- d = MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 + +-- -- -- i : _ +-- -- -- i = MV.i _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1 + +-- -- -- Δ : _ +-- -- -- Δ = MV.Δ _ _ (S₊ 1) (λ _ → tt) (λ _ → tt) 0 + + +-- -- -- d-surj : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x +-- -- -- d-surj = sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- -- -- λ x → MV.Ker-i⊂Im-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ +-- -- -- (sym (isContrHelper .snd _)) +-- -- -- where +-- -- -- isContrHelper : isContr (Group.type (×coHomGr 1 Unit Unit)) +-- -- -- isContrHelper = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) +-- -- -- , λ y → prodId (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₁ y)) +-- -- -- (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₂ y)) + +-- -- -- theFuns : (a : Int) → Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 +-- -- -- theFuns a (inl x) = 0ₖ +-- -- -- theFuns a (inr x) = 0ₖ +-- -- -- theFuns a (push north i) = Kn→ΩKn+1 0 a i +-- -- -- theFuns a (push south i) = 0ₖ + + +-- -- -- theFunsId2 : (a b : Int) (x : Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) +-- -- -- → Kn→ΩKn+1 1 (theFuns a x) ∙ Kn→ΩKn+1 1 (theFuns b x) ≡ Kn→ΩKn+1 1 (theFuns (a +ℤ b) x) +-- -- -- theFunsId2 a b (inl x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) +-- -- -- theFunsId2 a b (inr x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) +-- -- -- theFunsId2 a b (push north i) j = +-- -- -- hcomp (λ k → λ {(i = i0) → ((λ i₁ → +-- -- -- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- -- -- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) +-- -- -- j +-- -- -- ; (i = i1) → ((λ i₁ → +-- -- -- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- -- -- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) +-- -- -- j +-- -- -- ; (j = i0) → cong₂Funct2 (λ p q → Kn→ΩKn+1 1 p ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b) (~ k) i +-- -- -- ; (j = i1) → (helper2 a b) k i }) +-- -- -- (hcomp (λ k → λ { (j = i0) → compPath-filler (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) k i +-- -- -- ; (j = i1) → compPath-filler (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) k i +-- -- -- ; (i = i1) → RHS-filler b j k +-- -- -- ; (i = i0) → ((λ i₁ → +-- -- -- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- -- -- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) +-- -- -- j}) +-- -- -- (bottom-filler a j i)) + +-- -- -- where + +-- -- -- bottom-filler : (a : Int) → +-- -- -- PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- -- -- (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) +-- -- -- ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) +-- -- -- j ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- -- -- (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) +-- -- -- ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) +-- -- -- j) (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) +-- -- -- bottom-filler a j i = +-- -- -- hcomp (λ k → λ {(j = i0) → helper2 (~ k) i ; +-- -- -- (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 a) i}) +-- -- -- ((λ j₂ → ∣ rCancel (merid north) j j₂ ∣) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + +-- -- -- where +-- -- -- helper2 : Path (Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣ ≡ Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- -- -- (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- -- -- (λ i → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) +-- -- -- helper2 = cong₂Sym1 (Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) (~ i) j ∣) (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + +-- -- -- RHS-filler : (b : Int) → +-- -- -- PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j +-- -- -- ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j) +-- -- -- (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) +-- -- -- (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) +-- -- -- RHS-filler b j i = +-- -- -- hcomp (λ k → λ {(j = i0) → cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i ; +-- -- -- (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 b) i}) +-- -- -- (cong (λ q → (λ i → ∣ rCancel (merid north) j i ∣) ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i) + +-- -- -- helper2 : (a b : Int) → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b) ≡ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ℤ b)) +-- -- -- helper2 a b = +-- -- -- sym (congFunct (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b)) +-- -- -- ∙ (λ i → cong (Kn→ΩKn+1 1) (Iso.rightInv (Iso3-Kn-ΩKn+1 0) (Kn→ΩKn+1 0 a ∙ Kn→ΩKn+1 0 b) (~ i))) +-- -- -- ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ₖ b)) ) +-- -- -- ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (addLemma a b i))) + +-- -- -- theFunsId2 a b (push south i) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- -- -- ∙ sym (lUnit _) + +-- -- -- inj : (a : Int) → theFuns a ≡ (λ _ → ∣ north ∣) → a ≡ pos 0 +-- -- -- inj a id = +-- -- -- pRec (isSetInt _ _) +-- -- -- (sigmaElim (λ _ → isOfHLevelPath 2 isSetInt _ _) +-- -- -- (λ a p → pRec (isSetInt _ _) +-- -- -- (λ id2 → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 0) _) +-- -- -- ∙ cong (ΩKn+1→Kn) +-- -- -- (PathP→compPathR +-- -- -- (cong (λ f → cong f (push north)) id) +-- -- -- ∙ test)) +-- -- -- (Iso.fun PathIdTrunc₀Iso p))) (d-surj ∣ theFuns a ∣₀) +-- -- -- where + +-- -- -- test : (λ i → id i (inl tt)) ∙ (λ i → ∣ north ∣) ∙ sym (λ i → id i (inr tt)) ≡ refl +-- -- -- test = (λ i → cong (λ f → f (inl tt)) id +-- -- -- ∙ lUnit (sym (cong (λ f → f (inr tt)) id)) (~ i)) +-- -- -- ∙ (λ i → cong (λ f → f (push south i)) id +-- -- -- ∙ sym (cong (λ f → f (inr tt)) id)) +-- -- -- ∙ rCancel (cong (λ f → f (inr tt)) id) + + +-- -- -- consMember : (a : Int) → coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) +-- -- -- consMember a = d ∣ (λ _ → a) ∣₀ + +-- -- -- consMember≡0 : (a : Int) → consMember a ≡ 0ₕ +-- -- -- consMember≡0 a = +-- -- -- (MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ (λ _ → a) ∣₀ ∣ +-- -- -- (∣ (λ _ → a) ∣₀ , ∣ (λ _ → 0) ∣₀) +-- -- -- , cong ∣_∣₀ (λ i x → (rUnitₖ a i)) ∣₋₁) + +-- -- -- f+consMember' : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int × Int ] (f +ₕ (-ₕ (consMember (proj₁ x))) ≡ ∣ theFuns (proj₂ x) ∣₀) +-- -- -- f+consMember' = +-- -- -- sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- -- -- λ f → pRec propTruncIsProp +-- -- -- (sigmaElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- -- -- (λ g id → ∣ ((g south) , ((g north) +ₖ (-ₖ g south))) +-- -- -- , (pRec (setTruncIsSet _ _) +-- -- -- (λ id → (λ i → ∣ id (~ i) ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀) ∙ funId1 g) +-- -- -- (Iso.fun PathIdTrunc₀Iso id)) ∣₋₁)) +-- -- -- (d-surj ∣ f ∣₀) +-- -- -- where +-- -- -- funId1 : (g : S₊ 0 → Int) +-- -- -- → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀ ≡ +-- -- -- ∣ theFuns ((g north) +ₖ (-ₖ (g south))) ∣₀ +-- -- -- funId1 g = (λ i → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +-- -- -- +ₕ (morphMinus (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) d +-- -- -- (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ∣ (λ _ → g south) ∣₀ (~ i))) +-- -- -- ∙ sym (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ g ∣₀ (-ₕ ∣ (λ _ → g south) ∣₀)) +-- -- -- ∙ (cong (λ x → d ∣ x ∣₀) g'Id) +-- -- -- ∙ cong ∣_∣₀ helper +-- -- -- where +-- -- -- g' : S₊ 0 → Int +-- -- -- g' north = (g north) +ₖ (-ₖ (g south)) +-- -- -- g' south = 0 + +-- -- -- g'Id : (λ x → g x +ₖ (-ₖ (g south))) ≡ g' +-- -- -- g'Id = funExt (λ {north → refl +-- -- -- ; south → rCancelₖ (g south)}) + +-- -- -- helper : MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g' ≡ theFuns (g north +ₖ (-ₖ g south)) +-- -- -- helper = funExt λ {(inl tt) → refl +-- -- -- ; (inr tt) → refl +-- -- -- ; (push north i) → refl +-- -- -- ; (push south i) → refl} +-- -- -- finIm : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int ] (∣ theFuns x ∣₀ ≡ f) +-- -- -- finIm f = +-- -- -- pRec propTruncIsProp +-- -- -- (λ {((a , b) , id) → ∣ b , (sym id ∙ cong (λ x → f +ₕ x) ((λ i → (-ₕ (consMember≡0 a i))) ∙ sym (lUnitₕ (-ₕ 0ₕ)) ∙ rCancelₕ 0ₕ) ∙ (rUnitₕ f)) ∣₋₁}) +-- -- -- (f+consMember' f) +-- -- -- Hⁿ-Sⁿ≃Int (suc n) = +-- -- -- compGrIso (Hⁿ-Sⁿ≃Int n) +-- -- -- (transport (λ i → grIso {!!} {!coHomGr (suc (suc n)) (Pushout (λ _ → tt) (λ _ → tt))!}) {!!}) + + +-- -- -- {- +-- -- -- coHom1S1≃ℤ : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- grIso.fun coHom1S1≃ℤ = grIso.fun {!compGrIso coHom1Iso (invGrIso _ _ (ImIso _ _ (d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ?))!} +-- -- -- grIso.inv coHom1S1≃ℤ = {!!} +-- -- -- grIso.rightInv coHom1S1≃ℤ = {!!} +-- -- -- grIso.leftInv coHom1S1≃ℤ = {!!} +-- -- -- -} + +-- -- -- -- compGrIso coHom1Iso (invGrIso _ _ (ImIso _ _ {!d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0!} {!!})) + + +-- -- -- -- coHomGrIm : grIso (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- (imGroup (coHomGr 0 (S₊ 0)) +-- -- -- -- (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 +-- -- -- -- , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) +-- -- -- -- coHomGrIm = +-- -- -- -- ImIso _ +-- -- -- -- _ +-- -- -- -- _ +-- -- -- -- {!!} + + +-- -- -- -- -- coHom1RedS1 : Iso (coHom 1 (S₊ 1)) (coHomRed 1 (S₊ 1 , north)) +-- -- -- -- -- Iso.fun coHom1RedS1 = sRec setTruncIsSet λ f → ∣ f , (pRec {!!} {!!} ((transport (λ i → (b : truncIdempotent 3 {!S₊ 1!} (~ i)) → ∥ (transp (λ j → truncIdempotent {!3!} {!!} (~ i ∨ (~ j))) (~ i) north) ≡ b ∥₋₁) isConnectedS1) (f north)) ) ∣₀ +-- -- -- -- -- Iso.inv coHom1RedS1 = {!!} +-- -- -- -- -- Iso.rightInv coHom1RedS1 = {!setTruncIdempotent!} +-- -- -- -- -- Iso.leftInv coHom1RedS1 = {!!} + +-- -- -- -- -- coHom1Red : ∀ {ℓ} (A : Pointed ℓ) → Iso (coHom 1 (typ A)) (coHomRed 1 A) +-- -- -- -- -- Iso.fun (coHom1Red A) = sRec setTruncIsSet λ f → ∣ f , {!!} ∣₀ +-- -- -- -- -- Iso.inv (coHom1Red A) = {!!} +-- -- -- -- -- Iso.rightInv (coHom1Red A) = {!!} +-- -- -- -- -- Iso.leftInv (coHom1Red A) = {!!} + +-- -- -- -- -- -- morphtest : morph (coHomGr 1 (S₊ 1)) intGroup +-- -- -- -- -- -- fst morphtest = sRec isSetInt λ f → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (f (loop* i))))) +-- -- -- -- -- -- snd morphtest = sElim2 {!!} +-- -- -- -- -- -- (funRec 3 isGroupoidS1 +-- -- -- -- -- -- λ f → funRec 3 isGroupoidS1 +-- -- -- -- -- -- λ g → pRec (isSetInt _ _) +-- -- -- -- -- -- (λ n=fn → +-- -- -- -- -- -- pRec (isSetInt _ _) +-- -- -- -- -- -- (λ n=gn → (λ j → winding (basechange (SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (∣ f (north) ∣ +ₖ ∣ n=gn (~ j) ∣)))) (λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (∣ f (loop* i) ∣ +ₖ ∣ transp (λ i → n=gn ((~ i) ∨ (~ j)) ≡ n=gn ((~ i) ∨ (~ j))) (~ j) (λ i → g (loop* i)) i ∣)))))) +-- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- ∙ {!!}) +-- -- -- -- -- -- (isConnectedS1 (g north))) +-- -- -- -- -- -- (isConnectedS1 (f north))) + + +-- -- -- -- -- -- {- (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (∣ f (loop* i) ∣ +ₖ ∣ g (loop* i) ∣))))) +-- -- -- -- -- -- ∙ (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (Kn→ΩKn+1 1 ∣ f (loop* i) ∣ ∙ Kn→ΩKn+1 1 ∣ g (loop* i) ∣)))))) +-- -- -- -- -- -- ∙ (λ j → winding (basechange _ (cong (λ x → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (Kn→ΩKn+1 1 ∣ f x ∣ ∙ Kn→ΩKn+1 1 ∣ g x ∣))))) loop*)) ) +-- -- -- -- -- -- ∙ (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn ((cong ∣_∣ (merid (f (loop* i)) ∙ sym (merid north)) ∙ cong ∣_∣ (merid (g (loop* i)) ∙ sym (merid north))))))))) +-- -- -- -- -- -- ∙ (λ j → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (congFunct ∣_∣ (merid (f (loop* i)) ∙ sym (merid north)) (merid (g (loop* i)) ∙ sym (merid north)) (~ j))))))) +-- -- -- -- -- -- ∙ (λ j → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (cong ∣_∣ (({!!} ∙ {!!}) ∙ {!!}))))))) +-- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- ∙ {!!}) -} + +-- -- -- -- -- -- where +-- -- -- -- -- -- helper : ∀ {ℓ} {A : Type ℓ} (a : A) (f : A → S¹) (p q : a ≡ a) → winding (basechange (f a) (cong f (p ∙ q))) ≡ winding (basechange (f a) (cong f p ∙ cong f q)) +-- -- -- -- -- -- helper a f p q i = winding (basechange (f a) (congFunct f p q i)) +-- -- -- -- -- -- helper2 : (x : S¹) (p q : x ≡ x) → basechange x (p ∙ q) ≡ basechange x p ∙ basechange x q +-- -- -- -- -- -- helper2 base p q = refl +-- -- -- -- -- -- helper2 (loop i) p q = {!!} +-- -- -- -- -- -- helper4 : (x y z : S¹) (p : x ≡ z) (q r : x ≡ y) (s : y ≡ z) → basechange* x z p (q ∙ s) ≡ basechange* x y {!!} q ∙ {!!} +-- -- -- -- -- -- helper4 = {!!} +-- -- -- -- -- -- helper3 : (p q : ΩS¹) → winding (p ∙ q) ≡ (winding p +ℤ winding q) +-- -- -- -- -- -- helper3 = {!!} + + +-- -- -- -- -- -- -- fstmap : morph (dirProd intGroup intGroup) (coHomGr 0 (S₊ 0)) +-- -- -- -- -- -- -- fstmap = compMorph {F = dirProd intGroup intGroup} {G = ×coHomGr 0 Unit Unit} {H = coHomGr 0 (S₊ 0)} +-- -- -- -- -- -- -- (×morph (grIso.inv coHomGrUnit0) (grIso.inv coHomGrUnit0)) +-- -- -- -- -- -- -- (((MV.Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) zero)) , +-- -- -- -- -- -- -- {!MV.ΔIsHom _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) zero!}) + +-- -- -- -- -- -- -- fstMapId : (a : Int × Int) → fstmap .fst a ≡ ∣ (λ _ → proj₁ a +ℤ (0 - proj₂ a)) ∣₀ +-- -- -- -- -- -- -- fstMapId (a , b) = (λ i → ∣ (λ _ → a +ₖ (-ₖ b)) ∣₀) ∙ {!addLemma!} ∙ {!!} ∙ {!!} + +-- -- -- -- -- -- -- isoAgain : grIso intGroup (coHomGr 1 (S₊ 1)) +-- -- -- -- -- -- -- isoAgain = +-- -- -- -- -- -- -- Iso''→Iso _ _ +-- -- -- -- -- -- -- (iso'' ((λ a → ∣ (λ {north → 0ₖ ; south → 0ₖ ; (merid north i) → {!a!} ; (merid south i) → {!!}}) ∣₀) , {!!}) +-- -- -- -- -- -- -- {!!} +-- -- -- -- -- -- -- {!!}) + +-- -- -- -- -- -- -- -- -- test2 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → S¹) +-- -- -- -- -- -- -- -- -- Iso.fun test2 f = {!!} +-- -- -- -- -- -- -- -- -- Iso.inv test2 f north = ∣ transport (sym S¹≡S1) (f base) ∣ +-- -- -- -- -- -- -- -- -- Iso.inv test2 f south = ∣ transport (sym S¹≡S1) (f base) ∣ +-- -- -- -- -- -- -- -- -- Iso.inv test2 f (merid a i) = cong ∣_∣ {!transport (sym S¹≡S1) (f base)!} i +-- -- -- -- -- -- -- -- -- Iso.rightInv test2 = {!!} +-- -- -- -- -- -- -- -- -- Iso.leftInv test2 = {!!} + +-- -- -- -- -- -- -- -- -- F : winding (basechange base loop) ≡ 1 +-- -- -- -- -- -- -- -- -- F = refl + +-- -- -- -- -- -- -- -- -- another : (f g : Int) → winding (basechange {!!} {!!}) ≡ {!!} +-- -- -- -- -- -- -- -- -- another = {!!} -- -- -- -- -- -- -- -- -- test : Iso (S¹ → S¹) (S¹ × Int) --- -- -- -- -- -- -- -- -- Iso.fun test f = (f base) , transport (basedΩS¹≡Int (f base)) λ i → f (loop i) +-- -- -- -- -- -- -- -- -- Iso.fun test f = f base , winding (basechange (f base) (cong f loop)) -- -- -- -- -- -- -- -- -- Iso.inv test (x , int) base = x --- -- -- -- -- -- -- -- -- Iso.inv test (x , int) (loop i) = transport (sym (basedΩS¹≡Int x)) int i --- -- -- -- -- -- -- -- -- Iso.rightInv test (x , int) i = (x , transportTransport⁻ (basedΩS¹≡Int x) int i) --- -- -- -- -- -- -- -- -- Iso.leftInv test f = --- -- -- -- -- -- -- -- -- funExt λ { base → refl --- -- -- -- -- -- -- -- -- ; (loop i) j → transport⁻Transport (basedΩS¹≡Int (f base)) (λ i → f (loop i)) j i} - - --- -- -- -- -- -- -- -- -- lem : S¹ ≡ hLevelTrunc 3 (S₊ 1) --- -- -- -- -- -- -- -- -- lem = sym (truncIdempotent 3 isGroupoidS¹) ∙ λ i → hLevelTrunc 3 (S¹≡S1 (~ i)) - --- -- -- -- -- -- -- -- -- prodId : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (a b : A × B) → proj₁ a ≡ proj₁ b → proj₂ a ≡ proj₂ b → a ≡ b --- -- -- -- -- -- -- -- -- prodId (_ , _) (_ , _) id1 id2 i = (id1 i) , (id2 i) - --- -- -- -- -- -- -- -- -- test22 : Iso (S₊ 1 → coHomK 1) (S₊ 1 × Int) --- -- -- -- -- -- -- -- -- Iso.fun test22 f = {!f north!} , {!!} --- -- -- -- -- -- -- -- -- Iso.inv test22 = {!!} --- -- -- -- -- -- -- -- -- Iso.rightInv test22 = {!!} --- -- -- -- -- -- -- -- -- Iso.leftInv test22 = {!!} - - - - - - --- -- -- -- -- -- -- -- -- coHom1≃∥S1×ℤ∥₀ : Iso (coHom 1 (S₊ 1)) ∥ S₊ 1 × Int ∥₀ --- -- -- -- -- -- -- -- -- coHom1≃∥S1×ℤ∥₀ = setTruncIso test2 --- -- -- -- -- -- -- -- -- where --- -- -- -- -- -- -- -- -- test2 : Iso (S₊ 1 → coHomK 1) (S₊ 1 × Int) --- -- -- -- -- -- -- -- -- Iso.fun test2 f = transport (λ i → S¹≡S1 (~ i) × Int) (Iso.fun test (transport (λ i → (S¹≡S1 i → lem (~ i))) f)) --- -- -- -- -- -- -- -- -- Iso.inv test2 x = transport (λ i → (S¹≡S1 (~ i) → lem i)) (Iso.inv test (transport (λ i → S¹≡S1 i × Int) x)) --- -- -- -- -- -- -- -- -- Iso.rightInv test2 (s , int) = prodId _ _ {!!} {!!} --- -- -- -- -- -- -- -- -- Iso.leftInv test2 f = {!!} ∙ {!!} ∙ {!!} - --- -- -- -- -- -- -- -- -- test2Id : (a b : (S₊ 1 → coHomK 1)) → proj₂ (Iso.fun test2 (λ x → a x +ₖ b x)) ≡ (proj₂ (Iso.fun test2 a) +ₖ proj₂ (Iso.fun test2 a)) --- -- -- -- -- -- -- -- -- test2Id a b = {! --- -- -- -- -- -- -- -- -- transport --- -- -- -- -- -- -- -- -- (basedΩS¹≡Int --- -- -- -- -- -- -- -- -- (transport (λ i → S¹≡S1 i → lem (~ i)) (λ x → a x +ₖ b x) base)) --- -- -- -- -- -- -- -- -- (λ i → --- -- -- -- -- -- -- -- -- transport (λ i₁ → S¹≡S1 i₁ → lem (~ i₁)) (λ x → a x +ₖ b x) --- -- -- -- -- -- -- -- -- (loop i))!} ∙ {!transport (λ i → S¹≡S1 i → lem (~ i)) (λ x → a x +ₖ b x) base!} - - --- -- -- -- -- -- -- -- -- main : grIso intGroup (coHomGr 1 (S₊ 1)) --- -- -- -- -- -- -- -- -- main = Iso'→Iso --- -- -- -- -- -- -- -- -- (iso' {!!} --- -- -- -- -- -- -- -- -- {!!}) - - --- -- -- -- -- -- -- coHom1 : grIso (coHomGr 1 (S₊ 1)) intGroup --- -- -- -- -- -- -- coHom1 = --- -- -- -- -- -- -- Iso'→Iso --- -- -- -- -- -- -- (iso' (iso ({!!} ∘ {!!} ∘ {!!} ∘ {!!}) --- -- -- -- -- -- -- {!!} --- -- -- -- -- -- -- {!!} --- -- -- -- -- -- -- {!!}) --- -- -- -- -- -- -- {!!}) - - - --- -- -- -- -- -- -- schonf : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : (A × B) → Type ℓ''} --- -- -- -- -- -- -- → ((a : A) (b : B) → C (a , b)) --- -- -- -- -- -- -- → (x : A × B) → C x --- -- -- -- -- -- -- schonf f (a , b) = f a b - --- -- -- -- -- -- -- -- -- setTruncProdIso : ∀ {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') → Iso ∥ (A × B) ∥₀ (∥ A ∥₀ × ∥ B ∥₀) --- -- -- -- -- -- -- -- -- Iso.fun (setTruncProdIso A B) = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) λ {(a , b) → ∣ a ∣₀ , ∣ b ∣₀} --- -- -- -- -- -- -- -- -- Iso.inv (setTruncProdIso A B) (a , b) = sRec setTruncIsSet (λ a → sRec setTruncIsSet (λ b → ∣ a , b ∣₀) b) a --- -- -- -- -- -- -- -- -- Iso.rightInv (setTruncProdIso A B) = --- -- -- -- -- -- -- -- -- schonf (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) --- -- -- -- -- -- -- -- -- λ _ → sElim (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) --- -- -- -- -- -- -- -- -- λ _ → refl) --- -- -- -- -- -- -- -- -- Iso.leftInv (setTruncProdIso A B) = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ {(a , b) → refl} - --- -- -- -- -- -- -- -- -- setTruncProdLemma : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (a1 a2 : A) (b : B) → isHLevelConnected 2 A --- -- -- -- -- -- -- -- -- → Path ∥ A × B ∥₀ ∣ a1 , b ∣₀ ∣ a2 , b ∣₀ --- -- -- -- -- -- -- -- -- setTruncProdLemma {A = A} {B = B} a1 a2 b conA i = Iso.inv (setTruncProdIso A B) (Iso.inv setTruncTrunc0Iso ((sym (conA .snd ∣ a1 ∣) ∙ (conA .snd ∣ a2 ∣)) i) , ∣ b ∣₀) - --- -- -- -- -- -- -- -- -- test3 : Iso ∥ S₊ 1 × Int ∥₀ Int --- -- -- -- -- -- -- -- -- Iso.fun test3 = sRec isSetInt proj₂ --- -- -- -- -- -- -- -- -- Iso.inv test3 a = ∣ north , a ∣₀ --- -- -- -- -- -- -- -- -- Iso.rightInv test3 a = refl --- -- -- -- -- -- -- -- -- Iso.leftInv test3 = --- -- -- -- -- -- -- -- -- sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) --- -- -- -- -- -- -- -- -- λ {(s , int) → setTruncProdLemma north s int (sphereConnected 1)} - --- -- -- -- -- -- -- -- -- coHomGr0-S1 : grIso intGroup (coHomGr 1 (S₊ 1)) --- -- -- -- -- -- -- -- -- coHomGr0-S1 = --- -- -- -- -- -- -- -- -- Iso'→Iso --- -- -- -- -- -- -- -- -- (iso' (compIso (symIso test3) (symIso coHom1≃∥S1×ℤ∥₀)) --- -- -- -- -- -- -- -- -- (indIntGroup {G = coHomGr 1 (S₊ 1)} --- -- -- -- -- -- -- -- -- (Iso.fun (compIso (symIso test3) (symIso coHom1≃∥S1×ℤ∥₀))) --- -- -- -- -- -- -- -- -- ((λ i → ∣ transport (λ i → (S¹≡S1 (~ i) → lem i)) (Iso.inv test (base , 0)) ∣₀) --- -- -- -- -- -- -- -- -- ∙ (λ i → ∣ transport (λ i → (S¹≡S1 (~ i) → lem i)) (helper2 i) ∣₀) --- -- -- -- -- -- -- -- -- ∙ cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → {!!}})) --- -- -- -- -- -- -- -- -- {!!} --- -- -- -- -- -- -- -- -- {!!})) --- -- -- -- -- -- -- -- -- where --- -- -- -- -- -- -- -- -- helper : basedΩS¹≡ΩS¹ base ≡ {!basechange!} --- -- -- -- -- -- -- -- -- helper = {!substComposite!} - --- -- -- -- -- -- -- -- -- substComposite2 : ∀ {ℓ} {A B C : Type ℓ} --- -- -- -- -- -- -- -- -- (P : A ≡ B) (Q : B ≡ C) (a : A) --- -- -- -- -- -- -- -- -- → transport (P ∙ Q) a ≡ transport Q (transport P a) --- -- -- -- -- -- -- -- -- substComposite2 = {!!} - --- -- -- -- -- -- -- -- -- helper1 : transport (λ i → S¹≡S1 i × Int) (north , 0) ≡ (base , 0) --- -- -- -- -- -- -- -- -- helper1 = refl --- -- -- -- -- -- -- -- -- helper3 : transport (sym (basedΩS¹≡Int base)) 0 ≡ refl --- -- -- -- -- -- -- -- -- helper3 = (λ i → transport (symDistr (basedΩS¹≡ΩS¹ base) ΩS¹≡Int i) 0) --- -- -- -- -- -- -- -- -- ∙ substComposite2 (sym ΩS¹≡Int) (sym (basedΩS¹≡ΩS¹ base)) 0 --- -- -- -- -- -- -- -- -- ∙ (λ i → transport (λ i → basedΩS¹≡ΩS¹ base (~ i)) refl) -- --- -- -- -- -- -- -- -- -- ∙ transportRefl ((equiv-proof (basechange-isequiv base) refl) .fst .fst) --- -- -- -- -- -- -- -- -- ∙ (λ i → equiv-proof (transport (λ j → isEquiv (refl-conjugation j)) (basedΩS¹→ΩS¹-isequiv i0)) refl .fst .fst) --- -- -- -- -- -- -- -- -- ∙ (λ i → {!equiv-proof (transport (λ j → isEquiv (refl-conjugation j)) (basedΩS¹→ΩS¹-isequiv i0)) refl .fst!}) --- -- -- -- -- -- -- -- -- ∙ {!basedΩS¹→ΩS¹!} --- -- -- -- -- -- -- -- -- ∙ {!!} --- -- -- -- -- -- -- -- -- ∙ {!!} --- -- -- -- -- -- -- -- -- helper4 : (x : S¹) → Iso.inv test (base , 0) x ≡ x --- -- -- -- -- -- -- -- -- helper4 = {!!} --- -- -- -- -- -- -- -- -- helper2 : Iso.inv test (transport (λ i → S¹≡S1 i × Int) (north , 0)) ≡ λ x → x --- -- -- -- -- -- -- -- -- helper2 = (λ i → Iso.inv test (base , 0)) ∙ {!!} ∙ {!!} - --- -- -- -- -- -- -- prodId : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A × B} → proj₁ x ≡ proj₁ y → proj₂ x ≡ proj₂ y → x ≡ y --- -- -- -- -- -- -- prodId {x = (a , b)} {y = (c , d)} id1 id2 i = (id1 i) , (id2 i) - --- -- -- -- -- -- -- fstFunHelper : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- -- -- -- -- -- → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 _) (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x --- -- -- -- -- -- -- fstFunHelper a = MV.Ker-i⊂Im-d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a --- -- -- -- -- -- -- (sym (isContrH1Unit×H1Unit .snd _) ∙ (isContrH1Unit×H1Unit .snd _)) --- -- -- -- -- -- -- where --- -- -- -- -- -- -- isContrH1Unit×H1Unit : isContr (Group.type (×coHomGr 1 Unit Unit)) --- -- -- -- -- -- -- isContrH1Unit×H1Unit = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) --- -- -- -- -- -- -- , λ {(a , b) → sigmaProdElim {D = λ (x : Σ[ x ∈ Group.type (×coHomGr 1 Unit Unit)] Unit) → (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) ≡ fst x} --- -- -- -- -- -- -- (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) --- -- -- -- -- -- -- (λ a b _ → prodId (grIso.leftInv (coHomGrUnit≥1 0) ∣ a ∣₀) (grIso.leftInv (coHomGrUnit≥1 0) ∣ b ∣₀)) ((a , b) , tt)} - - - --- -- -- -- -- -- -- helperMorph : isMorph intGroup (dirProd intGroup intGroup) λ a → a , (0 - a) --- -- -- -- -- -- -- helperMorph = --- -- -- -- -- -- -- indIntGroup {G = dirProd intGroup intGroup} --- -- -- -- -- -- -- (λ a → a , (0 - a)) --- -- -- -- -- -- -- refl --- -- -- -- -- -- -- (λ a → prodId refl (helper2 a)) --- -- -- -- -- -- -- λ a → prodId refl (helper3 a) --- -- -- -- -- -- -- where --- -- -- -- -- -- -- helper1 : (a : Int) → predInt (sucInt a) ≡ a --- -- -- -- -- -- -- helper1 (pos zero) = refl --- -- -- -- -- -- -- helper1 (pos (suc n)) = refl --- -- -- -- -- -- -- helper1 (negsuc zero) = refl --- -- -- -- -- -- -- helper1 (negsuc (suc n)) = refl - --- -- -- -- -- -- -- helper4 : (a : Int) → sucInt (predInt a) ≡ a --- -- -- -- -- -- -- helper4 (pos zero) = refl --- -- -- -- -- -- -- helper4 (pos (suc n)) = refl --- -- -- -- -- -- -- helper4 (negsuc zero) = refl --- -- -- -- -- -- -- helper4 (negsuc (suc n)) = refl - --- -- -- -- -- -- -- helper2 : (a : Int) → (pos 0 - sucInt a) ≡ predInt (pos 0 - a) --- -- -- -- -- -- -- helper2 (pos zero) = refl --- -- -- -- -- -- -- helper2 (pos (suc n)) = refl --- -- -- -- -- -- -- helper2 (negsuc zero) = refl --- -- -- -- -- -- -- helper2 (negsuc (suc n)) = sym (helper1 _) - --- -- -- -- -- -- -- helper3 : (a : Int) → (pos 0 - predInt a) ≡ sucInt (pos 0 - a) --- -- -- -- -- -- -- helper3 (pos zero) = refl --- -- -- -- -- -- -- helper3 (pos (suc zero)) = refl --- -- -- -- -- -- -- helper3 (pos (suc (suc n))) = sym (helper4 _) --- -- -- -- -- -- -- helper3 (negsuc zero) = refl --- -- -- -- -- -- -- helper3 (negsuc (suc n)) = refl - --- -- -- -- -- -- -- helper : (a b : Int) → (pos 0 - (a +ℤ b)) ≡ ((pos 0 - a) +ℤ (pos 0 - b)) --- -- -- -- -- -- -- helper a (pos zero) = refl --- -- -- -- -- -- -- helper (pos zero) (pos (suc n)) = --- -- -- -- -- -- -- cong (λ x → pos 0 - sucInt x) (+ℤ-comm (pos zero) (pos n)) --- -- -- -- -- -- -- ∙ +ℤ-comm (pos 0 +negsuc n) (pos zero) --- -- -- -- -- -- -- helper (pos (suc n₁)) (pos (suc n)) = --- -- -- -- -- -- -- {!!} --- -- -- -- -- -- -- helper (negsuc n₁) (pos (suc n)) = {!!} --- -- -- -- -- -- -- helper a (negsuc n) = {!!} - --- -- -- -- -- -- -- fun : morph intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- -- -- -- -- -- fst fun = MV.d _ _ _ (λ _ → tt) (λ _ → tt) 0 ∘ grIso.inv coHom0-S0 .fst ∘ λ a → a , (0 - a) --- -- -- -- -- -- -- snd fun = {!!} --- -- -- -- -- -- -- {- compMorph {F = intGroup} {G = dirProd intGroup intGroup} {H = coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))} --- -- -- -- -- -- -- ((λ a → a , a) , (λ a b → refl)) --- -- -- -- -- -- -- (compMorph {F = dirProd intGroup intGroup} {G = coHomGr 0 (S₊ 0)} {H = coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))} (grIso.inv coHom0-S0) --- -- -- -- -- -- -- (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 --- -- -- -- -- -- -- , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) .snd -} --- -- -- -- -- -- -- {- theIso : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- -- -- -- -- -- theIso = Iso''→Iso _ _ --- -- -- -- -- -- -- (iso'' ((λ x → ∣ (λ {(inl tt) → 0ₖ ; (inr tt) → 0ₖ ; (push a i) → Kn→ΩKn+1 0 x i}) ∣₀) , {!!}) --- -- -- -- -- -- -- {!!} --- -- -- -- -- -- -- {!MV.d!}) --- -- -- -- -- -- -- -} - - - --- -- -- -- -- -- -- theIso : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- -- -- -- -- -- theIso = --- -- -- -- -- -- -- Iso''→Iso _ _ --- -- -- -- -- -- -- (iso'' fun --- -- -- -- -- -- -- (λ x inker → {!MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 --- -- -- -- -- -- -- (grIso.inv coHom0-S0 .fst (g , g))!}) --- -- -- -- -- -- -- (sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) --- -- -- -- -- -- -- λ x → pRec propTruncIsProp --- -- -- -- -- -- -- (λ {(a , b) → {!fun!} }) --- -- -- -- -- -- -- (fstFunHelper (∣ x ∣₀)))) --- -- -- -- -- -- -- where --- -- -- -- -- -- -- whoKnows : (a : S₊ 0 → Int) (x : MV.D Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt)) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (λ _ → a north) x --- -- -- -- -- -- -- ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 a x --- -- -- -- -- -- -- whoKnows a (inl x) = refl --- -- -- -- -- -- -- whoKnows a (inr x) = refl --- -- -- -- -- -- -- whoKnows a (push north i) = refl --- -- -- -- -- -- -- whoKnows a (push south i) j = {!!} - --- -- -- -- -- -- -- helper : (a : Int) → (grIso.inv coHom0-S0 .fst (a , a)) ≡ ∣ S0→Int (a , a) ∣₀ --- -- -- -- -- -- -- helper a = {!have : - --- -- -- -- -- -- -- ∣ --- -- -- -- -- -- -- MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 --- -- -- -- -- -- -- (S0→Int (x , x)) --- -- -- -- -- -- -- ∣₀ --- -- -- -- -- -- -- ≡ ∣ (λ _ → ∣ north ∣) ∣₀!} - --- -- -- -- -- -- -- helper2 : (a b : Int) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a)) ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (b , b)) --- -- -- -- -- -- -- → a ≡ b --- -- -- -- -- -- -- helper2 a b id = pRec (isSetInt a b) (λ {(pt , id) → {!!}}) (fstFunHelper ∣ (MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a))) ∣₀) - --- -- -- -- -- -- -- idFun : (a : Int) → MV.D Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 --- -- -- -- -- -- -- idFun a (inl x) = 0ₖ --- -- -- -- -- -- -- idFun a (inr x) = 0ₖ --- -- -- -- -- -- -- idFun a (push north i) = Kn→ΩKn+1 zero a i --- -- -- -- -- -- -- idFun a (push south i) = Kn→ΩKn+1 zero a i - --- -- -- -- -- -- -- helper3 : (a : Int) → (MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a))) ≡ idFun a --- -- -- -- -- -- -- helper3 a = funExt λ {(inl x) → refl ; (inr x) → refl ; (push north i) → refl ; (push south i) → refl } - --- -- -- -- -- -- -- helper4 : (a b : Int) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a)) ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (b , b)) --- -- -- -- -- -- -- → a ≡ b --- -- -- -- -- -- -- helper4 a b id = --- -- -- -- -- -- -- {!!} --- -- -- -- -- -- -- ∙ {!!} --- -- -- -- -- -- -- ∙ {!!} --- -- -- -- -- -- -- where --- -- -- -- -- -- -- helper5 : {!!} --PathP (λ k → id k (inl tt) ≡ id k (inr tt)) (Kn→ΩKn+1 zero a) (Kn→ΩKn+1 zero a) --- -- -- -- -- -- -- helper5 i j = {!id i!} - --- -- -- -- -- -- -- -- fun : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) → coHom 0 (S₊ 0) --- -- -- -- -- -- -- -- fun a = (pRec {P = Σ[ x ∈ coHom 0 (S₊ 0)] (MV.d _ _ _ (λ _ → tt) (λ _ → tt) 0 x ≡ a) × isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) (MV.Δ _ _ _ (λ _ → tt) (λ _ → tt) 0) x} --- -- -- -- -- -- -- -- (λ {(a1 , b) (c , d) → ΣProp≡ (λ x → isOfHLevelProd 1 (setTruncIsSet _ _) propTruncIsProp) --- -- -- -- -- -- -- -- (sElim2 {B = λ a1 c → (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a1 ≡ a) --- -- -- -- -- -- -- -- → MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 c ≡ a --- -- -- -- -- -- -- -- → isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) --- -- -- -- -- -- -- -- (λ z → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 z) a1 --- -- -- -- -- -- -- -- → isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) --- -- -- -- -- -- -- -- (λ z → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 z) c → a1 ≡ c} --- -- -- -- -- -- -- -- (λ _ _ → {!!}) --- -- -- -- -- -- -- -- (λ a c b1 d1 → pElim (λ _ → isOfHLevelΠ 1 λ _ → setTruncIsSet _ _) --- -- -- -- -- -- -- -- λ b2 → pElim (λ _ → setTruncIsSet _ _) --- -- -- -- -- -- -- -- λ d2 → {!d2!}) --- -- -- -- -- -- -- -- a1 c (proj₁ b) (proj₁ d) (proj₂ b) (proj₂ d))}) --- -- -- -- -- -- -- -- (λ {(a , b) → a , b , MV.Ker-d⊂Im-Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a {!!}}) --- -- -- -- -- -- -- -- (fstFunHelper a)) .fst -- pRec {!!} {!!} (fstFunHelper a) +-- -- -- -- -- -- -- -- -- Iso.inv test (x , int) (loop i) = {!!} +-- -- -- -- -- -- -- -- -- Iso.rightInv test = {!!} +-- -- -- -- -- -- -- -- -- Iso.leftInv test = {!!} + +-- -- -- -- -- -- -- -- -- -- test13 : Iso ∥ (S¹ → S¹) ∥₀ Int +-- -- -- -- -- -- -- -- -- -- Iso.fun test13 = sRec isSetInt λ f → winding (basechange (f base) (λ i → f (loop i))) +-- -- -- -- -- -- -- -- -- -- Iso.inv test13 a = ∣ (λ {base → {!!} ; (loop i) → {!!}}) ∣₀ +-- -- -- -- -- -- -- -- -- -- Iso.rightInv test13 = {!!} +-- -- -- -- -- -- -- -- -- -- Iso.leftInv test13 = {!!} + +-- -- -- -- -- -- -- -- -- -- test : Iso (S¹ → S¹) (S¹ × Int) +-- -- -- -- -- -- -- -- -- -- Iso.fun test f = (f base) , transport (basedΩS¹≡Int (f base)) λ i → f (loop i) +-- -- -- -- -- -- -- -- -- -- Iso.inv test (x , int) base = x +-- -- -- -- -- -- -- -- -- -- Iso.inv test (x , int) (loop i) = transport (sym (basedΩS¹≡Int x)) int i +-- -- -- -- -- -- -- -- -- -- Iso.rightInv test (x , int) i = (x , transportTransport⁻ (basedΩS¹≡Int x) int i) +-- -- -- -- -- -- -- -- -- -- Iso.leftInv test f = +-- -- -- -- -- -- -- -- -- -- funExt λ { base → refl +-- -- -- -- -- -- -- -- -- -- ; (loop i) j → transport⁻Transport (basedΩS¹≡Int (f base)) (λ i → f (loop i)) j i} + + +-- -- -- -- -- -- -- -- -- -- lem : S¹ ≡ hLevelTrunc 3 (S₊ 1) +-- -- -- -- -- -- -- -- -- -- lem = sym (truncIdempotent 3 isGroupoidS¹) ∙ λ i → hLevelTrunc 3 (S¹≡S1 (~ i)) + +-- -- -- -- -- -- -- -- -- -- prodId : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (a b : A × B) → proj₁ a ≡ proj₁ b → proj₂ a ≡ proj₂ b → a ≡ b +-- -- -- -- -- -- -- -- -- -- prodId (_ , _) (_ , _) id1 id2 i = (id1 i) , (id2 i) + +-- -- -- -- -- -- -- -- -- -- test22 : Iso (S₊ 1 → coHomK 1) (S₊ 1 × Int) +-- -- -- -- -- -- -- -- -- -- Iso.fun test22 f = {!f north!} , {!!} +-- -- -- -- -- -- -- -- -- -- Iso.inv test22 = {!!} +-- -- -- -- -- -- -- -- -- -- Iso.rightInv test22 = {!!} +-- -- -- -- -- -- -- -- -- -- Iso.leftInv test22 = {!!} + + + + + + +-- -- -- -- -- -- -- -- -- -- coHom1≃∥S1×ℤ∥₀ : Iso (coHom 1 (S₊ 1)) ∥ S₊ 1 × Int ∥₀ +-- -- -- -- -- -- -- -- -- -- coHom1≃∥S1×ℤ∥₀ = setTruncIso test2 +-- -- -- -- -- -- -- -- -- -- where +-- -- -- -- -- -- -- -- -- -- test2 : Iso (S₊ 1 → coHomK 1) (S₊ 1 × Int) +-- -- -- -- -- -- -- -- -- -- Iso.fun test2 f = transport (λ i → S¹≡S1 (~ i) × Int) (Iso.fun test (transport (λ i → (S¹≡S1 i → lem (~ i))) f)) +-- -- -- -- -- -- -- -- -- -- Iso.inv test2 x = transport (λ i → (S¹≡S1 (~ i) → lem i)) (Iso.inv test (transport (λ i → S¹≡S1 i × Int) x)) +-- -- -- -- -- -- -- -- -- -- Iso.rightInv test2 (s , int) = prodId _ _ {!!} {!!} +-- -- -- -- -- -- -- -- -- -- Iso.leftInv test2 f = {!!} ∙ {!!} ∙ {!!} + +-- -- -- -- -- -- -- -- -- -- test2Id : (a b : (S₊ 1 → coHomK 1)) → proj₂ (Iso.fun test2 (λ x → a x +ₖ b x)) ≡ (proj₂ (Iso.fun test2 a) +ₖ proj₂ (Iso.fun test2 a)) +-- -- -- -- -- -- -- -- -- -- test2Id a b = {! +-- -- -- -- -- -- -- -- -- -- transport +-- -- -- -- -- -- -- -- -- -- (basedΩS¹≡Int +-- -- -- -- -- -- -- -- -- -- (transport (λ i → S¹≡S1 i → lem (~ i)) (λ x → a x +ₖ b x) base)) +-- -- -- -- -- -- -- -- -- -- (λ i → +-- -- -- -- -- -- -- -- -- -- transport (λ i₁ → S¹≡S1 i₁ → lem (~ i₁)) (λ x → a x +ₖ b x) +-- -- -- -- -- -- -- -- -- -- (loop i))!} ∙ {!transport (λ i → S¹≡S1 i → lem (~ i)) (λ x → a x +ₖ b x) base!} + + +-- -- -- -- -- -- -- -- -- -- main : grIso intGroup (coHomGr 1 (S₊ 1)) +-- -- -- -- -- -- -- -- -- -- main = Iso'→Iso +-- -- -- -- -- -- -- -- -- -- (iso' {!!} +-- -- -- -- -- -- -- -- -- -- {!!}) + + +-- -- -- -- -- -- -- -- coHom1 : grIso (coHomGr 1 (S₊ 1)) intGroup +-- -- -- -- -- -- -- -- coHom1 = +-- -- -- -- -- -- -- -- Iso'→Iso +-- -- -- -- -- -- -- -- (iso' (iso ({!!} ∘ {!!} ∘ {!!} ∘ {!!}) +-- -- -- -- -- -- -- -- {!!} +-- -- -- -- -- -- -- -- {!!} +-- -- -- -- -- -- -- -- {!!}) +-- -- -- -- -- -- -- -- {!!}) + + + +-- -- -- -- -- -- -- -- schonf : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : (A × B) → Type ℓ''} +-- -- -- -- -- -- -- -- → ((a : A) (b : B) → C (a , b)) +-- -- -- -- -- -- -- -- → (x : A × B) → C x +-- -- -- -- -- -- -- -- schonf f (a , b) = f a b + +-- -- -- -- -- -- -- -- -- -- setTruncProdIso : ∀ {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') → Iso ∥ (A × B) ∥₀ (∥ A ∥₀ × ∥ B ∥₀) +-- -- -- -- -- -- -- -- -- -- Iso.fun (setTruncProdIso A B) = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) λ {(a , b) → ∣ a ∣₀ , ∣ b ∣₀} +-- -- -- -- -- -- -- -- -- -- Iso.inv (setTruncProdIso A B) (a , b) = sRec setTruncIsSet (λ a → sRec setTruncIsSet (λ b → ∣ a , b ∣₀) b) a +-- -- -- -- -- -- -- -- -- -- Iso.rightInv (setTruncProdIso A B) = +-- -- -- -- -- -- -- -- -- -- schonf (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- -- -- -- -- -- -- -- -- -- λ _ → sElim (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- -- -- -- -- -- -- -- -- -- λ _ → refl) +-- -- -- -- -- -- -- -- -- -- Iso.leftInv (setTruncProdIso A B) = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ {(a , b) → refl} + +-- -- -- -- -- -- -- -- -- -- setTruncProdLemma : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (a1 a2 : A) (b : B) → isHLevelConnected 2 A +-- -- -- -- -- -- -- -- -- -- → Path ∥ A × B ∥₀ ∣ a1 , b ∣₀ ∣ a2 , b ∣₀ +-- -- -- -- -- -- -- -- -- -- setTruncProdLemma {A = A} {B = B} a1 a2 b conA i = Iso.inv (setTruncProdIso A B) (Iso.inv setTruncTrunc0Iso ((sym (conA .snd ∣ a1 ∣) ∙ (conA .snd ∣ a2 ∣)) i) , ∣ b ∣₀) + +-- -- -- -- -- -- -- -- -- -- test3 : Iso ∥ S₊ 1 × Int ∥₀ Int +-- -- -- -- -- -- -- -- -- -- Iso.fun test3 = sRec isSetInt proj₂ +-- -- -- -- -- -- -- -- -- -- Iso.inv test3 a = ∣ north , a ∣₀ +-- -- -- -- -- -- -- -- -- -- Iso.rightInv test3 a = refl +-- -- -- -- -- -- -- -- -- -- Iso.leftInv test3 = +-- -- -- -- -- -- -- -- -- -- sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- -- -- -- -- -- -- -- -- -- λ {(s , int) → setTruncProdLemma north s int (sphereConnected 1)} + +-- -- -- -- -- -- -- -- -- -- coHomGr0-S1 : grIso intGroup (coHomGr 1 (S₊ 1)) +-- -- -- -- -- -- -- -- -- -- coHomGr0-S1 = +-- -- -- -- -- -- -- -- -- -- Iso'→Iso +-- -- -- -- -- -- -- -- -- -- (iso' (compIso (symIso test3) (symIso coHom1≃∥S1×ℤ∥₀)) +-- -- -- -- -- -- -- -- -- -- (indIntGroup {G = coHomGr 1 (S₊ 1)} +-- -- -- -- -- -- -- -- -- -- (Iso.fun (compIso (symIso test3) (symIso coHom1≃∥S1×ℤ∥₀))) +-- -- -- -- -- -- -- -- -- -- ((λ i → ∣ transport (λ i → (S¹≡S1 (~ i) → lem i)) (Iso.inv test (base , 0)) ∣₀) +-- -- -- -- -- -- -- -- -- -- ∙ (λ i → ∣ transport (λ i → (S¹≡S1 (~ i) → lem i)) (helper2 i) ∣₀) +-- -- -- -- -- -- -- -- -- -- ∙ cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → {!!}})) +-- -- -- -- -- -- -- -- -- -- {!!} +-- -- -- -- -- -- -- -- -- -- {!!})) +-- -- -- -- -- -- -- -- -- -- where +-- -- -- -- -- -- -- -- -- -- helper : basedΩS¹≡ΩS¹ base ≡ {!basechange!} +-- -- -- -- -- -- -- -- -- -- helper = {!substComposite!} + +-- -- -- -- -- -- -- -- -- -- substComposite2 : ∀ {ℓ} {A B C : Type ℓ} +-- -- -- -- -- -- -- -- -- -- (P : A ≡ B) (Q : B ≡ C) (a : A) +-- -- -- -- -- -- -- -- -- -- → transport (P ∙ Q) a ≡ transport Q (transport P a) +-- -- -- -- -- -- -- -- -- -- substComposite2 = {!!} + +-- -- -- -- -- -- -- -- -- -- helper1 : transport (λ i → S¹≡S1 i × Int) (north , 0) ≡ (base , 0) +-- -- -- -- -- -- -- -- -- -- helper1 = refl +-- -- -- -- -- -- -- -- -- -- helper3 : transport (sym (basedΩS¹≡Int base)) 0 ≡ refl +-- -- -- -- -- -- -- -- -- -- helper3 = (λ i → transport (symDistr (basedΩS¹≡ΩS¹ base) ΩS¹≡Int i) 0) +-- -- -- -- -- -- -- -- -- -- ∙ substComposite2 (sym ΩS¹≡Int) (sym (basedΩS¹≡ΩS¹ base)) 0 +-- -- -- -- -- -- -- -- -- -- ∙ (λ i → transport (λ i → basedΩS¹≡ΩS¹ base (~ i)) refl) -- +-- -- -- -- -- -- -- -- -- -- ∙ transportRefl ((equiv-proof (basechange-isequiv base) refl) .fst .fst) +-- -- -- -- -- -- -- -- -- -- ∙ (λ i → equiv-proof (transport (λ j → isEquiv (refl-conjugation j)) (basedΩS¹→ΩS¹-isequiv i0)) refl .fst .fst) +-- -- -- -- -- -- -- -- -- -- ∙ (λ i → {!equiv-proof (transport (λ j → isEquiv (refl-conjugation j)) (basedΩS¹→ΩS¹-isequiv i0)) refl .fst!}) +-- -- -- -- -- -- -- -- -- -- ∙ {!basedΩS¹→ΩS¹!} +-- -- -- -- -- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- -- -- -- -- helper4 : (x : S¹) → Iso.inv test (base , 0) x ≡ x +-- -- -- -- -- -- -- -- -- -- helper4 = {!!} +-- -- -- -- -- -- -- -- -- -- helper2 : Iso.inv test (transport (λ i → S¹≡S1 i × Int) (north , 0)) ≡ λ x → x +-- -- -- -- -- -- -- -- -- -- helper2 = (λ i → Iso.inv test (base , 0)) ∙ {!!} ∙ {!!} + +-- -- -- -- -- -- -- -- prodId : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A × B} → proj₁ x ≡ proj₁ y → proj₂ x ≡ proj₂ y → x ≡ y +-- -- -- -- -- -- -- -- prodId {x = (a , b)} {y = (c , d)} id1 id2 i = (id1 i) , (id2 i) + +-- -- -- -- -- -- -- -- fstFunHelper : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- -- -- -- → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 _) (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x +-- -- -- -- -- -- -- -- fstFunHelper a = MV.Ker-i⊂Im-d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a +-- -- -- -- -- -- -- -- (sym (isContrH1Unit×H1Unit .snd _) ∙ (isContrH1Unit×H1Unit .snd _)) +-- -- -- -- -- -- -- -- where +-- -- -- -- -- -- -- -- isContrH1Unit×H1Unit : isContr (Group.type (×coHomGr 1 Unit Unit)) +-- -- -- -- -- -- -- -- isContrH1Unit×H1Unit = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) +-- -- -- -- -- -- -- -- , λ {(a , b) → sigmaProdElim {D = λ (x : Σ[ x ∈ Group.type (×coHomGr 1 Unit Unit)] Unit) → (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) ≡ fst x} +-- -- -- -- -- -- -- -- (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- -- -- -- -- -- -- -- (λ a b _ → prodId (grIso.leftInv (coHomGrUnit≥1 0) ∣ a ∣₀) (grIso.leftInv (coHomGrUnit≥1 0) ∣ b ∣₀)) ((a , b) , tt)} + + + +-- -- -- -- -- -- -- -- helperMorph : isMorph intGroup (dirProd intGroup intGroup) λ a → a , (0 - a) +-- -- -- -- -- -- -- -- helperMorph = +-- -- -- -- -- -- -- -- indIntGroup {G = dirProd intGroup intGroup} +-- -- -- -- -- -- -- -- (λ a → a , (0 - a)) +-- -- -- -- -- -- -- -- refl +-- -- -- -- -- -- -- -- (λ a → prodId refl (helper2 a)) +-- -- -- -- -- -- -- -- λ a → prodId refl (helper3 a) +-- -- -- -- -- -- -- -- where +-- -- -- -- -- -- -- -- helper1 : (a : Int) → predInt (sucInt a) ≡ a +-- -- -- -- -- -- -- -- helper1 (pos zero) = refl +-- -- -- -- -- -- -- -- helper1 (pos (suc n)) = refl +-- -- -- -- -- -- -- -- helper1 (negsuc zero) = refl +-- -- -- -- -- -- -- -- helper1 (negsuc (suc n)) = refl + +-- -- -- -- -- -- -- -- helper4 : (a : Int) → sucInt (predInt a) ≡ a +-- -- -- -- -- -- -- -- helper4 (pos zero) = refl +-- -- -- -- -- -- -- -- helper4 (pos (suc n)) = refl +-- -- -- -- -- -- -- -- helper4 (negsuc zero) = refl +-- -- -- -- -- -- -- -- helper4 (negsuc (suc n)) = refl + +-- -- -- -- -- -- -- -- helper2 : (a : Int) → (pos 0 - sucInt a) ≡ predInt (pos 0 - a) +-- -- -- -- -- -- -- -- helper2 (pos zero) = refl +-- -- -- -- -- -- -- -- helper2 (pos (suc n)) = refl +-- -- -- -- -- -- -- -- helper2 (negsuc zero) = refl +-- -- -- -- -- -- -- -- helper2 (negsuc (suc n)) = sym (helper1 _) + +-- -- -- -- -- -- -- -- helper3 : (a : Int) → (pos 0 - predInt a) ≡ sucInt (pos 0 - a) +-- -- -- -- -- -- -- -- helper3 (pos zero) = refl +-- -- -- -- -- -- -- -- helper3 (pos (suc zero)) = refl +-- -- -- -- -- -- -- -- helper3 (pos (suc (suc n))) = sym (helper4 _) +-- -- -- -- -- -- -- -- helper3 (negsuc zero) = refl +-- -- -- -- -- -- -- -- helper3 (negsuc (suc n)) = refl + +-- -- -- -- -- -- -- -- helper : (a b : Int) → (pos 0 - (a +ℤ b)) ≡ ((pos 0 - a) +ℤ (pos 0 - b)) +-- -- -- -- -- -- -- -- helper a (pos zero) = refl +-- -- -- -- -- -- -- -- helper (pos zero) (pos (suc n)) = +-- -- -- -- -- -- -- -- cong (λ x → pos 0 - sucInt x) (+ℤ-comm (pos zero) (pos n)) +-- -- -- -- -- -- -- -- ∙ +ℤ-comm (pos 0 +negsuc n) (pos zero) +-- -- -- -- -- -- -- -- helper (pos (suc n₁)) (pos (suc n)) = +-- -- -- -- -- -- -- -- {!!} +-- -- -- -- -- -- -- -- helper (negsuc n₁) (pos (suc n)) = {!!} +-- -- -- -- -- -- -- -- helper a (negsuc n) = {!!} + +-- -- -- -- -- -- -- -- fun : morph intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- -- -- -- fst fun = MV.d _ _ _ (λ _ → tt) (λ _ → tt) 0 ∘ grIso.inv coHom0-S0 .fst ∘ λ a → a , (0 - a) +-- -- -- -- -- -- -- -- snd fun = {!!} +-- -- -- -- -- -- -- -- {- compMorph {F = intGroup} {G = dirProd intGroup intGroup} {H = coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))} +-- -- -- -- -- -- -- -- ((λ a → a , a) , (λ a b → refl)) +-- -- -- -- -- -- -- -- (compMorph {F = dirProd intGroup intGroup} {G = coHomGr 0 (S₊ 0)} {H = coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))} (grIso.inv coHom0-S0) +-- -- -- -- -- -- -- -- (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 +-- -- -- -- -- -- -- -- , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) .snd -} +-- -- -- -- -- -- -- -- {- theIso : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- -- -- -- theIso = Iso''→Iso _ _ +-- -- -- -- -- -- -- -- (iso'' ((λ x → ∣ (λ {(inl tt) → 0ₖ ; (inr tt) → 0ₖ ; (push a i) → Kn→ΩKn+1 0 x i}) ∣₀) , {!!}) +-- -- -- -- -- -- -- -- {!!} +-- -- -- -- -- -- -- -- {!MV.d!}) +-- -- -- -- -- -- -- -- -} + + + +-- -- -- -- -- -- -- -- theIso : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- -- -- -- theIso = +-- -- -- -- -- -- -- -- Iso''→Iso _ _ +-- -- -- -- -- -- -- -- (iso'' fun +-- -- -- -- -- -- -- -- (λ x inker → {!MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 +-- -- -- -- -- -- -- -- (grIso.inv coHom0-S0 .fst (g , g))!}) +-- -- -- -- -- -- -- -- (sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- -- -- -- -- -- -- -- λ x → pRec propTruncIsProp +-- -- -- -- -- -- -- -- (λ {(a , b) → {!fun!} }) +-- -- -- -- -- -- -- -- (fstFunHelper (∣ x ∣₀)))) +-- -- -- -- -- -- -- -- where +-- -- -- -- -- -- -- -- whoKnows : (a : S₊ 0 → Int) (x : MV.D Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt)) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (λ _ → a north) x +-- -- -- -- -- -- -- -- ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 a x +-- -- -- -- -- -- -- -- whoKnows a (inl x) = refl +-- -- -- -- -- -- -- -- whoKnows a (inr x) = refl +-- -- -- -- -- -- -- -- whoKnows a (push north i) = refl +-- -- -- -- -- -- -- -- whoKnows a (push south i) j = {!!} + +-- -- -- -- -- -- -- -- helper : (a : Int) → (grIso.inv coHom0-S0 .fst (a , a)) ≡ ∣ S0→Int (a , a) ∣₀ +-- -- -- -- -- -- -- -- helper a = {!have : + +-- -- -- -- -- -- -- -- ∣ +-- -- -- -- -- -- -- -- MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 +-- -- -- -- -- -- -- -- (S0→Int (x , x)) +-- -- -- -- -- -- -- -- ∣₀ +-- -- -- -- -- -- -- -- ≡ ∣ (λ _ → ∣ north ∣) ∣₀!} + +-- -- -- -- -- -- -- -- helper2 : (a b : Int) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a)) ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (b , b)) +-- -- -- -- -- -- -- -- → a ≡ b +-- -- -- -- -- -- -- -- helper2 a b id = pRec (isSetInt a b) (λ {(pt , id) → {!!}}) (fstFunHelper ∣ (MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a))) ∣₀) + +-- -- -- -- -- -- -- -- idFun : (a : Int) → MV.D Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 +-- -- -- -- -- -- -- -- idFun a (inl x) = 0ₖ +-- -- -- -- -- -- -- -- idFun a (inr x) = 0ₖ +-- -- -- -- -- -- -- -- idFun a (push north i) = Kn→ΩKn+1 zero a i +-- -- -- -- -- -- -- -- idFun a (push south i) = Kn→ΩKn+1 zero a i + +-- -- -- -- -- -- -- -- helper3 : (a : Int) → (MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a))) ≡ idFun a +-- -- -- -- -- -- -- -- helper3 a = funExt λ {(inl x) → refl ; (inr x) → refl ; (push north i) → refl ; (push south i) → refl } + +-- -- -- -- -- -- -- -- helper4 : (a b : Int) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a)) ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (b , b)) +-- -- -- -- -- -- -- -- → a ≡ b +-- -- -- -- -- -- -- -- helper4 a b id = +-- -- -- -- -- -- -- -- {!!} +-- -- -- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- -- -- where +-- -- -- -- -- -- -- -- helper5 : {!!} --PathP (λ k → id k (inl tt) ≡ id k (inr tt)) (Kn→ΩKn+1 zero a) (Kn→ΩKn+1 zero a) +-- -- -- -- -- -- -- -- helper5 i j = {!id i!} + +-- -- -- -- -- -- -- -- -- fun : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) → coHom 0 (S₊ 0) +-- -- -- -- -- -- -- -- -- fun a = (pRec {P = Σ[ x ∈ coHom 0 (S₊ 0)] (MV.d _ _ _ (λ _ → tt) (λ _ → tt) 0 x ≡ a) × isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) (MV.Δ _ _ _ (λ _ → tt) (λ _ → tt) 0) x} +-- -- -- -- -- -- -- -- -- (λ {(a1 , b) (c , d) → ΣProp≡ (λ x → isOfHLevelProd 1 (setTruncIsSet _ _) propTruncIsProp) +-- -- -- -- -- -- -- -- -- (sElim2 {B = λ a1 c → (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a1 ≡ a) +-- -- -- -- -- -- -- -- -- → MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 c ≡ a +-- -- -- -- -- -- -- -- -- → isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) +-- -- -- -- -- -- -- -- -- (λ z → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 z) a1 +-- -- -- -- -- -- -- -- -- → isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) +-- -- -- -- -- -- -- -- -- (λ z → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 z) c → a1 ≡ c} +-- -- -- -- -- -- -- -- -- (λ _ _ → {!!}) +-- -- -- -- -- -- -- -- -- (λ a c b1 d1 → pElim (λ _ → isOfHLevelΠ 1 λ _ → setTruncIsSet _ _) +-- -- -- -- -- -- -- -- -- λ b2 → pElim (λ _ → setTruncIsSet _ _) +-- -- -- -- -- -- -- -- -- λ d2 → {!d2!}) +-- -- -- -- -- -- -- -- -- a1 c (proj₁ b) (proj₁ d) (proj₂ b) (proj₂ d))}) +-- -- -- -- -- -- -- -- -- (λ {(a , b) → a , b , MV.Ker-d⊂Im-Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a {!!}}) +-- -- -- -- -- -- -- -- -- (fstFunHelper a)) .fst -- pRec {!!} {!!} (fstFunHelper a) From c8a15cde30646ab87db1fe9792523be6a840f0aa Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Wed, 10 Jun 2020 03:13:20 +0200 Subject: [PATCH 41/89] moreTorusStuff --- Cubical/Data/Group/Base.agda | 24 + Cubical/Foundations/Prelude.agda | 23 + Cubical/ZCohomology/MayerVietorisReduced.agda | 9 +- Cubical/ZCohomology/Sn/Sn.agda | 3000 +++++++++-------- 4 files changed, 1653 insertions(+), 1403 deletions(-) diff --git a/Cubical/Data/Group/Base.agda b/Cubical/Data/Group/Base.agda index ab8dd4fca..6b5409181 100644 --- a/Cubical/Data/Group/Base.agda +++ b/Cubical/Data/Group/Base.agda @@ -87,7 +87,25 @@ morph0→0 G' H' f ismorph = 0→0 comp H (f (id G)) (inv H (f (id G))) ≡⟨ rCancel H (f (id G)) ⟩ id H ∎ +morph0→0' : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (f : morph G H) + → morph.fun f (isGroup.id (Group.groupStruc G)) ≡ isGroup.id (Group.groupStruc H) +morph0→0' G' H' f' = 0→0 + where + open Group + open isGroup + G = Group.groupStruc G' + H = Group.groupStruc H' + f = morph.fun f' + ismorph = morph.ismorph f' + 0→0 : f (id G) ≡ id H + 0→0 = + f (id G) ≡⟨ sym (rUnit H (f (id G))) ⟩ + comp H (f (id G)) (id H) ≡⟨ (λ i → comp H (f (id G)) (rCancel H (f (id G)) (~ i))) ⟩ + comp H (f (id G)) (comp H (f (id G)) (inv H (f (id G)))) ≡⟨ sym (assoc H (f (id G)) (f (id G)) (inv H (f (id G)))) ⟩ + comp H (comp H (f (id G)) (f (id G))) (inv H (f (id G))) ≡⟨ sym (cong (λ x → comp H x (inv H (f (id G)))) (sym (cong f (lUnit G (id G))) ∙ ismorph (id G) (id G))) ⟩ + comp H (f (id G)) (inv H (f (id G))) ≡⟨ rCancel H (f (id G)) ⟩ + id H ∎ {- a morphism ϕ satisfies ϕ(- a) = - ϕ(a) -} morphMinus : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (f : (Group.type G → Group.type H)) @@ -341,6 +359,12 @@ Iso''→Iso {A = A} {B = B} is = comp A' (id A') a2 ≡⟨ lUnit A' a2 ⟩ a2 ∎ +groupIso→Iso : ∀ {ℓ ℓ'} {G : Group ℓ} {H : Group ℓ'} → Iso G H → I.Iso (Group.type G) (Group.type H) +I.Iso.fun (groupIso→Iso i) = morph.fun (Iso.fun i) +I.Iso.inv (groupIso→Iso i) = morph.fun (Iso.inv i) +I.Iso.rightInv (groupIso→Iso i) = Iso.rightInv i +I.Iso.leftInv (groupIso→Iso i) = Iso.leftInv i + -- -- Injectivity and surjectivity in terms of exact sequences diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index 330a07fe0..c69ebd6eb 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -443,3 +443,26 @@ PathP→compPathR {p = p} {q = q} {r = r} {s = s} P j i = ; (j = i0) → r i ; (j = i1) → doubleCompPath-filler∙ p s (sym q) (~ k) i}) (P j i) + + +-- otherdir + +compPathL→PathP : {a b c d : A} {p : a ≡ c} {q : b ≡ d} {r : a ≡ b} {s : c ≡ d} + → sym p ∙ r ∙ q ≡ s + → PathP (λ i → p i ≡ q i) r s +compPathL→PathP {p = p} {q = q} {r = r} {s = s} P j i = + hcomp (λ k → λ { (i = i0) → p (~ k ∨ j) + ; (i = i1) → q (~ k ∨ j) + ; (j = i0) → doubleCompPath-filler∙ (sym p) r q k i + ; (j = i1) → s i}) + (P j i) + +compPathR→PathP : {a b c d : A} {p : a ≡ c} {q : b ≡ d} {r : a ≡ b} {s : c ≡ d} + → r ≡ p ∙ s ∙ sym q + → PathP (λ i → p i ≡ q i) r s +compPathR→PathP {p = p} {q = q} {r = r} {s = s} P j i = + hcomp (λ k → λ { (i = i0) → p (k ∧ j) + ; (i = i1) → q (k ∧ j) + ; (j = i0) → r i + ; (j = i1) → doubleCompPath-filler∙ p s (sym q) k i}) + (P j i) diff --git a/Cubical/ZCohomology/MayerVietorisReduced.agda b/Cubical/ZCohomology/MayerVietorisReduced.agda index 580f79b21..c813e3898 100644 --- a/Cubical/ZCohomology/MayerVietorisReduced.agda +++ b/Cubical/ZCohomology/MayerVietorisReduced.agda @@ -356,12 +356,15 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : d-morph : ∀ {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) (n : ℕ) → morph (coHomGr n C) (coHomGr (suc n) (Pushout f g)) -d-morph A B C f g n = mph (MV.d A B C f g n) (MV.dIsHom A B C f g n) +morph.fun (d-morph A B C f g n) = MV.d A B C f g n +morph.ismorph (d-morph A B C f g n) = MV.dIsHom A B C f g n i-morph : ∀ {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) (n : ℕ) → morph (coHomGr n (Pushout f g)) (×coHomGr n A B) -i-morph A B C f g n = mph (MV.i A B C f g n) (MV.iIsHom A B C f g n) +morph.fun (i-morph A B C f g n) = MV.i A B C f g n +morph.ismorph (i-morph A B C f g n) = MV.iIsHom A B C f g n Δ-morph : ∀ {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) (n : ℕ) → morph (×coHomGr n A B) (coHomGr n C) -Δ-morph A B C f g n = mph (MV.Δ A B C f g n) (MV.ΔIsHom A B C f g n) +morph.fun (Δ-morph A B C f g n) = MV.Δ A B C f g n +morph.ismorph (Δ-morph A B C f g n) = MV.ΔIsHom A B C f g n diff --git a/Cubical/ZCohomology/Sn/Sn.agda b/Cubical/ZCohomology/Sn/Sn.agda index bb093ed34..e22dd3637 100644 --- a/Cubical/ZCohomology/Sn/Sn.agda +++ b/Cubical/ZCohomology/Sn/Sn.agda @@ -38,6 +38,8 @@ open import Cubical.Data.Group.GroupLibrary open import Cubical.ZCohomology.coHomZero-map open import Cubical.HITs.Wedge +open import Cubical.Foundations.Equiv.HalfAdjoint + open import Cubical.ZCohomology.KcompPrelims @@ -49,6 +51,8 @@ open import Cubical.Data.Bool hiding (_⊕_) + + funIso : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} → Iso (A → B × C) ((A → B) × (A → C)) Iso.fun funIso = λ f → (λ a → proj₁ (f a)) , (λ a → proj₂ (f a)) Iso.inv funIso (f , g) = λ a → (f a) , (g a) @@ -135,7 +139,7 @@ Iso.rightInv test13 (a , b) = ×≡ (rUnitₖ a) ((cong ΩKn+1→Kn (congHelper++ (Kn→ΩKn+1 1 b) (λ x → (a +ₖ x) +ₖ (-ₖ (a +ₖ 0ₖ))) (funExt (λ x → sym (cancelHelper a x))) (rCancelₖ (a +ₖ 0ₖ)))) ∙ Iso.leftInv (Iso3-Kn-ΩKn+1 1) b) - where + module _ where cancelHelper : (a b : hLevelTrunc 4 (S₊ 2)) → (a +ₖ b) +ₖ (-ₖ (a +ₖ 0ₖ)) ≡ b cancelHelper a b = (a +ₖ b) +ₖ (-ₖ (a +ₖ 0ₖ)) ≡⟨ (λ i → (a +ₖ b) +ₖ (-ₖ (rUnitₖ a i))) ⟩ @@ -168,10 +172,6 @@ Iso.rightInv test13 (a , b) = ∙ cong (_∙ mid) (rCancel (sym q ∙ p)) ∙ sym (lUnit mid) - - - --- (cong (λ x → (a +ₖ 0ₖ) +ₖ x) (-distrₖ a 0ₖ) congHelper : ∀ {ℓ} {A : Type ℓ} {a1 : A} (p : a1 ≡ a1) (f : A → A) (id : (λ x → x) ≡ f) → cong f p ≡ sym (funExt⁻ id a1) ∙ p ∙ funExt⁻ id a1 congHelper {a1 = a1} p f id = @@ -189,1405 +189,1605 @@ Iso.rightInv test13 (a , b) = Iso.leftInv test13 a = funExt λ {north → rUnitₖ (a north) ; south → rUnitₖ (a north) ∙ cong a (merid north) - ; (merid south i) j → {!!} - ; (merid north i) → {!!}} - - - - --- S1→S¹ : S₊ 1 → S¹ --- S1→S¹ x = SuspBool→S¹ (S1→SuspBool x) - --- S¹→S1 : S¹ → S₊ 1 --- S¹→S1 x = SuspBool→S1 (S¹→SuspBool x) + ; (merid south i) → test i + ; (merid north i) → test2 i} + where + test : PathP (λ i → a north +ₖ Kn→ΩKn+1 1 (ΩKn+1→Kn (sym (rCancelₖ (a north)) + ∙ (λ i → a (merid south i) +ₖ (-ₖ a (merid north i))) + ∙ rCancelₖ (a south))) i ≡ a (merid south i)) (rUnitₖ (a north)) (rUnitₖ (a north) ∙ cong a (merid north)) + test j i = + hcomp (λ k → λ { (i = i0) → (helper ∙ together ∙ congFunct (_+ₖ 0ₖ) (cong a (merid south)) (cong a (sym (merid north)))) (~ k) j + ; (i = i1) → a (merid south j) + ; (j = i0) → rUnitₖ (a north) i + ; (j = i1) → ((λ j → rUnitₖ (a (merid north ((~ k) ∧ j))) (j ∧ k)) ∙ λ j → rUnitₖ (a (merid north ((~ k) ∨ j))) (k ∨ j)) i }) + (hcomp (λ k → λ { (i = i1) → a (merid south j) + ; (i = i0) → compPath-filler ((cong (λ x → a x +ₖ 0ₖ) (merid south))) (cong (λ x → a x +ₖ 0ₖ) (sym (merid north))) k j + ; (j = i0) → rUnitₖ (a north) i + ; (j = i1) → pathFiller2 (cong (_+ₖ 0ₖ) (cong a (merid north))) (rUnitₖ (a south)) k i}) + (rUnitₖ (a (merid south j)) i)) + + + where + pathFiller : (rUnitₖ (a north) ∙ cong a (merid north)) ≡ cong (_+ₖ 0ₖ) (cong a (merid north)) ∙ rUnitₖ (a south) + pathFiller = (λ i → (λ j → rUnitₖ (a (merid north (i ∧ j))) (j ∧ (~ i))) ∙ λ j → rUnitₖ (a (merid north (i ∨ j))) (~ i ∨ j)) + + pathFiller2 : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) → PathP (λ i → p (~ i) ≡ z) q (p ∙ q) + pathFiller2 p q i = + hcomp (λ k → λ {(i = i0) → lUnit q (~ k) + ; (i = i1) → p ∙ q}) + ((λ j → p (~ i ∨ j)) ∙ q) + + helper : Path (_ ≡ _) (λ i → a north +ₖ Kn→ΩKn+1 1 (ΩKn+1→Kn (sym (rCancelₖ (a north)) + ∙ (λ i → a (merid south i) +ₖ (-ₖ a (merid north i))) + ∙ rCancelₖ (a south))) i) + --- + (cong (a north +ₖ_) (sym (rCancelₖ (a north))) + ∙ ((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north))) + ∙ cong (a north +ₖ_) (rCancelₖ (a north))) + helper = + (λ j i → a north +ₖ Iso.rightInv (Iso3-Kn-ΩKn+1 1) ((sym (rCancelₖ (a north)) + ∙ (λ i → a (merid south i) +ₖ (-ₖ a (merid north i))) + ∙ rCancelₖ (a south))) j i) ∙ + congFunct (a north +ₖ_) (sym (rCancelₖ (a north))) ((λ i → a (merid south i) +ₖ (-ₖ a (merid north i))) ∙ rCancelₖ (a south)) ∙ + (λ j → cong (a north +ₖ_) (sym (rCancelₖ (a north))) ∙ congFunct (a north +ₖ_) ((λ i → a (merid south i) +ₖ (-ₖ a (merid north i)))) (rCancelₖ (a south)) j) ∙ + (λ j → cong (a north +ₖ_) (sym (rCancelₖ (a north))) ∙ cong₂Funct2 (λ x y → a north +ₖ a x +ₖ (-ₖ (a y))) (merid south) (merid north) j ∙ cong (a north +ₖ_) ((rCancelₖ (a south)))) ∙ + (λ i → cong (a north +ₖ_) (sym (rCancelₖ (a north))) ∙ ((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ i ∨ (~ j))) +ₖ (-ₖ (a (merid north (j ∧ (~ i)))))) ∙ cong (a north +ₖ_) ((rCancelₖ (a (merid north (~ i)))))) --- S1→S¹-sect : section S1→S¹ S¹→S1 --- S1→S¹-sect x = --- cong SuspBool→S¹ (SuspBool→S1-retr (S¹→SuspBool x)) --- ∙ S¹→SuspBool→S¹ x + abstract + pathHelper : (a b : hLevelTrunc 4 (S₊ 2)) → a +ₖ b +ₖ (-ₖ a) ≡ b +ₖ 0ₖ + pathHelper a b = + a +ₖ b +ₖ (-ₖ a) ≡⟨ commₖ a (b +ₖ (-ₖ a)) ⟩ + (b +ₖ (-ₖ a)) +ₖ a ≡⟨ assocₖ b (-ₖ a) a ⟩ + b +ₖ (-ₖ a) +ₖ a ≡⟨ cong (b +ₖ_) (lCancelₖ a) ⟩ + b +ₖ 0ₖ ∎ + + + helperFun : ∀ {ℓ} {A : Type ℓ} {a b : A} (p p' : a ≡ b) (q q' : b ≡ b) (r : a ≡ a) + → ((p q : a ≡ a) → p ∙ q ≡ q ∙ p) + → q ≡ q' + → PathP (λ i → p' (~ i) ≡ p' (~ i)) q' r + → p ∙ q ∙ sym p ≡ r + helperFun p p' q q' r comm qid dep = + p ∙ q ∙ sym p ≡⟨ cong (λ x → p ∙ x ∙ sym p) qid ⟩ + p ∙ q' ∙ sym p ≡⟨ cong (λ x → p ∙ x ∙ sym p) (λ i → lUnit (rUnit q' i) i) ⟩ + p ∙ (refl ∙ q' ∙ refl) ∙ sym p ≡⟨ cong (λ x → p ∙ x ∙ sym p) (λ i → (λ j → p' (~ i ∨ ~ j)) ∙ dep i ∙ λ j → p' (~ i ∨ j)) ⟩ + p ∙ (sym p' ∙ r ∙ p') ∙ sym p ≡⟨ assoc p (sym p' ∙ r ∙ p') (sym p) ⟩ + (p ∙ sym p' ∙ r ∙ p') ∙ sym p ≡⟨ cong (_∙ sym p) (assoc p (sym p') (r ∙ p')) ⟩ + ((p ∙ sym p') ∙ r ∙ p') ∙ sym p ≡⟨ sym (assoc (p ∙ sym p') (r ∙ p') (sym p)) ⟩ + (p ∙ sym p') ∙ (r ∙ p') ∙ sym p ≡⟨ cong ((p ∙ sym p') ∙_) (sym (assoc r p' (sym p))) ⟩ + (p ∙ sym p') ∙ r ∙ (p' ∙ sym p) ≡⟨ cong (λ x → (p ∙ sym p') ∙ r ∙ x) (sym (symDistr p (sym p'))) ⟩ + (p ∙ sym p') ∙ r ∙ sym (p ∙ sym p') ≡⟨ cong ((p ∙ sym p') ∙_) (comm r (sym (p ∙ sym p'))) ⟩ + (p ∙ sym p') ∙ sym (p ∙ sym p') ∙ r ≡⟨ assoc (p ∙ sym p') (sym (p ∙ sym p')) r ⟩ + ((p ∙ sym p') ∙ sym (p ∙ sym p')) ∙ r ≡⟨ cong (_∙ r) (rCancel (p ∙ sym p')) ⟩ + refl ∙ r ≡⟨ sym (lUnit r) ⟩ + r ∎ + together : Path (_ ≡ _) (cong (a north +ₖ_) (sym (rCancelₖ (a north))) + ∙ ((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north))) + ∙ cong (a north +ₖ_) (rCancelₖ (a north))) + (cong(_+ₖ 0ₖ) (cong a (merid south) ∙ cong a (sym (merid north)))) + together = + helperFun (cong (a north +ₖ_) (sym (rCancelₖ (a north)))) + (λ i → pathHelper (a north) (a north) (~ i)) + ((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north))) + (((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north)))) + (cong(_+ₖ 0ₖ) (cong a (merid south) ∙ cong a (sym (merid north)))) + (transport (λ i → (p q : rUnitₖ (a north) (~ i) ≡ rUnitₖ (a north) (~ i)) → p ∙ q ≡ q ∙ p) + (λ p q → isCommS2 _ p q)) + refl + λ i → hcomp (λ k → λ {(i = i0) → (λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north)) + ; (i = i1) → congFunct (_+ₖ 0ₖ) (cong a (merid south)) (cong a (sym (merid north))) (~ k)}) + ((λ j → pathHelper (a north) (a (merid south j)) i) ∙ (λ j → pathHelper (a north) (a (merid north (~ j))) i)) + where + isCommS2 : (x : hLevelTrunc 4 (S₊ 2)) (p q : x ≡ x) → p ∙ q ≡ q ∙ p + isCommS2 x p q = sym (transport⁻Transport (typId x) (p ∙ q)) + ∙ cong (transport (λ i → typId x (~ i))) (typIdHelper p q) + ∙ (λ i → transport (λ i → typId x (~ i)) (commHelper ∣ north ∣ ∣ north ∣ (transport (λ i → typId x i) p) (transport (λ i → typId x i) q) i)) + ∙ cong (transport (λ i → typId x (~ i))) (sym (typIdHelper q p)) + ∙ transport⁻Transport (typId x) (q ∙ p) + where + congIsoHelper : (x : hLevelTrunc 4 (S₊ 2)) → Iso (hLevelTrunc 4 (S₊ 2)) (hLevelTrunc 4 (S₊ 2)) + Iso.fun (congIsoHelper x) = _+ₖ x + Iso.inv (congIsoHelper x) = _+ₖ (-ₖ x) + Iso.rightInv (congIsoHelper x) a = assocₖ a (-ₖ x) x ∙ cong (a +ₖ_) (lCancelₖ x) ∙ rUnitₖ a + Iso.leftInv (congIsoHelper x) a = assocₖ a x (-ₖ x) ∙ cong (a +ₖ_) (rCancelₖ x) ∙ rUnitₖ a + typId : (x : hLevelTrunc 4 (S₊ 2)) → (x ≡ x) ≡ Path (hLevelTrunc 4 (S₊ 2)) 0ₖ 0ₖ + typId x = isoToPath ((congIso (isoToEquiv (symIso (congIsoHelper x))))) ∙ λ i → rCancelₖ x i ≡ rCancelₖ x i + where + + typIdHelper : (p q : (x ≡ x)) → transport (λ i → typId x i) (p ∙ q) ≡ transport (λ i → typId x i) p ∙ transport (λ i → typId x i) q + typIdHelper p q = + (substComposite (λ x → x) (isoToPath ((congIso (isoToEquiv (symIso (congIsoHelper x)))))) (λ i → rCancelₖ x i ≡ rCancelₖ x i) (p ∙ q)) + ∙ (λ i → transport (λ i → rCancelₖ x i ≡ rCancelₖ x i) (transport (isoToPath (congIso (isoToEquiv (symIso (congIsoHelper x))))) (p ∙ q))) + ∙ (λ i → transport (λ i → rCancelₖ x i ≡ rCancelₖ x i) (transportRefl (congFunct (_+ₖ (-ₖ x)) p q i) i)) + ∙ (λ i → transport (λ i → rCancelₖ x i ≡ rCancelₖ x i) ((lUnit (rUnit (transportRefl (cong (_+ₖ (-ₖ x)) p) (~ i)) i) i) ∙ (lUnit (rUnit (transportRefl (cong (_+ₖ (-ₖ x)) q) (~ i)) i) i))) + ∙ (λ i → transp (λ j → rCancelₖ x (i ∨ j) ≡ rCancelₖ x (i ∨ j)) i (((λ j → rCancelₖ x (i ∧ (~ j))) ∙ transport refl (cong (_+ₖ (-ₖ x)) p) ∙ λ j → rCancelₖ x (i ∧ j)) ∙ (λ j → rCancelₖ x (i ∧ (~ j))) ∙ transport refl (cong (_+ₖ (-ₖ x)) q) ∙ λ j → rCancelₖ x (i ∧ j))) + ∙ (λ i → transp (λ j → rCancelₖ x ((~ i) ∨ j) ≡ rCancelₖ x ((~ i) ∨ j)) (~ i) ((λ j → rCancelₖ x ((~ i) ∧ (~ j))) ∙ transport refl (cong (_+ₖ (-ₖ x)) p) ∙ λ j → rCancelₖ x ((~ i) ∧ j)) + ∙ transp (λ j → rCancelₖ x ((~ i) ∨ j) ≡ rCancelₖ x ((~ i) ∨ j)) (~ i) ((λ j → rCancelₖ x ((~ i) ∧ (~ j))) ∙ transport refl (cong (_+ₖ (-ₖ x)) q) ∙ λ j → rCancelₖ x ((~ i) ∧ j))) + ∙ (λ i → transport (λ j → rCancelₖ x j ≡ rCancelₖ x j) (lUnit (rUnit (transport refl (cong (_+ₖ (-ₖ x)) p)) (~ i)) (~ i)) + ∙ transport (λ j → rCancelₖ x j ≡ rCancelₖ x j) (lUnit (rUnit (transport refl (cong (_+ₖ (-ₖ x)) q)) (~ i)) (~ i))) + ∙ cong₂ (_∙_) (sym (substComposite (λ x → x) (isoToPath ((congIso (isoToEquiv (symIso (congIsoHelper x)))))) (λ i → rCancelₖ x i ≡ rCancelₖ x i) p)) + (sym (substComposite (λ x → x) (isoToPath ((congIso (isoToEquiv (symIso (congIsoHelper x)))))) (λ i → rCancelₖ x i ≡ rCancelₖ x i) q)) + + + test2 : PathP (λ i → a north +ₖ ∣ north ∣ ≡ a (merid north i)) (rUnitₖ (a north)) (rUnitₖ (a north) ∙ cong a (merid north)) + test2 i j = + hcomp (λ k → λ { (i = i0) → rUnitₖ (a north) j + ; (j = i1) → a (merid north (i ∧ k)) + ; (j = i0) → a north +ₖ ∣ north ∣}) + (rUnitₖ (a north) j) + + +S1→S¹ : S₊ 1 → S¹ +S1→S¹ x = SuspBool→S¹ (S1→SuspBool x) + +S¹→S1 : S¹ → S₊ 1 +S¹→S1 x = SuspBool→S1 (S¹→SuspBool x) + +S1→S¹-sect : section S1→S¹ S¹→S1 +S1→S¹-sect x = + cong SuspBool→S¹ (SuspBool→S1-retr (S¹→SuspBool x)) + ∙ S¹→SuspBool→S¹ x + +S1→S¹-retr : retract S1→S¹ S¹→S1 +S1→S¹-retr x = + cong SuspBool→S1 (SuspBool→S¹→SuspBool (S1→SuspBool x)) + ∙ SuspBool→S1-sect x + + +prodElim : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : ∥ A ∥₀ × ∥ B ∥₀ → Type ℓ''} + → ((x : ∥ A ∥₀ × ∥ B ∥₀) → isOfHLevel 2 (C x)) + → ((a : A) (b : B) → C (∣ a ∣₀ , ∣ b ∣₀)) + → (x : ∥ A ∥₀ × ∥ B ∥₀) → C x +prodElim {A = A} {B = B} {C = C} hlevel ind (a , b) = schonf a b + where + schonf : (a : ∥ A ∥₀) (b : ∥ B ∥₀) → C (a , b) + schonf = sElim (λ x → isOfHLevelΠ 2 λ y → hlevel (_ , _)) λ a → sElim (λ x → hlevel (_ , _)) + λ b → ind a b + +setTruncOfProdIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → Iso ∥ A × B ∥₀ (∥ A ∥₀ × ∥ B ∥₀) +Iso.fun setTruncOfProdIso = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) λ { (a , b) → ∣ a ∣₀ , ∣ b ∣₀ } +Iso.inv setTruncOfProdIso = prodElim (λ _ → setTruncIsSet) λ a b → ∣ a , b ∣₀ +Iso.rightInv setTruncOfProdIso = prodElim (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ _ _ → refl +Iso.leftInv setTruncOfProdIso = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ {(a , b) → refl} + + +isGroupoidS1 : isGroupoid (S₊ 1) +isGroupoidS1 = transport (λ i → isGroupoid (S¹≡S1 (~ i))) isGroupoidS¹ +isConnectedS1 : (x : S₊ 1) → ∥ north ≡ x ∥₋₁ +isConnectedS1 x = pRec propTruncIsProp + (λ p → ∣ cong (transport (sym (S¹≡S1))) p ∙ transport⁻Transport (S¹≡S1) x ∣₋₁) + (isConnectedS¹ (transport S¹≡S1 x)) + + +open import Cubical.HITs.S2 + +test : (S₊ 2) → S₊ 1 +test north = north +test south = south +test (merid a i) = merid north i + +test2 : (S₊ 1) → (S₊ 2) +test2 north = north +test2 south = south +test2 (merid a i) = merid (merid a i) i + + +testS² : S² → S¹ +testS² base = base +testS² (surf i i₁) = base + +test4 : S₊ 1 → hLevelTrunc 4 (S₊ 2) +test4 north = ∣ north ∣ +test4 south = ∣ north ∣ +test4 (merid a i) = (Kn→ΩKn+1 1 ∣ south ∣) i + + + + +coHom0Torus : grIso (coHomGr 0 (S₊ 1 × S₊ 1)) intGroup +coHom0Torus = + Iso'→Iso + (iso' + (iso (sRec isSetInt (λ f → f (north , north))) + (λ a → ∣ (λ x → a) ∣₀) + (λ a → refl) + (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ + (funExt λ {(x , y) → toProdElimSuspElim2 + {B = λ x y → f (north , north) ≡ f (x , y)} + north + (λ _ _ → isSetInt _ _) + refl + x y}))) + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) λ a b → addLemma (a (north , north)) (b (north , north)))) + +private + S¹map : hLevelTrunc 3 S¹ → S¹ + S¹map = trElim (λ _ → isGroupoidS¹) (idfun S¹) + + S¹map-id : (x : hLevelTrunc 3 S¹) → Path (hLevelTrunc 3 S¹) ∣ S¹map x ∣ x + S¹map-id = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → refl + + S1map : hLevelTrunc 3 (S₊ 1) → (S₊ 1) + S1map = trElim (λ _ → isGroupoidS1) (idfun _) + + +S¹→S¹ : Iso (S¹ → hLevelTrunc 3 S¹) (S¹ × Int) +Iso.fun S¹→S¹ f = S¹map (f base) + , winding (basechange2⁻ (S¹map (f base)) λ i → S¹map (f (loop i))) +Iso.inv S¹→S¹ (s , int) base = ∣ s ∣ +Iso.inv S¹→S¹ (s , int) (loop i) = ∣ basechange2 s (intLoop int) i ∣ +Iso.rightInv S¹→S¹ (s , int) = ×≡ refl ((λ i → winding (basechange2-retr s (λ i → intLoop int i) i)) + ∙ windingIntLoop int) +Iso.leftInv S¹→S¹ f = funExt λ { base → S¹map-id (f base) + ; (loop i) j → helper2 j i} + where + helper2 : PathP (λ i → S¹map-id (f base) i ≡ S¹map-id (f base) i) + (λ i → ∣ basechange2 (S¹map (f base)) (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁)))))) i ∣) + (cong f loop) + helper2 i j = + hcomp (λ k → λ { (i = i0) → cong ∣_∣ (basechange2 (S¹map (f base)) (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁))))))) j + ; (i = i1) → S¹map-id (f (loop j)) k + ; (j = i0) → S¹map-id (f base) (i ∧ k) + ; (j = i1) → S¹map-id (f base) (i ∧ k)}) + (helper4 i j) + where + helper4 : Path (Path (hLevelTrunc 3 _) _ _) + (cong ∣_∣ (basechange2 (S¹map (f base)) + (intLoop + (winding + (basechange2⁻ (S¹map (f base)) + (λ i₁ → S¹map (f (loop i₁)))))))) + λ i → ∣ S¹map (f (loop i)) ∣ + helper4 i = + cong ∣_∣ + ((cong (basechange2 (S¹map (f base))) + (decodeEncode base (basechange2⁻ (S¹map (f base)) + (λ i₁ → S¹map (f (loop i₁))))) + ∙ basechange2-sect (S¹map (f base)) + (λ i → S¹map (f (loop i)))) i) + +S1→S1→S1×Int : Iso ((S₊ 1) → hLevelTrunc 3 (S₊ 1)) ((hLevelTrunc 3 (S₊ 1)) × Int) +S1→S1→S1×Int = compIso helper2 (compIso S¹→S¹ helper) + where + helper : Iso (S¹ × Int) (hLevelTrunc 3 (S₊ 1) × Int) + Iso.fun helper (s , int) = ∣ S¹→S1 s ∣ , int + Iso.inv helper (s , int) = (S1→S¹ (S1map s)) , int + Iso.rightInv helper (s , int) = + trElim {B = λ s → (∣ S¹→S1 (S1→S¹ (S1map s)) ∣ , int) ≡ (s , int)} + (λ _ → isOfHLevelPath 3 (isOfHLevelProd 3 (isOfHLevelTrunc 3) (isOfHLevelSuc 2 isSetInt)) _ _) + (λ a → ×≡ (cong ∣_∣ (S1→S¹-retr a)) refl) + s + Iso.leftInv helper (s , int) = ×≡ (S1→S¹-sect s) refl + + helper2 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → hLevelTrunc 3 S¹) + Iso.fun helper2 f x = trMap S1→S¹ (f (S¹→S1 x)) + Iso.inv helper2 f x = trMap S¹→S1 (f (S1→S¹ x)) + Iso.rightInv helper2 f = funExt λ x i → helper3 (f (S1→S¹-sect x i)) i + where + helper3 : (x : hLevelTrunc 3 S¹) → trMap S1→S¹ (trMap S¹→S1 x) ≡ x + helper3 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → cong ∣_∣ (S1→S¹-sect a) + Iso.leftInv helper2 f = funExt λ x i → helper3 (f (S1→S¹-retr x i)) i + where + helper3 : (x : hLevelTrunc 3 (S₊ 1)) → trMap S¹→S1 (trMap S1→S¹ x) ≡ x + helper3 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → cong ∣_∣ (S1→S¹-retr a) + +S¹→S¹≡S1→S1 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → hLevelTrunc 3 S¹) +Iso.fun S¹→S¹≡S1→S1 f x = trMap S1→S¹ (f (S¹→S1 x)) +Iso.inv S¹→S¹≡S1→S1 f x = trMap S¹→S1 (f (S1→S¹ x)) +Iso.rightInv S¹→S¹≡S1→S1 F = funExt λ x i → helper2 (F (S1→S¹-sect x i)) i + where + helper2 : (x : hLevelTrunc 3 S¹) → trMap S1→S¹ (trMap S¹→S1 x) ≡ x + helper2 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → cong ∣_∣ (S1→S¹-sect a) +Iso.leftInv S¹→S¹≡S1→S1 F = funExt λ x i → helper2 (F (S1→S¹-retr x i)) i + where + helper2 : (x : hLevelTrunc 3 (S₊ 1)) → trMap S¹→S1 (trMap S1→S¹ x) ≡ x + helper2 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → cong ∣_∣ (S1→S¹-retr a) + + + + +basechange-lemma : ∀ {ℓ} {A : Type ℓ} {a : A} (x y : S¹) (F : a ≡ a → S¹) (f : S¹ → a ≡ a) (g : S¹ → a ≡ a) + → (f base ≡ refl) + → (g base ≡ refl) + → basechange2⁻ (F (f base ∙ g base)) (cong₂ {A = S¹} {B = λ x → S¹} (λ x y → F (f x ∙ g y)) loop loop) + ≡ basechange2⁻ (F (f base)) (cong (λ x → F (f x)) loop) ∙ basechange2⁻ (F (g base)) (cong (λ x → F (g x)) loop) +basechange-lemma x y F f g frefl grefl = + (λ i → basechange2⁻ (F (f base ∙ g base)) (cong₂Funct2 (λ x y → F (f x ∙ g y)) loop loop i)) + ∙ (λ i → basechange2⁻ (F (f base ∙ g base)) (cong (λ x₁ → F (f x₁ ∙ g base)) loop ∙ cong (λ y₁ → F (f base ∙ g y₁)) loop)) + ∙ basechange2⁻-morph (F (f base ∙ g base)) _ _ + ∙ (λ j → basechange2⁻ (F (f base ∙ grefl j)) + (λ i → F (f (loop i) ∙ grefl j)) + ∙ basechange2⁻ (F (frefl j ∙ g base)) + (λ i → F (frefl j ∙ g (loop i)))) + ∙ ((λ j → basechange2⁻ (F (rUnit (f base) (~ j))) + (λ i → F (rUnit (f (loop i)) (~ j))) + ∙ basechange2⁻ (F (lUnit (g base) (~ j))) + (λ i → F (lUnit (g (loop i)) (~ j))))) + + +basechange-lemma2 : (f g : S¹ → hLevelTrunc 3 (S₊ 1)) (F : hLevelTrunc 3 (S₊ 1) → S¹) + → ((basechange2⁻ (F (f base +ₖ g base)) λ i → F ((f (loop i)) +ₖ g (loop i))) + ≡ basechange2⁻ (F (f base)) (cong (F ∘ f) loop) + ∙ basechange2⁻ (F (g base)) (cong (F ∘ g) loop)) +basechange-lemma2 f g F = coInd f g F (f base) (g base) refl refl + where + coInd : (f g : S¹ → hLevelTrunc 3 (S₊ 1)) (F : hLevelTrunc 3 (S₊ 1) → S¹) (x y : hLevelTrunc 3 (S₊ 1)) + → f base ≡ x + → g base ≡ y + → ((basechange2⁻ (F (f base +ₖ g base)) λ i → F ((f (loop i)) +ₖ g (loop i))) + ≡ basechange2⁻ (F (f base)) (cong (F ∘ f) loop) + ∙ basechange2⁻ (F (g base)) (cong (F ∘ g) loop)) + coInd f g F = + elim2 (λ _ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 (isGroupoidS¹ base base)) _ _ ) + (toProdElimSuspElim2 + north + (λ _ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → isGroupoidS¹ _ _ _ _) + λ fid gid → + (λ i → basechange2⁻ (F (f base +ₖ g base)) (cong₂Funct2 (λ x y → F (f x +ₖ g y)) loop loop i)) -- (λ i → F (f (loop i) +ₖ g (loop i)))) + ∙ (basechange2⁻-morph (F (f base +ₖ g base)) + (cong (λ x → F (f x +ₖ g base)) loop) + (cong (λ x → F (f base +ₖ g x)) loop)) + ∙ (λ i → basechange2⁻ (F (f base +ₖ gid i)) (cong (λ x → F (f x +ₖ gid i)) loop) + ∙ basechange2⁻ (F (fid i +ₖ g base)) (cong (λ x → F (fid i +ₖ g x)) loop)) + ∙ (λ i → basechange2⁻ (F (rUnitₖ (f base) i)) (cong (λ x → F (rUnitₖ (f x) i)) loop) + ∙ basechange2⁻ (F (lUnitₖ (g base) i)) (cong (λ x → F (lUnitₖ (g x) i)) loop))) + + + +coHom1S1≃ℤ : grIso (coHomGr 1 (S₊ 1)) intGroup +coHom1S1≃ℤ = + Iso'→Iso + (iso' + (iso + (sRec isSetInt (λ f → proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) + (λ a → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , a)) ∣₀) + (λ a → (λ i → proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 (Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , a)))))) + ∙ (λ i → proj₂ (Iso.fun S¹→S¹ (Iso.rightInv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , a)) i))) + ∙ λ i → proj₂ (Iso.rightInv S¹→S¹ (base , a) i)) + (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → (λ i → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) ∣₀) + ∙ (λ i → sRec setTruncIsSet + (λ x → ∣ Iso.inv S¹→S¹≡S1→S1 x ∣₀) + (sRec setTruncIsSet + (λ x → ∣ Iso.inv S¹→S¹ (x , (proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) ∣₀) + ∣ base ∣₀)) + ∙ (λ i → sRec setTruncIsSet + (λ x → ∣ Iso.inv S¹→S¹≡S1→S1 x ∣₀) + (sRec setTruncIsSet + (λ x → ∣ Iso.inv S¹→S¹ (x , (proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) ∣₀) + (Iso.inv PathIdTrunc₀Iso (isConnectedS¹ (proj₁ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) i))) + ∙ (λ i → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.leftInv S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f) i) ∣₀) + ∙ (λ i → ∣ Iso.leftInv S¹→S¹≡S1→S1 f i ∣₀))) + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) + λ f g → (λ i → winding (basechange2⁻ (S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)) ∙ Kn→ΩKn+1 1 (g (S¹→S1 base)))))) + (λ i → S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 (loop i))) ∙ Kn→ΩKn+1 1 (g (S¹→S1 (loop i))))))))) + ∙ cong winding (helper2 (f (S¹→S1 base)) (g (S¹→S1 base)) f g refl refl) + ∙ winding-hom ((basechange2⁻ (S¹map (trMap S1→S¹ (f north))) + (λ i → S¹map (trMap S1→S¹ (f (S¹→S1 (loop i))))))) + ((basechange2⁻ (S¹map (trMap S1→S¹ (g north))) + (λ i → S¹map (trMap S1→S¹ (g (S¹→S1 (loop i))))))))) + + + where + helper2 : (x y : hLevelTrunc 3 (S₊ 1)) (f g : S₊ 1 → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1) + → (f (S¹→S1 base)) ≡ x + → (g (S¹→S1 base)) ≡ y + → (basechange2⁻ (S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)) ∙ Kn→ΩKn+1 1 (g (S¹→S1 base))))))) + (λ i → S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 (loop i))) ∙ Kn→ΩKn+1 1 (g (S¹→S1 (loop i))))))) + ≡ (basechange2⁻ (S¹map (trMap S1→S¹ ((f (S¹→S1 base)))))) + (λ i → S¹map (trMap S1→S¹ (f (S¹→S1 (loop i))))) + ∙ (basechange2⁻ (S¹map (trMap S1→S¹ ((g (S¹→S1 base))))) + (λ i → S¹map (trMap S1→S¹ ((g (S¹→S1 (loop i))))))) + helper2 = elim2 + (λ _ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 + λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 + λ _ → isOfHLevelPath 3 (isOfHLevelSuc 3 (isGroupoidS¹) base base) _ _) + (toProdElimSuspElim2 {A = S₊ 0} north + (λ _ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → (isGroupoidS¹) _ _ _ _ ) + λ f g reflf reflg → + (basechange-lemma + base + base + (λ x → S¹map (trMap S1→S¹ (ΩKn+1→Kn x))) + (λ x → Kn→ΩKn+1 1 (f (S¹→S1 x))) ((λ x → Kn→ΩKn+1 1 (g (S¹→S1 x)))) + (cong (Kn→ΩKn+1 1) reflf ∙ Kn→ΩKn+10ₖ 1) + (cong (Kn→ΩKn+1 1) reflg ∙ Kn→ΩKn+10ₖ 1)) + ∙ λ j → basechange2⁻ (S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 base)) j))) + (λ i → S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 (loop i))) j))) + ∙ basechange2⁻ (S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 base)) j))) + (λ i → S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 (loop i))) j)))) + + + + + + + + + +pushoutSn : (n : ℕ) → Iso (S₊ (suc n)) (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt) +Iso.fun (pushoutSn n) north = inl tt +Iso.fun (pushoutSn n) south = inr tt +Iso.fun (pushoutSn n) (merid a i) = push a i +Iso.inv (pushoutSn n) (inl x) = north +Iso.inv (pushoutSn n) (inr x) = south +Iso.inv (pushoutSn n) (push a i) = merid a i +Iso.rightInv (pushoutSn n) (inl x) = refl +Iso.rightInv (pushoutSn n) (inr x) = refl +Iso.rightInv (pushoutSn n) (push a i) = refl +Iso.leftInv (pushoutSn n) north = refl +Iso.leftInv (pushoutSn n) south = refl +Iso.leftInv (pushoutSn n) (merid a i) = refl + +Sn≡Pushout : (n : ℕ) → (S₊ (suc n)) ≡ (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt) +Sn≡Pushout n = isoToPath (pushoutSn n) + +coHomPushout≡coHomSn' : (n m : ℕ) → grIso (coHomGr m (S₊ (suc n))) (coHomGr m (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt)) +morph.fun (grIso.fun (coHomPushout≡coHomSn' n m)) = + sRec setTruncIsSet + λ f → ∣ (λ {(inl x) → f north ; (inr x) → f south ; (push a i) → f (merid a i)}) ∣₀ +morph.ismorph (grIso.fun (coHomPushout≡coHomSn' n m)) = + sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ a b → cong ∣_∣₀ (funExt λ {(inl x) → refl ; (inr x) → refl ; (push a i) → refl }) +morph.fun (grIso.inv (coHomPushout≡coHomSn' n m)) = sRec setTruncIsSet (λ f → ∣ (λ {north → f (inl tt) ; south → f (inr tt) ; (merid a i) → f (push a i)}) ∣₀) +morph.ismorph (grIso.inv (coHomPushout≡coHomSn' n m)) = + sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ a b → cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → refl }) +grIso.rightInv (coHomPushout≡coHomSn' n m) = + sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ (funExt λ {(inl x) → refl ; (inr x) → refl ; (push a i) → refl }) +grIso.leftInv (coHomPushout≡coHomSn' n m) = + sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → refl }) + + +isContr→≡Unit : {A : Type₀} → isContr A → A ≡ Unit +isContr→≡Unit contr = isoToPath (iso (λ _ → tt) (λ _ → fst contr) (λ _ → refl) λ _ → snd contr _) + +isContr→isContrTrunc : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → isContr A → isContr (hLevelTrunc n A) +isContr→isContrTrunc n contr = ∣ fst contr ∣ , (trElim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → cong ∣_∣ (snd contr a)) +isContr→isContrSetTrunc : ∀ {ℓ} {A : Type ℓ} → isContr A → isContr (∥ A ∥₀) +isContr→isContrSetTrunc contr = ∣ fst contr ∣₀ , sElim (λ _ → isOfHLevelPath 2 (setTruncIsSet) _ _) λ a → cong ∣_∣₀ (snd contr a) + +coHomGrUnit0 : grIso (coHomGr 0 Unit) intGroup +grIso.fun coHomGrUnit0 = mph (sRec isSetInt (λ f → f tt)) + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) + (λ a b → addLemma (a tt) (b tt))) +grIso.inv coHomGrUnit0 = mph (λ a → ∣ (λ _ → a) ∣₀) (λ a b i → ∣ (λ _ → addLemma a b (~ i)) ∣₀) +grIso.rightInv coHomGrUnit0 a = refl +grIso.leftInv coHomGrUnit0 = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ a → refl + +isContrCohomUnit : (n : ℕ) → isContr (coHom (suc n) Unit) +isContrCohomUnit n = subst isContr (λ i → ∥ UnitToTypeId (coHomK (suc n)) (~ i) ∥₀) (helper n) + where + helper : (n : ℕ) → isContr (∥ coHomK (suc n) ∥₀) + helper n = subst isContr ((isoToPath (truncOfTruncIso {A = S₊ (1 + n)} 2 (1 + n))) ∙ sym propTrunc≡Trunc0 ∙ λ i → ∥ hLevelTrunc (suc (+-comm n 2 i)) (S₊ (1 + n)) ∥₀) + (isConnectedSubtr 2 (helper2 n .fst) (subst (λ x → isHLevelConnected x (S₊ (suc n))) (sym (helper2 n .snd)) (sphereConnected (suc n))) ) + where + helper2 : (n : ℕ) → Σ[ m ∈ ℕ ] m + 2 ≡ 2 + n + helper2 zero = 0 , refl + helper2 (suc n) = (suc n) , λ i → suc (+-comm n 2 i) + +coHomGrUnit≥1 : (n : ℕ) → grIso (coHomGr (suc n) Unit) trivialGroup +grIso.fun (coHomGrUnit≥1 n) = mph (λ _ → tt) (λ _ _ → refl) +grIso.inv (coHomGrUnit≥1 n) = mph (λ _ → ∣ (λ _ → ∣ north ∣) ∣₀) (λ _ _ → sym (rUnitₕ 0ₕ)) +grIso.rightInv (coHomGrUnit≥1 n) a = refl +grIso.leftInv (coHomGrUnit≥1 n) a = sym (isContrCohomUnit n .snd 0ₕ) ∙ isContrCohomUnit n .snd a + +S0→Int : (a : Int × Int) → S₊ 0 → Int +S0→Int a north = proj₁ a +S0→Int a south = proj₂ a + +coHom0-S0 : grIso (coHomGr 0 (S₊ 0)) (dirProd intGroup intGroup) +coHom0-S0 = + Iso'→Iso + (iso' + (iso (sRec (isOfHLevelProd 2 isSetInt isSetInt) + λ f → (f north) , (f south)) + (λ a → ∣ S0→Int a ∣₀) + (λ { (a , b) → refl}) + (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ f → cong ∣_∣₀ (funExt (λ {north → refl ; south → refl})))) + (sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 isSetInt isSetInt) _ _) + λ a b i → addLemma (a north) (b north) i , addLemma (a south) (b south) i)) + +×morph : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group ℓ} {B : Group ℓ'} {C : Group ℓ''} {D : Group ℓ'''} → morph A B → morph C D → morph (dirProd A C) (dirProd B D) +morph.fun (×morph mf1 mf2) = + (λ {(a , b) → (morph.fun mf1 a) , morph.fun mf2 b}) +morph.ismorph (×morph mf1 mf2) = + (λ {(a , b) (c , d) i → morph.ismorph mf1 a c i , morph.ismorph mf2 b d i}) + + +coHom0S1 : grIso intGroup (coHomGr 0 (S₊ 1)) +coHom0S1 = + Iso'→Iso + (iso' + (iso + (λ a → ∣ (λ x → a) ∣₀) + (sRec isSetInt (λ f → f north)) + (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ (funExt (toProdElimSuspRec north (λ _ → isSetInt _ _) refl))) + (λ a → refl)) + λ a b → cong ∣_∣₀ (funExt λ x → sym (addLemma a b))) + +coHom1S1 : grIso intGroup (coHomGr 1 (S₊ 1)) +coHom1S1 = + compGrIso + (diagonalIso1 + _ + (coHomGr 0 (S₊ 0)) + _ + (invGrIso coHom0-S0) + (d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) + (λ x → MV.Ker-i⊂Im-d _ _(S₊ 0) (λ _ → tt) (λ _ → tt) 0 x + (×≡ (isOfHLevelSuc 0 (isContrCohomUnit 0) _ _) + (isOfHLevelSuc 0 (isContrCohomUnit 0) _ _))) + (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) + (λ x inker + → pRec propTruncIsProp + (λ {((f , g) , id') → helper x f g id' inker}) + ((MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ inker)))) + (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ F surj + → pRec (setTruncIsSet _ _) (λ { (x , id) → MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ F ∣₀ ∣ (∣ (λ _ → x) ∣₀ , ∣ (λ _ → 0) ∣₀) , + (cong ∣_∣₀ (funExt (surjHelper x))) ∙ sym id ∣₋₁ }) surj) ) + (invGrIso (coHomPushout≡coHomSn' 0 1)) + + where + surjHelper : (x : Int) (x₁ : S₊ 0) → x +ₖ (-ₖ pos 0) ≡ S0→Int (x , x) x₁ + surjHelper x north = cong (x +ₖ_) (-0ₖ) ∙ rUnitₖ x + surjHelper x south = cong (x +ₖ_) (-0ₖ) ∙ rUnitₖ x + + helper : (F : S₊ 0 → Int) (f g : ∥ (Unit → Int) ∥₀) (id : MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (f , g) ≡ ∣ F ∣₀) + → isInKer (coHomGr 0 (S₊ 0)) + (coHomGr 1 (Pushout (λ _ → tt) (λ _ → tt))) + (morph.fun (d-morph Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) + ∣ F ∣₀ + → ∃[ x ∈ Int ] ∣ F ∣₀ ≡ morph.fun (grIso.inv coHom0-S0) (x , x) + helper F = + sElim2 (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ f g id inker + → pRec propTruncIsProp + (λ {((a , b) , id2) + → sElim2 {B = λ f g → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (f , g) ≡ ∣ F ∣₀ → _ } + (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) + (λ f g id → ∣ (helper2 f g .fst) , (sym id ∙ sym (helper2 f g .snd)) ∣₋₁) + a b id2}) + (MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ F ∣₀ inker) + where + helper2 : (f g : Unit → Int) + → Σ[ x ∈ Int ] morph.fun (grIso.inv coHom0-S0) (x , x) + ≡ MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (∣ f ∣₀ , ∣ g ∣₀) + helper2 f g = (f _ +ₖ (-ₖ g _) ) , cong ∣_∣₀ (funExt (λ {north → refl ; south → refl})) + + + +coHom-n-Sn : (n : ℕ) → grIso intGroup (coHomGr (suc n) (S₊ (suc n))) +coHom-n-Sn zero = coHom1S1 +coHom-n-Sn (suc n) = + compGrIso + (compGrIso + (coHom-n-Sn n) + theIso) + (invGrIso (coHomPushout≡coHomSn' (suc n) (suc (suc n)))) + where + theIso : grIso (coHomGr (suc n) (S₊ (suc n))) (coHomGr (suc (suc n)) + (Pushout {A = S₊ (suc n)} (λ _ → tt) (λ _ → tt))) + theIso = + SES→Iso + (×coHomGr (suc n) Unit Unit) + (×coHomGr (suc (suc n)) Unit Unit) + (ses (λ p q → ×≡ (isOfHLevelSuc 0 (isContrCohomUnit n) (proj₁ p) (proj₁ q)) + (isOfHLevelSuc 0 (isContrCohomUnit n) (proj₂ p) (proj₂ q))) + (λ p q → ×≡ (isOfHLevelSuc 0 (isContrCohomUnit (suc n)) (proj₁ p) (proj₁ q)) + (isOfHLevelSuc 0 (isContrCohomUnit (suc n)) (proj₂ p) (proj₂ q))) + (Δ-morph _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) + (i-morph _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (2 + n)) + (d-morph _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) + (MV.Ker-d⊂Im-Δ _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) + (MV.Ker-i⊂Im-d _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n))) + + +coHom1-s0 : grIso (coHomGr 1 (S₊ 0)) trivialGroup +coHom1-s0 = + Iso'→Iso + (iso' + (iso (λ _ → tt) + (λ _ → 0ₕ) + (λ _ → refl) + (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → pRec (setTruncIsSet _ _) + (λ id → cong ∣_∣₀ (sym id)) + (helper f (f north) (f south) refl refl))) + λ _ _ → refl) + where + helper : (f : S₊ 0 → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1) → (x y : ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1) → (f north ≡ x) → (f south ≡ y) → ∥ f ≡ (λ _ → 0ₖ) ∥₋₁ + helper f = + elim2 (λ _ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelPlus {n = 1} 2 propTruncIsProp) + (toProdElimSuspElim2 north (λ _ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → propTruncIsProp) + λ id id2 → ∣ funExt (λ {north → id ; south → id2}) ∣₋₁) + +coHom2-S1 : grIso (coHomGr 2 (S₊ 1)) trivialGroup +coHom2-S1 = + compGrIso + (coHomPushout≡coHomSn' 0 2) + (compGrIso + (invGrIso + (SES→Iso + (×coHomGr 1 Unit Unit) + (×coHomGr 2 Unit Unit) + (ses (isOfHLevelProd 1 (isOfHLevelSuc 0 (isContrCohomUnit 0)) (isOfHLevelSuc 0 (isContrCohomUnit 0))) + (isOfHLevelProd 1 (isOfHLevelSuc 0 (isContrCohomUnit 1)) (isOfHLevelSuc 0 (isContrCohomUnit 1))) + (Δ-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1) + (i-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 2) + (d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1) + (MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1) + ({! MV.Ker-i⊂Im-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1 !})))) + coHom1-s0) + +setTruncIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → Iso A B → Iso ∥ A ∥₀ ∥ B ∥₀ +Iso.fun (setTruncIso is) = sRec setTruncIsSet (λ x → ∣ Iso.fun is x ∣₀) +Iso.inv (setTruncIso is) = sRec setTruncIsSet (λ x → ∣ Iso.inv is x ∣₀) +Iso.rightInv (setTruncIso is) = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ a → cong ∣_∣₀ (Iso.rightInv is a) +Iso.leftInv (setTruncIso is) = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ a → cong ∣_∣₀ (Iso.leftInv is a) + +targetIso : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} → Iso B C → Iso (A → B) (A → C) +Iso.fun (targetIso is) f a = Iso.fun is (f a) +Iso.inv (targetIso is) f a = Iso.inv is (f a) +Iso.rightInv (targetIso is) f = funExt λ a → Iso.rightInv is (f a) +Iso.leftInv (targetIso is) f = funExt λ a → Iso.leftInv is (f a) + +prodIso : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} + → Iso A C + → Iso B D + → Iso (A × B) (C × D) +Iso.fun (prodIso iAC iBD) (a , b) = (Iso.fun iAC a) , Iso.fun iBD b +Iso.inv (prodIso iAC iBD) (c , d) = (Iso.inv iAC c) , Iso.inv iBD d +Iso.rightInv (prodIso iAC iBD) (c , d) = ×≡ (Iso.rightInv iAC c) (Iso.rightInv iBD d) +Iso.leftInv (prodIso iAC iBD) (a , b) = ×≡ (Iso.leftInv iAC a) (Iso.leftInv iBD b) + + +coHom1Torus : grIso (coHomGr 1 ((S₊ 1) × (S₊ 1))) (dirProd intGroup intGroup) +coHom1Torus = + compGrIso + (Iso'→Iso + (iso' (compIso + (setTruncIso (compIso + schonfIso + (compIso + (targetIso S1→S1→S1×Int) + funIso))) + (setTruncOfProdIso)) + (sElim2 + (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ f g → ×≡ (cong ∣_∣₀ + (funExt (λ x → helper (f (x , S¹→S1 base) +ₖ g (x , S¹→S1 base)) + ∙ sym (cong₂ (λ x y → x +ₖ y) + (helper (f (x , S¹→S1 base))) + (helper (g (x , S¹→S1 base))))))) + (cong ∣_∣₀ + (funExt + (toProdElimSuspRec + north + (λ _ → isSetInt _ _) + (cong winding + (basechange-lemma2 + (λ x → f (north , S¹→S1 x)) + (λ x → g (north , S¹→S1 x)) + λ x → S¹map (trMap S1→S¹ x)) + ∙ winding-hom + (basechange2⁻ + (S¹map (trMap S1→S¹ (f (north , S¹→S1 base)))) + (λ i → S¹map (trMap S1→S¹ (f (north , S¹→S1 (loop i)))))) + (basechange2⁻ + (S¹map (trMap S1→S¹ (g (north , S¹→S1 base)))) + (λ i → S¹map (trMap S1→S¹ (g (north , S¹→S1 (loop i)))))) + ∙ sym (addLemma + (winding + (basechange2⁻ + (S¹map (trMap S1→S¹ (f (north , S¹→S1 base)))) + (λ i → S¹map (trMap S1→S¹ (f (north , S¹→S1 (loop i))))))) + (winding + (basechange2⁻ + (S¹map (trMap S1→S¹ (g (north , S¹→S1 base)))) + (λ i → S¹map (trMap S1→S¹ (g (north , S¹→S1 (loop i))))))))))))))) + (dirProdIso (invGrIso (coHom-n-Sn 0)) (invGrIso coHom0S1)) + + where + helper : (x : hLevelTrunc 3 (S₊ 1)) → ∣ S¹→S1 (S¹map (trMap S1→S¹ x)) ∣ ≡ x + helper = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ a → cong ∣_∣ (S1→S¹-retr a) + + +coHom2Torus : grIso (coHomGr 2 ((S₊ 1) × (S₊ 1))) intGroup +coHom2Torus = + invGrIso + (Iso'→Iso + (iso' + (compIso + helper' + (compIso + (symIso (prodIso (groupIso→Iso coHom2-S1) + (groupIso→Iso (invGrIso (coHom-n-Sn 0))))) + (compIso + (symIso setTruncOfProdIso) + (symIso + (setTruncIso + (compIso + schonfIso + (compIso + (targetIso test13) + funIso))))))) + λ a b → {!!} {- (λ i → Iso.fun + (symIso + (setTruncIso + (compIso schonfIso (compIso (targetIso test13) funIso)))) + (Iso.fun (symIso setTruncOfProdIso) + (helper'' (Iso.inv (groupIso→Iso coHom2-S1) tt) 0ₕ i , morph.ismorph (grIso.inv (invGrIso coHom1S1)) a b i))) + ∙ (λ i → Iso.fun + (symIso + (setTruncIso + (compIso schonfIso (compIso (targetIso test13) funIso)))) + ∣ (λ _ → ∣ north ∣) , {!!} ∣₀) + ∙ {!funIso!} -})) + where + helper'' : isProp ∥ (S₊ 1 → hLevelTrunc 4 (S₊ 2)) ∥₀ + helper'' = {!!} + + helper2' : (f g : (S₊ 1 → hLevelTrunc 3 (S₊ 1))) → + Path (coHom 2 ((S₊ 1) × (S₊ 1))) (Iso.fun + (symIso + (setTruncIso + (compIso schonfIso (compIso (targetIso test13) funIso)))) + (Iso.fun (symIso setTruncOfProdIso) (0ₕ , ∣ f ∣₀ +ₕ ∣ g ∣₀))) + (Iso.fun (symIso + (setTruncIso + (compIso schonfIso (compIso (targetIso test13) funIso)))) + (Iso.fun (symIso setTruncOfProdIso) (0ₕ , ∣ f ∣₀)) +ₕ Iso.fun (symIso + (setTruncIso + (compIso schonfIso (compIso (targetIso test13) funIso)))) + (Iso.fun (symIso setTruncOfProdIso) (0ₕ , ∣ g ∣₀))) + helper2' = {!!} + + helper' : Iso Int (Unit × Int) + Iso.inv helper' = proj₂ + Iso.fun helper' x = tt , x + Iso.leftInv helper' x = refl + Iso.rightInv helper' x = ×≡ refl refl + + helper : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} → Iso C Unit → Iso A B → Iso (A × C) B + Iso.fun (helper cUnit iAB) x = Iso.fun iAB (proj₁ x) + Iso.inv (helper cUnit iAB) x = (Iso.inv iAB x) , (Iso.inv cUnit tt ) + Iso.rightInv (helper cUnit iAB) = Iso.rightInv iAB + Iso.leftInv (helper cUnit iAB) b = ×≡ (Iso.leftInv iAB (proj₁ b)) (Iso.leftInv cUnit (proj₂ b)) + + helper2 : Iso (coHom 2 ((S₊ 1) × (S₊ 1))) (coHom 2 ((S₊ 1) × hLevelTrunc 3 (S₊ 1))) + Iso.fun helper2 = sRec setTruncIsSet (λ f → ∣ (λ {(x , y) → f (x , trRec isGroupoidS1 (idfun (S₊ 1)) y)}) ∣₀) + Iso.inv helper2 = sRec setTruncIsSet (λ f → ∣ (λ {(x , y) → f (x , ∣ y ∣)}) ∣₀) + Iso.rightInv helper2 = + sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + (λ f → cong ∣_∣₀ + (funExt λ {(x , y) → + trElim {B = λ y → f (x , ∣ trRec isGroupoidS1 (λ x₁ → x₁) y ∣) ≡ f (x , y)} + (λ _ → isOfHLevelPath' 3 (isOfHLevelTrunc 4) _ _) + (λ a → refl) y})) + Iso.leftInv helper2 = + sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ (funExt λ {(x , y) → refl}) + +-- basechange* : (x y : S¹) → x ≡ y → x ≡ y → ΩS¹ +-- basechange* x y = J (λ y p → (x ≡ y) → ΩS¹) (basechange x) + + +-- test1 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S₊ 1 × Int) +-- Iso.fun test1 f = (trRec isGroupoidS1 (λ a → a) (f north)) +-- , winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (f (loop* i))))) +-- Iso.inv test1 (north , b) x = ∣ x ∣ +-- Iso.inv test1 (south , b) x = ∣ x ∣ +-- Iso.inv test1 (merid a i , b) x = {!!} +-- Iso.rightInv test1 = {!!} +-- Iso.leftInv test1 = {!!} + +-- funRec : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) {C : (A → hLevelTrunc n B) → Type ℓ''} +-- → isOfHLevel n B +-- → ((f : A → B) → C (λ a → ∣ f a ∣)) +-- → (f : A → hLevelTrunc n B) → C f +-- funRec {A = A} {B = B} n {C = C} hLev ind f = subst C (helper f) (ind (λ a → trRec hLev (λ x → x) (f a))) +-- where +-- helper : retract {A = A → hLevelTrunc n B} {B = A → B} (λ f₁ a → trRec hLev (λ x → x) (f₁ a)) (λ f₁ a → ∣ f₁ a ∣) +-- helper f = funExt λ a → helper2 (f a) +-- where +-- helper2 : (x : hLevelTrunc n B) → ∣ trRec hLev (λ x → x) x ∣ ≡ x +-- helper2 = trElim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → refl + +-- invMapSurj : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (ϕ : morph G H) → ((x : Group.type H) → isInIm G H (fst ϕ) x) +-- → morph H G +-- fst (invMapSurj G H (ϕ , pf) surj) a = {!pRec!} +-- snd (invMapSurj G H (ϕ , pf) surj) = {!!} + +{- +ImIso : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (ϕ : morph G H) → ((x : Group.type H) → isInIm G H (fst ϕ) x) + → grIso H (imGroup G H ϕ) +ImIso G H (ϕ , mf) surj = + let idH = isGroup.id (Group.groupStruc H) + idG = isGroup.id (Group.groupStruc G) + _+G_ = isGroup.comp (Group.groupStruc G) + _+H_ = isGroup.comp (Group.groupStruc H) + _+Im_ = isGroup.comp (Group.groupStruc (imGroup G H (ϕ , mf))) + invG = isGroup.inv (Group.groupStruc G) + invH = isGroup.inv (Group.groupStruc H) + lUnit = isGroup.lUnit (Group.groupStruc H) + lCancel = isGroup.rCancel (Group.groupStruc H) + in + Iso''→Iso _ _ + (iso'' ((λ x → x , pRec propTruncIsProp (λ (a , b) → ∣ a , b ∣₋₁) (surj x)) + , λ a b → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) + (λ surja → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) + (λ surjb → + pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) + (λ surja+b → + (λ i → (a +H b) , (pRec (propTruncIsProp) + (λ (a , b) → ∣ a , b ∣₋₁) + (propTruncIsProp (surj (isGroup.comp (Group.groupStruc H) a b)) ∣ surja+b ∣₋₁ i))) ∙ + (λ i → (a +H b) , ∣ (fst surja+b) , (snd surja+b) ∣₋₁) ∙ + ΣProp≡ (λ _ → propTruncIsProp) refl ∙ + λ i → (a +H b) , pRec (propTruncIsProp) + (λ p1 → + pRec (λ x y → squash x y) + (λ p2 → + ∣ + isGroup.comp (Group.groupStruc G) (fst p1) (fst p2) , + mf (fst p1) (fst p2) ∙ + cong₂ (isGroup.comp (Group.groupStruc H)) (snd p1) (snd p2) + ∣₋₁) + (pRec (propTruncIsProp) + ∣_∣₋₁ (propTruncIsProp ∣ surjb ∣₋₁ (surj b) i))) + (pRec (propTruncIsProp) + ∣_∣₋₁ (propTruncIsProp ∣ surja ∣₋₁ (surj a) i ))) + (surj (isGroup.comp (Group.groupStruc H) a b))) + (surj b)) + (surj a)) + (λ x inker → cong fst inker) + λ x → pRec propTruncIsProp (λ inimx → ∣ (ϕ (inimx .fst)) , ΣProp≡ (λ _ → propTruncIsProp) (inimx .snd) ∣₋₁) (snd x)) +-} + + +{- +H¹-S¹≃Int : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +H¹-S¹≃Int = + Iso''→Iso _ _ + (iso'' ((λ x → ∣ theFuns x ∣₀) , λ a b → cong ∣_∣₀ (funExt λ x → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 1) _) ∙ sym (cong (ΩKn+1→Kn) (theFunsId2 a b x)))) + (λ x inker → pRec (isSetInt _ _) (inj x) (Iso.fun PathIdTrunc₀Iso inker)) + finIm) + where + d : _ + d = MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 + + i : _ + i = MV.i _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1 + + Δ : _ + Δ = MV.Δ _ _ (S₊ 1) (λ _ → tt) (λ _ → tt) 0 + + + d-surj : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) + → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x + d-surj = sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ x → MV.Ker-i⊂Im-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ + (sym (isContrHelper .snd _)) + where + isContrHelper : isContr (Group.type (×coHomGr 1 Unit Unit)) + isContrHelper = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) + , λ y → prodId (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₁ y)) + (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₂ y)) + + theFuns : (a : Int) → Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 + theFuns a (inl x) = 0ₖ + theFuns a (inr x) = 0ₖ + theFuns a (push north i) = Kn→ΩKn+1 0 a i + theFuns a (push south i) = 0ₖ + + + theFunsId2 : (a b : Int) (x : Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) + → Kn→ΩKn+1 1 (theFuns a x) ∙ Kn→ΩKn+1 1 (theFuns b x) ≡ Kn→ΩKn+1 1 (theFuns (a +ℤ b) x) + theFunsId2 a b (inl x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) + theFunsId2 a b (inr x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) + theFunsId2 a b (push north i) j = + hcomp (λ k → λ {(i = i0) → ((λ i₁ → + (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) + j + ; (i = i1) → ((λ i₁ → + (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) + j + ; (j = i0) → cong₂Funct2 (λ p q → Kn→ΩKn+1 1 p ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b) (~ k) i + ; (j = i1) → (helper2 a b) k i }) + (hcomp (λ k → λ { (j = i0) → compPath-filler (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) k i + ; (j = i1) → compPath-filler (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) k i + ; (i = i1) → RHS-filler b j k + ; (i = i0) → ((λ i₁ → + (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) + j}) + (bottom-filler a j i)) --- S1→S¹-retr : retract S1→S¹ S¹→S1 --- S1→S¹-retr x = --- cong SuspBool→S1 (SuspBool→S¹→SuspBool (S1→SuspBool x)) --- ∙ SuspBool→S1-sect x + where + bottom-filler : (a : Int) → + PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) + (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) + ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) + j ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) + (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) + ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) + j) (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) + bottom-filler a j i = + hcomp (λ k → λ {(j = i0) → helper2 (~ k) i ; + (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 a) i}) + ((λ j₂ → ∣ rCancel (merid north) j j₂ ∣) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + + where + helper2 : Path (Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣ ≡ Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣) + (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i) ∙ Kn→ΩKn+1 1 ∣ north ∣) + (λ i → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + helper2 = cong₂Sym1 (Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) (~ i) j ∣) (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + + RHS-filler : (b : Int) → + PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j + ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j) + (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) + (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) + RHS-filler b j i = + hcomp (λ k → λ {(j = i0) → cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i ; + (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 b) i}) + (cong (λ q → (λ i → ∣ rCancel (merid north) j i ∣) ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i) + + helper2 : (a b : Int) → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b) ≡ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ℤ b)) + helper2 a b = + sym (congFunct (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b)) + ∙ (λ i → cong (Kn→ΩKn+1 1) (Iso.rightInv (Iso3-Kn-ΩKn+1 0) (Kn→ΩKn+1 0 a ∙ Kn→ΩKn+1 0 b) (~ i))) + ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ₖ b)) ) + ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (addLemma a b i))) + + theFunsId2 a b (push south i) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ sym (lUnit _) + + inj : (a : Int) → theFuns a ≡ (λ _ → ∣ north ∣) → a ≡ pos 0 + inj a id = + pRec (isSetInt _ _) + (sigmaElim (λ _ → isOfHLevelPath 2 isSetInt _ _) + (λ a p → pRec (isSetInt _ _) + (λ id2 → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 0) _) + ∙ cong (ΩKn+1→Kn) + (PathP→compPathR + (cong (λ f → cong f (push north)) id) + ∙ test)) + (Iso.fun PathIdTrunc₀Iso p))) (d-surj ∣ theFuns a ∣₀) + where --- prodElim : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : ∥ A ∥₀ × ∥ B ∥₀ → Type ℓ''} --- → ((x : ∥ A ∥₀ × ∥ B ∥₀) → isOfHLevel 2 (C x)) --- → ((a : A) (b : B) → C (∣ a ∣₀ , ∣ b ∣₀)) --- → (x : ∥ A ∥₀ × ∥ B ∥₀) → C x --- prodElim {A = A} {B = B} {C = C} hlevel ind (a , b) = schonf a b + test : (λ i → id i (inl tt)) ∙ (λ i → ∣ north ∣) ∙ sym (λ i → id i (inr tt)) ≡ refl + test = (λ i → cong (λ f → f (inl tt)) id + ∙ lUnit (sym (cong (λ f → f (inr tt)) id)) (~ i)) + ∙ (λ i → cong (λ f → f (push south i)) id + ∙ sym (cong (λ f → f (inr tt)) id)) + ∙ rCancel (cong (λ f → f (inr tt)) id) + + + consMember : (a : Int) → coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) + consMember a = d ∣ (λ _ → a) ∣₀ + + consMember≡0 : (a : Int) → consMember a ≡ 0ₕ + consMember≡0 a = + (MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ (λ _ → a) ∣₀ ∣ + (∣ (λ _ → a) ∣₀ , ∣ (λ _ → 0) ∣₀) + , cong ∣_∣₀ (λ i x → (rUnitₖ a i)) ∣₋₁) + + f+consMember' : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int × Int ] (f +ₕ (-ₕ (consMember (proj₁ x))) ≡ ∣ theFuns (proj₂ x) ∣₀) + f+consMember' = + sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ f → pRec propTruncIsProp + (sigmaElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) + (λ g id → ∣ ((g south) , ((g north) +ₖ (-ₖ g south))) + , (pRec (setTruncIsSet _ _) + (λ id → (λ i → ∣ id (~ i) ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀) ∙ funId1 g) + (Iso.fun PathIdTrunc₀Iso id)) ∣₋₁)) + (d-surj ∣ f ∣₀) + where + funId1 : (g : S₊ 0 → Int) + → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀ ≡ + ∣ theFuns ((g north) +ₖ (-ₖ (g south))) ∣₀ + funId1 g = (λ i → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ + +ₕ (morphMinus (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) d + (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ∣ (λ _ → g south) ∣₀ (~ i))) + ∙ sym (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ g ∣₀ (-ₕ ∣ (λ _ → g south) ∣₀)) + ∙ (cong (λ x → d ∣ x ∣₀) g'Id) + ∙ cong ∣_∣₀ helper + where + g' : S₊ 0 → Int + g' north = (g north) +ₖ (-ₖ (g south)) + g' south = 0 + + g'Id : (λ x → g x +ₖ (-ₖ (g south))) ≡ g' + g'Id = funExt (λ {north → refl + ; south → rCancelₖ (g south)}) + + helper : MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g' ≡ theFuns (g north +ₖ (-ₖ g south)) + helper = funExt λ {(inl tt) → refl + ; (inr tt) → refl + ; (push north i) → refl + ; (push south i) → refl} + finIm : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int ] (∣ theFuns x ∣₀ ≡ f) + finIm f = + pRec propTruncIsProp + (λ {((a , b) , id) → ∣ b , (sym id ∙ cong (λ x → f +ₕ x) ((λ i → (-ₕ (consMember≡0 a i))) ∙ sym (lUnitₕ (-ₕ 0ₕ)) ∙ rCancelₕ 0ₕ) ∙ (rUnitₕ f)) ∣₋₁}) + (f+consMember' f) +-} + + + +-- Hⁿ-Sⁿ≃Int : (n : ℕ) → grIso intGroup (coHomGr (suc n) (S₊ (suc n))) +-- Hⁿ-Sⁿ≃Int zero = +-- compGrIso {F = intGroup} {G = {!!}} {H = {!coHomGr 1 (S₊ 1)!}} +-- (Iso''→Iso +-- (iso'' ((λ x → ∣ theFuns x ∣₀) , λ a b → cong ∣_∣₀ (funExt λ x → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 1) _) ∙ sym (cong (ΩKn+1→Kn) (theFunsId2 a b x)))) +-- (λ x inker → pRec (isSetInt _ _) (inj x) (Iso.fun PathIdTrunc₀Iso inker)) +-- finIm)) +-- {!invGrIso _ _ (coHomPushout≡coHomSn 0 1)!} -- where --- schonf : (a : ∥ A ∥₀) (b : ∥ B ∥₀) → C (a , b) --- schonf = sElim (λ x → isOfHLevelΠ 2 λ y → hlevel (_ , _)) λ a → sElim (λ x → hlevel (_ , _)) --- λ b → ind a b - --- setTruncOfProdIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → Iso ∥ A × B ∥₀ (∥ A ∥₀ × ∥ B ∥₀) --- Iso.fun setTruncOfProdIso = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) λ { (a , b) → ∣ a ∣₀ , ∣ b ∣₀ } --- Iso.inv setTruncOfProdIso = prodElim (λ _ → setTruncIsSet) λ a b → ∣ a , b ∣₀ --- Iso.rightInv setTruncOfProdIso = prodElim (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) --- λ _ _ → refl --- Iso.leftInv setTruncOfProdIso = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) --- λ {(a , b) → refl} - - --- isGroupoidS1 : isGroupoid (S₊ 1) --- isGroupoidS1 = transport (λ i → isGroupoid (S¹≡S1 (~ i))) isGroupoidS¹ --- isConnectedS1 : (x : S₊ 1) → ∥ north ≡ x ∥₋₁ --- isConnectedS1 x = pRec propTruncIsProp --- (λ p → ∣ cong (transport (sym (S¹≡S1))) p ∙ transport⁻Transport (S¹≡S1) x ∣₋₁) --- (isConnectedS¹ (transport S¹≡S1 x)) - - --- open import Cubical.HITs.S2 - --- test : (S₊ 2) → S₊ 1 --- test north = north --- test south = south --- test (merid a i) = merid north i - --- test2 : (S₊ 1) → (S₊ 2) --- test2 north = north --- test2 south = south --- test2 (merid a i) = merid (merid a i) i - --- S1test : ∀ {ℓ} {A : S¹ → Type ℓ} → (Abase : A base) → subst A loop Abase ≡ Abase → (x : S¹) → A x --- S1test = {!!} - --- testS² : S² → S¹ --- testS² base = base --- testS² (surf i i₁) = base - --- test4 : S₊ 1 → hLevelTrunc 4 (S₊ 2) --- test4 north = ∣ north ∣ --- test4 south = ∣ north ∣ --- test4 (merid a i) = (Kn→ΩKn+1 1 ∣ south ∣) i - --- test3 : hLevelTrunc 4 (S₊ 2) → S₊ 1 --- test3 = --- trElim (λ _ → {!!}) --- λ {north → north ; south → north ; (merid a i) → loop* i} - --- testIso2 : Iso ((S₊ 1) → hLevelTrunc 4 (S₊ 2)) ((S₊ 1) × hLevelTrunc 4 (S₊ 2)) --- Iso.fun testIso2 f = (test3 (f north)) , trElim (λ _ → isOfHLevelTrunc 4) (λ s → ΩKn+1→Kn (cong ∣_∣ (merid s ∙ sym (merid north)))) (f north) --- Iso.inv testIso2 (x , p) y = (test4 x) +ₖ (test4 y) +ₖ p --- Iso.rightInv testIso2 (s , s2) = trElim {B = λ s2 → Iso.fun testIso2 (Iso.inv testIso2 (s , s2)) ≡ (s , s2)} --- {!!} --- (λ s3 → ×≡ {!!} {!!}) --- s2 --- Iso.leftInv testIso2 a = funExt λ x → {!!} - --- testIso : Iso (S¹ → hLevelTrunc 4 S²) (S¹ × hLevelTrunc 4 S²) --- Iso.fun testIso f = {!(f base)!} , {!!} -- trElim (λ _ → isOfHLevelSuc 3 (isGroupoidS1)) test (f north) , (f north) --- Iso.inv testIso (s , tr) x = tr --- Iso.rightInv testIso (x , tr) = {!!} --- Iso.leftInv testIso a = funExt (S1test (cong a loop) {!!}) -- funExt (toPropElim {B = λ x → a base ≡ a x} {!a base!} refl) - - - - --- coHom0Torus : grIso (coHomGr 0 (S₊ 1 × S₊ 1)) intGroup --- coHom0Torus = --- Iso'→Iso --- (iso' --- (iso (sRec isSetInt (λ f → f (north , north))) --- (λ a → ∣ (λ x → a) ∣₀) --- (λ a → refl) --- (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) --- λ f → cong ∣_∣₀ --- (funExt λ {(x , y) → toProdElimSuspElim2 --- {B = λ x y → f (north , north) ≡ f (x , y)} --- north --- (λ _ _ → isSetInt _ _) --- refl --- x y}))) --- (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) λ a b → addLemma (a (north , north)) (b (north , north)))) - --- -- private --- -- S¹map : hLevelTrunc 3 S¹ → S¹ --- -- S¹map = trElim (λ _ → isGroupoidS¹) (idfun S¹) - --- -- S¹map-id : (x : hLevelTrunc 3 S¹) → Path (hLevelTrunc 3 S¹) ∣ S¹map x ∣ x --- -- S¹map-id = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) --- -- λ a → refl - --- -- S1map : hLevelTrunc 3 (S₊ 1) → (S₊ 1) --- -- S1map = trElim (λ _ → isGroupoidS1) (idfun _) - - --- -- S¹→S¹ : Iso (S¹ → hLevelTrunc 3 S¹) (S¹ × Int) --- -- Iso.fun S¹→S¹ f = S¹map (f base) --- -- , winding (basechange2⁻ (S¹map (f base)) λ i → S¹map (f (loop i))) --- -- Iso.inv S¹→S¹ (s , int) base = ∣ s ∣ --- -- Iso.inv S¹→S¹ (s , int) (loop i) = ∣ basechange2 s (intLoop int) i ∣ --- -- Iso.rightInv S¹→S¹ (s , int) = ×≡ refl ((λ i → winding (basechange2-retr s (λ i → intLoop int i) i)) --- -- ∙ windingIntLoop int) --- -- Iso.leftInv S¹→S¹ f = funExt λ { base → S¹map-id (f base) --- -- ; (loop i) j → helper2 j i} --- -- where --- -- helper2 : PathP (λ i → S¹map-id (f base) i ≡ S¹map-id (f base) i) --- -- (λ i → ∣ basechange2 (S¹map (f base)) (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁)))))) i ∣) --- -- (cong f loop) --- -- helper2 i j = --- -- hcomp (λ k → λ { (i = i0) → cong ∣_∣ (basechange2 (S¹map (f base)) (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁))))))) j --- -- ; (i = i1) → S¹map-id (f (loop j)) k --- -- ; (j = i0) → S¹map-id (f base) (i ∧ k) --- -- ; (j = i1) → S¹map-id (f base) (i ∧ k)}) --- -- (helper4 i j) --- -- where --- -- helper4 : Path (Path (hLevelTrunc 3 _) _ _) --- -- (cong ∣_∣ (basechange2 (S¹map (f base)) --- -- (intLoop --- -- (winding --- -- (basechange2⁻ (S¹map (f base)) --- -- (λ i₁ → S¹map (f (loop i₁)))))))) --- -- λ i → ∣ S¹map (f (loop i)) ∣ --- -- helper4 i = --- -- cong ∣_∣ --- -- ((cong (basechange2 (S¹map (f base))) --- -- (decodeEncode base (basechange2⁻ (S¹map (f base)) --- -- (λ i₁ → S¹map (f (loop i₁))))) --- -- ∙ basechange2-sect (S¹map (f base)) --- -- (λ i → S¹map (f (loop i)))) i) - --- -- S1→S1→S1×Int : Iso ((S₊ 1) → hLevelTrunc 3 (S₊ 1)) ((hLevelTrunc 3 (S₊ 1)) × Int) --- -- S1→S1→S1×Int = compIso helper2 (compIso S¹→S¹ helper) --- -- where --- -- helper : Iso (S¹ × Int) (hLevelTrunc 3 (S₊ 1) × Int) --- -- Iso.fun helper (s , int) = ∣ S¹→S1 s ∣ , int --- -- Iso.inv helper (s , int) = (S1→S¹ (S1map s)) , int --- -- Iso.rightInv helper (s , int) = --- -- trElim {B = λ s → (∣ S¹→S1 (S1→S¹ (S1map s)) ∣ , int) ≡ (s , int)} --- -- (λ _ → isOfHLevelPath 3 (isOfHLevelProd 3 (isOfHLevelTrunc 3) (isOfHLevelSuc 2 isSetInt)) _ _) --- -- (λ a → ×≡ (cong ∣_∣ (S1→S¹-retr a)) refl) --- -- s --- -- Iso.leftInv helper (s , int) = ×≡ (S1→S¹-sect s) refl - --- -- helper2 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → hLevelTrunc 3 S¹) --- -- Iso.fun helper2 f x = trMap S1→S¹ (f (S¹→S1 x)) --- -- Iso.inv helper2 f x = trMap S¹→S1 (f (S1→S¹ x)) --- -- Iso.rightInv helper2 f = funExt λ x i → helper3 (f (S1→S¹-sect x i)) i --- -- where --- -- helper3 : (x : hLevelTrunc 3 S¹) → trMap S1→S¹ (trMap S¹→S1 x) ≡ x --- -- helper3 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) --- -- λ a → cong ∣_∣ (S1→S¹-sect a) --- -- Iso.leftInv helper2 f = funExt λ x i → helper3 (f (S1→S¹-retr x i)) i --- -- where --- -- helper3 : (x : hLevelTrunc 3 (S₊ 1)) → trMap S¹→S1 (trMap S1→S¹ x) ≡ x --- -- helper3 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) --- -- λ a → cong ∣_∣ (S1→S¹-retr a) - --- -- S¹→S¹≡S1→S1 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → hLevelTrunc 3 S¹) --- -- Iso.fun S¹→S¹≡S1→S1 f x = trMap S1→S¹ (f (S¹→S1 x)) --- -- Iso.inv S¹→S¹≡S1→S1 f x = trMap S¹→S1 (f (S1→S¹ x)) --- -- Iso.rightInv S¹→S¹≡S1→S1 F = funExt λ x i → helper2 (F (S1→S¹-sect x i)) i --- -- where --- -- helper2 : (x : hLevelTrunc 3 S¹) → trMap S1→S¹ (trMap S¹→S1 x) ≡ x --- -- helper2 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) --- -- λ a → cong ∣_∣ (S1→S¹-sect a) --- -- Iso.leftInv S¹→S¹≡S1→S1 F = funExt λ x i → helper2 (F (S1→S¹-retr x i)) i --- -- where --- -- helper2 : (x : hLevelTrunc 3 (S₊ 1)) → trMap S¹→S1 (trMap S1→S¹ x) ≡ x --- -- helper2 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) --- -- λ a → cong ∣_∣ (S1→S¹-retr a) - - - - --- -- basechange-lemma : ∀ {ℓ} {A : Type ℓ} {a : A} (x y : S¹) (F : a ≡ a → S¹) (f : S¹ → a ≡ a) (g : S¹ → a ≡ a) --- -- → (f base ≡ refl) --- -- → (g base ≡ refl) --- -- → basechange2⁻ (F (f base ∙ g base)) (cong₂ {A = S¹} {B = λ x → S¹} (λ x y → F (f x ∙ g y)) loop loop) --- -- ≡ basechange2⁻ (F (f base)) (cong (λ x → F (f x)) loop) ∙ basechange2⁻ (F (g base)) (cong (λ x → F (g x)) loop) --- -- basechange-lemma x y F f g frefl grefl = --- -- (λ i → basechange2⁻ (F (f base ∙ g base)) (cong₂Funct2 (λ x y → F (f x ∙ g y)) loop loop i)) --- -- ∙ (λ i → basechange2⁻ (F (f base ∙ g base)) (cong (λ x₁ → F (f x₁ ∙ g base)) loop ∙ cong (λ y₁ → F (f base ∙ g y₁)) loop)) --- -- ∙ basechange2⁻-morph (F (f base ∙ g base)) _ _ --- -- ∙ (λ j → basechange2⁻ (F (f base ∙ grefl j)) --- -- (λ i → F (f (loop i) ∙ grefl j)) --- -- ∙ basechange2⁻ (F (frefl j ∙ g base)) --- -- (λ i → F (frefl j ∙ g (loop i)))) --- -- ∙ ((λ j → basechange2⁻ (F (rUnit (f base) (~ j))) --- -- (λ i → F (rUnit (f (loop i)) (~ j))) --- -- ∙ basechange2⁻ (F (lUnit (g base) (~ j))) --- -- (λ i → F (lUnit (g (loop i)) (~ j))))) - - --- -- basechange-lemma2 : (f g : S¹ → hLevelTrunc 3 (S₊ 1)) (F : hLevelTrunc 3 (S₊ 1) → S¹) --- -- → ((basechange2⁻ (F (f base +ₖ g base)) λ i → F ((f (loop i)) +ₖ g (loop i))) --- -- ≡ basechange2⁻ (F (f base)) (cong (F ∘ f) loop) --- -- ∙ basechange2⁻ (F (g base)) (cong (F ∘ g) loop)) --- -- basechange-lemma2 f g F = coInd f g F (f base) (g base) refl refl --- -- where --- -- coInd : (f g : S¹ → hLevelTrunc 3 (S₊ 1)) (F : hLevelTrunc 3 (S₊ 1) → S¹) (x y : hLevelTrunc 3 (S₊ 1)) --- -- → f base ≡ x --- -- → g base ≡ y --- -- → ((basechange2⁻ (F (f base +ₖ g base)) λ i → F ((f (loop i)) +ₖ g (loop i))) --- -- ≡ basechange2⁻ (F (f base)) (cong (F ∘ f) loop) --- -- ∙ basechange2⁻ (F (g base)) (cong (F ∘ g) loop)) --- -- coInd f g F = --- -- elim2 (λ _ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 (isGroupoidS¹ base base)) _ _ ) --- -- (toProdElimSuspElim2 --- -- north --- -- (λ _ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → isGroupoidS¹ _ _ _ _) --- -- λ fid gid → --- -- (λ i → basechange2⁻ (F (f base +ₖ g base)) (cong₂Funct2 (λ x y → F (f x +ₖ g y)) loop loop i)) -- (λ i → F (f (loop i) +ₖ g (loop i)))) --- -- ∙ (basechange2⁻-morph (F (f base +ₖ g base)) --- -- (cong (λ x → F (f x +ₖ g base)) loop) --- -- (cong (λ x → F (f base +ₖ g x)) loop)) --- -- ∙ (λ i → basechange2⁻ (F (f base +ₖ gid i)) (cong (λ x → F (f x +ₖ gid i)) loop) --- -- ∙ basechange2⁻ (F (fid i +ₖ g base)) (cong (λ x → F (fid i +ₖ g x)) loop)) --- -- ∙ (λ i → basechange2⁻ (F (rUnitₖ (f base) i)) (cong (λ x → F (rUnitₖ (f x) i)) loop) --- -- ∙ basechange2⁻ (F (lUnitₖ (g base) i)) (cong (λ x → F (lUnitₖ (g x) i)) loop))) - - - --- -- coHom1S1≃ℤ : grIso (coHomGr 1 (S₊ 1)) intGroup --- -- coHom1S1≃ℤ = --- -- Iso'→Iso --- -- (iso' --- -- (iso --- -- (sRec isSetInt (λ f → proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) --- -- (λ a → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , a)) ∣₀) --- -- (λ a → (λ i → proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 (Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , a)))))) --- -- ∙ (λ i → proj₂ (Iso.fun S¹→S¹ (Iso.rightInv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , a)) i))) --- -- ∙ λ i → proj₂ (Iso.rightInv S¹→S¹ (base , a) i)) --- -- (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) --- -- λ f → (λ i → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) ∣₀) --- -- ∙ (λ i → sRec setTruncIsSet --- -- (λ x → ∣ Iso.inv S¹→S¹≡S1→S1 x ∣₀) --- -- (sRec setTruncIsSet --- -- (λ x → ∣ Iso.inv S¹→S¹ (x , (proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) ∣₀) --- -- ∣ base ∣₀)) --- -- ∙ (λ i → sRec setTruncIsSet --- -- (λ x → ∣ Iso.inv S¹→S¹≡S1→S1 x ∣₀) --- -- (sRec setTruncIsSet --- -- (λ x → ∣ Iso.inv S¹→S¹ (x , (proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) ∣₀) --- -- (Iso.inv PathIdTrunc₀Iso (isConnectedS¹ (proj₁ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) i))) --- -- ∙ (λ i → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.leftInv S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f) i) ∣₀) --- -- ∙ (λ i → ∣ Iso.leftInv S¹→S¹≡S1→S1 f i ∣₀))) --- -- (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) --- -- λ f g → (λ i → winding (basechange2⁻ (S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)) ∙ Kn→ΩKn+1 1 (g (S¹→S1 base)))))) --- -- (λ i → S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 (loop i))) ∙ Kn→ΩKn+1 1 (g (S¹→S1 (loop i))))))))) --- -- ∙ cong winding (helper2 (f (S¹→S1 base)) (g (S¹→S1 base)) f g refl refl) --- -- ∙ winding-hom ((basechange2⁻ (S¹map (trMap S1→S¹ (f north))) --- -- (λ i → S¹map (trMap S1→S¹ (f (S¹→S1 (loop i))))))) --- -- ((basechange2⁻ (S¹map (trMap S1→S¹ (g north))) --- -- (λ i → S¹map (trMap S1→S¹ (g (S¹→S1 (loop i))))))))) - - --- -- where --- -- helper2 : (x y : hLevelTrunc 3 (S₊ 1)) (f g : S₊ 1 → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1) --- -- → (f (S¹→S1 base)) ≡ x --- -- → (g (S¹→S1 base)) ≡ y --- -- → (basechange2⁻ (S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)) ∙ Kn→ΩKn+1 1 (g (S¹→S1 base))))))) --- -- (λ i → S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 (loop i))) ∙ Kn→ΩKn+1 1 (g (S¹→S1 (loop i))))))) --- -- ≡ (basechange2⁻ (S¹map (trMap S1→S¹ ((f (S¹→S1 base)))))) --- -- (λ i → S¹map (trMap S1→S¹ (f (S¹→S1 (loop i))))) --- -- ∙ (basechange2⁻ (S¹map (trMap S1→S¹ ((g (S¹→S1 base))))) --- -- (λ i → S¹map (trMap S1→S¹ ((g (S¹→S1 (loop i))))))) --- -- helper2 = elim2 --- -- (λ _ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 --- -- λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 --- -- λ _ → isOfHLevelPath 3 (isOfHLevelSuc 3 (isGroupoidS¹) base base) _ _) --- -- (toProdElimSuspElim2 {A = S₊ 0} north --- -- (λ _ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → (isGroupoidS¹) _ _ _ _ ) --- -- λ f g reflf reflg → --- -- (basechange-lemma --- -- base --- -- base --- -- (λ x → S¹map (trMap S1→S¹ (ΩKn+1→Kn x))) --- -- (λ x → Kn→ΩKn+1 1 (f (S¹→S1 x))) ((λ x → Kn→ΩKn+1 1 (g (S¹→S1 x)))) --- -- (cong (Kn→ΩKn+1 1) reflf ∙ Kn→ΩKn+10ₖ 1) --- -- (cong (Kn→ΩKn+1 1) reflg ∙ Kn→ΩKn+10ₖ 1)) --- -- ∙ λ j → basechange2⁻ (S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 base)) j))) --- -- (λ i → S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 (loop i))) j))) --- -- ∙ basechange2⁻ (S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 base)) j))) --- -- (λ i → S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 (loop i))) j)))) - - - - - - - - --- -- indIntGroup : ∀ {ℓ} {G : Group ℓ} → (ϕ : Int → Group.type G) --- -- → ϕ 0 ≡ isGroup.id (Group.groupStruc G) --- -- → ((a : Int) → ϕ (a +ℤ 1) ≡ isGroup.comp (Group.groupStruc G) (ϕ a) (ϕ 1)) --- -- → ((n : Int) → ϕ (predInt n) ≡ isGroup.comp (Group.groupStruc G) (ϕ n) (ϕ (negsuc zero))) --- -- → isMorph intGroup G ϕ --- -- indIntGroup {G = group G gSet (group-struct _ _ _+G_ _ rUnit₁ _ _ _)} ϕ zeroId _ _ n (pos zero) = --- -- sym (rUnit₁ (ϕ n)) ∙ cong (λ x → ϕ n +G x) (sym zeroId) --- -- indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId oneId minOneId n (pos (suc m)) = --- -- (λ i → ϕ ((n +pos m) +ℤ 1)) --- -- ∙ oneId (n +pos m) --- -- ∙ cong (λ x → x +G ϕ (pos 1)) --- -- (indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} --- -- ϕ zeroId oneId minOneId n (pos m)) --- -- ∙ assoc₁ (ϕ n) (ϕ (pos m)) (ϕ (pos 1)) --- -- ∙ sym (cong (λ x → ϕ n +G x) (oneId (pos m))) --- -- indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId _ minOneId n (negsuc zero) = minOneId n --- -- indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId a minOneId n (negsuc (suc m)) = --- -- (λ i → ϕ ((n +negsuc m) +ℤ (negsuc zero))) --- -- ∙ minOneId (n +negsuc m) --- -- ∙ cong (λ x → x +G ϕ (negsuc zero)) (indIntGroup {G = group G gSet (group-struct id inv₁ _+G_ lUnit₁ rUnit₁ assoc₁ lCancel₁ rCancel₁)} ϕ zeroId a minOneId n (negsuc m)) --- -- ∙ assoc₁ (ϕ n) (ϕ (negsuc m)) (ϕ (negsuc zero)) --- -- ∙ cong (λ x → ϕ n +G x) (sym (minOneId (negsuc m))) - --- -- pushoutSn : (n : ℕ) → Iso (S₊ (suc n)) (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt) --- -- Iso.fun (pushoutSn n) north = inl tt --- -- Iso.fun (pushoutSn n) south = inr tt --- -- Iso.fun (pushoutSn n) (merid a i) = push a i --- -- Iso.inv (pushoutSn n) (inl x) = north --- -- Iso.inv (pushoutSn n) (inr x) = south --- -- Iso.inv (pushoutSn n) (push a i) = merid a i --- -- Iso.rightInv (pushoutSn n) (inl x) = refl --- -- Iso.rightInv (pushoutSn n) (inr x) = refl --- -- Iso.rightInv (pushoutSn n) (push a i) = refl --- -- Iso.leftInv (pushoutSn n) north = refl --- -- Iso.leftInv (pushoutSn n) south = refl --- -- Iso.leftInv (pushoutSn n) (merid a i) = refl - --- -- Sn≡Pushout : (n : ℕ) → (S₊ (suc n)) ≡ (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt) --- -- Sn≡Pushout n = isoToPath (pushoutSn n) - --- -- coHomPushout≡coHomSn' : (n m : ℕ) → grIso (coHomGr m (S₊ (suc n))) (coHomGr m (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt)) --- -- morph.fun (grIso.fun (coHomPushout≡coHomSn' n m)) = --- -- sRec setTruncIsSet --- -- λ f → ∣ (λ {(inl x) → f north ; (inr x) → f south ; (push a i) → f (merid a i)}) ∣₀ --- -- morph.ismorph (grIso.fun (coHomPushout≡coHomSn' n m)) = --- -- sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) --- -- λ a b → cong ∣_∣₀ (funExt λ {(inl x) → refl ; (inr x) → refl ; (push a i) → refl }) --- -- morph.fun (grIso.inv (coHomPushout≡coHomSn' n m)) = sRec setTruncIsSet (λ f → ∣ (λ {north → f (inl tt) ; south → f (inr tt) ; (merid a i) → f (push a i)}) ∣₀) --- -- morph.ismorph (grIso.inv (coHomPushout≡coHomSn' n m)) = --- -- sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) --- -- λ a b → cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → refl }) --- -- grIso.rightInv (coHomPushout≡coHomSn' n m) = --- -- sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) --- -- λ f → cong ∣_∣₀ (funExt λ {(inl x) → refl ; (inr x) → refl ; (push a i) → refl }) --- -- grIso.leftInv (coHomPushout≡coHomSn' n m) = --- -- sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) --- -- λ f → cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → refl }) - - --- -- isContr→≡Unit : {A : Type₀} → isContr A → A ≡ Unit --- -- isContr→≡Unit contr = isoToPath (iso (λ _ → tt) (λ _ → fst contr) (λ _ → refl) λ _ → snd contr _) - --- -- isContr→isContrTrunc : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → isContr A → isContr (hLevelTrunc n A) --- -- isContr→isContrTrunc n contr = ∣ fst contr ∣ , (trElim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → cong ∣_∣ (snd contr a)) --- -- isContr→isContrSetTrunc : ∀ {ℓ} {A : Type ℓ} → isContr A → isContr (∥ A ∥₀) --- -- isContr→isContrSetTrunc contr = ∣ fst contr ∣₀ , sElim (λ _ → isOfHLevelPath 2 (setTruncIsSet) _ _) λ a → cong ∣_∣₀ (snd contr a) - --- -- coHomGrUnit0 : grIso (coHomGr 0 Unit) intGroup --- -- grIso.fun coHomGrUnit0 = mph (sRec isSetInt (λ f → f tt)) --- -- (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) --- -- (λ a b → addLemma (a tt) (b tt))) --- -- grIso.inv coHomGrUnit0 = mph (λ a → ∣ (λ _ → a) ∣₀) (λ a b i → ∣ (λ _ → addLemma a b (~ i)) ∣₀) --- -- grIso.rightInv coHomGrUnit0 a = refl --- -- grIso.leftInv coHomGrUnit0 = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ a → refl - --- -- isContrCohomUnit : (n : ℕ) → isContr (coHom (suc n) Unit) --- -- isContrCohomUnit n = subst isContr (λ i → ∥ UnitToTypeId (coHomK (suc n)) (~ i) ∥₀) (helper n) --- -- where --- -- helper : (n : ℕ) → isContr (∥ coHomK (suc n) ∥₀) --- -- helper n = subst isContr ((isoToPath (truncOfTruncIso {A = S₊ (1 + n)} 2 (1 + n))) ∙ sym propTrunc≡Trunc0 ∙ λ i → ∥ hLevelTrunc (suc (+-comm n 2 i)) (S₊ (1 + n)) ∥₀) --- -- (isConnectedSubtr 2 (helper2 n .fst) (subst (λ x → isHLevelConnected x (S₊ (suc n))) (sym (helper2 n .snd)) (sphereConnected (suc n))) ) --- -- where --- -- helper2 : (n : ℕ) → Σ[ m ∈ ℕ ] m + 2 ≡ 2 + n --- -- helper2 zero = 0 , refl --- -- helper2 (suc n) = (suc n) , λ i → suc (+-comm n 2 i) - --- -- coHomGrUnit≥1 : (n : ℕ) → grIso (coHomGr (suc n) Unit) trivialGroup --- -- grIso.fun (coHomGrUnit≥1 n) = mph (λ _ → tt) (λ _ _ → refl) --- -- grIso.inv (coHomGrUnit≥1 n) = mph (λ _ → ∣ (λ _ → ∣ north ∣) ∣₀) (λ _ _ → sym (rUnitₕ 0ₕ)) --- -- grIso.rightInv (coHomGrUnit≥1 n) a = refl --- -- grIso.leftInv (coHomGrUnit≥1 n) a = sym (isContrCohomUnit n .snd 0ₕ) ∙ isContrCohomUnit n .snd a - --- -- S0→Int : (a : Int × Int) → S₊ 0 → Int --- -- S0→Int a north = proj₁ a --- -- S0→Int a south = proj₂ a - --- -- coHom0-S0 : grIso (coHomGr 0 (S₊ 0)) (dirProd intGroup intGroup) --- -- coHom0-S0 = --- -- Iso'→Iso --- -- (iso' --- -- (iso (sRec (isOfHLevelProd 2 isSetInt isSetInt) --- -- λ f → (f north) , (f south)) --- -- (λ a → ∣ S0→Int a ∣₀) --- -- (λ { (a , b) → refl}) --- -- (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ f → cong ∣_∣₀ (funExt (λ {north → refl ; south → refl})))) --- -- (sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 isSetInt isSetInt) _ _) --- -- λ a b i → addLemma (a north) (b north) i , addLemma (a south) (b south) i)) - --- -- ×morph : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group ℓ} {B : Group ℓ'} {C : Group ℓ''} {D : Group ℓ'''} → morph A B → morph C D → morph (dirProd A C) (dirProd B D) --- -- morph.fun (×morph mf1 mf2) = --- -- (λ {(a , b) → (morph.fun mf1 a) , morph.fun mf2 b}) --- -- morph.ismorph (×morph mf1 mf2) = --- -- (λ {(a , b) (c , d) i → morph.ismorph mf1 a c i , morph.ismorph mf2 b d i}) - - --- -- coHom0S1 : grIso intGroup (coHomGr 0 (S₊ 1)) --- -- coHom0S1 = --- -- Iso'→Iso --- -- (iso' --- -- (iso --- -- (λ a → ∣ (λ x → a) ∣₀) --- -- (sRec isSetInt (λ f → f north)) --- -- (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) --- -- λ f → cong ∣_∣₀ (funExt (toProdElimSuspRec north (λ _ → isSetInt _ _) refl))) --- -- (λ a → refl)) --- -- λ a b → cong ∣_∣₀ (funExt λ x → sym (addLemma a b))) - --- -- coHom1S1 : grIso intGroup (coHomGr 1 (S₊ 1)) --- -- coHom1S1 = --- -- compGrIso --- -- (diagonalIso1 --- -- _ --- -- (coHomGr 0 (S₊ 0)) --- -- _ --- -- (invGrIso coHom0-S0) --- -- (d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) --- -- (λ x → MV.Ker-i⊂Im-d _ _(S₊ 0) (λ _ → tt) (λ _ → tt) 0 x --- -- (×≡ (isOfHLevelSuc 0 (isContrCohomUnit 0) _ _) --- -- (isOfHLevelSuc 0 (isContrCohomUnit 0) _ _))) --- -- (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) --- -- (λ x inker --- -- → pRec propTruncIsProp --- -- (λ {((f , g) , id') → helper x f g id' inker}) --- -- ((MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ inker)))) --- -- (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) --- -- λ F surj --- -- → pRec (setTruncIsSet _ _) (λ { (x , id) → MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ F ∣₀ ∣ (∣ (λ _ → x) ∣₀ , ∣ (λ _ → 0) ∣₀) , --- -- (cong ∣_∣₀ (funExt (surjHelper x))) ∙ sym id ∣₋₁ }) surj) ) --- -- (invGrIso (coHomPushout≡coHomSn' 0 1)) - --- -- where --- -- surjHelper : (x : Int) (x₁ : S₊ 0) → x +ₖ (-ₖ pos 0) ≡ S0→Int (x , x) x₁ --- -- surjHelper x north = cong (x +ₖ_) (-0ₖ) ∙ rUnitₖ x --- -- surjHelper x south = cong (x +ₖ_) (-0ₖ) ∙ rUnitₖ x - --- -- helper : (F : S₊ 0 → Int) (f g : ∥ (Unit → Int) ∥₀) (id : MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (f , g) ≡ ∣ F ∣₀) --- -- → isInKer (coHomGr 0 (S₊ 0)) --- -- (coHomGr 1 (Pushout (λ _ → tt) (λ _ → tt))) --- -- (morph.fun (d-morph Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) --- -- ∣ F ∣₀ --- -- → ∃[ x ∈ Int ] ∣ F ∣₀ ≡ morph.fun (grIso.inv coHom0-S0) (x , x) --- -- helper F = --- -- sElim2 (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) --- -- λ f g id inker --- -- → pRec propTruncIsProp --- -- (λ {((a , b) , id2) --- -- → sElim2 {B = λ f g → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (f , g) ≡ ∣ F ∣₀ → _ } --- -- (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) --- -- (λ f g id → ∣ (helper2 f g .fst) , (sym id ∙ sym (helper2 f g .snd)) ∣₋₁) --- -- a b id2}) --- -- (MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ F ∣₀ inker) --- -- where --- -- helper2 : (f g : Unit → Int) --- -- → Σ[ x ∈ Int ] morph.fun (grIso.inv coHom0-S0) (x , x) --- -- ≡ MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (∣ f ∣₀ , ∣ g ∣₀) --- -- helper2 f g = (f _ +ₖ (-ₖ g _) ) , cong ∣_∣₀ (funExt (λ {north → refl ; south → refl})) - - --- -- coHom-n-Sn : (n : ℕ) → grIso intGroup (coHomGr (suc n) (S₊ (suc n))) --- -- coHom-n-Sn zero = coHom1S1 --- -- coHom-n-Sn (suc n) = --- -- compGrIso --- -- (compGrIso --- -- (coHom-n-Sn n) --- -- theIso) --- -- (invGrIso (coHomPushout≡coHomSn' (suc n) (suc (suc n)))) --- -- where --- -- theIso : grIso (coHomGr (suc n) (S₊ (suc n))) (coHomGr (suc (suc n)) --- -- (Pushout {A = S₊ (suc n)} (λ _ → tt) (λ _ → tt))) --- -- theIso = --- -- SES→Iso --- -- (×coHomGr (suc n) Unit Unit) --- -- (×coHomGr (suc (suc n)) Unit Unit) --- -- (ses (λ p q → ×≡ (isOfHLevelSuc 0 (isContrCohomUnit n) (proj₁ p) (proj₁ q)) --- -- (isOfHLevelSuc 0 (isContrCohomUnit n) (proj₂ p) (proj₂ q))) --- -- (λ p q → ×≡ (isOfHLevelSuc 0 (isContrCohomUnit (suc n)) (proj₁ p) (proj₁ q)) --- -- (isOfHLevelSuc 0 (isContrCohomUnit (suc n)) (proj₂ p) (proj₂ q))) --- -- (Δ-morph _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) --- -- (i-morph _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (2 + n)) --- -- (d-morph _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) --- -- (MV.Ker-d⊂Im-Δ _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) --- -- (MV.Ker-i⊂Im-d _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n))) - - --- -- setTruncIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → Iso A B → Iso ∥ A ∥₀ ∥ B ∥₀ --- -- Iso.fun (setTruncIso is) = sRec setTruncIsSet (λ x → ∣ Iso.fun is x ∣₀) --- -- Iso.inv (setTruncIso is) = sRec setTruncIsSet (λ x → ∣ Iso.inv is x ∣₀) --- -- Iso.rightInv (setTruncIso is) = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ a → cong ∣_∣₀ (Iso.rightInv is a) --- -- Iso.leftInv (setTruncIso is) = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ a → cong ∣_∣₀ (Iso.leftInv is a) - --- -- targetIso : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} → Iso B C → Iso (A → B) (A → C) --- -- Iso.fun (targetIso is) f a = Iso.fun is (f a) --- -- Iso.inv (targetIso is) f a = Iso.inv is (f a) --- -- Iso.rightInv (targetIso is) f = funExt λ a → Iso.rightInv is (f a) --- -- Iso.leftInv (targetIso is) f = funExt λ a → Iso.leftInv is (f a) - --- -- coHom1Torus : grIso (coHomGr 1 ((S₊ 1) × (S₊ 1))) (dirProd intGroup intGroup) --- -- coHom1Torus = --- -- compGrIso --- -- (Iso'→Iso --- -- (iso' (compIso --- -- (setTruncIso (compIso --- -- schonfIso --- -- (compIso --- -- (targetIso S1→S1→S1×Int) --- -- funIso))) --- -- (setTruncOfProdIso)) --- -- (sElim2 --- -- (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) --- -- λ f g → ×≡ (cong ∣_∣₀ --- -- (funExt (λ x → helper (f (x , S¹→S1 base) +ₖ g (x , S¹→S1 base)) --- -- ∙ sym (cong₂ (λ x y → x +ₖ y) --- -- (helper (f (x , S¹→S1 base))) --- -- (helper (g (x , S¹→S1 base))))))) --- -- (cong ∣_∣₀ --- -- (funExt --- -- (toProdElimSuspRec --- -- north --- -- (λ _ → isSetInt _ _) --- -- (cong winding --- -- (basechange-lemma2 --- -- (λ x → f (north , S¹→S1 x)) --- -- (λ x → g (north , S¹→S1 x)) --- -- λ x → S¹map (trMap S1→S¹ x)) --- -- ∙ winding-hom --- -- (basechange2⁻ --- -- (S¹map (trMap S1→S¹ (f (north , S¹→S1 base)))) --- -- (λ i → S¹map (trMap S1→S¹ (f (north , S¹→S1 (loop i)))))) --- -- (basechange2⁻ --- -- (S¹map (trMap S1→S¹ (g (north , S¹→S1 base)))) --- -- (λ i → S¹map (trMap S1→S¹ (g (north , S¹→S1 (loop i)))))) --- -- ∙ sym (addLemma --- -- (winding --- -- (basechange2⁻ --- -- (S¹map (trMap S1→S¹ (f (north , S¹→S1 base)))) --- -- (λ i → S¹map (trMap S1→S¹ (f (north , S¹→S1 (loop i))))))) --- -- (winding --- -- (basechange2⁻ --- -- (S¹map (trMap S1→S¹ (g (north , S¹→S1 base)))) --- -- (λ i → S¹map (trMap S1→S¹ (g (north , S¹→S1 (loop i))))))))))))))) --- -- (dirProdIso (invGrIso (coHom-n-Sn 0)) (invGrIso coHom0S1)) - --- -- where --- -- helper : (x : hLevelTrunc 3 (S₊ 1)) → ∣ S¹→S1 (S¹map (trMap S1→S¹ x)) ∣ ≡ x --- -- helper = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ a → cong ∣_∣ (S1→S¹-retr a) - - - - --- -- -- basechange* : (x y : S¹) → x ≡ y → x ≡ y → ΩS¹ --- -- -- basechange* x y = J (λ y p → (x ≡ y) → ΩS¹) (basechange x) - - --- -- -- test1 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S₊ 1 × Int) --- -- -- Iso.fun test1 f = (trRec isGroupoidS1 (λ a → a) (f north)) --- -- -- , winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (f (loop* i))))) --- -- -- Iso.inv test1 (north , b) x = ∣ x ∣ --- -- -- Iso.inv test1 (south , b) x = ∣ x ∣ --- -- -- Iso.inv test1 (merid a i , b) x = {!!} --- -- -- Iso.rightInv test1 = {!!} --- -- -- Iso.leftInv test1 = {!!} - --- -- -- funRec : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) {C : (A → hLevelTrunc n B) → Type ℓ''} --- -- -- → isOfHLevel n B --- -- -- → ((f : A → B) → C (λ a → ∣ f a ∣)) --- -- -- → (f : A → hLevelTrunc n B) → C f --- -- -- funRec {A = A} {B = B} n {C = C} hLev ind f = subst C (helper f) (ind (λ a → trRec hLev (λ x → x) (f a))) --- -- -- where --- -- -- helper : retract {A = A → hLevelTrunc n B} {B = A → B} (λ f₁ a → trRec hLev (λ x → x) (f₁ a)) (λ f₁ a → ∣ f₁ a ∣) --- -- -- helper f = funExt λ a → helper2 (f a) --- -- -- where --- -- -- helper2 : (x : hLevelTrunc n B) → ∣ trRec hLev (λ x → x) x ∣ ≡ x --- -- -- helper2 = trElim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → refl - --- -- -- invMapSurj : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (ϕ : morph G H) → ((x : Group.type H) → isInIm G H (fst ϕ) x) --- -- -- → morph H G --- -- -- fst (invMapSurj G H (ϕ , pf) surj) a = {!pRec!} --- -- -- snd (invMapSurj G H (ϕ , pf) surj) = {!!} - --- -- {- --- -- ImIso : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (ϕ : morph G H) → ((x : Group.type H) → isInIm G H (fst ϕ) x) --- -- → grIso H (imGroup G H ϕ) --- -- ImIso G H (ϕ , mf) surj = --- -- let idH = isGroup.id (Group.groupStruc H) --- -- idG = isGroup.id (Group.groupStruc G) --- -- _+G_ = isGroup.comp (Group.groupStruc G) --- -- _+H_ = isGroup.comp (Group.groupStruc H) --- -- _+Im_ = isGroup.comp (Group.groupStruc (imGroup G H (ϕ , mf))) --- -- invG = isGroup.inv (Group.groupStruc G) --- -- invH = isGroup.inv (Group.groupStruc H) --- -- lUnit = isGroup.lUnit (Group.groupStruc H) --- -- lCancel = isGroup.rCancel (Group.groupStruc H) --- -- in --- -- Iso''→Iso _ _ --- -- (iso'' ((λ x → x , pRec propTruncIsProp (λ (a , b) → ∣ a , b ∣₋₁) (surj x)) --- -- , λ a b → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) --- -- (λ surja → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) --- -- (λ surjb → --- -- pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) --- -- (λ surja+b → --- -- (λ i → (a +H b) , (pRec (propTruncIsProp) --- -- (λ (a , b) → ∣ a , b ∣₋₁) --- -- (propTruncIsProp (surj (isGroup.comp (Group.groupStruc H) a b)) ∣ surja+b ∣₋₁ i))) ∙ --- -- (λ i → (a +H b) , ∣ (fst surja+b) , (snd surja+b) ∣₋₁) ∙ --- -- ΣProp≡ (λ _ → propTruncIsProp) refl ∙ --- -- λ i → (a +H b) , pRec (propTruncIsProp) --- -- (λ p1 → --- -- pRec (λ x y → squash x y) --- -- (λ p2 → --- -- ∣ --- -- isGroup.comp (Group.groupStruc G) (fst p1) (fst p2) , --- -- mf (fst p1) (fst p2) ∙ --- -- cong₂ (isGroup.comp (Group.groupStruc H)) (snd p1) (snd p2) --- -- ∣₋₁) --- -- (pRec (propTruncIsProp) --- -- ∣_∣₋₁ (propTruncIsProp ∣ surjb ∣₋₁ (surj b) i))) --- -- (pRec (propTruncIsProp) --- -- ∣_∣₋₁ (propTruncIsProp ∣ surja ∣₋₁ (surj a) i ))) --- -- (surj (isGroup.comp (Group.groupStruc H) a b))) --- -- (surj b)) --- -- (surj a)) --- -- (λ x inker → cong fst inker) --- -- λ x → pRec propTruncIsProp (λ inimx → ∣ (ϕ (inimx .fst)) , ΣProp≡ (λ _ → propTruncIsProp) (inimx .snd) ∣₋₁) (snd x)) --- -- -} - - --- -- {- --- -- H¹-S¹≃Int : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- H¹-S¹≃Int = --- -- Iso''→Iso _ _ --- -- (iso'' ((λ x → ∣ theFuns x ∣₀) , λ a b → cong ∣_∣₀ (funExt λ x → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 1) _) ∙ sym (cong (ΩKn+1→Kn) (theFunsId2 a b x)))) --- -- (λ x inker → pRec (isSetInt _ _) (inj x) (Iso.fun PathIdTrunc₀Iso inker)) --- -- finIm) --- -- where --- -- d : _ --- -- d = MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 - --- -- i : _ --- -- i = MV.i _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1 - --- -- Δ : _ --- -- Δ = MV.Δ _ _ (S₊ 1) (λ _ → tt) (λ _ → tt) 0 - - --- -- d-surj : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x --- -- d-surj = sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) --- -- λ x → MV.Ker-i⊂Im-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ --- -- (sym (isContrHelper .snd _)) --- -- where --- -- isContrHelper : isContr (Group.type (×coHomGr 1 Unit Unit)) --- -- isContrHelper = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) --- -- , λ y → prodId (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₁ y)) --- -- (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₂ y)) - --- -- theFuns : (a : Int) → Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 --- -- theFuns a (inl x) = 0ₖ --- -- theFuns a (inr x) = 0ₖ --- -- theFuns a (push north i) = Kn→ΩKn+1 0 a i --- -- theFuns a (push south i) = 0ₖ - - --- -- theFunsId2 : (a b : Int) (x : Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) --- -- → Kn→ΩKn+1 1 (theFuns a x) ∙ Kn→ΩKn+1 1 (theFuns b x) ≡ Kn→ΩKn+1 1 (theFuns (a +ℤ b) x) --- -- theFunsId2 a b (inl x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) --- -- theFunsId2 a b (inr x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) --- -- theFunsId2 a b (push north i) j = --- -- hcomp (λ k → λ {(i = i0) → ((λ i₁ → --- -- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) --- -- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) --- -- j --- -- ; (i = i1) → ((λ i₁ → --- -- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) --- -- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) --- -- j --- -- ; (j = i0) → cong₂Funct2 (λ p q → Kn→ΩKn+1 1 p ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b) (~ k) i --- -- ; (j = i1) → (helper2 a b) k i }) --- -- (hcomp (λ k → λ { (j = i0) → compPath-filler (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) k i --- -- ; (j = i1) → compPath-filler (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) k i --- -- ; (i = i1) → RHS-filler b j k --- -- ; (i = i0) → ((λ i₁ → --- -- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) --- -- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) --- -- j}) --- -- (bottom-filler a j i)) - --- -- where - --- -- bottom-filler : (a : Int) → --- -- PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) --- -- (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) --- -- ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) --- -- j ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) --- -- (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) --- -- ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) --- -- j) (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) --- -- bottom-filler a j i = --- -- hcomp (λ k → λ {(j = i0) → helper2 (~ k) i ; --- -- (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 a) i}) --- -- ((λ j₂ → ∣ rCancel (merid north) j j₂ ∣) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) - --- -- where --- -- helper2 : Path (Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣ ≡ Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣) --- -- (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i) ∙ Kn→ΩKn+1 1 ∣ north ∣) --- -- (λ i → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) --- -- helper2 = cong₂Sym1 (Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) (~ i) j ∣) (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) - --- -- RHS-filler : (b : Int) → --- -- PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j --- -- ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j) --- -- (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) --- -- (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) --- -- RHS-filler b j i = --- -- hcomp (λ k → λ {(j = i0) → cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i ; --- -- (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 b) i}) --- -- (cong (λ q → (λ i → ∣ rCancel (merid north) j i ∣) ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i) - --- -- helper2 : (a b : Int) → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b) ≡ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ℤ b)) --- -- helper2 a b = --- -- sym (congFunct (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b)) --- -- ∙ (λ i → cong (Kn→ΩKn+1 1) (Iso.rightInv (Iso3-Kn-ΩKn+1 0) (Kn→ΩKn+1 0 a ∙ Kn→ΩKn+1 0 b) (~ i))) --- -- ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ₖ b)) ) --- -- ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (addLemma a b i))) - --- -- theFunsId2 a b (push south i) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) --- -- ∙ sym (lUnit _) - --- -- inj : (a : Int) → theFuns a ≡ (λ _ → ∣ north ∣) → a ≡ pos 0 --- -- inj a id = --- -- pRec (isSetInt _ _) --- -- (sigmaElim (λ _ → isOfHLevelPath 2 isSetInt _ _) --- -- (λ a p → pRec (isSetInt _ _) --- -- (λ id2 → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 0) _) --- -- ∙ cong (ΩKn+1→Kn) --- -- (PathP→compPathR --- -- (cong (λ f → cong f (push north)) id) --- -- ∙ test)) --- -- (Iso.fun PathIdTrunc₀Iso p))) (d-surj ∣ theFuns a ∣₀) --- -- where - --- -- test : (λ i → id i (inl tt)) ∙ (λ i → ∣ north ∣) ∙ sym (λ i → id i (inr tt)) ≡ refl --- -- test = (λ i → cong (λ f → f (inl tt)) id --- -- ∙ lUnit (sym (cong (λ f → f (inr tt)) id)) (~ i)) --- -- ∙ (λ i → cong (λ f → f (push south i)) id --- -- ∙ sym (cong (λ f → f (inr tt)) id)) --- -- ∙ rCancel (cong (λ f → f (inr tt)) id) - - --- -- consMember : (a : Int) → coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) --- -- consMember a = d ∣ (λ _ → a) ∣₀ - --- -- consMember≡0 : (a : Int) → consMember a ≡ 0ₕ --- -- consMember≡0 a = --- -- (MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ (λ _ → a) ∣₀ ∣ --- -- (∣ (λ _ → a) ∣₀ , ∣ (λ _ → 0) ∣₀) --- -- , cong ∣_∣₀ (λ i x → (rUnitₖ a i)) ∣₋₁) - --- -- f+consMember' : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int × Int ] (f +ₕ (-ₕ (consMember (proj₁ x))) ≡ ∣ theFuns (proj₂ x) ∣₀) --- -- f+consMember' = --- -- sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) --- -- λ f → pRec propTruncIsProp --- -- (sigmaElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) --- -- (λ g id → ∣ ((g south) , ((g north) +ₖ (-ₖ g south))) --- -- , (pRec (setTruncIsSet _ _) --- -- (λ id → (λ i → ∣ id (~ i) ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀) ∙ funId1 g) --- -- (Iso.fun PathIdTrunc₀Iso id)) ∣₋₁)) --- -- (d-surj ∣ f ∣₀) --- -- where --- -- funId1 : (g : S₊ 0 → Int) --- -- → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀ ≡ --- -- ∣ theFuns ((g north) +ₖ (-ₖ (g south))) ∣₀ --- -- funId1 g = (λ i → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ --- -- +ₕ (morphMinus (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) d --- -- (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ∣ (λ _ → g south) ∣₀ (~ i))) --- -- ∙ sym (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ g ∣₀ (-ₕ ∣ (λ _ → g south) ∣₀)) --- -- ∙ (cong (λ x → d ∣ x ∣₀) g'Id) --- -- ∙ cong ∣_∣₀ helper --- -- where --- -- g' : S₊ 0 → Int --- -- g' north = (g north) +ₖ (-ₖ (g south)) --- -- g' south = 0 - --- -- g'Id : (λ x → g x +ₖ (-ₖ (g south))) ≡ g' --- -- g'Id = funExt (λ {north → refl --- -- ; south → rCancelₖ (g south)}) - --- -- helper : MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g' ≡ theFuns (g north +ₖ (-ₖ g south)) --- -- helper = funExt λ {(inl tt) → refl --- -- ; (inr tt) → refl --- -- ; (push north i) → refl --- -- ; (push south i) → refl} --- -- finIm : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int ] (∣ theFuns x ∣₀ ≡ f) --- -- finIm f = --- -- pRec propTruncIsProp --- -- (λ {((a , b) , id) → ∣ b , (sym id ∙ cong (λ x → f +ₕ x) ((λ i → (-ₕ (consMember≡0 a i))) ∙ sym (lUnitₕ (-ₕ 0ₕ)) ∙ rCancelₕ 0ₕ) ∙ (rUnitₕ f)) ∣₋₁}) --- -- (f+consMember' f) --- -- -} - - - --- -- -- Hⁿ-Sⁿ≃Int : (n : ℕ) → grIso intGroup (coHomGr (suc n) (S₊ (suc n))) --- -- -- Hⁿ-Sⁿ≃Int zero = --- -- -- compGrIso {F = intGroup} {G = {!!}} {H = {!coHomGr 1 (S₊ 1)!}} --- -- -- (Iso''→Iso --- -- -- (iso'' ((λ x → ∣ theFuns x ∣₀) , λ a b → cong ∣_∣₀ (funExt λ x → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 1) _) ∙ sym (cong (ΩKn+1→Kn) (theFunsId2 a b x)))) --- -- -- (λ x inker → pRec (isSetInt _ _) (inj x) (Iso.fun PathIdTrunc₀Iso inker)) --- -- -- finIm)) --- -- -- {!invGrIso _ _ (coHomPushout≡coHomSn 0 1)!} --- -- -- where --- -- -- d : _ --- -- -- d = MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 - --- -- -- i : _ --- -- -- i = MV.i _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1 - --- -- -- Δ : _ --- -- -- Δ = MV.Δ _ _ (S₊ 1) (λ _ → tt) (λ _ → tt) 0 - - --- -- -- d-surj : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- -- → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x --- -- -- d-surj = sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) --- -- -- λ x → MV.Ker-i⊂Im-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ --- -- -- (sym (isContrHelper .snd _)) --- -- -- where --- -- -- isContrHelper : isContr (Group.type (×coHomGr 1 Unit Unit)) --- -- -- isContrHelper = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) --- -- -- , λ y → prodId (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₁ y)) --- -- -- (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₂ y)) - --- -- -- theFuns : (a : Int) → Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 --- -- -- theFuns a (inl x) = 0ₖ --- -- -- theFuns a (inr x) = 0ₖ --- -- -- theFuns a (push north i) = Kn→ΩKn+1 0 a i --- -- -- theFuns a (push south i) = 0ₖ - - --- -- -- theFunsId2 : (a b : Int) (x : Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) --- -- -- → Kn→ΩKn+1 1 (theFuns a x) ∙ Kn→ΩKn+1 1 (theFuns b x) ≡ Kn→ΩKn+1 1 (theFuns (a +ℤ b) x) --- -- -- theFunsId2 a b (inl x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) --- -- -- theFunsId2 a b (inr x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) --- -- -- theFunsId2 a b (push north i) j = --- -- -- hcomp (λ k → λ {(i = i0) → ((λ i₁ → --- -- -- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) --- -- -- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) --- -- -- j --- -- -- ; (i = i1) → ((λ i₁ → --- -- -- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) --- -- -- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) --- -- -- j --- -- -- ; (j = i0) → cong₂Funct2 (λ p q → Kn→ΩKn+1 1 p ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b) (~ k) i --- -- -- ; (j = i1) → (helper2 a b) k i }) --- -- -- (hcomp (λ k → λ { (j = i0) → compPath-filler (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) k i --- -- -- ; (j = i1) → compPath-filler (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) k i --- -- -- ; (i = i1) → RHS-filler b j k --- -- -- ; (i = i0) → ((λ i₁ → --- -- -- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) --- -- -- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) --- -- -- j}) --- -- -- (bottom-filler a j i)) - --- -- -- where - --- -- -- bottom-filler : (a : Int) → --- -- -- PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) --- -- -- (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) --- -- -- ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) --- -- -- j ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) --- -- -- (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) --- -- -- ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) --- -- -- j) (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) --- -- -- bottom-filler a j i = --- -- -- hcomp (λ k → λ {(j = i0) → helper2 (~ k) i ; --- -- -- (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 a) i}) --- -- -- ((λ j₂ → ∣ rCancel (merid north) j j₂ ∣) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) - --- -- -- where --- -- -- helper2 : Path (Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣ ≡ Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣) --- -- -- (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i) ∙ Kn→ΩKn+1 1 ∣ north ∣) --- -- -- (λ i → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) --- -- -- helper2 = cong₂Sym1 (Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) (~ i) j ∣) (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) - --- -- -- RHS-filler : (b : Int) → --- -- -- PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j --- -- -- ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j) --- -- -- (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) --- -- -- (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) --- -- -- RHS-filler b j i = --- -- -- hcomp (λ k → λ {(j = i0) → cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i ; --- -- -- (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 b) i}) --- -- -- (cong (λ q → (λ i → ∣ rCancel (merid north) j i ∣) ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i) - --- -- -- helper2 : (a b : Int) → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b) ≡ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ℤ b)) --- -- -- helper2 a b = --- -- -- sym (congFunct (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b)) --- -- -- ∙ (λ i → cong (Kn→ΩKn+1 1) (Iso.rightInv (Iso3-Kn-ΩKn+1 0) (Kn→ΩKn+1 0 a ∙ Kn→ΩKn+1 0 b) (~ i))) --- -- -- ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ₖ b)) ) --- -- -- ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (addLemma a b i))) - --- -- -- theFunsId2 a b (push south i) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) --- -- -- ∙ sym (lUnit _) - --- -- -- inj : (a : Int) → theFuns a ≡ (λ _ → ∣ north ∣) → a ≡ pos 0 --- -- -- inj a id = --- -- -- pRec (isSetInt _ _) --- -- -- (sigmaElim (λ _ → isOfHLevelPath 2 isSetInt _ _) --- -- -- (λ a p → pRec (isSetInt _ _) --- -- -- (λ id2 → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 0) _) --- -- -- ∙ cong (ΩKn+1→Kn) --- -- -- (PathP→compPathR --- -- -- (cong (λ f → cong f (push north)) id) --- -- -- ∙ test)) --- -- -- (Iso.fun PathIdTrunc₀Iso p))) (d-surj ∣ theFuns a ∣₀) --- -- -- where - --- -- -- test : (λ i → id i (inl tt)) ∙ (λ i → ∣ north ∣) ∙ sym (λ i → id i (inr tt)) ≡ refl --- -- -- test = (λ i → cong (λ f → f (inl tt)) id --- -- -- ∙ lUnit (sym (cong (λ f → f (inr tt)) id)) (~ i)) --- -- -- ∙ (λ i → cong (λ f → f (push south i)) id --- -- -- ∙ sym (cong (λ f → f (inr tt)) id)) --- -- -- ∙ rCancel (cong (λ f → f (inr tt)) id) - - --- -- -- consMember : (a : Int) → coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) --- -- -- consMember a = d ∣ (λ _ → a) ∣₀ - --- -- -- consMember≡0 : (a : Int) → consMember a ≡ 0ₕ --- -- -- consMember≡0 a = --- -- -- (MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ (λ _ → a) ∣₀ ∣ --- -- -- (∣ (λ _ → a) ∣₀ , ∣ (λ _ → 0) ∣₀) --- -- -- , cong ∣_∣₀ (λ i x → (rUnitₖ a i)) ∣₋₁) - --- -- -- f+consMember' : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int × Int ] (f +ₕ (-ₕ (consMember (proj₁ x))) ≡ ∣ theFuns (proj₂ x) ∣₀) --- -- -- f+consMember' = --- -- -- sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) --- -- -- λ f → pRec propTruncIsProp --- -- -- (sigmaElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) --- -- -- (λ g id → ∣ ((g south) , ((g north) +ₖ (-ₖ g south))) --- -- -- , (pRec (setTruncIsSet _ _) --- -- -- (λ id → (λ i → ∣ id (~ i) ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀) ∙ funId1 g) --- -- -- (Iso.fun PathIdTrunc₀Iso id)) ∣₋₁)) --- -- -- (d-surj ∣ f ∣₀) --- -- -- where --- -- -- funId1 : (g : S₊ 0 → Int) --- -- -- → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀ ≡ --- -- -- ∣ theFuns ((g north) +ₖ (-ₖ (g south))) ∣₀ --- -- -- funId1 g = (λ i → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ --- -- -- +ₕ (morphMinus (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) d --- -- -- (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ∣ (λ _ → g south) ∣₀ (~ i))) --- -- -- ∙ sym (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ g ∣₀ (-ₕ ∣ (λ _ → g south) ∣₀)) --- -- -- ∙ (cong (λ x → d ∣ x ∣₀) g'Id) --- -- -- ∙ cong ∣_∣₀ helper --- -- -- where --- -- -- g' : S₊ 0 → Int --- -- -- g' north = (g north) +ₖ (-ₖ (g south)) --- -- -- g' south = 0 - --- -- -- g'Id : (λ x → g x +ₖ (-ₖ (g south))) ≡ g' --- -- -- g'Id = funExt (λ {north → refl --- -- -- ; south → rCancelₖ (g south)}) - --- -- -- helper : MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g' ≡ theFuns (g north +ₖ (-ₖ g south)) --- -- -- helper = funExt λ {(inl tt) → refl --- -- -- ; (inr tt) → refl --- -- -- ; (push north i) → refl --- -- -- ; (push south i) → refl} --- -- -- finIm : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int ] (∣ theFuns x ∣₀ ≡ f) --- -- -- finIm f = --- -- -- pRec propTruncIsProp --- -- -- (λ {((a , b) , id) → ∣ b , (sym id ∙ cong (λ x → f +ₕ x) ((λ i → (-ₕ (consMember≡0 a i))) ∙ sym (lUnitₕ (-ₕ 0ₕ)) ∙ rCancelₕ 0ₕ) ∙ (rUnitₕ f)) ∣₋₁}) --- -- -- (f+consMember' f) --- -- -- Hⁿ-Sⁿ≃Int (suc n) = --- -- -- compGrIso (Hⁿ-Sⁿ≃Int n) --- -- -- (transport (λ i → grIso {!!} {!coHomGr (suc (suc n)) (Pushout (λ _ → tt) (λ _ → tt))!}) {!!}) - - --- -- -- {- --- -- -- coHom1S1≃ℤ : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- -- grIso.fun coHom1S1≃ℤ = grIso.fun {!compGrIso coHom1Iso (invGrIso _ _ (ImIso _ _ (d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ?))!} --- -- -- grIso.inv coHom1S1≃ℤ = {!!} --- -- -- grIso.rightInv coHom1S1≃ℤ = {!!} --- -- -- grIso.leftInv coHom1S1≃ℤ = {!!} --- -- -- -} - --- -- -- -- compGrIso coHom1Iso (invGrIso _ _ (ImIso _ _ {!d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0!} {!!})) - - --- -- -- -- coHomGrIm : grIso (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- -- -- (imGroup (coHomGr 0 (S₊ 0)) --- -- -- -- (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- -- -- (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 --- -- -- -- , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) --- -- -- -- coHomGrIm = --- -- -- -- ImIso _ --- -- -- -- _ --- -- -- -- _ --- -- -- -- {!!} - - --- -- -- -- -- coHom1RedS1 : Iso (coHom 1 (S₊ 1)) (coHomRed 1 (S₊ 1 , north)) --- -- -- -- -- Iso.fun coHom1RedS1 = sRec setTruncIsSet λ f → ∣ f , (pRec {!!} {!!} ((transport (λ i → (b : truncIdempotent 3 {!S₊ 1!} (~ i)) → ∥ (transp (λ j → truncIdempotent {!3!} {!!} (~ i ∨ (~ j))) (~ i) north) ≡ b ∥₋₁) isConnectedS1) (f north)) ) ∣₀ --- -- -- -- -- Iso.inv coHom1RedS1 = {!!} --- -- -- -- -- Iso.rightInv coHom1RedS1 = {!setTruncIdempotent!} --- -- -- -- -- Iso.leftInv coHom1RedS1 = {!!} - --- -- -- -- -- coHom1Red : ∀ {ℓ} (A : Pointed ℓ) → Iso (coHom 1 (typ A)) (coHomRed 1 A) --- -- -- -- -- Iso.fun (coHom1Red A) = sRec setTruncIsSet λ f → ∣ f , {!!} ∣₀ --- -- -- -- -- Iso.inv (coHom1Red A) = {!!} --- -- -- -- -- Iso.rightInv (coHom1Red A) = {!!} --- -- -- -- -- Iso.leftInv (coHom1Red A) = {!!} - --- -- -- -- -- -- morphtest : morph (coHomGr 1 (S₊ 1)) intGroup --- -- -- -- -- -- fst morphtest = sRec isSetInt λ f → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (f (loop* i))))) --- -- -- -- -- -- snd morphtest = sElim2 {!!} --- -- -- -- -- -- (funRec 3 isGroupoidS1 --- -- -- -- -- -- λ f → funRec 3 isGroupoidS1 --- -- -- -- -- -- λ g → pRec (isSetInt _ _) --- -- -- -- -- -- (λ n=fn → --- -- -- -- -- -- pRec (isSetInt _ _) --- -- -- -- -- -- (λ n=gn → (λ j → winding (basechange (SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (∣ f (north) ∣ +ₖ ∣ n=gn (~ j) ∣)))) (λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (∣ f (loop* i) ∣ +ₖ ∣ transp (λ i → n=gn ((~ i) ∨ (~ j)) ≡ n=gn ((~ i) ∨ (~ j))) (~ j) (λ i → g (loop* i)) i ∣)))))) --- -- -- -- -- -- ∙ {!!} --- -- -- -- -- -- ∙ {!!}) --- -- -- -- -- -- (isConnectedS1 (g north))) --- -- -- -- -- -- (isConnectedS1 (f north))) - - --- -- -- -- -- -- {- (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (∣ f (loop* i) ∣ +ₖ ∣ g (loop* i) ∣))))) --- -- -- -- -- -- ∙ (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (Kn→ΩKn+1 1 ∣ f (loop* i) ∣ ∙ Kn→ΩKn+1 1 ∣ g (loop* i) ∣)))))) --- -- -- -- -- -- ∙ (λ j → winding (basechange _ (cong (λ x → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (Kn→ΩKn+1 1 ∣ f x ∣ ∙ Kn→ΩKn+1 1 ∣ g x ∣))))) loop*)) ) --- -- -- -- -- -- ∙ (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn ((cong ∣_∣ (merid (f (loop* i)) ∙ sym (merid north)) ∙ cong ∣_∣ (merid (g (loop* i)) ∙ sym (merid north))))))))) --- -- -- -- -- -- ∙ (λ j → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (congFunct ∣_∣ (merid (f (loop* i)) ∙ sym (merid north)) (merid (g (loop* i)) ∙ sym (merid north)) (~ j))))))) --- -- -- -- -- -- ∙ (λ j → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (cong ∣_∣ (({!!} ∙ {!!}) ∙ {!!}))))))) --- -- -- -- -- -- ∙ {!!} --- -- -- -- -- -- ∙ {!!} --- -- -- -- -- -- ∙ {!!}) -} - --- -- -- -- -- -- where --- -- -- -- -- -- helper : ∀ {ℓ} {A : Type ℓ} (a : A) (f : A → S¹) (p q : a ≡ a) → winding (basechange (f a) (cong f (p ∙ q))) ≡ winding (basechange (f a) (cong f p ∙ cong f q)) --- -- -- -- -- -- helper a f p q i = winding (basechange (f a) (congFunct f p q i)) --- -- -- -- -- -- helper2 : (x : S¹) (p q : x ≡ x) → basechange x (p ∙ q) ≡ basechange x p ∙ basechange x q --- -- -- -- -- -- helper2 base p q = refl --- -- -- -- -- -- helper2 (loop i) p q = {!!} --- -- -- -- -- -- helper4 : (x y z : S¹) (p : x ≡ z) (q r : x ≡ y) (s : y ≡ z) → basechange* x z p (q ∙ s) ≡ basechange* x y {!!} q ∙ {!!} --- -- -- -- -- -- helper4 = {!!} --- -- -- -- -- -- helper3 : (p q : ΩS¹) → winding (p ∙ q) ≡ (winding p +ℤ winding q) --- -- -- -- -- -- helper3 = {!!} - - --- -- -- -- -- -- -- fstmap : morph (dirProd intGroup intGroup) (coHomGr 0 (S₊ 0)) --- -- -- -- -- -- -- fstmap = compMorph {F = dirProd intGroup intGroup} {G = ×coHomGr 0 Unit Unit} {H = coHomGr 0 (S₊ 0)} --- -- -- -- -- -- -- (×morph (grIso.inv coHomGrUnit0) (grIso.inv coHomGrUnit0)) --- -- -- -- -- -- -- (((MV.Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) zero)) , --- -- -- -- -- -- -- {!MV.ΔIsHom _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) zero!}) - --- -- -- -- -- -- -- fstMapId : (a : Int × Int) → fstmap .fst a ≡ ∣ (λ _ → proj₁ a +ℤ (0 - proj₂ a)) ∣₀ --- -- -- -- -- -- -- fstMapId (a , b) = (λ i → ∣ (λ _ → a +ₖ (-ₖ b)) ∣₀) ∙ {!addLemma!} ∙ {!!} ∙ {!!} - --- -- -- -- -- -- -- isoAgain : grIso intGroup (coHomGr 1 (S₊ 1)) --- -- -- -- -- -- -- isoAgain = --- -- -- -- -- -- -- Iso''→Iso _ _ --- -- -- -- -- -- -- (iso'' ((λ a → ∣ (λ {north → 0ₖ ; south → 0ₖ ; (merid north i) → {!a!} ; (merid south i) → {!!}}) ∣₀) , {!!}) --- -- -- -- -- -- -- {!!} --- -- -- -- -- -- -- {!!}) - --- -- -- -- -- -- -- -- -- test2 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → S¹) --- -- -- -- -- -- -- -- -- Iso.fun test2 f = {!!} --- -- -- -- -- -- -- -- -- Iso.inv test2 f north = ∣ transport (sym S¹≡S1) (f base) ∣ --- -- -- -- -- -- -- -- -- Iso.inv test2 f south = ∣ transport (sym S¹≡S1) (f base) ∣ --- -- -- -- -- -- -- -- -- Iso.inv test2 f (merid a i) = cong ∣_∣ {!transport (sym S¹≡S1) (f base)!} i --- -- -- -- -- -- -- -- -- Iso.rightInv test2 = {!!} --- -- -- -- -- -- -- -- -- Iso.leftInv test2 = {!!} - --- -- -- -- -- -- -- -- -- F : winding (basechange base loop) ≡ 1 --- -- -- -- -- -- -- -- -- F = refl - --- -- -- -- -- -- -- -- -- another : (f g : Int) → winding (basechange {!!} {!!}) ≡ {!!} --- -- -- -- -- -- -- -- -- another = {!!} - --- -- -- -- -- -- -- -- -- test : Iso (S¹ → S¹) (S¹ × Int) --- -- -- -- -- -- -- -- -- Iso.fun test f = f base , winding (basechange (f base) (cong f loop)) --- -- -- -- -- -- -- -- -- Iso.inv test (x , int) base = x --- -- -- -- -- -- -- -- -- Iso.inv test (x , int) (loop i) = {!!} --- -- -- -- -- -- -- -- -- Iso.rightInv test = {!!} --- -- -- -- -- -- -- -- -- Iso.leftInv test = {!!} - --- -- -- -- -- -- -- -- -- -- test13 : Iso ∥ (S¹ → S¹) ∥₀ Int --- -- -- -- -- -- -- -- -- -- Iso.fun test13 = sRec isSetInt λ f → winding (basechange (f base) (λ i → f (loop i))) --- -- -- -- -- -- -- -- -- -- Iso.inv test13 a = ∣ (λ {base → {!!} ; (loop i) → {!!}}) ∣₀ --- -- -- -- -- -- -- -- -- -- Iso.rightInv test13 = {!!} --- -- -- -- -- -- -- -- -- -- Iso.leftInv test13 = {!!} - --- -- -- -- -- -- -- -- -- -- test : Iso (S¹ → S¹) (S¹ × Int) --- -- -- -- -- -- -- -- -- -- Iso.fun test f = (f base) , transport (basedΩS¹≡Int (f base)) λ i → f (loop i) --- -- -- -- -- -- -- -- -- -- Iso.inv test (x , int) base = x --- -- -- -- -- -- -- -- -- -- Iso.inv test (x , int) (loop i) = transport (sym (basedΩS¹≡Int x)) int i --- -- -- -- -- -- -- -- -- -- Iso.rightInv test (x , int) i = (x , transportTransport⁻ (basedΩS¹≡Int x) int i) --- -- -- -- -- -- -- -- -- -- Iso.leftInv test f = --- -- -- -- -- -- -- -- -- -- funExt λ { base → refl --- -- -- -- -- -- -- -- -- -- ; (loop i) j → transport⁻Transport (basedΩS¹≡Int (f base)) (λ i → f (loop i)) j i} - - --- -- -- -- -- -- -- -- -- -- lem : S¹ ≡ hLevelTrunc 3 (S₊ 1) --- -- -- -- -- -- -- -- -- -- lem = sym (truncIdempotent 3 isGroupoidS¹) ∙ λ i → hLevelTrunc 3 (S¹≡S1 (~ i)) - --- -- -- -- -- -- -- -- -- -- prodId : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (a b : A × B) → proj₁ a ≡ proj₁ b → proj₂ a ≡ proj₂ b → a ≡ b --- -- -- -- -- -- -- -- -- -- prodId (_ , _) (_ , _) id1 id2 i = (id1 i) , (id2 i) - --- -- -- -- -- -- -- -- -- -- test22 : Iso (S₊ 1 → coHomK 1) (S₊ 1 × Int) --- -- -- -- -- -- -- -- -- -- Iso.fun test22 f = {!f north!} , {!!} --- -- -- -- -- -- -- -- -- -- Iso.inv test22 = {!!} --- -- -- -- -- -- -- -- -- -- Iso.rightInv test22 = {!!} --- -- -- -- -- -- -- -- -- -- Iso.leftInv test22 = {!!} - - - - - - --- -- -- -- -- -- -- -- -- -- coHom1≃∥S1×ℤ∥₀ : Iso (coHom 1 (S₊ 1)) ∥ S₊ 1 × Int ∥₀ --- -- -- -- -- -- -- -- -- -- coHom1≃∥S1×ℤ∥₀ = setTruncIso test2 --- -- -- -- -- -- -- -- -- -- where --- -- -- -- -- -- -- -- -- -- test2 : Iso (S₊ 1 → coHomK 1) (S₊ 1 × Int) --- -- -- -- -- -- -- -- -- -- Iso.fun test2 f = transport (λ i → S¹≡S1 (~ i) × Int) (Iso.fun test (transport (λ i → (S¹≡S1 i → lem (~ i))) f)) --- -- -- -- -- -- -- -- -- -- Iso.inv test2 x = transport (λ i → (S¹≡S1 (~ i) → lem i)) (Iso.inv test (transport (λ i → S¹≡S1 i × Int) x)) --- -- -- -- -- -- -- -- -- -- Iso.rightInv test2 (s , int) = prodId _ _ {!!} {!!} --- -- -- -- -- -- -- -- -- -- Iso.leftInv test2 f = {!!} ∙ {!!} ∙ {!!} - --- -- -- -- -- -- -- -- -- -- test2Id : (a b : (S₊ 1 → coHomK 1)) → proj₂ (Iso.fun test2 (λ x → a x +ₖ b x)) ≡ (proj₂ (Iso.fun test2 a) +ₖ proj₂ (Iso.fun test2 a)) --- -- -- -- -- -- -- -- -- -- test2Id a b = {! --- -- -- -- -- -- -- -- -- -- transport --- -- -- -- -- -- -- -- -- -- (basedΩS¹≡Int --- -- -- -- -- -- -- -- -- -- (transport (λ i → S¹≡S1 i → lem (~ i)) (λ x → a x +ₖ b x) base)) --- -- -- -- -- -- -- -- -- -- (λ i → --- -- -- -- -- -- -- -- -- -- transport (λ i₁ → S¹≡S1 i₁ → lem (~ i₁)) (λ x → a x +ₖ b x) --- -- -- -- -- -- -- -- -- -- (loop i))!} ∙ {!transport (λ i → S¹≡S1 i → lem (~ i)) (λ x → a x +ₖ b x) base!} - - --- -- -- -- -- -- -- -- -- -- main : grIso intGroup (coHomGr 1 (S₊ 1)) --- -- -- -- -- -- -- -- -- -- main = Iso'→Iso --- -- -- -- -- -- -- -- -- -- (iso' {!!} --- -- -- -- -- -- -- -- -- -- {!!}) - - --- -- -- -- -- -- -- -- coHom1 : grIso (coHomGr 1 (S₊ 1)) intGroup --- -- -- -- -- -- -- -- coHom1 = --- -- -- -- -- -- -- -- Iso'→Iso --- -- -- -- -- -- -- -- (iso' (iso ({!!} ∘ {!!} ∘ {!!} ∘ {!!}) --- -- -- -- -- -- -- -- {!!} --- -- -- -- -- -- -- -- {!!} --- -- -- -- -- -- -- -- {!!}) --- -- -- -- -- -- -- -- {!!}) - - - --- -- -- -- -- -- -- -- schonf : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : (A × B) → Type ℓ''} --- -- -- -- -- -- -- -- → ((a : A) (b : B) → C (a , b)) --- -- -- -- -- -- -- -- → (x : A × B) → C x --- -- -- -- -- -- -- -- schonf f (a , b) = f a b - --- -- -- -- -- -- -- -- -- -- setTruncProdIso : ∀ {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') → Iso ∥ (A × B) ∥₀ (∥ A ∥₀ × ∥ B ∥₀) --- -- -- -- -- -- -- -- -- -- Iso.fun (setTruncProdIso A B) = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) λ {(a , b) → ∣ a ∣₀ , ∣ b ∣₀} --- -- -- -- -- -- -- -- -- -- Iso.inv (setTruncProdIso A B) (a , b) = sRec setTruncIsSet (λ a → sRec setTruncIsSet (λ b → ∣ a , b ∣₀) b) a --- -- -- -- -- -- -- -- -- -- Iso.rightInv (setTruncProdIso A B) = --- -- -- -- -- -- -- -- -- -- schonf (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) --- -- -- -- -- -- -- -- -- -- λ _ → sElim (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) --- -- -- -- -- -- -- -- -- -- λ _ → refl) --- -- -- -- -- -- -- -- -- -- Iso.leftInv (setTruncProdIso A B) = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ {(a , b) → refl} - --- -- -- -- -- -- -- -- -- -- setTruncProdLemma : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (a1 a2 : A) (b : B) → isHLevelConnected 2 A --- -- -- -- -- -- -- -- -- -- → Path ∥ A × B ∥₀ ∣ a1 , b ∣₀ ∣ a2 , b ∣₀ --- -- -- -- -- -- -- -- -- -- setTruncProdLemma {A = A} {B = B} a1 a2 b conA i = Iso.inv (setTruncProdIso A B) (Iso.inv setTruncTrunc0Iso ((sym (conA .snd ∣ a1 ∣) ∙ (conA .snd ∣ a2 ∣)) i) , ∣ b ∣₀) - --- -- -- -- -- -- -- -- -- -- test3 : Iso ∥ S₊ 1 × Int ∥₀ Int --- -- -- -- -- -- -- -- -- -- Iso.fun test3 = sRec isSetInt proj₂ --- -- -- -- -- -- -- -- -- -- Iso.inv test3 a = ∣ north , a ∣₀ --- -- -- -- -- -- -- -- -- -- Iso.rightInv test3 a = refl --- -- -- -- -- -- -- -- -- -- Iso.leftInv test3 = --- -- -- -- -- -- -- -- -- -- sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) --- -- -- -- -- -- -- -- -- -- λ {(s , int) → setTruncProdLemma north s int (sphereConnected 1)} - --- -- -- -- -- -- -- -- -- -- coHomGr0-S1 : grIso intGroup (coHomGr 1 (S₊ 1)) --- -- -- -- -- -- -- -- -- -- coHomGr0-S1 = --- -- -- -- -- -- -- -- -- -- Iso'→Iso --- -- -- -- -- -- -- -- -- -- (iso' (compIso (symIso test3) (symIso coHom1≃∥S1×ℤ∥₀)) --- -- -- -- -- -- -- -- -- -- (indIntGroup {G = coHomGr 1 (S₊ 1)} --- -- -- -- -- -- -- -- -- -- (Iso.fun (compIso (symIso test3) (symIso coHom1≃∥S1×ℤ∥₀))) --- -- -- -- -- -- -- -- -- -- ((λ i → ∣ transport (λ i → (S¹≡S1 (~ i) → lem i)) (Iso.inv test (base , 0)) ∣₀) --- -- -- -- -- -- -- -- -- -- ∙ (λ i → ∣ transport (λ i → (S¹≡S1 (~ i) → lem i)) (helper2 i) ∣₀) --- -- -- -- -- -- -- -- -- -- ∙ cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → {!!}})) --- -- -- -- -- -- -- -- -- -- {!!} --- -- -- -- -- -- -- -- -- -- {!!})) --- -- -- -- -- -- -- -- -- -- where --- -- -- -- -- -- -- -- -- -- helper : basedΩS¹≡ΩS¹ base ≡ {!basechange!} --- -- -- -- -- -- -- -- -- -- helper = {!substComposite!} - --- -- -- -- -- -- -- -- -- -- substComposite2 : ∀ {ℓ} {A B C : Type ℓ} --- -- -- -- -- -- -- -- -- -- (P : A ≡ B) (Q : B ≡ C) (a : A) --- -- -- -- -- -- -- -- -- -- → transport (P ∙ Q) a ≡ transport Q (transport P a) --- -- -- -- -- -- -- -- -- -- substComposite2 = {!!} - --- -- -- -- -- -- -- -- -- -- helper1 : transport (λ i → S¹≡S1 i × Int) (north , 0) ≡ (base , 0) --- -- -- -- -- -- -- -- -- -- helper1 = refl --- -- -- -- -- -- -- -- -- -- helper3 : transport (sym (basedΩS¹≡Int base)) 0 ≡ refl --- -- -- -- -- -- -- -- -- -- helper3 = (λ i → transport (symDistr (basedΩS¹≡ΩS¹ base) ΩS¹≡Int i) 0) --- -- -- -- -- -- -- -- -- -- ∙ substComposite2 (sym ΩS¹≡Int) (sym (basedΩS¹≡ΩS¹ base)) 0 --- -- -- -- -- -- -- -- -- -- ∙ (λ i → transport (λ i → basedΩS¹≡ΩS¹ base (~ i)) refl) -- --- -- -- -- -- -- -- -- -- -- ∙ transportRefl ((equiv-proof (basechange-isequiv base) refl) .fst .fst) --- -- -- -- -- -- -- -- -- -- ∙ (λ i → equiv-proof (transport (λ j → isEquiv (refl-conjugation j)) (basedΩS¹→ΩS¹-isequiv i0)) refl .fst .fst) --- -- -- -- -- -- -- -- -- -- ∙ (λ i → {!equiv-proof (transport (λ j → isEquiv (refl-conjugation j)) (basedΩS¹→ΩS¹-isequiv i0)) refl .fst!}) --- -- -- -- -- -- -- -- -- -- ∙ {!basedΩS¹→ΩS¹!} --- -- -- -- -- -- -- -- -- -- ∙ {!!} --- -- -- -- -- -- -- -- -- -- ∙ {!!} --- -- -- -- -- -- -- -- -- -- helper4 : (x : S¹) → Iso.inv test (base , 0) x ≡ x --- -- -- -- -- -- -- -- -- -- helper4 = {!!} --- -- -- -- -- -- -- -- -- -- helper2 : Iso.inv test (transport (λ i → S¹≡S1 i × Int) (north , 0)) ≡ λ x → x --- -- -- -- -- -- -- -- -- -- helper2 = (λ i → Iso.inv test (base , 0)) ∙ {!!} ∙ {!!} - --- -- -- -- -- -- -- -- prodId : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A × B} → proj₁ x ≡ proj₁ y → proj₂ x ≡ proj₂ y → x ≡ y --- -- -- -- -- -- -- -- prodId {x = (a , b)} {y = (c , d)} id1 id2 i = (id1 i) , (id2 i) - --- -- -- -- -- -- -- -- fstFunHelper : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- -- -- -- -- -- -- → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 _) (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x --- -- -- -- -- -- -- -- fstFunHelper a = MV.Ker-i⊂Im-d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a --- -- -- -- -- -- -- -- (sym (isContrH1Unit×H1Unit .snd _) ∙ (isContrH1Unit×H1Unit .snd _)) --- -- -- -- -- -- -- -- where --- -- -- -- -- -- -- -- isContrH1Unit×H1Unit : isContr (Group.type (×coHomGr 1 Unit Unit)) --- -- -- -- -- -- -- -- isContrH1Unit×H1Unit = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) --- -- -- -- -- -- -- -- , λ {(a , b) → sigmaProdElim {D = λ (x : Σ[ x ∈ Group.type (×coHomGr 1 Unit Unit)] Unit) → (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) ≡ fst x} --- -- -- -- -- -- -- -- (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) --- -- -- -- -- -- -- -- (λ a b _ → prodId (grIso.leftInv (coHomGrUnit≥1 0) ∣ a ∣₀) (grIso.leftInv (coHomGrUnit≥1 0) ∣ b ∣₀)) ((a , b) , tt)} - - - --- -- -- -- -- -- -- -- helperMorph : isMorph intGroup (dirProd intGroup intGroup) λ a → a , (0 - a) --- -- -- -- -- -- -- -- helperMorph = --- -- -- -- -- -- -- -- indIntGroup {G = dirProd intGroup intGroup} --- -- -- -- -- -- -- -- (λ a → a , (0 - a)) --- -- -- -- -- -- -- -- refl --- -- -- -- -- -- -- -- (λ a → prodId refl (helper2 a)) --- -- -- -- -- -- -- -- λ a → prodId refl (helper3 a) +-- d : _ +-- d = MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 + +-- i : _ +-- i = MV.i _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1 + +-- Δ : _ +-- Δ = MV.Δ _ _ (S₊ 1) (λ _ → tt) (λ _ → tt) 0 + + +-- d-surj : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x +-- d-surj = sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- λ x → MV.Ker-i⊂Im-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ +-- (sym (isContrHelper .snd _)) +-- where +-- isContrHelper : isContr (Group.type (×coHomGr 1 Unit Unit)) +-- isContrHelper = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) +-- , λ y → prodId (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₁ y)) +-- (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₂ y)) + +-- theFuns : (a : Int) → Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 +-- theFuns a (inl x) = 0ₖ +-- theFuns a (inr x) = 0ₖ +-- theFuns a (push north i) = Kn→ΩKn+1 0 a i +-- theFuns a (push south i) = 0ₖ + + +-- theFunsId2 : (a b : Int) (x : Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) +-- → Kn→ΩKn+1 1 (theFuns a x) ∙ Kn→ΩKn+1 1 (theFuns b x) ≡ Kn→ΩKn+1 1 (theFuns (a +ℤ b) x) +-- theFunsId2 a b (inl x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) +-- theFunsId2 a b (inr x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) +-- theFunsId2 a b (push north i) j = +-- hcomp (λ k → λ {(i = i0) → ((λ i₁ → +-- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) +-- j +-- ; (i = i1) → ((λ i₁ → +-- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) +-- j +-- ; (j = i0) → cong₂Funct2 (λ p q → Kn→ΩKn+1 1 p ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b) (~ k) i +-- ; (j = i1) → (helper2 a b) k i }) +-- (hcomp (λ k → λ { (j = i0) → compPath-filler (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) k i +-- ; (j = i1) → compPath-filler (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) k i +-- ; (i = i1) → RHS-filler b j k +-- ; (i = i0) → ((λ i₁ → +-- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) +-- j}) +-- (bottom-filler a j i)) + +-- where + +-- bottom-filler : (a : Int) → +-- PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) +-- ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) +-- j ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) +-- ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) +-- j) (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) +-- bottom-filler a j i = +-- hcomp (λ k → λ {(j = i0) → helper2 (~ k) i ; +-- (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 a) i}) +-- ((λ j₂ → ∣ rCancel (merid north) j j₂ ∣) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + +-- where +-- helper2 : Path (Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣ ≡ Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- (λ i → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) +-- helper2 = cong₂Sym1 (Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) (~ i) j ∣) (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + +-- RHS-filler : (b : Int) → +-- PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j +-- ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j) +-- (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) +-- (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) +-- RHS-filler b j i = +-- hcomp (λ k → λ {(j = i0) → cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i ; +-- (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 b) i}) +-- (cong (λ q → (λ i → ∣ rCancel (merid north) j i ∣) ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i) + +-- helper2 : (a b : Int) → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b) ≡ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ℤ b)) +-- helper2 a b = +-- sym (congFunct (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b)) +-- ∙ (λ i → cong (Kn→ΩKn+1 1) (Iso.rightInv (Iso3-Kn-ΩKn+1 0) (Kn→ΩKn+1 0 a ∙ Kn→ΩKn+1 0 b) (~ i))) +-- ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ₖ b)) ) +-- ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (addLemma a b i))) + +-- theFunsId2 a b (push south i) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- ∙ sym (lUnit _) + +-- inj : (a : Int) → theFuns a ≡ (λ _ → ∣ north ∣) → a ≡ pos 0 +-- inj a id = +-- pRec (isSetInt _ _) +-- (sigmaElim (λ _ → isOfHLevelPath 2 isSetInt _ _) +-- (λ a p → pRec (isSetInt _ _) +-- (λ id2 → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 0) _) +-- ∙ cong (ΩKn+1→Kn) +-- (PathP→compPathR +-- (cong (λ f → cong f (push north)) id) +-- ∙ test)) +-- (Iso.fun PathIdTrunc₀Iso p))) (d-surj ∣ theFuns a ∣₀) +-- where + +-- test : (λ i → id i (inl tt)) ∙ (λ i → ∣ north ∣) ∙ sym (λ i → id i (inr tt)) ≡ refl +-- test = (λ i → cong (λ f → f (inl tt)) id +-- ∙ lUnit (sym (cong (λ f → f (inr tt)) id)) (~ i)) +-- ∙ (λ i → cong (λ f → f (push south i)) id +-- ∙ sym (cong (λ f → f (inr tt)) id)) +-- ∙ rCancel (cong (λ f → f (inr tt)) id) + + +-- consMember : (a : Int) → coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) +-- consMember a = d ∣ (λ _ → a) ∣₀ + +-- consMember≡0 : (a : Int) → consMember a ≡ 0ₕ +-- consMember≡0 a = +-- (MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ (λ _ → a) ∣₀ ∣ +-- (∣ (λ _ → a) ∣₀ , ∣ (λ _ → 0) ∣₀) +-- , cong ∣_∣₀ (λ i x → (rUnitₖ a i)) ∣₋₁) + +-- f+consMember' : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int × Int ] (f +ₕ (-ₕ (consMember (proj₁ x))) ≡ ∣ theFuns (proj₂ x) ∣₀) +-- f+consMember' = +-- sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- λ f → pRec propTruncIsProp +-- (sigmaElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- (λ g id → ∣ ((g south) , ((g north) +ₖ (-ₖ g south))) +-- , (pRec (setTruncIsSet _ _) +-- (λ id → (λ i → ∣ id (~ i) ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀) ∙ funId1 g) +-- (Iso.fun PathIdTrunc₀Iso id)) ∣₋₁)) +-- (d-surj ∣ f ∣₀) +-- where +-- funId1 : (g : S₊ 0 → Int) +-- → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀ ≡ +-- ∣ theFuns ((g north) +ₖ (-ₖ (g south))) ∣₀ +-- funId1 g = (λ i → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +-- +ₕ (morphMinus (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) d +-- (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ∣ (λ _ → g south) ∣₀ (~ i))) +-- ∙ sym (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ g ∣₀ (-ₕ ∣ (λ _ → g south) ∣₀)) +-- ∙ (cong (λ x → d ∣ x ∣₀) g'Id) +-- ∙ cong ∣_∣₀ helper +-- where +-- g' : S₊ 0 → Int +-- g' north = (g north) +ₖ (-ₖ (g south)) +-- g' south = 0 + +-- g'Id : (λ x → g x +ₖ (-ₖ (g south))) ≡ g' +-- g'Id = funExt (λ {north → refl +-- ; south → rCancelₖ (g south)}) + +-- helper : MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g' ≡ theFuns (g north +ₖ (-ₖ g south)) +-- helper = funExt λ {(inl tt) → refl +-- ; (inr tt) → refl +-- ; (push north i) → refl +-- ; (push south i) → refl} +-- finIm : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int ] (∣ theFuns x ∣₀ ≡ f) +-- finIm f = +-- pRec propTruncIsProp +-- (λ {((a , b) , id) → ∣ b , (sym id ∙ cong (λ x → f +ₕ x) ((λ i → (-ₕ (consMember≡0 a i))) ∙ sym (lUnitₕ (-ₕ 0ₕ)) ∙ rCancelₕ 0ₕ) ∙ (rUnitₕ f)) ∣₋₁}) +-- (f+consMember' f) +-- Hⁿ-Sⁿ≃Int (suc n) = +-- compGrIso (Hⁿ-Sⁿ≃Int n) +-- (transport (λ i → grIso {!!} {!coHomGr (suc (suc n)) (Pushout (λ _ → tt) (λ _ → tt))!}) {!!}) + + +-- {- +-- coHom1S1≃ℤ : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- grIso.fun coHom1S1≃ℤ = grIso.fun {!compGrIso coHom1Iso (invGrIso _ _ (ImIso _ _ (d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ?))!} +-- grIso.inv coHom1S1≃ℤ = {!!} +-- grIso.rightInv coHom1S1≃ℤ = {!!} +-- grIso.leftInv coHom1S1≃ℤ = {!!} +-- -} + +-- -- compGrIso coHom1Iso (invGrIso _ _ (ImIso _ _ {!d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0!} {!!})) + + +-- -- coHomGrIm : grIso (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- (imGroup (coHomGr 0 (S₊ 0)) +-- -- (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 +-- -- , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) +-- -- coHomGrIm = +-- -- ImIso _ +-- -- _ +-- -- _ +-- -- {!!} + + +-- -- -- coHom1RedS1 : Iso (coHom 1 (S₊ 1)) (coHomRed 1 (S₊ 1 , north)) +-- -- -- Iso.fun coHom1RedS1 = sRec setTruncIsSet λ f → ∣ f , (pRec {!!} {!!} ((transport (λ i → (b : truncIdempotent 3 {!S₊ 1!} (~ i)) → ∥ (transp (λ j → truncIdempotent {!3!} {!!} (~ i ∨ (~ j))) (~ i) north) ≡ b ∥₋₁) isConnectedS1) (f north)) ) ∣₀ +-- -- -- Iso.inv coHom1RedS1 = {!!} +-- -- -- Iso.rightInv coHom1RedS1 = {!setTruncIdempotent!} +-- -- -- Iso.leftInv coHom1RedS1 = {!!} + +-- -- -- coHom1Red : ∀ {ℓ} (A : Pointed ℓ) → Iso (coHom 1 (typ A)) (coHomRed 1 A) +-- -- -- Iso.fun (coHom1Red A) = sRec setTruncIsSet λ f → ∣ f , {!!} ∣₀ +-- -- -- Iso.inv (coHom1Red A) = {!!} +-- -- -- Iso.rightInv (coHom1Red A) = {!!} +-- -- -- Iso.leftInv (coHom1Red A) = {!!} + +-- -- -- -- morphtest : morph (coHomGr 1 (S₊ 1)) intGroup +-- -- -- -- fst morphtest = sRec isSetInt λ f → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (f (loop* i))))) +-- -- -- -- snd morphtest = sElim2 {!!} +-- -- -- -- (funRec 3 isGroupoidS1 +-- -- -- -- λ f → funRec 3 isGroupoidS1 +-- -- -- -- λ g → pRec (isSetInt _ _) +-- -- -- -- (λ n=fn → +-- -- -- -- pRec (isSetInt _ _) +-- -- -- -- (λ n=gn → (λ j → winding (basechange (SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (∣ f (north) ∣ +ₖ ∣ n=gn (~ j) ∣)))) (λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (∣ f (loop* i) ∣ +ₖ ∣ transp (λ i → n=gn ((~ i) ∨ (~ j)) ≡ n=gn ((~ i) ∨ (~ j))) (~ j) (λ i → g (loop* i)) i ∣)))))) +-- -- -- -- ∙ {!!} +-- -- -- -- ∙ {!!}) +-- -- -- -- (isConnectedS1 (g north))) +-- -- -- -- (isConnectedS1 (f north))) + + +-- -- -- -- {- (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (∣ f (loop* i) ∣ +ₖ ∣ g (loop* i) ∣))))) +-- -- -- -- ∙ (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (Kn→ΩKn+1 1 ∣ f (loop* i) ∣ ∙ Kn→ΩKn+1 1 ∣ g (loop* i) ∣)))))) +-- -- -- -- ∙ (λ j → winding (basechange _ (cong (λ x → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (Kn→ΩKn+1 1 ∣ f x ∣ ∙ Kn→ΩKn+1 1 ∣ g x ∣))))) loop*)) ) +-- -- -- -- ∙ (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn ((cong ∣_∣ (merid (f (loop* i)) ∙ sym (merid north)) ∙ cong ∣_∣ (merid (g (loop* i)) ∙ sym (merid north))))))))) +-- -- -- -- ∙ (λ j → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (congFunct ∣_∣ (merid (f (loop* i)) ∙ sym (merid north)) (merid (g (loop* i)) ∙ sym (merid north)) (~ j))))))) +-- -- -- -- ∙ (λ j → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (cong ∣_∣ (({!!} ∙ {!!}) ∙ {!!}))))))) +-- -- -- -- ∙ {!!} +-- -- -- -- ∙ {!!} +-- -- -- -- ∙ {!!}) -} + +-- -- -- -- where +-- -- -- -- helper : ∀ {ℓ} {A : Type ℓ} (a : A) (f : A → S¹) (p q : a ≡ a) → winding (basechange (f a) (cong f (p ∙ q))) ≡ winding (basechange (f a) (cong f p ∙ cong f q)) +-- -- -- -- helper a f p q i = winding (basechange (f a) (congFunct f p q i)) +-- -- -- -- helper2 : (x : S¹) (p q : x ≡ x) → basechange x (p ∙ q) ≡ basechange x p ∙ basechange x q +-- -- -- -- helper2 base p q = refl +-- -- -- -- helper2 (loop i) p q = {!!} +-- -- -- -- helper4 : (x y z : S¹) (p : x ≡ z) (q r : x ≡ y) (s : y ≡ z) → basechange* x z p (q ∙ s) ≡ basechange* x y {!!} q ∙ {!!} +-- -- -- -- helper4 = {!!} +-- -- -- -- helper3 : (p q : ΩS¹) → winding (p ∙ q) ≡ (winding p +ℤ winding q) +-- -- -- -- helper3 = {!!} + + +-- -- -- -- -- fstmap : morph (dirProd intGroup intGroup) (coHomGr 0 (S₊ 0)) +-- -- -- -- -- fstmap = compMorph {F = dirProd intGroup intGroup} {G = ×coHomGr 0 Unit Unit} {H = coHomGr 0 (S₊ 0)} +-- -- -- -- -- (×morph (grIso.inv coHomGrUnit0) (grIso.inv coHomGrUnit0)) +-- -- -- -- -- (((MV.Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) zero)) , +-- -- -- -- -- {!MV.ΔIsHom _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) zero!}) + +-- -- -- -- -- fstMapId : (a : Int × Int) → fstmap .fst a ≡ ∣ (λ _ → proj₁ a +ℤ (0 - proj₂ a)) ∣₀ +-- -- -- -- -- fstMapId (a , b) = (λ i → ∣ (λ _ → a +ₖ (-ₖ b)) ∣₀) ∙ {!addLemma!} ∙ {!!} ∙ {!!} + +-- -- -- -- -- isoAgain : grIso intGroup (coHomGr 1 (S₊ 1)) +-- -- -- -- -- isoAgain = +-- -- -- -- -- Iso''→Iso _ _ +-- -- -- -- -- (iso'' ((λ a → ∣ (λ {north → 0ₖ ; south → 0ₖ ; (merid north i) → {!a!} ; (merid south i) → {!!}}) ∣₀) , {!!}) +-- -- -- -- -- {!!} +-- -- -- -- -- {!!}) + +-- -- -- -- -- -- -- test2 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → S¹) +-- -- -- -- -- -- -- Iso.fun test2 f = {!!} +-- -- -- -- -- -- -- Iso.inv test2 f north = ∣ transport (sym S¹≡S1) (f base) ∣ +-- -- -- -- -- -- -- Iso.inv test2 f south = ∣ transport (sym S¹≡S1) (f base) ∣ +-- -- -- -- -- -- -- Iso.inv test2 f (merid a i) = cong ∣_∣ {!transport (sym S¹≡S1) (f base)!} i +-- -- -- -- -- -- -- Iso.rightInv test2 = {!!} +-- -- -- -- -- -- -- Iso.leftInv test2 = {!!} + +-- -- -- -- -- -- -- F : winding (basechange base loop) ≡ 1 +-- -- -- -- -- -- -- F = refl + +-- -- -- -- -- -- -- another : (f g : Int) → winding (basechange {!!} {!!}) ≡ {!!} +-- -- -- -- -- -- -- another = {!!} + +-- -- -- -- -- -- -- test : Iso (S¹ → S¹) (S¹ × Int) +-- -- -- -- -- -- -- Iso.fun test f = f base , winding (basechange (f base) (cong f loop)) +-- -- -- -- -- -- -- Iso.inv test (x , int) base = x +-- -- -- -- -- -- -- Iso.inv test (x , int) (loop i) = {!!} +-- -- -- -- -- -- -- Iso.rightInv test = {!!} +-- -- -- -- -- -- -- Iso.leftInv test = {!!} + +-- -- -- -- -- -- -- -- test13 : Iso ∥ (S¹ → S¹) ∥₀ Int +-- -- -- -- -- -- -- -- Iso.fun test13 = sRec isSetInt λ f → winding (basechange (f base) (λ i → f (loop i))) +-- -- -- -- -- -- -- -- Iso.inv test13 a = ∣ (λ {base → {!!} ; (loop i) → {!!}}) ∣₀ +-- -- -- -- -- -- -- -- Iso.rightInv test13 = {!!} +-- -- -- -- -- -- -- -- Iso.leftInv test13 = {!!} + +-- -- -- -- -- -- -- -- test : Iso (S¹ → S¹) (S¹ × Int) +-- -- -- -- -- -- -- -- Iso.fun test f = (f base) , transport (basedΩS¹≡Int (f base)) λ i → f (loop i) +-- -- -- -- -- -- -- -- Iso.inv test (x , int) base = x +-- -- -- -- -- -- -- -- Iso.inv test (x , int) (loop i) = transport (sym (basedΩS¹≡Int x)) int i +-- -- -- -- -- -- -- -- Iso.rightInv test (x , int) i = (x , transportTransport⁻ (basedΩS¹≡Int x) int i) +-- -- -- -- -- -- -- -- Iso.leftInv test f = +-- -- -- -- -- -- -- -- funExt λ { base → refl +-- -- -- -- -- -- -- -- ; (loop i) j → transport⁻Transport (basedΩS¹≡Int (f base)) (λ i → f (loop i)) j i} + + +-- -- -- -- -- -- -- -- lem : S¹ ≡ hLevelTrunc 3 (S₊ 1) +-- -- -- -- -- -- -- -- lem = sym (truncIdempotent 3 isGroupoidS¹) ∙ λ i → hLevelTrunc 3 (S¹≡S1 (~ i)) + +-- -- -- -- -- -- -- -- prodId : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (a b : A × B) → proj₁ a ≡ proj₁ b → proj₂ a ≡ proj₂ b → a ≡ b +-- -- -- -- -- -- -- -- prodId (_ , _) (_ , _) id1 id2 i = (id1 i) , (id2 i) + +-- -- -- -- -- -- -- -- test22 : Iso (S₊ 1 → coHomK 1) (S₊ 1 × Int) +-- -- -- -- -- -- -- -- Iso.fun test22 f = {!f north!} , {!!} +-- -- -- -- -- -- -- -- Iso.inv test22 = {!!} +-- -- -- -- -- -- -- -- Iso.rightInv test22 = {!!} +-- -- -- -- -- -- -- -- Iso.leftInv test22 = {!!} + + + + + + +-- -- -- -- -- -- -- -- coHom1≃∥S1×ℤ∥₀ : Iso (coHom 1 (S₊ 1)) ∥ S₊ 1 × Int ∥₀ +-- -- -- -- -- -- -- -- coHom1≃∥S1×ℤ∥₀ = setTruncIso test2 -- -- -- -- -- -- -- -- where --- -- -- -- -- -- -- -- helper1 : (a : Int) → predInt (sucInt a) ≡ a --- -- -- -- -- -- -- -- helper1 (pos zero) = refl --- -- -- -- -- -- -- -- helper1 (pos (suc n)) = refl --- -- -- -- -- -- -- -- helper1 (negsuc zero) = refl --- -- -- -- -- -- -- -- helper1 (negsuc (suc n)) = refl - --- -- -- -- -- -- -- -- helper4 : (a : Int) → sucInt (predInt a) ≡ a --- -- -- -- -- -- -- -- helper4 (pos zero) = refl --- -- -- -- -- -- -- -- helper4 (pos (suc n)) = refl --- -- -- -- -- -- -- -- helper4 (negsuc zero) = refl --- -- -- -- -- -- -- -- helper4 (negsuc (suc n)) = refl - --- -- -- -- -- -- -- -- helper2 : (a : Int) → (pos 0 - sucInt a) ≡ predInt (pos 0 - a) --- -- -- -- -- -- -- -- helper2 (pos zero) = refl --- -- -- -- -- -- -- -- helper2 (pos (suc n)) = refl --- -- -- -- -- -- -- -- helper2 (negsuc zero) = refl --- -- -- -- -- -- -- -- helper2 (negsuc (suc n)) = sym (helper1 _) - --- -- -- -- -- -- -- -- helper3 : (a : Int) → (pos 0 - predInt a) ≡ sucInt (pos 0 - a) --- -- -- -- -- -- -- -- helper3 (pos zero) = refl --- -- -- -- -- -- -- -- helper3 (pos (suc zero)) = refl --- -- -- -- -- -- -- -- helper3 (pos (suc (suc n))) = sym (helper4 _) --- -- -- -- -- -- -- -- helper3 (negsuc zero) = refl --- -- -- -- -- -- -- -- helper3 (negsuc (suc n)) = refl - --- -- -- -- -- -- -- -- helper : (a b : Int) → (pos 0 - (a +ℤ b)) ≡ ((pos 0 - a) +ℤ (pos 0 - b)) --- -- -- -- -- -- -- -- helper a (pos zero) = refl --- -- -- -- -- -- -- -- helper (pos zero) (pos (suc n)) = --- -- -- -- -- -- -- -- cong (λ x → pos 0 - sucInt x) (+ℤ-comm (pos zero) (pos n)) --- -- -- -- -- -- -- -- ∙ +ℤ-comm (pos 0 +negsuc n) (pos zero) --- -- -- -- -- -- -- -- helper (pos (suc n₁)) (pos (suc n)) = --- -- -- -- -- -- -- -- {!!} --- -- -- -- -- -- -- -- helper (negsuc n₁) (pos (suc n)) = {!!} --- -- -- -- -- -- -- -- helper a (negsuc n) = {!!} - --- -- -- -- -- -- -- -- fun : morph intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- -- -- -- -- -- -- fst fun = MV.d _ _ _ (λ _ → tt) (λ _ → tt) 0 ∘ grIso.inv coHom0-S0 .fst ∘ λ a → a , (0 - a) --- -- -- -- -- -- -- -- snd fun = {!!} --- -- -- -- -- -- -- -- {- compMorph {F = intGroup} {G = dirProd intGroup intGroup} {H = coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))} --- -- -- -- -- -- -- -- ((λ a → a , a) , (λ a b → refl)) --- -- -- -- -- -- -- -- (compMorph {F = dirProd intGroup intGroup} {G = coHomGr 0 (S₊ 0)} {H = coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))} (grIso.inv coHom0-S0) --- -- -- -- -- -- -- -- (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 --- -- -- -- -- -- -- -- , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) .snd -} --- -- -- -- -- -- -- -- {- theIso : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- -- -- -- -- -- -- theIso = Iso''→Iso _ _ --- -- -- -- -- -- -- -- (iso'' ((λ x → ∣ (λ {(inl tt) → 0ₖ ; (inr tt) → 0ₖ ; (push a i) → Kn→ΩKn+1 0 x i}) ∣₀) , {!!}) --- -- -- -- -- -- -- -- {!!} --- -- -- -- -- -- -- -- {!MV.d!}) --- -- -- -- -- -- -- -- -} - - - --- -- -- -- -- -- -- -- theIso : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- -- -- -- -- -- -- theIso = --- -- -- -- -- -- -- -- Iso''→Iso _ _ --- -- -- -- -- -- -- -- (iso'' fun --- -- -- -- -- -- -- -- (λ x inker → {!MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 --- -- -- -- -- -- -- -- (grIso.inv coHom0-S0 .fst (g , g))!}) --- -- -- -- -- -- -- -- (sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) --- -- -- -- -- -- -- -- λ x → pRec propTruncIsProp --- -- -- -- -- -- -- -- (λ {(a , b) → {!fun!} }) --- -- -- -- -- -- -- -- (fstFunHelper (∣ x ∣₀)))) --- -- -- -- -- -- -- -- where --- -- -- -- -- -- -- -- whoKnows : (a : S₊ 0 → Int) (x : MV.D Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt)) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (λ _ → a north) x --- -- -- -- -- -- -- -- ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 a x --- -- -- -- -- -- -- -- whoKnows a (inl x) = refl --- -- -- -- -- -- -- -- whoKnows a (inr x) = refl --- -- -- -- -- -- -- -- whoKnows a (push north i) = refl --- -- -- -- -- -- -- -- whoKnows a (push south i) j = {!!} - --- -- -- -- -- -- -- -- helper : (a : Int) → (grIso.inv coHom0-S0 .fst (a , a)) ≡ ∣ S0→Int (a , a) ∣₀ --- -- -- -- -- -- -- -- helper a = {!have : - --- -- -- -- -- -- -- -- ∣ --- -- -- -- -- -- -- -- MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 --- -- -- -- -- -- -- -- (S0→Int (x , x)) --- -- -- -- -- -- -- -- ∣₀ --- -- -- -- -- -- -- -- ≡ ∣ (λ _ → ∣ north ∣) ∣₀!} - --- -- -- -- -- -- -- -- helper2 : (a b : Int) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a)) ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (b , b)) --- -- -- -- -- -- -- -- → a ≡ b --- -- -- -- -- -- -- -- helper2 a b id = pRec (isSetInt a b) (λ {(pt , id) → {!!}}) (fstFunHelper ∣ (MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a))) ∣₀) - --- -- -- -- -- -- -- -- idFun : (a : Int) → MV.D Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 --- -- -- -- -- -- -- -- idFun a (inl x) = 0ₖ --- -- -- -- -- -- -- -- idFun a (inr x) = 0ₖ --- -- -- -- -- -- -- -- idFun a (push north i) = Kn→ΩKn+1 zero a i --- -- -- -- -- -- -- -- idFun a (push south i) = Kn→ΩKn+1 zero a i - --- -- -- -- -- -- -- -- helper3 : (a : Int) → (MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a))) ≡ idFun a --- -- -- -- -- -- -- -- helper3 a = funExt λ {(inl x) → refl ; (inr x) → refl ; (push north i) → refl ; (push south i) → refl } - --- -- -- -- -- -- -- -- helper4 : (a b : Int) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a)) ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (b , b)) --- -- -- -- -- -- -- -- → a ≡ b --- -- -- -- -- -- -- -- helper4 a b id = --- -- -- -- -- -- -- -- {!!} --- -- -- -- -- -- -- -- ∙ {!!} --- -- -- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- -- -- test2 : Iso (S₊ 1 → coHomK 1) (S₊ 1 × Int) +-- -- -- -- -- -- -- -- Iso.fun test2 f = transport (λ i → S¹≡S1 (~ i) × Int) (Iso.fun test (transport (λ i → (S¹≡S1 i → lem (~ i))) f)) +-- -- -- -- -- -- -- -- Iso.inv test2 x = transport (λ i → (S¹≡S1 (~ i) → lem i)) (Iso.inv test (transport (λ i → S¹≡S1 i × Int) x)) +-- -- -- -- -- -- -- -- Iso.rightInv test2 (s , int) = prodId _ _ {!!} {!!} +-- -- -- -- -- -- -- -- Iso.leftInv test2 f = {!!} ∙ {!!} ∙ {!!} + +-- -- -- -- -- -- -- -- test2Id : (a b : (S₊ 1 → coHomK 1)) → proj₂ (Iso.fun test2 (λ x → a x +ₖ b x)) ≡ (proj₂ (Iso.fun test2 a) +ₖ proj₂ (Iso.fun test2 a)) +-- -- -- -- -- -- -- -- test2Id a b = {! +-- -- -- -- -- -- -- -- transport +-- -- -- -- -- -- -- -- (basedΩS¹≡Int +-- -- -- -- -- -- -- -- (transport (λ i → S¹≡S1 i → lem (~ i)) (λ x → a x +ₖ b x) base)) +-- -- -- -- -- -- -- -- (λ i → +-- -- -- -- -- -- -- -- transport (λ i₁ → S¹≡S1 i₁ → lem (~ i₁)) (λ x → a x +ₖ b x) +-- -- -- -- -- -- -- -- (loop i))!} ∙ {!transport (λ i → S¹≡S1 i → lem (~ i)) (λ x → a x +ₖ b x) base!} + + +-- -- -- -- -- -- -- -- main : grIso intGroup (coHomGr 1 (S₊ 1)) +-- -- -- -- -- -- -- -- main = Iso'→Iso +-- -- -- -- -- -- -- -- (iso' {!!} +-- -- -- -- -- -- -- -- {!!}) + + +-- -- -- -- -- -- coHom1 : grIso (coHomGr 1 (S₊ 1)) intGroup +-- -- -- -- -- -- coHom1 = +-- -- -- -- -- -- Iso'→Iso +-- -- -- -- -- -- (iso' (iso ({!!} ∘ {!!} ∘ {!!} ∘ {!!}) +-- -- -- -- -- -- {!!} +-- -- -- -- -- -- {!!} +-- -- -- -- -- -- {!!}) +-- -- -- -- -- -- {!!}) + + + +-- -- -- -- -- -- schonf : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : (A × B) → Type ℓ''} +-- -- -- -- -- -- → ((a : A) (b : B) → C (a , b)) +-- -- -- -- -- -- → (x : A × B) → C x +-- -- -- -- -- -- schonf f (a , b) = f a b + +-- -- -- -- -- -- -- -- setTruncProdIso : ∀ {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') → Iso ∥ (A × B) ∥₀ (∥ A ∥₀ × ∥ B ∥₀) +-- -- -- -- -- -- -- -- Iso.fun (setTruncProdIso A B) = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) λ {(a , b) → ∣ a ∣₀ , ∣ b ∣₀} +-- -- -- -- -- -- -- -- Iso.inv (setTruncProdIso A B) (a , b) = sRec setTruncIsSet (λ a → sRec setTruncIsSet (λ b → ∣ a , b ∣₀) b) a +-- -- -- -- -- -- -- -- Iso.rightInv (setTruncProdIso A B) = +-- -- -- -- -- -- -- -- schonf (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- -- -- -- -- -- -- -- λ _ → sElim (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- -- -- -- -- -- -- -- λ _ → refl) +-- -- -- -- -- -- -- -- Iso.leftInv (setTruncProdIso A B) = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ {(a , b) → refl} + +-- -- -- -- -- -- -- -- setTruncProdLemma : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (a1 a2 : A) (b : B) → isHLevelConnected 2 A +-- -- -- -- -- -- -- -- → Path ∥ A × B ∥₀ ∣ a1 , b ∣₀ ∣ a2 , b ∣₀ +-- -- -- -- -- -- -- -- setTruncProdLemma {A = A} {B = B} a1 a2 b conA i = Iso.inv (setTruncProdIso A B) (Iso.inv setTruncTrunc0Iso ((sym (conA .snd ∣ a1 ∣) ∙ (conA .snd ∣ a2 ∣)) i) , ∣ b ∣₀) + +-- -- -- -- -- -- -- -- test3 : Iso ∥ S₊ 1 × Int ∥₀ Int +-- -- -- -- -- -- -- -- Iso.fun test3 = sRec isSetInt proj₂ +-- -- -- -- -- -- -- -- Iso.inv test3 a = ∣ north , a ∣₀ +-- -- -- -- -- -- -- -- Iso.rightInv test3 a = refl +-- -- -- -- -- -- -- -- Iso.leftInv test3 = +-- -- -- -- -- -- -- -- sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- -- -- -- -- -- -- -- λ {(s , int) → setTruncProdLemma north s int (sphereConnected 1)} + +-- -- -- -- -- -- -- -- coHomGr0-S1 : grIso intGroup (coHomGr 1 (S₊ 1)) +-- -- -- -- -- -- -- -- coHomGr0-S1 = +-- -- -- -- -- -- -- -- Iso'→Iso +-- -- -- -- -- -- -- -- (iso' (compIso (symIso test3) (symIso coHom1≃∥S1×ℤ∥₀)) +-- -- -- -- -- -- -- -- (indIntGroup {G = coHomGr 1 (S₊ 1)} +-- -- -- -- -- -- -- -- (Iso.fun (compIso (symIso test3) (symIso coHom1≃∥S1×ℤ∥₀))) +-- -- -- -- -- -- -- -- ((λ i → ∣ transport (λ i → (S¹≡S1 (~ i) → lem i)) (Iso.inv test (base , 0)) ∣₀) +-- -- -- -- -- -- -- -- ∙ (λ i → ∣ transport (λ i → (S¹≡S1 (~ i) → lem i)) (helper2 i) ∣₀) +-- -- -- -- -- -- -- -- ∙ cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → {!!}})) +-- -- -- -- -- -- -- -- {!!} +-- -- -- -- -- -- -- -- {!!})) -- -- -- -- -- -- -- -- where --- -- -- -- -- -- -- -- helper5 : {!!} --PathP (λ k → id k (inl tt) ≡ id k (inr tt)) (Kn→ΩKn+1 zero a) (Kn→ΩKn+1 zero a) --- -- -- -- -- -- -- -- helper5 i j = {!id i!} - --- -- -- -- -- -- -- -- -- fun : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) → coHom 0 (S₊ 0) --- -- -- -- -- -- -- -- -- fun a = (pRec {P = Σ[ x ∈ coHom 0 (S₊ 0)] (MV.d _ _ _ (λ _ → tt) (λ _ → tt) 0 x ≡ a) × isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) (MV.Δ _ _ _ (λ _ → tt) (λ _ → tt) 0) x} --- -- -- -- -- -- -- -- -- (λ {(a1 , b) (c , d) → ΣProp≡ (λ x → isOfHLevelProd 1 (setTruncIsSet _ _) propTruncIsProp) --- -- -- -- -- -- -- -- -- (sElim2 {B = λ a1 c → (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a1 ≡ a) --- -- -- -- -- -- -- -- -- → MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 c ≡ a --- -- -- -- -- -- -- -- -- → isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) --- -- -- -- -- -- -- -- -- (λ z → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 z) a1 --- -- -- -- -- -- -- -- -- → isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) --- -- -- -- -- -- -- -- -- (λ z → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 z) c → a1 ≡ c} --- -- -- -- -- -- -- -- -- (λ _ _ → {!!}) --- -- -- -- -- -- -- -- -- (λ a c b1 d1 → pElim (λ _ → isOfHLevelΠ 1 λ _ → setTruncIsSet _ _) --- -- -- -- -- -- -- -- -- λ b2 → pElim (λ _ → setTruncIsSet _ _) --- -- -- -- -- -- -- -- -- λ d2 → {!d2!}) --- -- -- -- -- -- -- -- -- a1 c (proj₁ b) (proj₁ d) (proj₂ b) (proj₂ d))}) --- -- -- -- -- -- -- -- -- (λ {(a , b) → a , b , MV.Ker-d⊂Im-Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a {!!}}) --- -- -- -- -- -- -- -- -- (fstFunHelper a)) .fst -- pRec {!!} {!!} (fstFunHelper a) +-- -- -- -- -- -- -- -- helper : basedΩS¹≡ΩS¹ base ≡ {!basechange!} +-- -- -- -- -- -- -- -- helper = {!substComposite!} + +-- -- -- -- -- -- -- -- substComposite2 : ∀ {ℓ} {A B C : Type ℓ} +-- -- -- -- -- -- -- -- (P : A ≡ B) (Q : B ≡ C) (a : A) +-- -- -- -- -- -- -- -- → transport (P ∙ Q) a ≡ transport Q (transport P a) +-- -- -- -- -- -- -- -- substComposite2 = {!!} + +-- -- -- -- -- -- -- -- helper1 : transport (λ i → S¹≡S1 i × Int) (north , 0) ≡ (base , 0) +-- -- -- -- -- -- -- -- helper1 = refl +-- -- -- -- -- -- -- -- helper3 : transport (sym (basedΩS¹≡Int base)) 0 ≡ refl +-- -- -- -- -- -- -- -- helper3 = (λ i → transport (symDistr (basedΩS¹≡ΩS¹ base) ΩS¹≡Int i) 0) +-- -- -- -- -- -- -- -- ∙ substComposite2 (sym ΩS¹≡Int) (sym (basedΩS¹≡ΩS¹ base)) 0 +-- -- -- -- -- -- -- -- ∙ (λ i → transport (λ i → basedΩS¹≡ΩS¹ base (~ i)) refl) -- +-- -- -- -- -- -- -- -- ∙ transportRefl ((equiv-proof (basechange-isequiv base) refl) .fst .fst) +-- -- -- -- -- -- -- -- ∙ (λ i → equiv-proof (transport (λ j → isEquiv (refl-conjugation j)) (basedΩS¹→ΩS¹-isequiv i0)) refl .fst .fst) +-- -- -- -- -- -- -- -- ∙ (λ i → {!equiv-proof (transport (λ j → isEquiv (refl-conjugation j)) (basedΩS¹→ΩS¹-isequiv i0)) refl .fst!}) +-- -- -- -- -- -- -- -- ∙ {!basedΩS¹→ΩS¹!} +-- -- -- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- -- -- helper4 : (x : S¹) → Iso.inv test (base , 0) x ≡ x +-- -- -- -- -- -- -- -- helper4 = {!!} +-- -- -- -- -- -- -- -- helper2 : Iso.inv test (transport (λ i → S¹≡S1 i × Int) (north , 0)) ≡ λ x → x +-- -- -- -- -- -- -- -- helper2 = (λ i → Iso.inv test (base , 0)) ∙ {!!} ∙ {!!} + +-- -- -- -- -- -- prodId : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A × B} → proj₁ x ≡ proj₁ y → proj₂ x ≡ proj₂ y → x ≡ y +-- -- -- -- -- -- prodId {x = (a , b)} {y = (c , d)} id1 id2 i = (id1 i) , (id2 i) + +-- -- -- -- -- -- fstFunHelper : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- -- → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 _) (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x +-- -- -- -- -- -- fstFunHelper a = MV.Ker-i⊂Im-d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a +-- -- -- -- -- -- (sym (isContrH1Unit×H1Unit .snd _) ∙ (isContrH1Unit×H1Unit .snd _)) +-- -- -- -- -- -- where +-- -- -- -- -- -- isContrH1Unit×H1Unit : isContr (Group.type (×coHomGr 1 Unit Unit)) +-- -- -- -- -- -- isContrH1Unit×H1Unit = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) +-- -- -- -- -- -- , λ {(a , b) → sigmaProdElim {D = λ (x : Σ[ x ∈ Group.type (×coHomGr 1 Unit Unit)] Unit) → (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) ≡ fst x} +-- -- -- -- -- -- (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- -- -- -- -- -- (λ a b _ → prodId (grIso.leftInv (coHomGrUnit≥1 0) ∣ a ∣₀) (grIso.leftInv (coHomGrUnit≥1 0) ∣ b ∣₀)) ((a , b) , tt)} + + + +-- -- -- -- -- -- helperMorph : isMorph intGroup (dirProd intGroup intGroup) λ a → a , (0 - a) +-- -- -- -- -- -- helperMorph = +-- -- -- -- -- -- indIntGroup {G = dirProd intGroup intGroup} +-- -- -- -- -- -- (λ a → a , (0 - a)) +-- -- -- -- -- -- refl +-- -- -- -- -- -- (λ a → prodId refl (helper2 a)) +-- -- -- -- -- -- λ a → prodId refl (helper3 a) +-- -- -- -- -- -- where +-- -- -- -- -- -- helper1 : (a : Int) → predInt (sucInt a) ≡ a +-- -- -- -- -- -- helper1 (pos zero) = refl +-- -- -- -- -- -- helper1 (pos (suc n)) = refl +-- -- -- -- -- -- helper1 (negsuc zero) = refl +-- -- -- -- -- -- helper1 (negsuc (suc n)) = refl + +-- -- -- -- -- -- helper4 : (a : Int) → sucInt (predInt a) ≡ a +-- -- -- -- -- -- helper4 (pos zero) = refl +-- -- -- -- -- -- helper4 (pos (suc n)) = refl +-- -- -- -- -- -- helper4 (negsuc zero) = refl +-- -- -- -- -- -- helper4 (negsuc (suc n)) = refl + +-- -- -- -- -- -- helper2 : (a : Int) → (pos 0 - sucInt a) ≡ predInt (pos 0 - a) +-- -- -- -- -- -- helper2 (pos zero) = refl +-- -- -- -- -- -- helper2 (pos (suc n)) = refl +-- -- -- -- -- -- helper2 (negsuc zero) = refl +-- -- -- -- -- -- helper2 (negsuc (suc n)) = sym (helper1 _) + +-- -- -- -- -- -- helper3 : (a : Int) → (pos 0 - predInt a) ≡ sucInt (pos 0 - a) +-- -- -- -- -- -- helper3 (pos zero) = refl +-- -- -- -- -- -- helper3 (pos (suc zero)) = refl +-- -- -- -- -- -- helper3 (pos (suc (suc n))) = sym (helper4 _) +-- -- -- -- -- -- helper3 (negsuc zero) = refl +-- -- -- -- -- -- helper3 (negsuc (suc n)) = refl + +-- -- -- -- -- -- helper : (a b : Int) → (pos 0 - (a +ℤ b)) ≡ ((pos 0 - a) +ℤ (pos 0 - b)) +-- -- -- -- -- -- helper a (pos zero) = refl +-- -- -- -- -- -- helper (pos zero) (pos (suc n)) = +-- -- -- -- -- -- cong (λ x → pos 0 - sucInt x) (+ℤ-comm (pos zero) (pos n)) +-- -- -- -- -- -- ∙ +ℤ-comm (pos 0 +negsuc n) (pos zero) +-- -- -- -- -- -- helper (pos (suc n₁)) (pos (suc n)) = +-- -- -- -- -- -- {!!} +-- -- -- -- -- -- helper (negsuc n₁) (pos (suc n)) = {!!} +-- -- -- -- -- -- helper a (negsuc n) = {!!} + +-- -- -- -- -- -- fun : morph intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- -- fst fun = MV.d _ _ _ (λ _ → tt) (λ _ → tt) 0 ∘ grIso.inv coHom0-S0 .fst ∘ λ a → a , (0 - a) +-- -- -- -- -- -- snd fun = {!!} +-- -- -- -- -- -- {- compMorph {F = intGroup} {G = dirProd intGroup intGroup} {H = coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))} +-- -- -- -- -- -- ((λ a → a , a) , (λ a b → refl)) +-- -- -- -- -- -- (compMorph {F = dirProd intGroup intGroup} {G = coHomGr 0 (S₊ 0)} {H = coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))} (grIso.inv coHom0-S0) +-- -- -- -- -- -- (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 +-- -- -- -- -- -- , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) .snd -} +-- -- -- -- -- -- {- theIso : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- -- theIso = Iso''→Iso _ _ +-- -- -- -- -- -- (iso'' ((λ x → ∣ (λ {(inl tt) → 0ₖ ; (inr tt) → 0ₖ ; (push a i) → Kn→ΩKn+1 0 x i}) ∣₀) , {!!}) +-- -- -- -- -- -- {!!} +-- -- -- -- -- -- {!MV.d!}) +-- -- -- -- -- -- -} + + + +-- -- -- -- -- -- theIso : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- -- theIso = +-- -- -- -- -- -- Iso''→Iso _ _ +-- -- -- -- -- -- (iso'' fun +-- -- -- -- -- -- (λ x inker → {!MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 +-- -- -- -- -- -- (grIso.inv coHom0-S0 .fst (g , g))!}) +-- -- -- -- -- -- (sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- -- -- -- -- -- λ x → pRec propTruncIsProp +-- -- -- -- -- -- (λ {(a , b) → {!fun!} }) +-- -- -- -- -- -- (fstFunHelper (∣ x ∣₀)))) +-- -- -- -- -- -- where +-- -- -- -- -- -- whoKnows : (a : S₊ 0 → Int) (x : MV.D Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt)) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (λ _ → a north) x +-- -- -- -- -- -- ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 a x +-- -- -- -- -- -- whoKnows a (inl x) = refl +-- -- -- -- -- -- whoKnows a (inr x) = refl +-- -- -- -- -- -- whoKnows a (push north i) = refl +-- -- -- -- -- -- whoKnows a (push south i) j = {!!} + +-- -- -- -- -- -- helper : (a : Int) → (grIso.inv coHom0-S0 .fst (a , a)) ≡ ∣ S0→Int (a , a) ∣₀ +-- -- -- -- -- -- helper a = {!have : + +-- -- -- -- -- -- ∣ +-- -- -- -- -- -- MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 +-- -- -- -- -- -- (S0→Int (x , x)) +-- -- -- -- -- -- ∣₀ +-- -- -- -- -- -- ≡ ∣ (λ _ → ∣ north ∣) ∣₀!} + +-- -- -- -- -- -- helper2 : (a b : Int) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a)) ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (b , b)) +-- -- -- -- -- -- → a ≡ b +-- -- -- -- -- -- helper2 a b id = pRec (isSetInt a b) (λ {(pt , id) → {!!}}) (fstFunHelper ∣ (MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a))) ∣₀) + +-- -- -- -- -- -- idFun : (a : Int) → MV.D Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 +-- -- -- -- -- -- idFun a (inl x) = 0ₖ +-- -- -- -- -- -- idFun a (inr x) = 0ₖ +-- -- -- -- -- -- idFun a (push north i) = Kn→ΩKn+1 zero a i +-- -- -- -- -- -- idFun a (push south i) = Kn→ΩKn+1 zero a i + +-- -- -- -- -- -- helper3 : (a : Int) → (MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a))) ≡ idFun a +-- -- -- -- -- -- helper3 a = funExt λ {(inl x) → refl ; (inr x) → refl ; (push north i) → refl ; (push south i) → refl } + +-- -- -- -- -- -- helper4 : (a b : Int) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a)) ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (b , b)) +-- -- -- -- -- -- → a ≡ b +-- -- -- -- -- -- helper4 a b id = +-- -- -- -- -- -- {!!} +-- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- where +-- -- -- -- -- -- helper5 : {!!} --PathP (λ k → id k (inl tt) ≡ id k (inr tt)) (Kn→ΩKn+1 zero a) (Kn→ΩKn+1 zero a) +-- -- -- -- -- -- helper5 i j = {!id i!} + +-- -- -- -- -- -- -- fun : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) → coHom 0 (S₊ 0) +-- -- -- -- -- -- -- fun a = (pRec {P = Σ[ x ∈ coHom 0 (S₊ 0)] (MV.d _ _ _ (λ _ → tt) (λ _ → tt) 0 x ≡ a) × isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) (MV.Δ _ _ _ (λ _ → tt) (λ _ → tt) 0) x} +-- -- -- -- -- -- -- (λ {(a1 , b) (c , d) → ΣProp≡ (λ x → isOfHLevelProd 1 (setTruncIsSet _ _) propTruncIsProp) +-- -- -- -- -- -- -- (sElim2 {B = λ a1 c → (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a1 ≡ a) +-- -- -- -- -- -- -- → MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 c ≡ a +-- -- -- -- -- -- -- → isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) +-- -- -- -- -- -- -- (λ z → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 z) a1 +-- -- -- -- -- -- -- → isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) +-- -- -- -- -- -- -- (λ z → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 z) c → a1 ≡ c} +-- -- -- -- -- -- -- (λ _ _ → {!!}) +-- -- -- -- -- -- -- (λ a c b1 d1 → pElim (λ _ → isOfHLevelΠ 1 λ _ → setTruncIsSet _ _) +-- -- -- -- -- -- -- λ b2 → pElim (λ _ → setTruncIsSet _ _) +-- -- -- -- -- -- -- λ d2 → {!d2!}) +-- -- -- -- -- -- -- a1 c (proj₁ b) (proj₁ d) (proj₂ b) (proj₂ d))}) +-- -- -- -- -- -- -- (λ {(a , b) → a , b , MV.Ker-d⊂Im-Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a {!!}}) +-- -- -- -- -- -- -- (fstFunHelper a)) .fst -- pRec {!!} {!!} (fstFunHelper a) From 8da07727009d977c8488788962653e6ea34d5ae4 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Wed, 10 Jun 2020 16:01:10 +0200 Subject: [PATCH 42/89] newBranchBackup --- Cubical/ZCohomology/Sn/Sn.agda | 66 +++++++++++++++++++++++++++++++++- 1 file changed, 65 insertions(+), 1 deletion(-) diff --git a/Cubical/ZCohomology/Sn/Sn.agda b/Cubical/ZCohomology/Sn/Sn.agda index e22dd3637..c43262385 100644 --- a/Cubical/ZCohomology/Sn/Sn.agda +++ b/Cubical/ZCohomology/Sn/Sn.agda @@ -953,7 +953,71 @@ coHom2Torus = (setTruncIso (compIso schonfIso (compIso (targetIso test13) funIso)))) (Iso.fun (symIso setTruncOfProdIso) (0ₕ , ∣ g ∣₀))) - helper2' = {!!} + helper2' f g = (λ i → Iso.fun (symIso + (setTruncIso + (compIso schonfIso ((targetIso test13))))) ∣ (λ x → 0ₖ , (f x +ₖ g x)) ∣₀) + ∙ (λ i → ∣ Iso.fun (symIso (compIso schonfIso (targetIso test13))) (λ x → 0ₖ , (f x +ₖ g x)) ∣₀) + ∙ (λ i → ∣ Iso.inv schonfIso (Iso.inv (targetIso test13) ((λ x → 0ₖ , (f x +ₖ g x)))) ∣₀) + ∙ (λ i → ∣ {!!} ∣₀) + ∙ {!!} + ∙ {!!} + where + fun : S₊ 1 × S₊ 1 → hLevelTrunc 4 (S₊ 2) + fun (x , north) = 0ₖ +ₖ 0ₖ + fun (x , south) = 0ₖ +ₖ 0ₖ + fun (x , (merid south i)) = 0ₖ +ₖ (Kn→ΩKn+1 1 (f x +ₖ g x) i) + fun (x , (merid north i)) = 0ₖ +ₖ 0ₖ + + helper : Iso.inv schonfIso (Iso.inv (targetIso test13) ((λ x → 0ₖ , (f x +ₖ g x)))) ≡ fun + helper = funExt λ {(x , north) → refl ; (x , south) → refl ; (x , (merid north i)) → refl ; (x , (merid south i)) → refl} + + helper2 : ∣ Iso.inv schonfIso (Iso.inv (targetIso test13) ((λ x → 0ₖ , (f x +ₖ g x)))) ∣₀ + ≡ (∣ Iso.inv schonfIso (Iso.inv (targetIso test13) ((λ x → 0ₖ , f x))) ∣₀ +ₕ ∣ Iso.inv schonfIso (Iso.inv (targetIso test13) ((λ x → 0ₖ , g x))) ∣₀) + helper2 = + cong ∣_∣₀ + (funExt λ {(x , north) → ((λ i → (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x +ₖ g x) i)) + ∙∙ sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ cong (λ x → ((0ₖ +ₖ 0ₖ) +ₖ x)) (rUnitₖ 0ₖ) + ∙∙ refl) + ; (x , south) → refl + ∙∙ sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ cong (λ x → ((0ₖ +ₖ 0ₖ) +ₖ x)) (rUnitₖ 0ₖ) + ∙∙ (λ i → (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x) i) +ₖ ∣ north ∣ +ₖ Kn→ΩKn+1 1 (g x) i) + ; (x , (merid south i)) j → hcomp (λ k → λ {(j = i0) → ∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x +ₖ g x) (i ∨ (~ k)) + ; (j = i1) → (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x) (i ∧ k)) +ₖ ∣ north ∣ +ₖ Kn→ΩKn+1 1 (g x) (i ∧ k)}) + ((sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ cong (λ x → ((0ₖ +ₖ 0ₖ) +ₖ x)) (rUnitₖ 0ₖ)) j) + ; (x , (merid north u)) → {!!}}) + where + anotherone : (x : _) → cong (0ₖ +ₖ_) (Kn→ΩKn+1 1 (f x +ₖ g x)) ≡ {!!} + anotherone x = {!!} + + helper5 : (x : _) → Iso.inv schonfIso (Iso.inv (targetIso test13) ((λ x → 0ₖ , f x +ₖ g x))) x + ≡ (Iso.inv schonfIso (Iso.inv (targetIso test13) ((λ x → 0ₖ , f x))) x) +ₖ (Iso.inv schonfIso (Iso.inv (targetIso test13) ((λ x → 0ₖ , g x))) x) + helper5 (x , north) = sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i))) + helper5 (x , south) = sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i))) + helper5 (x , merid north i) = sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i))) + helper5 (x , merid south i) j = + {!!} + where + helper13 : sym (sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i)))) + ∙ (λ i → ∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x +ₖ g x) i) ∙ (sym (rUnitₖ (0ₖ +ₖ 0ₖ)) + ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i)))) + ≡ λ i → (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x) i) +ₖ + ∣ north ∣ +ₖ Kn→ΩKn+1 1 (g x) i + helper13 = (λ j → sym ((λ i → (rUnitₖ (lUnitₖ 0ₖ (j ∧ (~ i))) (~ i))) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ lUnitₖ 0ₖ (~ i))) + ∙ lUnit (rUnit ((λ i → lUnitₖ (Kn→ΩKn+1 1 (f x +ₖ g x) i) j)) j) j + ∙ ((λ i → (rUnitₖ (lUnitₖ 0ₖ (j ∧ (~ i))) (~ i)))) ∙ λ i → (0ₖ +ₖ 0ₖ) +ₖ lUnitₖ 0ₖ (~ i)) + ∙ (λ j → sym ((λ i → (rUnitₖ (lUnitₖ 0ₖ (~ i)) (~ i))) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ lUnitₖ 0ₖ (~ i))) ∙ + {!!} ∙ + (λ i → (rUnitₖ (lUnitₖ 0ₖ (~ i)) (~ i))) ∙ λ i → (0ₖ +ₖ 0ₖ) +ₖ lUnitₖ 0ₖ (~ i)) + ∙ {!!} + ∙ {!!} +{- +i = i0 ⊢ (sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i)))) j +i = i1 ⊢ (sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i)))) j +j = i0 ⊢ ∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x +ₖ g x) i +j = i1 ⊢ (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x) i) +ₖ + ∣ north ∣ +ₖ Kn→ΩKn+1 1 (g x) i +-} + helper' : Iso Int (Unit × Int) Iso.inv helper' = proj₂ From bc62fb7f4fccfad37d01f76d49c7b300a3964358 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Wed, 10 Jun 2020 17:56:45 +0200 Subject: [PATCH 43/89] moreCopatterns --- Cubical/Data/Group/Base.agda | 13 +- Cubical/Data/Group/GroupLibrary.agda | 24 +- Cubical/Data/Prod/Properties.agda | 26 + Cubical/Data/Unit/Properties.agda | 16 +- Cubical/Foundations/GroupoidLaws.agda | 2 +- Cubical/Foundations/Isomorphism.agda | 8 + Cubical/HITs/SetTruncation/Properties.agda | 37 + Cubical/HITs/Susp/Properties.agda | 23 + Cubical/HITs/Truncation/Properties.agda | 16 +- Cubical/ZCohomology/Groups/Sn.agda | 1758 +++++++++++++++++ Cubical/ZCohomology/MayerVietoris.agda | 235 --- ...duced.agda => MayerVietorisUnreduced.agda} | 226 ++- Cubical/ZCohomology/altToFreudenthal.agda | 201 -- 13 files changed, 2032 insertions(+), 553 deletions(-) create mode 100644 Cubical/ZCohomology/Groups/Sn.agda delete mode 100644 Cubical/ZCohomology/MayerVietoris.agda rename Cubical/ZCohomology/{MayerVietorisReduced.agda => MayerVietorisUnreduced.agda} (67%) delete mode 100644 Cubical/ZCohomology/altToFreudenthal.agda diff --git a/Cubical/Data/Group/Base.agda b/Cubical/Data/Group/Base.agda index 6b5409181..ccf7397f0 100644 --- a/Cubical/Data/Group/Base.agda +++ b/Cubical/Data/Group/Base.agda @@ -43,7 +43,7 @@ record morph {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') : Type (ℓ-max ℓ ℓ field fun : Group.type G → Group.type H ismorph : isMorph G H fun - + open import Cubical.Data.Sigma hiding (_×_ ; comp) @@ -108,10 +108,9 @@ morph0→0' G' H' f' = 0→0 id H ∎ {- a morphism ϕ satisfies ϕ(- a) = - ϕ(a) -} -morphMinus : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (f : (Group.type G → Group.type H)) - → isMorph G H f - → (g : Group.type G) → f (isGroup.inv (Group.groupStruc G) g) ≡ isGroup.inv (Group.groupStruc H) (f g) -morphMinus G H f mf g = +morphMinus : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (ϕ : morph G H) + → (g : Group.type G) → morph.fun ϕ (isGroup.inv (Group.groupStruc G) g) ≡ isGroup.inv (Group.groupStruc H) (morph.fun ϕ g) +morphMinus G H ϕ g = let idG = isGroup.id (Group.groupStruc G) idH = isGroup.id (Group.groupStruc H) invG = isGroup.inv (Group.groupStruc G) @@ -123,6 +122,8 @@ morphMinus G H f mf g = assocH = isGroup.assoc (Group.groupStruc H) compG = isGroup.comp (Group.groupStruc G) compH = isGroup.comp (Group.groupStruc H) + f = morph.fun ϕ + mf = morph.ismorph ϕ helper : compH (f (invG g)) (f g) ≡ idH helper = sym (mf (invG g) g) ∙ (λ i → f (lCancelG g i)) ∙ morph0→0 G H f mf -- sym (morph (invG g) g) ∙ (λ i → f (lCancelG g i)) ∙ morph0→0 G H f morph in f (invG g) ≡⟨ sym (rUnitH (f (invG g))) ⟩ @@ -347,7 +348,7 @@ Iso''→Iso {A = A} {B = B} is = _+B_ = isGroup.comp (Group.groupStruc B) in Iso''.inj is _ (ϕmf a1 (inv A' a2) - ∙ cong (λ x → comp B' (ϕ a1) x) (morphMinus A B ϕ ϕmf a2) + ∙ cong (λ x → comp B' (ϕ a1) x) (morphMinus A B (Iso''.ϕ is) a2) ∙ cong (λ x → comp B' x (inv B' (ϕ a2))) (pf1 ∙ sym pf2) ∙ rCancel B' (ϕ a2)) fstId : a1 ≡ a2 diff --git a/Cubical/Data/Group/GroupLibrary.agda b/Cubical/Data/Group/GroupLibrary.agda index 75ef82d0c..73efef0f0 100644 --- a/Cubical/Data/Group/GroupLibrary.agda +++ b/Cubical/Data/Group/GroupLibrary.agda @@ -195,21 +195,23 @@ isGroup.lCancel (Group.groupStruc (dirProd A B)) (a1 , b1) i = isGroup.rCancel (Group.groupStruc (dirProd A B)) (a1 , b1) i = (isGroup.rCancel (Group.groupStruc A) a1 i) , (isGroup.rCancel (Group.groupStruc B) b1 i) +×morph : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group ℓ} {B : Group ℓ'} {C : Group ℓ''} {D : Group ℓ'''} → morph A B → morph C D → morph (dirProd A C) (dirProd B D) +morph.fun (×morph mf1 mf2) = + (λ {(a , b) → (morph.fun mf1 a) , morph.fun mf2 b}) +morph.ismorph (×morph mf1 mf2) = + (λ {(a , b) (c , d) i → morph.ismorph mf1 a c i , morph.ismorph mf2 b d i}) dirProdIso : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group ℓ} {B : Group ℓ'} {C : Group ℓ''} {D : Group ℓ'''} → Iso A C → Iso B D → Iso (dirProd A B) (dirProd C D) -dirProdIso {A = A'} {B = B'} {C = C'} {D = D'} isoAC isoBD = - Iso'→Iso - (iso' - (I.iso - (λ a → (morph.fun (Iso.fun isoAC) (proj₁ a)) , (morph.fun (Iso.fun isoBD) (proj₂ a))) - (λ a → (morph.fun (Iso.inv isoAC) (proj₁ a)) , (morph.fun (Iso.inv isoBD) (proj₂ a))) - (λ a → ×≡ (Iso.rightInv isoAC (proj₁ a)) (Iso.rightInv isoBD (proj₂ a))) - λ a → ×≡ (Iso.leftInv isoAC (proj₁ a)) (Iso.leftInv isoBD (proj₂ a))) - λ {(_ , _) (_ , _) → ×≡ (morph.ismorph (Iso.fun isoAC) _ _) (morph.ismorph (Iso.fun isoBD) _ _)}) - - +Iso.fun (dirProdIso {A = A'} {B = B'} {C = C'} {D = D'} isoAC isoBD) = + ×morph (Iso.fun isoAC) (Iso.fun isoBD) +Iso.inv (dirProdIso {A = A'} {B = B'} {C = C'} {D = D'} isoAC isoBD) = + ×morph (Iso.inv isoAC) (Iso.inv isoBD) +Iso.rightInv (dirProdIso {A = A'} {B = B'} {C = C'} {D = D'} isoAC isoBD) (a , b) = + ×≡ (Iso.rightInv isoAC a) (Iso.rightInv isoBD b) +Iso.leftInv (dirProdIso {A = A'} {B = B'} {C = C'} {D = D'} isoAC isoBD) (a , b) = + ×≡ (Iso.leftInv isoAC a) (Iso.leftInv isoBD b) {- Given the following diagram diff --git a/Cubical/Data/Prod/Properties.agda b/Cubical/Data/Prod/Properties.agda index af78ac624..bc1bd8499 100644 --- a/Cubical/Data/Prod/Properties.agda +++ b/Cubical/Data/Prod/Properties.agda @@ -83,3 +83,29 @@ isOfHLevelProd {A = A} {B = B} n h1 h2 = ε : retract φ ψ ε (a , b) i = secEq f a i , secEq g b i + + +{- Some simple ismorphisms -} + +prodIso : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} + → Iso A C + → Iso B D + → Iso (A × B) (C × D) +Iso.fun (prodIso iAC iBD) (a , b) = (Iso.fun iAC a) , Iso.fun iBD b +Iso.inv (prodIso iAC iBD) (c , d) = (Iso.inv iAC c) , Iso.inv iBD d +Iso.rightInv (prodIso iAC iBD) (c , d) = ×≡ (Iso.rightInv iAC c) (Iso.rightInv iBD d) +Iso.leftInv (prodIso iAC iBD) (a , b) = ×≡ (Iso.leftInv iAC a) (Iso.leftInv iBD b) + +toProdIso : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + → Iso (A → B × C) ((A → B) × (A → C)) +Iso.fun toProdIso = λ f → (λ a → proj₁ (f a)) , (λ a → proj₂ (f a)) +Iso.inv toProdIso (f , g) = λ a → (f a) , (g a) +Iso.rightInv toProdIso (f , g) = refl +Iso.leftInv toProdIso b = funExt λ a → sym (×-η _) + +schönfinkelIso : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + → Iso (A × B → C) (A → B → C) +Iso.fun schönfinkelIso f a b = f (a , b) +Iso.inv schönfinkelIso f (a , b) = f a b +Iso.rightInv schönfinkelIso a = refl +Iso.leftInv schönfinkelIso f = funExt λ {(a , b) → refl} diff --git a/Cubical/Data/Unit/Properties.agda b/Cubical/Data/Unit/Properties.agda index 5cd706783..f2aac251a 100644 --- a/Cubical/Data/Unit/Properties.agda +++ b/Cubical/Data/Unit/Properties.agda @@ -30,9 +30,13 @@ diagonal-unit : Unit ≡ Unit × Unit diagonal-unit = isoToPath (iso (λ x → tt , tt) (λ x → tt) (λ {(tt , tt) i → tt , tt}) λ {tt i → tt}) fibId : ∀ {ℓ} (A : Type ℓ) → (fiber (λ (x : A) → tt) tt) ≡ A -fibId A = isoToPath - (iso fst - (λ a → a , refl) - (λ _ → refl) - (λ a i → fst a - , isOfHLevelSuc 1 isPropUnit _ _ (snd a) refl i)) +fibId A = + isoToPath + (iso fst + (λ a → a , refl) + (λ _ → refl) + (λ a i → fst a + , isOfHLevelSuc 1 isPropUnit _ _ (snd a) refl i)) + +isContr→≡Unit : {A : Type₀} → isContr A → A ≡ Unit +isContr→≡Unit contr = isoToPath (iso (λ _ → tt) (λ _ → fst contr) (λ _ → refl) λ _ → snd contr _) diff --git a/Cubical/Foundations/GroupoidLaws.agda b/Cubical/Foundations/GroupoidLaws.agda index f6b280384..15bfa3c03 100644 --- a/Cubical/Foundations/GroupoidLaws.agda +++ b/Cubical/Foundations/GroupoidLaws.agda @@ -299,7 +299,7 @@ congFunct-dep : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} {x y z : A} → PathP (λ i → PathP (λ j → B (compPath-filler p q i j)) (f x) (f (q i))) (cong f p) (cong f (p ∙ q)) congFunct-dep {B = B} {x = x} f p q i j = f (compPath-filler p q i j) -cong₂Funct : (f : A → A → A) → +cong₂Funct : ∀ {ℓ ℓ'} {A : Type ℓ} {x y : A} {B : Type ℓ'} (f : A → A → B) → (p : x ≡ y) → {u v : A} (q : u ≡ v) → cong₂ f p q ≡ cong (λ x → f x u) p ∙ cong (f y) q diff --git a/Cubical/Foundations/Isomorphism.agda b/Cubical/Foundations/Isomorphism.agda index 9604036b1..b87b15cb5 100644 --- a/Cubical/Foundations/Isomorphism.agda +++ b/Cubical/Foundations/Isomorphism.agda @@ -143,3 +143,11 @@ Iso.fun (invIso isom) = Iso.inv isom Iso.inv (invIso isom) = Iso.fun isom Iso.rightInv (invIso isom) = Iso.leftInv isom Iso.leftInv (invIso isom) = Iso.rightInv isom + +codomainIso : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + → Iso B C + → Iso (A → B) (A → C) +Iso.fun (codomainIso is) f a = Iso.fun is (f a) +Iso.inv (codomainIso is) f a = Iso.inv is (f a) +Iso.rightInv (codomainIso is) f = funExt λ a → Iso.rightInv is (f a) +Iso.leftInv (codomainIso is) f = funExt λ a → Iso.leftInv is (f a) diff --git a/Cubical/HITs/SetTruncation/Properties.agda b/Cubical/HITs/SetTruncation/Properties.agda index abf36a839..51c6173e8 100644 --- a/Cubical/HITs/SetTruncation/Properties.agda +++ b/Cubical/HITs/SetTruncation/Properties.agda @@ -77,6 +77,21 @@ setTruncIdempotent≃ {A = A} hA = isoToEquiv f setTruncIdempotent : isSet A → ∥ A ∥₀ ≡ A setTruncIdempotent hA = ua (setTruncIdempotent≃ hA) +isContr→isContrSetTrunc : ∀ {ℓ} {A : Type ℓ} → isContr A → isContr (∥ A ∥₀) +isContr→isContrSetTrunc contr = ∣ fst contr ∣₀ + , elim (λ _ → isOfHLevelPath 2 (setTruncIsSet) _ _) + λ a → cong ∣_∣₀ (snd contr a) + + +setTruncIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → Iso A B → Iso ∥ A ∥₀ ∥ B ∥₀ +Iso.fun (setTruncIso is) = rec setTruncIsSet (λ x → ∣ Iso.fun is x ∣₀) +Iso.inv (setTruncIso is) = rec setTruncIsSet (λ x → ∣ Iso.inv is x ∣₀) +Iso.rightInv (setTruncIso is) = + elim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ a → cong ∣_∣₀ (Iso.rightInv is a) +Iso.leftInv (setTruncIso is) = + elim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ a → cong ∣_∣₀ (Iso.leftInv is a) setSigmaIso : ∀ {ℓ} {B : A → Type ℓ} → Iso ∥ Σ A B ∥₀ ∥ Σ A (λ x → ∥ B x ∥₀) ∥₀ setSigmaIso {A = A} {B = B} = iso fun funinv sect retr @@ -112,3 +127,25 @@ sigmaProdElim {B = B} {C = C} {D = D} set g ((x , y) , c) = (λ x → elim (λ _ → isOfHLevelΠ 2 λ _ → set _) λ y c → g x y c) x y c + + +prodElim : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : ∥ A ∥₀ × ∥ B ∥₀ → Type ℓ''} + → ((x : ∥ A ∥₀ × ∥ B ∥₀) → isOfHLevel 2 (C x)) + → ((a : A) (b : B) → C (∣ a ∣₀ , ∣ b ∣₀)) + → (x : ∥ A ∥₀ × ∥ B ∥₀) → C x +prodElim {A = A} {B = B} {C = C} hlevel ind (a , b) = schonf a b + where + schonf : (a : ∥ A ∥₀) (b : ∥ B ∥₀) → C (a , b) + schonf = elim (λ x → isOfHLevelΠ 2 λ y → hlevel (_ , _)) λ a → elim (λ x → hlevel (_ , _)) + λ b → ind a b + +setTruncOfProdIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → Iso ∥ A × B ∥₀ (∥ A ∥₀ × ∥ B ∥₀) +Iso.fun setTruncOfProdIso = rec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) λ { (a , b) → ∣ a ∣₀ , ∣ b ∣₀ } +Iso.inv setTruncOfProdIso = prodElim (λ _ → setTruncIsSet) λ a b → ∣ a , b ∣₀ +Iso.rightInv setTruncOfProdIso = + prodElim (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ _ _ → refl +Iso.leftInv setTruncOfProdIso = + elim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ {(a , b) → refl} + diff --git a/Cubical/HITs/Susp/Properties.agda b/Cubical/HITs/Susp/Properties.agda index ed7e9c05e..b26684c27 100644 --- a/Cubical/HITs/Susp/Properties.agda +++ b/Cubical/HITs/Susp/Properties.agda @@ -2,6 +2,7 @@ module Cubical.HITs.Susp.Properties where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv @@ -56,3 +57,25 @@ congSuspEquiv {ℓ} {A} {B} h = isoToEquiv isom Iso.leftInv isom north = refl Iso.leftInv isom south = refl Iso.leftInv isom (merid a i) j = merid (secEq h a j) i + +suspToPropRec : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Susp A → Type ℓ'} (a : A) + → ((x : Susp A) → isProp (B x)) + → B north + → (x : Susp A) → B x +suspToPropRec a isProp Bnorth north = Bnorth +suspToPropRec {B = B} a isProp Bnorth south = subst B (merid a) Bnorth +suspToPropRec {B = B} a isProp Bnorth (merid a₁ i) = + hcomp (λ k → λ {(i = i0) → Bnorth ; + (i = i1) → isProp + south + (subst B (merid a₁) Bnorth) + (subst B (merid a) Bnorth) k}) + (transp (λ j → B (merid a₁ (j ∧ i))) (~ i) Bnorth) + +suspToPropRec2 : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Susp A → Susp A → Type ℓ'} (a : A) + → ((x y : Susp A) → isProp (B x y)) + → B north north + → (x y : Susp A) → B x y +suspToPropRec2 a isProp Bnorth = + suspToPropRec a (λ x → isOfHLevelΠ 1 λ y → isProp x y) + (suspToPropRec a (λ x → isProp north x) Bnorth) diff --git a/Cubical/HITs/Truncation/Properties.agda b/Cubical/HITs/Truncation/Properties.agda index 64726af9e..e3f45c243 100644 --- a/Cubical/HITs/Truncation/Properties.agda +++ b/Cubical/HITs/Truncation/Properties.agda @@ -198,8 +198,12 @@ map g = rec (isOfHLevelTrunc _) (λ a → ∣ g a ∣) mapCompIso : {n : ℕ} {B : Type ℓ'} → (Iso A B) → Iso (hLevelTrunc n A) (hLevelTrunc n B) Iso.fun (mapCompIso g) = recElim (isOfHLevelTrunc _) λ a → ∣ Iso.fun g a ∣ Iso.inv (mapCompIso g) = recElim (isOfHLevelTrunc _) λ b → ∣ Iso.inv g b ∣ -Iso.rightInv (mapCompIso g) = elim (λ x → isOfHLevelPath _ (isOfHLevelTrunc _) _ _) λ b → cong ∣_∣ (Iso.rightInv g b) -Iso.leftInv (mapCompIso g) = elim (λ x → isOfHLevelPath _ (isOfHLevelTrunc _) _ _) λ a → cong ∣_∣ (Iso.leftInv g a) +Iso.rightInv (mapCompIso g) = + elim (λ x → isOfHLevelPath _ (isOfHLevelTrunc _) _ _) + λ b → cong ∣_∣ (Iso.rightInv g b) +Iso.leftInv (mapCompIso g) = + elim (λ x → isOfHLevelPath _ (isOfHLevelTrunc _) _ _) + λ a → cong ∣_∣ (Iso.leftInv g a) mapId : {n : ℕ} → ∀ t → map {n = n} (idfun A) t ≡ t @@ -262,6 +266,12 @@ Iso.leftInv 2GroupoidTrunc≃Trunc2Iso = 2GpdTrunc.elim (λ _ → isOfHLevelPath 2GroupoidTrunc≡Trunc2 : ∥ A ∥₂ ≡ ∥ A ∥ 2 2GroupoidTrunc≡Trunc2 = ua 2GroupoidTrunc≃Trunc2 + +isContr→isContrTrunc : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → isContr A → isContr (hLevelTrunc n A) +isContr→isContrTrunc n contr = ∣ fst contr ∣ , (elim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → cong ∣_∣ (snd contr a)) + + + ---- ∥ Ω A ∥ ₙ ≡ Ω ∥ A ∥ₙ₊₁ ---- abstract @@ -382,3 +392,5 @@ Iso.leftInv (truncOfTruncIso n m) = elim (λ x → isOfHLevelPath n (isOfHLevelT truncOfTruncEq : (n m : ℕ) → (hLevelTrunc n A) ≃ (hLevelTrunc n (hLevelTrunc (m + n) A)) truncOfTruncEq n m = isoToEquiv (truncOfTruncIso n m) + + diff --git a/Cubical/ZCohomology/Groups/Sn.agda b/Cubical/ZCohomology/Groups/Sn.agda new file mode 100644 index 000000000..ac11a8579 --- /dev/null +++ b/Cubical/ZCohomology/Groups/Sn.agda @@ -0,0 +1,1758 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.Groups.Sn where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.S1.S1 +open import Cubical.HITs.S1 +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.MayerVietorisUnreduced +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence +open import Cubical.Data.NatMinusTwo.Base +open import Cubical.Data.Empty +open import Cubical.Data.Sigma hiding (_×_) +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₋₁ ; ∣_∣ to ∣_∣₋₁) +open import Cubical.HITs.Nullification +open import Cubical.Data.Int renaming (_+_ to _+ℤ_; +-comm to +ℤ-comm ; +-assoc to +ℤ-assoc) +open import Cubical.Data.Nat +open import Cubical.Data.Prod +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; recElim to trRec ; rec to trRec2 ; elim3 to trElim3) +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Freudenthal +open import Cubical.HITs.SmashProduct.Base +open import Cubical.Data.Unit +open import Cubical.Data.Group.Base renaming (Iso to grIso ; compIso to compGrIso ; invIso to invGrIso ; idIso to idGrIso) +open import Cubical.Data.Group.GroupLibrary +open import Cubical.ZCohomology.coHomZero-map +open import Cubical.HITs.Wedge + +open import Cubical.Foundations.Equiv.HalfAdjoint + + +open import Cubical.ZCohomology.KcompPrelims + + +open import Cubical.HITs.Pushout +open import Cubical.Data.Sum.Base +open import Cubical.Data.HomotopyGroup +open import Cubical.Data.Bool hiding (_⊕_) + +{- +cong₂Sym : ∀ {ℓ} {A : Type ℓ} {x : A} + (p : x ≡ x) → + (refl ≡ p) → + (P : p ≡ p) → + cong₂ (λ x y → x ∙ y) P (sym P) ≡ refl +cong₂Sym {x = x} p = J (λ p _ → (P : p ≡ p) → cong₂ (λ x y → x ∙ y) P (sym P) ≡ refl) + λ P → cong₂Funct (λ x y → x ∙ y) P (sym P) + ∙ PathP→compPathR (λ j → cong (λ x → rUnit x (~ j)) P ∙ cong (λ x → lUnit x (~ j)) (sym P)) + ∙ (λ j → (sym (rUnit refl)) ∙ rCancel P j ∙ lUnit refl) + ∙ (λ j → sym (rUnit refl) ∙ lUnit (lUnit refl) (~ j) ) + ∙ rCancel (sym (rUnit refl)) + +abstract + cong₂Sym1 : ∀ {ℓ} {A : Type ℓ} {x : A} + (p : x ≡ x) → + (refl ≡ p) → + (P : p ≡ p) → + cong (λ x → x ∙ p) P ≡ cong (λ x → p ∙ x) P + cong₂Sym1 {x = x} p id P = rUnit _ + ∙ (λ i → cong (λ x → x ∙ p) P ∙ lCancel (λ i → p ∙ P i) (~ i)) -- cong (λ x → cong (λ x → x ∙ p) P ∙ x) {!!} + ∙ assoc _ _ _ + ∙ (λ j → cong₂Funct (λ x y → x ∙ y) P (sym P) (~ j) ∙ (λ i → p ∙ P i)) + ∙ (λ j → cong₂Sym p id P j ∙ (λ i → p ∙ P i)) + ∙ sym (lUnit _) +-} + + +test13 : Iso (S₊ 1 → hLevelTrunc 4 (S₊ 2)) (hLevelTrunc 4 (S₊ 2) × hLevelTrunc 3 (S₊ 1)) +Iso.fun test13 f = f north , ΩKn+1→Kn (sym (rCancelₖ (f north)) + ∙ (λ i → f (merid south i) +ₖ (-ₖ f (merid north i))) + ∙ rCancelₖ (f south)) +Iso.inv test13 (a , b) north = a +ₖ 0ₖ +Iso.inv test13 (a , b) south = a +ₖ 0ₖ +Iso.inv test13 (a , b) (merid south i) = a +ₖ (Kn→ΩKn+1 1 b i) +Iso.inv test13 (a , b) (merid north i) = a +ₖ 0ₖ +Iso.rightInv test13 (a , b) = + ×≡ (rUnitₖ a) + ((cong ΩKn+1→Kn (congHelper++ (Kn→ΩKn+1 1 b) (λ x → (a +ₖ x) +ₖ (-ₖ (a +ₖ 0ₖ))) (funExt (λ x → sym (cancelHelper a x))) (rCancelₖ (a +ₖ 0ₖ)))) + ∙ Iso.leftInv (Iso3-Kn-ΩKn+1 1) b) + module _ where + cancelHelper : (a b : hLevelTrunc 4 (S₊ 2)) → (a +ₖ b) +ₖ (-ₖ (a +ₖ 0ₖ)) ≡ b + cancelHelper a b = + (a +ₖ b) +ₖ (-ₖ (a +ₖ 0ₖ)) ≡⟨ (λ i → (a +ₖ b) +ₖ (-ₖ (rUnitₖ a i))) ⟩ + (a +ₖ b) +ₖ (-ₖ a) ≡⟨ cong (λ x → x +ₖ (-ₖ a)) (commₖ a b) ⟩ + (b +ₖ a) +ₖ (-ₖ a) ≡⟨ assocₖ b a (-ₖ a) ⟩ + b +ₖ a +ₖ (-ₖ a) ≡⟨ cong (λ x → b +ₖ x) (rCancelₖ a) ⟩ + b +ₖ 0ₖ ≡⟨ rUnitₖ b ⟩ + b ∎ + + abstract + commHelper : (p q : Path (hLevelTrunc 4 (S₊ 2)) 0ₖ 0ₖ) → p ∙ q ≡ q ∙ p + commHelper p q = + cong₂ _∙_ (sym (Iso.rightInv (Iso3-Kn-ΩKn+1 1) p)) + (sym (Iso.rightInv (Iso3-Kn-ΩKn+1 1) q)) + ∙ sym (Iso.rightInv (Iso3-Kn-ΩKn+1 1) (Kn→ΩKn+1 1 (ΩKn+1→Kn p) ∙ Kn→ΩKn+1 1 (ΩKn+1→Kn q))) + ∙ cong (Kn→ΩKn+1 1) (commₖ (ΩKn+1→Kn p) (ΩKn+1→Kn q)) + ∙ Iso.rightInv (Iso3-Kn-ΩKn+1 1) (Kn→ΩKn+1 1 (ΩKn+1→Kn q) ∙ Kn→ΩKn+1 1 (ΩKn+1→Kn p)) + ∙ sym (cong₂ _∙_ (sym (Iso.rightInv (Iso3-Kn-ΩKn+1 1) q)) + (sym (Iso.rightInv (Iso3-Kn-ΩKn+1 1) p))) + + moveabout : {x : hLevelTrunc 4 (S₊ 2)} (p q : x ≡ 0ₖ) (mid : 0ₖ ≡ 0ₖ) + → sym q ∙ (p ∙ mid ∙ sym p) ∙ q ≡ mid + moveabout p q mid = assoc (sym q) _ _ + ∙ cong (_∙ q) (assoc (sym q) p (mid ∙ sym p)) + ∙ sym (assoc (sym q ∙ p) (mid ∙ sym p) q) + ∙ cong ((sym q ∙ p) ∙_) (sym (assoc mid (sym p) q)) + ∙ cong (λ x → (sym q ∙ p) ∙ (mid ∙ x)) (sym (symDistr (sym q) p)) + ∙ cong ((sym q ∙ p)∙_) (commHelper mid _) + ∙ assoc _ _ _ + ∙ cong (_∙ mid) (rCancel (sym q ∙ p)) + ∙ sym (lUnit mid) + + congHelper : ∀ {ℓ} {A : Type ℓ} {a1 : A} (p : a1 ≡ a1) (f : A → A) (id : (λ x → x) ≡ f) + → cong f p ≡ sym (funExt⁻ id a1) ∙ p ∙ funExt⁻ id a1 + congHelper {a1 = a1} p f id = + (λ i → lUnit (rUnit (cong f p) i) i) + ∙ (λ i → (λ j → id ((~ i) ∨ (~ j)) a1) ∙ cong (id (~ i)) p ∙ λ j → id (~ i ∨ j) a1) + + + congHelper++ : (p : 0ₖ ≡ 0ₖ) (f : hLevelTrunc 4 (S₊ 2) → hLevelTrunc 4 (S₊ 2)) (id : (λ x → x) ≡ f) + → (q : (f 0ₖ) ≡ 0ₖ) + → (sym q) ∙ cong f p ∙ q ≡ p + congHelper++ p f id q = + cong (λ x → sym q ∙ x ∙ q) (congHelper p f id) ∙ + moveabout (sym (funExt⁻ id ∣ north ∣)) q p + +Iso.leftInv test13 a = + funExt λ {north → rUnitₖ (a north) + ; south → rUnitₖ (a north) ∙ cong a (merid north) + ; (merid south i) → test i + ; (merid north i) → test2 i} + where + test : PathP (λ i → a north +ₖ Kn→ΩKn+1 1 (ΩKn+1→Kn (sym (rCancelₖ (a north)) + ∙ (λ i → a (merid south i) +ₖ (-ₖ a (merid north i))) + ∙ rCancelₖ (a south))) i ≡ a (merid south i)) (rUnitₖ (a north)) (rUnitₖ (a north) ∙ cong a (merid north)) + test j i = + hcomp (λ k → λ { (i = i0) → (helper ∙ together ∙ congFunct (_+ₖ 0ₖ) (cong a (merid south)) (cong a (sym (merid north)))) (~ k) j + ; (i = i1) → a (merid south j) + ; (j = i0) → rUnitₖ (a north) i + ; (j = i1) → ((λ j → rUnitₖ (a (merid north ((~ k) ∧ j))) (j ∧ k)) ∙ λ j → rUnitₖ (a (merid north ((~ k) ∨ j))) (k ∨ j)) i }) + (hcomp (λ k → λ { (i = i1) → a (merid south j) + ; (i = i0) → compPath-filler ((cong (λ x → a x +ₖ 0ₖ) (merid south))) (cong (λ x → a x +ₖ 0ₖ) (sym (merid north))) k j + ; (j = i0) → rUnitₖ (a north) i + ; (j = i1) → pathFiller2 (cong (_+ₖ 0ₖ) (cong a (merid north))) (rUnitₖ (a south)) k i}) + (rUnitₖ (a (merid south j)) i)) + + + where + pathFiller : (rUnitₖ (a north) ∙ cong a (merid north)) ≡ cong (_+ₖ 0ₖ) (cong a (merid north)) ∙ rUnitₖ (a south) + pathFiller = (λ i → (λ j → rUnitₖ (a (merid north (i ∧ j))) (j ∧ (~ i))) ∙ λ j → rUnitₖ (a (merid north (i ∨ j))) (~ i ∨ j)) + + pathFiller2 : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) → PathP (λ i → p (~ i) ≡ z) q (p ∙ q) + pathFiller2 p q i = + hcomp (λ k → λ {(i = i0) → lUnit q (~ k) + ; (i = i1) → p ∙ q}) + ((λ j → p (~ i ∨ j)) ∙ q) + + helper : Path (_ ≡ _) (λ i → a north +ₖ Kn→ΩKn+1 1 (ΩKn+1→Kn (sym (rCancelₖ (a north)) + ∙ (λ i → a (merid south i) +ₖ (-ₖ a (merid north i))) + ∙ rCancelₖ (a south))) i) + --- + (cong (a north +ₖ_) (sym (rCancelₖ (a north))) + ∙ ((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north))) + ∙ cong (a north +ₖ_) (rCancelₖ (a north))) + helper = + (λ j i → a north +ₖ Iso.rightInv (Iso3-Kn-ΩKn+1 1) ((sym (rCancelₖ (a north)) + ∙ (λ i → a (merid south i) +ₖ (-ₖ a (merid north i))) + ∙ rCancelₖ (a south))) j i) ∙ + congFunct (a north +ₖ_) (sym (rCancelₖ (a north))) ((λ i → a (merid south i) +ₖ (-ₖ a (merid north i))) ∙ rCancelₖ (a south)) ∙ + (λ j → cong (a north +ₖ_) (sym (rCancelₖ (a north))) ∙ congFunct (a north +ₖ_) ((λ i → a (merid south i) +ₖ (-ₖ a (merid north i)))) (rCancelₖ (a south)) j) ∙ + (λ j → cong (a north +ₖ_) (sym (rCancelₖ (a north))) ∙ cong₂Funct (λ x y → a north +ₖ a x +ₖ (-ₖ (a y))) (merid south) (merid north) j ∙ cong (a north +ₖ_) ((rCancelₖ (a south)))) ∙ + (λ i → cong (a north +ₖ_) (sym (rCancelₖ (a north))) ∙ ((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ i ∨ (~ j))) +ₖ (-ₖ (a (merid north (j ∧ (~ i)))))) ∙ cong (a north +ₖ_) ((rCancelₖ (a (merid north (~ i)))))) + + abstract + pathHelper : (a b : hLevelTrunc 4 (S₊ 2)) → a +ₖ b +ₖ (-ₖ a) ≡ b +ₖ 0ₖ + pathHelper a b = + a +ₖ b +ₖ (-ₖ a) ≡⟨ commₖ a (b +ₖ (-ₖ a)) ⟩ + (b +ₖ (-ₖ a)) +ₖ a ≡⟨ assocₖ b (-ₖ a) a ⟩ + b +ₖ (-ₖ a) +ₖ a ≡⟨ cong (b +ₖ_) (lCancelₖ a) ⟩ + b +ₖ 0ₖ ∎ + + + helperFun : ∀ {ℓ} {A : Type ℓ} {a b : A} (p p' : a ≡ b) (q q' : b ≡ b) (r : a ≡ a) + → ((p q : a ≡ a) → p ∙ q ≡ q ∙ p) + → q ≡ q' + → PathP (λ i → p' (~ i) ≡ p' (~ i)) q' r + → p ∙ q ∙ sym p ≡ r + helperFun p p' q q' r comm qid dep = + p ∙ q ∙ sym p ≡⟨ cong (λ x → p ∙ x ∙ sym p) qid ⟩ + p ∙ q' ∙ sym p ≡⟨ cong (λ x → p ∙ x ∙ sym p) (λ i → lUnit (rUnit q' i) i) ⟩ + p ∙ (refl ∙ q' ∙ refl) ∙ sym p ≡⟨ cong (λ x → p ∙ x ∙ sym p) (λ i → (λ j → p' (~ i ∨ ~ j)) ∙ dep i ∙ λ j → p' (~ i ∨ j)) ⟩ + p ∙ (sym p' ∙ r ∙ p') ∙ sym p ≡⟨ assoc p (sym p' ∙ r ∙ p') (sym p) ⟩ + (p ∙ sym p' ∙ r ∙ p') ∙ sym p ≡⟨ cong (_∙ sym p) (assoc p (sym p') (r ∙ p')) ⟩ + ((p ∙ sym p') ∙ r ∙ p') ∙ sym p ≡⟨ sym (assoc (p ∙ sym p') (r ∙ p') (sym p)) ⟩ + (p ∙ sym p') ∙ (r ∙ p') ∙ sym p ≡⟨ cong ((p ∙ sym p') ∙_) (sym (assoc r p' (sym p))) ⟩ + (p ∙ sym p') ∙ r ∙ (p' ∙ sym p) ≡⟨ cong (λ x → (p ∙ sym p') ∙ r ∙ x) (sym (symDistr p (sym p'))) ⟩ + (p ∙ sym p') ∙ r ∙ sym (p ∙ sym p') ≡⟨ cong ((p ∙ sym p') ∙_) (comm r (sym (p ∙ sym p'))) ⟩ + (p ∙ sym p') ∙ sym (p ∙ sym p') ∙ r ≡⟨ assoc (p ∙ sym p') (sym (p ∙ sym p')) r ⟩ + ((p ∙ sym p') ∙ sym (p ∙ sym p')) ∙ r ≡⟨ cong (_∙ r) (rCancel (p ∙ sym p')) ⟩ + refl ∙ r ≡⟨ sym (lUnit r) ⟩ + r ∎ + together : Path (_ ≡ _) (cong (a north +ₖ_) (sym (rCancelₖ (a north))) + ∙ ((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north))) + ∙ cong (a north +ₖ_) (rCancelₖ (a north))) + (cong(_+ₖ 0ₖ) (cong a (merid south) ∙ cong a (sym (merid north)))) + together = + helperFun (cong (a north +ₖ_) (sym (rCancelₖ (a north)))) + (λ i → pathHelper (a north) (a north) (~ i)) + ((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north))) + (((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north)))) + (cong(_+ₖ 0ₖ) (cong a (merid south) ∙ cong a (sym (merid north)))) + (transport (λ i → (p q : rUnitₖ (a north) (~ i) ≡ rUnitₖ (a north) (~ i)) → p ∙ q ≡ q ∙ p) + (λ p q → isCommS2 _ p q)) + refl + λ i → hcomp (λ k → λ {(i = i0) → (λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north)) + ; (i = i1) → congFunct (_+ₖ 0ₖ) (cong a (merid south)) (cong a (sym (merid north))) (~ k)}) + ((λ j → pathHelper (a north) (a (merid south j)) i) ∙ (λ j → pathHelper (a north) (a (merid north (~ j))) i)) + where + isCommS2 : (x : hLevelTrunc 4 (S₊ 2)) (p q : x ≡ x) → p ∙ q ≡ q ∙ p + isCommS2 x p q = sym (transport⁻Transport (typId x) (p ∙ q)) + ∙ cong (transport (λ i → typId x (~ i))) (typIdHelper p q) + ∙ (λ i → transport (λ i → typId x (~ i)) (commHelper ∣ north ∣ ∣ north ∣ (transport (λ i → typId x i) p) (transport (λ i → typId x i) q) i)) + ∙ cong (transport (λ i → typId x (~ i))) (sym (typIdHelper q p)) + ∙ transport⁻Transport (typId x) (q ∙ p) + where + congIsoHelper : (x : hLevelTrunc 4 (S₊ 2)) → Iso (hLevelTrunc 4 (S₊ 2)) (hLevelTrunc 4 (S₊ 2)) + Iso.fun (congIsoHelper x) = _+ₖ x + Iso.inv (congIsoHelper x) = _+ₖ (-ₖ x) + Iso.rightInv (congIsoHelper x) a = assocₖ a (-ₖ x) x ∙ cong (a +ₖ_) (lCancelₖ x) ∙ rUnitₖ a + Iso.leftInv (congIsoHelper x) a = assocₖ a x (-ₖ x) ∙ cong (a +ₖ_) (rCancelₖ x) ∙ rUnitₖ a + typId : (x : hLevelTrunc 4 (S₊ 2)) → (x ≡ x) ≡ Path (hLevelTrunc 4 (S₊ 2)) 0ₖ 0ₖ + typId x = isoToPath ((congIso (isoToEquiv (symIso (congIsoHelper x))))) ∙ λ i → rCancelₖ x i ≡ rCancelₖ x i + where + + typIdHelper : (p q : (x ≡ x)) → transport (λ i → typId x i) (p ∙ q) ≡ transport (λ i → typId x i) p ∙ transport (λ i → typId x i) q + typIdHelper p q = + (substComposite (λ x → x) (isoToPath ((congIso (isoToEquiv (symIso (congIsoHelper x)))))) (λ i → rCancelₖ x i ≡ rCancelₖ x i) (p ∙ q)) + ∙ (λ i → transport (λ i → rCancelₖ x i ≡ rCancelₖ x i) (transport (isoToPath (congIso (isoToEquiv (symIso (congIsoHelper x))))) (p ∙ q))) + ∙ (λ i → transport (λ i → rCancelₖ x i ≡ rCancelₖ x i) (transportRefl (congFunct (_+ₖ (-ₖ x)) p q i) i)) + ∙ (λ i → transport (λ i → rCancelₖ x i ≡ rCancelₖ x i) ((lUnit (rUnit (transportRefl (cong (_+ₖ (-ₖ x)) p) (~ i)) i) i) ∙ (lUnit (rUnit (transportRefl (cong (_+ₖ (-ₖ x)) q) (~ i)) i) i))) + ∙ (λ i → transp (λ j → rCancelₖ x (i ∨ j) ≡ rCancelₖ x (i ∨ j)) i (((λ j → rCancelₖ x (i ∧ (~ j))) ∙ transport refl (cong (_+ₖ (-ₖ x)) p) ∙ λ j → rCancelₖ x (i ∧ j)) ∙ (λ j → rCancelₖ x (i ∧ (~ j))) ∙ transport refl (cong (_+ₖ (-ₖ x)) q) ∙ λ j → rCancelₖ x (i ∧ j))) + ∙ (λ i → transp (λ j → rCancelₖ x ((~ i) ∨ j) ≡ rCancelₖ x ((~ i) ∨ j)) (~ i) ((λ j → rCancelₖ x ((~ i) ∧ (~ j))) ∙ transport refl (cong (_+ₖ (-ₖ x)) p) ∙ λ j → rCancelₖ x ((~ i) ∧ j)) + ∙ transp (λ j → rCancelₖ x ((~ i) ∨ j) ≡ rCancelₖ x ((~ i) ∨ j)) (~ i) ((λ j → rCancelₖ x ((~ i) ∧ (~ j))) ∙ transport refl (cong (_+ₖ (-ₖ x)) q) ∙ λ j → rCancelₖ x ((~ i) ∧ j))) + ∙ (λ i → transport (λ j → rCancelₖ x j ≡ rCancelₖ x j) (lUnit (rUnit (transport refl (cong (_+ₖ (-ₖ x)) p)) (~ i)) (~ i)) + ∙ transport (λ j → rCancelₖ x j ≡ rCancelₖ x j) (lUnit (rUnit (transport refl (cong (_+ₖ (-ₖ x)) q)) (~ i)) (~ i))) + ∙ cong₂ (_∙_) (sym (substComposite (λ x → x) (isoToPath ((congIso (isoToEquiv (symIso (congIsoHelper x)))))) (λ i → rCancelₖ x i ≡ rCancelₖ x i) p)) + (sym (substComposite (λ x → x) (isoToPath ((congIso (isoToEquiv (symIso (congIsoHelper x)))))) (λ i → rCancelₖ x i ≡ rCancelₖ x i) q)) + + + test2 : PathP (λ i → a north +ₖ ∣ north ∣ ≡ a (merid north i)) (rUnitₖ (a north)) (rUnitₖ (a north) ∙ cong a (merid north)) + test2 i j = + hcomp (λ k → λ { (i = i0) → rUnitₖ (a north) j + ; (j = i1) → a (merid north (i ∧ k)) + ; (j = i0) → a north +ₖ ∣ north ∣}) + (rUnitₖ (a north) j) + + +S1→S¹ : S₊ 1 → S¹ +S1→S¹ x = SuspBool→S¹ (S1→SuspBool x) + +S¹→S1 : S¹ → S₊ 1 +S¹→S1 x = SuspBool→S1 (S¹→SuspBool x) + +S1→S¹-sect : section S1→S¹ S¹→S1 +S1→S¹-sect x = + cong SuspBool→S¹ (SuspBool→S1-retr (S¹→SuspBool x)) + ∙ S¹→SuspBool→S¹ x + +S1→S¹-retr : retract S1→S¹ S¹→S1 +S1→S¹-retr x = + cong SuspBool→S1 (SuspBool→S¹→SuspBool (S1→SuspBool x)) + ∙ SuspBool→S1-sect x + + + + + + + +isGroupoidS1 : isGroupoid (S₊ 1) +isGroupoidS1 = transport (λ i → isGroupoid (S¹≡S1 (~ i))) isGroupoidS¹ +isConnectedS1 : (x : S₊ 1) → ∥ north ≡ x ∥₋₁ +isConnectedS1 x = pRec propTruncIsProp + (λ p → ∣ cong (transport (sym (S¹≡S1))) p ∙ transport⁻Transport (S¹≡S1) x ∣₋₁) + (isConnectedS¹ (transport S¹≡S1 x)) + + +open import Cubical.HITs.S2 + +test : (S₊ 2) → S₊ 1 +test north = north +test south = south +test (merid a i) = merid north i + +test2 : (S₊ 1) → (S₊ 2) +test2 north = north +test2 south = south +test2 (merid a i) = merid (merid a i) i + + +testS² : S² → S¹ +testS² base = base +testS² (surf i i₁) = base + +test4 : S₊ 1 → hLevelTrunc 4 (S₊ 2) +test4 north = ∣ north ∣ +test4 south = ∣ north ∣ +test4 (merid a i) = (Kn→ΩKn+1 1 ∣ south ∣) i + + + + +coHom0Torus : grIso (coHomGr 0 (S₊ 1 × S₊ 1)) intGroup +coHom0Torus = + Iso'→Iso + (iso' + (iso (sRec isSetInt (λ f → f (north , north))) + (λ a → ∣ (λ x → a) ∣₀) + (λ a → refl) + (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ + (funExt λ {(x , y) → suspToPropRec2 + {B = λ x y → f (north , north) ≡ f (x , y)} + north + (λ _ _ → isSetInt _ _) + refl + x y}))) + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) λ a b → addLemma (a (north , north)) (b (north , north)))) + +private + S¹map : hLevelTrunc 3 S¹ → S¹ + S¹map = trElim (λ _ → isGroupoidS¹) (idfun S¹) + + S¹map-id : (x : hLevelTrunc 3 S¹) → Path (hLevelTrunc 3 S¹) ∣ S¹map x ∣ x + S¹map-id = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → refl + + S1map : hLevelTrunc 3 (S₊ 1) → (S₊ 1) + S1map = trElim (λ _ → isGroupoidS1) (idfun _) + + +S¹→S¹ : Iso (S¹ → hLevelTrunc 3 S¹) (S¹ × Int) +Iso.fun S¹→S¹ f = S¹map (f base) + , winding (basechange2⁻ (S¹map (f base)) λ i → S¹map (f (loop i))) +Iso.inv S¹→S¹ (s , int) base = ∣ s ∣ +Iso.inv S¹→S¹ (s , int) (loop i) = ∣ basechange2 s (intLoop int) i ∣ +Iso.rightInv S¹→S¹ (s , int) = ×≡ refl ((λ i → winding (basechange2-retr s (λ i → intLoop int i) i)) + ∙ windingIntLoop int) +Iso.leftInv S¹→S¹ f = funExt λ { base → S¹map-id (f base) + ; (loop i) j → helper2 j i} + where + helper2 : PathP (λ i → S¹map-id (f base) i ≡ S¹map-id (f base) i) + (λ i → ∣ basechange2 (S¹map (f base)) (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁)))))) i ∣) + (cong f loop) + helper2 i j = + hcomp (λ k → λ { (i = i0) → cong ∣_∣ (basechange2 (S¹map (f base)) (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁))))))) j + ; (i = i1) → S¹map-id (f (loop j)) k + ; (j = i0) → S¹map-id (f base) (i ∧ k) + ; (j = i1) → S¹map-id (f base) (i ∧ k)}) + (helper4 i j) + where + helper4 : Path (Path (hLevelTrunc 3 _) _ _) + (cong ∣_∣ (basechange2 (S¹map (f base)) + (intLoop + (winding + (basechange2⁻ (S¹map (f base)) + (λ i₁ → S¹map (f (loop i₁)))))))) + λ i → ∣ S¹map (f (loop i)) ∣ + helper4 i = + cong ∣_∣ + ((cong (basechange2 (S¹map (f base))) + (decodeEncode base (basechange2⁻ (S¹map (f base)) + (λ i₁ → S¹map (f (loop i₁))))) + ∙ basechange2-sect (S¹map (f base)) + (λ i → S¹map (f (loop i)))) i) + +S1→S1→S1×Int : Iso ((S₊ 1) → hLevelTrunc 3 (S₊ 1)) ((hLevelTrunc 3 (S₊ 1)) × Int) +S1→S1→S1×Int = compIso helper2 (compIso S¹→S¹ helper) + where + helper : Iso (S¹ × Int) (hLevelTrunc 3 (S₊ 1) × Int) + Iso.fun helper (s , int) = ∣ S¹→S1 s ∣ , int + Iso.inv helper (s , int) = (S1→S¹ (S1map s)) , int + Iso.rightInv helper (s , int) = + trElim {B = λ s → (∣ S¹→S1 (S1→S¹ (S1map s)) ∣ , int) ≡ (s , int)} + (λ _ → isOfHLevelPath 3 (isOfHLevelProd 3 (isOfHLevelTrunc 3) (isOfHLevelSuc 2 isSetInt)) _ _) + (λ a → ×≡ (cong ∣_∣ (S1→S¹-retr a)) refl) + s + Iso.leftInv helper (s , int) = ×≡ (S1→S¹-sect s) refl + + helper2 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → hLevelTrunc 3 S¹) + Iso.fun helper2 f x = trMap S1→S¹ (f (S¹→S1 x)) + Iso.inv helper2 f x = trMap S¹→S1 (f (S1→S¹ x)) + Iso.rightInv helper2 f = funExt λ x i → helper3 (f (S1→S¹-sect x i)) i + where + helper3 : (x : hLevelTrunc 3 S¹) → trMap S1→S¹ (trMap S¹→S1 x) ≡ x + helper3 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → cong ∣_∣ (S1→S¹-sect a) + Iso.leftInv helper2 f = funExt λ x i → helper3 (f (S1→S¹-retr x i)) i + where + helper3 : (x : hLevelTrunc 3 (S₊ 1)) → trMap S¹→S1 (trMap S1→S¹ x) ≡ x + helper3 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → cong ∣_∣ (S1→S¹-retr a) + +S¹→S¹≡S1→S1 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → hLevelTrunc 3 S¹) +Iso.fun S¹→S¹≡S1→S1 f x = trMap S1→S¹ (f (S¹→S1 x)) +Iso.inv S¹→S¹≡S1→S1 f x = trMap S¹→S1 (f (S1→S¹ x)) +Iso.rightInv S¹→S¹≡S1→S1 F = funExt λ x i → helper2 (F (S1→S¹-sect x i)) i + where + helper2 : (x : hLevelTrunc 3 S¹) → trMap S1→S¹ (trMap S¹→S1 x) ≡ x + helper2 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → cong ∣_∣ (S1→S¹-sect a) +Iso.leftInv S¹→S¹≡S1→S1 F = funExt λ x i → helper2 (F (S1→S¹-retr x i)) i + where + helper2 : (x : hLevelTrunc 3 (S₊ 1)) → trMap S¹→S1 (trMap S1→S¹ x) ≡ x + helper2 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → cong ∣_∣ (S1→S¹-retr a) + + + + +basechange-lemma : ∀ {ℓ} {A : Type ℓ} {a : A} (x y : S¹) (F : a ≡ a → S¹) (f : S¹ → a ≡ a) (g : S¹ → a ≡ a) + → (f base ≡ refl) + → (g base ≡ refl) + → basechange2⁻ (F (f base ∙ g base)) (cong₂ {A = S¹} {B = λ x → S¹} (λ x y → F (f x ∙ g y)) loop loop) + ≡ basechange2⁻ (F (f base)) (cong (λ x → F (f x)) loop) ∙ basechange2⁻ (F (g base)) (cong (λ x → F (g x)) loop) +basechange-lemma x y F f g frefl grefl = + (λ i → basechange2⁻ (F (f base ∙ g base)) (cong₂Funct (λ x y → F (f x ∙ g y)) loop loop i)) + ∙ (λ i → basechange2⁻ (F (f base ∙ g base)) (cong (λ x₁ → F (f x₁ ∙ g base)) loop ∙ cong (λ y₁ → F (f base ∙ g y₁)) loop)) + ∙ basechange2⁻-morph (F (f base ∙ g base)) _ _ + ∙ (λ j → basechange2⁻ (F (f base ∙ grefl j)) + (λ i → F (f (loop i) ∙ grefl j)) + ∙ basechange2⁻ (F (frefl j ∙ g base)) + (λ i → F (frefl j ∙ g (loop i)))) + ∙ ((λ j → basechange2⁻ (F (rUnit (f base) (~ j))) + (λ i → F (rUnit (f (loop i)) (~ j))) + ∙ basechange2⁻ (F (lUnit (g base) (~ j))) + (λ i → F (lUnit (g (loop i)) (~ j))))) + + +basechange-lemma2 : (f g : S¹ → hLevelTrunc 3 (S₊ 1)) (F : hLevelTrunc 3 (S₊ 1) → S¹) + → ((basechange2⁻ (F (f base +ₖ g base)) λ i → F ((f (loop i)) +ₖ g (loop i))) + ≡ basechange2⁻ (F (f base)) (cong (F ∘ f) loop) + ∙ basechange2⁻ (F (g base)) (cong (F ∘ g) loop)) +basechange-lemma2 f g F = coInd f g F (f base) (g base) refl refl + where + coInd : (f g : S¹ → hLevelTrunc 3 (S₊ 1)) (F : hLevelTrunc 3 (S₊ 1) → S¹) (x y : hLevelTrunc 3 (S₊ 1)) + → f base ≡ x + → g base ≡ y + → ((basechange2⁻ (F (f base +ₖ g base)) λ i → F ((f (loop i)) +ₖ g (loop i))) + ≡ basechange2⁻ (F (f base)) (cong (F ∘ f) loop) + ∙ basechange2⁻ (F (g base)) (cong (F ∘ g) loop)) + coInd f g F = + elim2 (λ _ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 (isGroupoidS¹ base base)) _ _ ) + (suspToPropRec2 + north + (λ _ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → isGroupoidS¹ _ _ _ _) + λ fid gid → + (λ i → basechange2⁻ (F (f base +ₖ g base)) (cong₂Funct (λ x y → F (f x +ₖ g y)) loop loop i)) + ∙ (basechange2⁻-morph (F (f base +ₖ g base)) + (cong (λ x → F (f x +ₖ g base)) loop) + (cong (λ x → F (f base +ₖ g x)) loop)) + ∙ (λ i → basechange2⁻ (F (f base +ₖ gid i)) (cong (λ x → F (f x +ₖ gid i)) loop) + ∙ basechange2⁻ (F (fid i +ₖ g base)) (cong (λ x → F (fid i +ₖ g x)) loop)) + ∙ (λ i → basechange2⁻ (F (rUnitₖ (f base) i)) (cong (λ x → F (rUnitₖ (f x) i)) loop) + ∙ basechange2⁻ (F (lUnitₖ (g base) i)) (cong (λ x → F (lUnitₖ (g x) i)) loop))) + + + +coHom1S1≃ℤ : grIso (coHomGr 1 (S₊ 1)) intGroup +coHom1S1≃ℤ = + Iso'→Iso + (iso' + (iso + (sRec isSetInt (λ f → proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) + (λ a → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , a)) ∣₀) + (λ a → (λ i → proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 (Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , a)))))) + ∙ (λ i → proj₂ (Iso.fun S¹→S¹ (Iso.rightInv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , a)) i))) + ∙ λ i → proj₂ (Iso.rightInv S¹→S¹ (base , a) i)) + (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → (λ i → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) ∣₀) + ∙ (λ i → sRec setTruncIsSet + (λ x → ∣ Iso.inv S¹→S¹≡S1→S1 x ∣₀) + (sRec setTruncIsSet + (λ x → ∣ Iso.inv S¹→S¹ (x , (proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) ∣₀) + ∣ base ∣₀)) + ∙ (λ i → sRec setTruncIsSet + (λ x → ∣ Iso.inv S¹→S¹≡S1→S1 x ∣₀) + (sRec setTruncIsSet + (λ x → ∣ Iso.inv S¹→S¹ (x , (proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) ∣₀) + (Iso.inv PathIdTrunc₀Iso (isConnectedS¹ (proj₁ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) i))) + ∙ (λ i → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.leftInv S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f) i) ∣₀) + ∙ (λ i → ∣ Iso.leftInv S¹→S¹≡S1→S1 f i ∣₀))) + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) + λ f g → (λ i → winding (basechange2⁻ (S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)) ∙ Kn→ΩKn+1 1 (g (S¹→S1 base)))))) + (λ i → S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 (loop i))) ∙ Kn→ΩKn+1 1 (g (S¹→S1 (loop i))))))))) + ∙ cong winding (helper2 (f (S¹→S1 base)) (g (S¹→S1 base)) f g refl refl) + ∙ winding-hom ((basechange2⁻ (S¹map (trMap S1→S¹ (f north))) + (λ i → S¹map (trMap S1→S¹ (f (S¹→S1 (loop i))))))) + ((basechange2⁻ (S¹map (trMap S1→S¹ (g north))) + (λ i → S¹map (trMap S1→S¹ (g (S¹→S1 (loop i))))))))) + + + where + helper2 : (x y : hLevelTrunc 3 (S₊ 1)) (f g : S₊ 1 → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1) + → (f (S¹→S1 base)) ≡ x + → (g (S¹→S1 base)) ≡ y + → (basechange2⁻ (S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)) ∙ Kn→ΩKn+1 1 (g (S¹→S1 base))))))) + (λ i → S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 (loop i))) ∙ Kn→ΩKn+1 1 (g (S¹→S1 (loop i))))))) + ≡ (basechange2⁻ (S¹map (trMap S1→S¹ ((f (S¹→S1 base)))))) + (λ i → S¹map (trMap S1→S¹ (f (S¹→S1 (loop i))))) + ∙ (basechange2⁻ (S¹map (trMap S1→S¹ ((g (S¹→S1 base))))) + (λ i → S¹map (trMap S1→S¹ ((g (S¹→S1 (loop i))))))) + helper2 = elim2 + (λ _ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 + λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 + λ _ → isOfHLevelPath 3 (isOfHLevelSuc 3 (isGroupoidS¹) base base) _ _) + (suspToPropRec2 {A = S₊ 0} north + (λ _ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → (isGroupoidS¹) _ _ _ _ ) + λ f g reflf reflg → + (basechange-lemma + base + base + (λ x → S¹map (trMap S1→S¹ (ΩKn+1→Kn x))) + (λ x → Kn→ΩKn+1 1 (f (S¹→S1 x))) ((λ x → Kn→ΩKn+1 1 (g (S¹→S1 x)))) + (cong (Kn→ΩKn+1 1) reflf ∙ Kn→ΩKn+10ₖ 1) + (cong (Kn→ΩKn+1 1) reflg ∙ Kn→ΩKn+10ₖ 1)) + ∙ λ j → basechange2⁻ (S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 base)) j))) + (λ i → S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 (loop i))) j))) + ∙ basechange2⁻ (S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 base)) j))) + (λ i → S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 (loop i))) j)))) + +pushoutSn : (n : ℕ) → Iso (S₊ (suc n)) (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt) +Iso.fun (pushoutSn n) north = inl tt +Iso.fun (pushoutSn n) south = inr tt +Iso.fun (pushoutSn n) (merid a i) = push a i +Iso.inv (pushoutSn n) (inl x) = north +Iso.inv (pushoutSn n) (inr x) = south +Iso.inv (pushoutSn n) (push a i) = merid a i +Iso.rightInv (pushoutSn n) (inl x) = refl +Iso.rightInv (pushoutSn n) (inr x) = refl +Iso.rightInv (pushoutSn n) (push a i) = refl +Iso.leftInv (pushoutSn n) north = refl +Iso.leftInv (pushoutSn n) south = refl +Iso.leftInv (pushoutSn n) (merid a i) = refl + +Sn≡Pushout : (n : ℕ) → (S₊ (suc n)) ≡ (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt) +Sn≡Pushout n = isoToPath (pushoutSn n) + +coHomPushout≡coHomSn' : (n m : ℕ) → grIso (coHomGr m (S₊ (suc n))) (coHomGr m (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt)) +morph.fun (grIso.fun (coHomPushout≡coHomSn' n m)) = + sRec setTruncIsSet + λ f → ∣ (λ {(inl x) → f north ; (inr x) → f south ; (push a i) → f (merid a i)}) ∣₀ +morph.ismorph (grIso.fun (coHomPushout≡coHomSn' n m)) = + sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ a b → cong ∣_∣₀ (funExt λ {(inl x) → refl ; (inr x) → refl ; (push a i) → refl }) +morph.fun (grIso.inv (coHomPushout≡coHomSn' n m)) = sRec setTruncIsSet (λ f → ∣ (λ {north → f (inl tt) ; south → f (inr tt) ; (merid a i) → f (push a i)}) ∣₀) +morph.ismorph (grIso.inv (coHomPushout≡coHomSn' n m)) = + sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ a b → cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → refl }) +grIso.rightInv (coHomPushout≡coHomSn' n m) = + sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ (funExt λ {(inl x) → refl ; (inr x) → refl ; (push a i) → refl }) +grIso.leftInv (coHomPushout≡coHomSn' n m) = + sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → refl }) + + + + + + +coHomGrUnit0 : grIso (coHomGr 0 Unit) intGroup +grIso.fun coHomGrUnit0 = mph (sRec isSetInt (λ f → f tt)) + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) + (λ a b → addLemma (a tt) (b tt))) +grIso.inv coHomGrUnit0 = mph (λ a → ∣ (λ _ → a) ∣₀) (λ a b i → ∣ (λ _ → addLemma a b (~ i)) ∣₀) +grIso.rightInv coHomGrUnit0 a = refl +grIso.leftInv coHomGrUnit0 = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ a → refl + +isContrCohomUnit : (n : ℕ) → isContr (coHom (suc n) Unit) +isContrCohomUnit n = subst isContr (λ i → ∥ UnitToTypeId (coHomK (suc n)) (~ i) ∥₀) (helper n) + where + helper : (n : ℕ) → isContr (∥ coHomK (suc n) ∥₀) + helper n = subst isContr ((isoToPath (truncOfTruncIso {A = S₊ (1 + n)} 2 (1 + n))) ∙ sym propTrunc≡Trunc0 ∙ λ i → ∥ hLevelTrunc (suc (+-comm n 2 i)) (S₊ (1 + n)) ∥₀) + (isConnectedSubtr 2 (helper2 n .fst) (subst (λ x → isHLevelConnected x (S₊ (suc n))) (sym (helper2 n .snd)) (sphereConnected (suc n))) ) + where + helper2 : (n : ℕ) → Σ[ m ∈ ℕ ] m + 2 ≡ 2 + n + helper2 zero = 0 , refl + helper2 (suc n) = (suc n) , λ i → suc (+-comm n 2 i) + +coHomGrUnit≥1 : (n : ℕ) → grIso (coHomGr (suc n) Unit) trivialGroup +grIso.fun (coHomGrUnit≥1 n) = mph (λ _ → tt) (λ _ _ → refl) +grIso.inv (coHomGrUnit≥1 n) = mph (λ _ → ∣ (λ _ → ∣ north ∣) ∣₀) (λ _ _ → sym (rUnitₕ 0ₕ)) +grIso.rightInv (coHomGrUnit≥1 n) a = refl +grIso.leftInv (coHomGrUnit≥1 n) a = sym (isContrCohomUnit n .snd 0ₕ) ∙ isContrCohomUnit n .snd a + +S0→Int : (a : Int × Int) → S₊ 0 → Int +S0→Int a north = proj₁ a +S0→Int a south = proj₂ a + +coHom0-S0 : grIso (coHomGr 0 (S₊ 0)) (dirProd intGroup intGroup) +coHom0-S0 = + Iso'→Iso + (iso' + (iso (sRec (isOfHLevelProd 2 isSetInt isSetInt) + λ f → (f north) , (f south)) + (λ a → ∣ S0→Int a ∣₀) + (λ { (a , b) → refl}) + (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ f → cong ∣_∣₀ (funExt (λ {north → refl ; south → refl})))) + (sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 isSetInt isSetInt) _ _) + λ a b i → addLemma (a north) (b north) i , addLemma (a south) (b south) i)) + + + + +coHom0S1 : grIso intGroup (coHomGr 0 (S₊ 1)) +coHom0S1 = + Iso'→Iso + (iso' + (iso + (λ a → ∣ (λ x → a) ∣₀) + (sRec isSetInt (λ f → f north)) + (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ (funExt (suspToPropRec north (λ _ → isSetInt _ _) refl))) + (λ a → refl)) + λ a b → cong ∣_∣₀ (funExt λ x → sym (addLemma a b))) + +coHom1S1 : grIso intGroup (coHomGr 1 (S₊ 1)) +coHom1S1 = + compGrIso + (diagonalIso1 + _ + (coHomGr 0 (S₊ 0)) + _ + (invGrIso coHom0-S0) + (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) + (λ x → MV.Ker-i⊂Im-d _ _(S₊ 0) (λ _ → tt) (λ _ → tt) 0 x + (×≡ (isOfHLevelSuc 0 (isContrCohomUnit 0) _ _) + (isOfHLevelSuc 0 (isContrCohomUnit 0) _ _))) + (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) + (λ x inker + → pRec propTruncIsProp + (λ {((f , g) , id') → helper x f g id' inker}) + ((MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ inker)))) + (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ F surj + → pRec (setTruncIsSet _ _) (λ { (x , id) → MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ F ∣₀ ∣ (∣ (λ _ → x) ∣₀ , ∣ (λ _ → 0) ∣₀) , + (cong ∣_∣₀ (funExt (surjHelper x))) ∙ sym id ∣₋₁ }) surj) ) + (invGrIso (coHomPushout≡coHomSn' 0 1)) + + where + surjHelper : (x : Int) (x₁ : S₊ 0) → x +ₖ (-ₖ pos 0) ≡ S0→Int (x , x) x₁ + surjHelper x north = cong (x +ₖ_) (-0ₖ) ∙ rUnitₖ x + surjHelper x south = cong (x +ₖ_) (-0ₖ) ∙ rUnitₖ x + + helper : (F : S₊ 0 → Int) (f g : ∥ (Unit → Int) ∥₀) (id : morph.fun (MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) (f , g) ≡ ∣ F ∣₀) + → isInKer (coHomGr 0 (S₊ 0)) + (coHomGr 1 (Pushout (λ _ → tt) (λ _ → tt))) + (morph.fun (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) + ∣ F ∣₀ + → ∃[ x ∈ Int ] ∣ F ∣₀ ≡ morph.fun (grIso.inv coHom0-S0) (x , x) + helper F = + sElim2 (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ f g id inker + → pRec propTruncIsProp + (λ {((a , b) , id2) + → sElim2 {B = λ f g → morph.fun (MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) (f , g) ≡ ∣ F ∣₀ → _ } + (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) + (λ f g id → ∣ (helper2 f g .fst) , (sym id ∙ sym (helper2 f g .snd)) ∣₋₁) + a b id2}) + (MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ F ∣₀ inker) + where + helper2 : (f g : Unit → Int) + → Σ[ x ∈ Int ] morph.fun (grIso.inv coHom0-S0) (x , x) + ≡ morph.fun (MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) (∣ f ∣₀ , ∣ g ∣₀) + helper2 f g = (f _ +ₖ (-ₖ g _) ) , cong ∣_∣₀ (funExt (λ {north → refl ; south → refl})) + + + +coHom-n-Sn : (n : ℕ) → grIso intGroup (coHomGr (suc n) (S₊ (suc n))) +coHom-n-Sn zero = coHom1S1 +coHom-n-Sn (suc n) = + compGrIso + (compGrIso + (coHom-n-Sn n) + theIso) + (invGrIso (coHomPushout≡coHomSn' (suc n) (suc (suc n)))) + where + theIso : grIso (coHomGr (suc n) (S₊ (suc n))) (coHomGr (suc (suc n)) + (Pushout {A = S₊ (suc n)} (λ _ → tt) (λ _ → tt))) + theIso = + SES→Iso + (×coHomGr (suc n) Unit Unit) + (×coHomGr (suc (suc n)) Unit Unit) + (ses (λ p q → ×≡ (isOfHLevelSuc 0 (isContrCohomUnit n) (proj₁ p) (proj₁ q)) + (isOfHLevelSuc 0 (isContrCohomUnit n) (proj₂ p) (proj₂ q))) + (λ p q → ×≡ (isOfHLevelSuc 0 (isContrCohomUnit (suc n)) (proj₁ p) (proj₁ q)) + (isOfHLevelSuc 0 (isContrCohomUnit (suc n)) (proj₂ p) (proj₂ q))) + (MV.Δ _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) + (MV.i _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (2 + n)) + (MV.d _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) + (MV.Ker-d⊂Im-Δ _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) + (MV.Ker-i⊂Im-d _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n))) + + +coHom1-s0 : grIso (coHomGr 1 (S₊ 0)) trivialGroup +coHom1-s0 = + Iso'→Iso + (iso' + (iso (λ _ → tt) + (λ _ → 0ₕ) + (λ _ → refl) + (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → pRec (setTruncIsSet _ _) + (λ id → cong ∣_∣₀ (sym id)) + (helper f (f north) (f south) refl refl))) + λ _ _ → refl) + where + helper : (f : S₊ 0 → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1) → (x y : ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1) → (f north ≡ x) → (f south ≡ y) → ∥ f ≡ (λ _ → 0ₖ) ∥₋₁ + helper f = + elim2 (λ _ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelPlus {n = 1} 2 propTruncIsProp) + (suspToPropRec2 north (λ _ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → propTruncIsProp) + λ id id2 → ∣ funExt (λ {north → id ; south → id2}) ∣₋₁) + +coHom2-S1 : grIso (coHomGr 2 (S₊ 1)) trivialGroup +coHom2-S1 = + compGrIso + (coHomPushout≡coHomSn' 0 2) + (compGrIso + (invGrIso + (SES→Iso + (×coHomGr 1 Unit Unit) + (×coHomGr 2 Unit Unit) + (ses (isOfHLevelProd 1 (isOfHLevelSuc 0 (isContrCohomUnit 0)) (isOfHLevelSuc 0 (isContrCohomUnit 0))) + (isOfHLevelProd 1 (isOfHLevelSuc 0 (isContrCohomUnit 1)) (isOfHLevelSuc 0 (isContrCohomUnit 1))) + (MV.Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1) + (MV.i _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 2) + (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1) + (MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1) + ({! MV.Ker-i⊂Im-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1 !})))) + coHom1-s0) + + + + +coHom1Torus : grIso (coHomGr 1 ((S₊ 1) × (S₊ 1))) (dirProd intGroup intGroup) +coHom1Torus = + compGrIso + (Iso'→Iso + (iso' (compIso + (setTruncIso (compIso + schönfinkelIso + (compIso + (codomainIso S1→S1→S1×Int) + toProdIso))) + (setTruncOfProdIso)) + (sElim2 + (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ f g → ×≡ (cong ∣_∣₀ + (funExt (λ x → helper (f (x , S¹→S1 base) +ₖ g (x , S¹→S1 base)) + ∙ sym (cong₂ (λ x y → x +ₖ y) + (helper (f (x , S¹→S1 base))) + (helper (g (x , S¹→S1 base))))))) + (cong ∣_∣₀ + (funExt + (suspToPropRec + north + (λ _ → isSetInt _ _) + (cong winding + (basechange-lemma2 + (λ x → f (north , S¹→S1 x)) + (λ x → g (north , S¹→S1 x)) + λ x → S¹map (trMap S1→S¹ x)) + ∙ winding-hom + (basechange2⁻ + (S¹map (trMap S1→S¹ (f (north , S¹→S1 base)))) + (λ i → S¹map (trMap S1→S¹ (f (north , S¹→S1 (loop i)))))) + (basechange2⁻ + (S¹map (trMap S1→S¹ (g (north , S¹→S1 base)))) + (λ i → S¹map (trMap S1→S¹ (g (north , S¹→S1 (loop i)))))) + ∙ sym (addLemma + (winding + (basechange2⁻ + (S¹map (trMap S1→S¹ (f (north , S¹→S1 base)))) + (λ i → S¹map (trMap S1→S¹ (f (north , S¹→S1 (loop i))))))) + (winding + (basechange2⁻ + (S¹map (trMap S1→S¹ (g (north , S¹→S1 base)))) + (λ i → S¹map (trMap S1→S¹ (g (north , S¹→S1 (loop i))))))))))))))) + (dirProdIso (invGrIso (coHom-n-Sn 0)) (invGrIso coHom0S1)) + + where + helper : (x : hLevelTrunc 3 (S₊ 1)) → ∣ S¹→S1 (S¹map (trMap S1→S¹ x)) ∣ ≡ x + helper = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ a → cong ∣_∣ (S1→S¹-retr a) + + +coHom2Torus : grIso (coHomGr 2 ((S₊ 1) × (S₊ 1))) intGroup +coHom2Torus = + invGrIso + (Iso'→Iso + (iso' + (compIso + helper' + (compIso + (symIso (prodIso (groupIso→Iso coHom2-S1) + (groupIso→Iso (invGrIso (coHom-n-Sn 0))))) + (compIso + (symIso setTruncOfProdIso) + (symIso + (setTruncIso + (compIso + schönfinkelIso + (compIso + (codomainIso test13) + toProdIso))))))) + λ a b → {!!} {- (λ i → Iso.fun + (symIso + (setTruncIso + (compIso schönfinkelIso (compIso (codomainIso test13) toProdIso)))) + (Iso.fun (symIso setTruncOfProdIso) + (helper'' (Iso.inv (groupIso→Iso coHom2-S1) tt) 0ₕ i , morph.ismorph (grIso.inv (invGrIso coHom1S1)) a b i))) + ∙ (λ i → Iso.fun + (symIso + (setTruncIso + (compIso schönfinkelIso (compIso (codomainIso test13) toProdIso)))) + ∣ (λ _ → ∣ north ∣) , {!!} ∣₀) + ∙ {!toProdIso!} -})) + where + helper'' : isProp ∥ (S₊ 1 → hLevelTrunc 4 (S₊ 2)) ∥₀ + helper'' = {!!} + + helper2' : (f g : (S₊ 1 → hLevelTrunc 3 (S₊ 1))) → + Path (coHom 2 ((S₊ 1) × (S₊ 1))) (Iso.fun + (symIso + (setTruncIso + (compIso schönfinkelIso (compIso (codomainIso test13) toProdIso)))) + (Iso.fun (symIso setTruncOfProdIso) (0ₕ , ∣ f ∣₀ +ₕ ∣ g ∣₀))) + (Iso.fun (symIso + (setTruncIso + (compIso schönfinkelIso (compIso (codomainIso test13) toProdIso)))) + (Iso.fun (symIso setTruncOfProdIso) (0ₕ , ∣ f ∣₀)) +ₕ Iso.fun (symIso + (setTruncIso + (compIso schönfinkelIso (compIso (codomainIso test13) toProdIso)))) + (Iso.fun (symIso setTruncOfProdIso) (0ₕ , ∣ g ∣₀))) + helper2' f g = (λ i → Iso.fun (symIso + (setTruncIso + (compIso schönfinkelIso ((codomainIso test13))))) ∣ (λ x → 0ₖ , (f x +ₖ g x)) ∣₀) + ∙ (λ i → ∣ Iso.fun (symIso (compIso schönfinkelIso (codomainIso test13))) (λ x → 0ₖ , (f x +ₖ g x)) ∣₀) + ∙ (λ i → ∣ Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , (f x +ₖ g x)))) ∣₀) + ∙ (λ i → ∣ {!!} ∣₀) + ∙ {!!} + ∙ {!!} + where + fun : S₊ 1 × S₊ 1 → hLevelTrunc 4 (S₊ 2) + fun (x , north) = 0ₖ +ₖ 0ₖ + fun (x , south) = 0ₖ +ₖ 0ₖ + fun (x , (merid south i)) = 0ₖ +ₖ (Kn→ΩKn+1 1 (f x +ₖ g x) i) + fun (x , (merid north i)) = 0ₖ +ₖ 0ₖ + + helper : Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , (f x +ₖ g x)))) ≡ fun + helper = funExt λ {(x , north) → refl ; (x , south) → refl ; (x , (merid north i)) → refl ; (x , (merid south i)) → refl} + + helper2 : ∣ Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , (f x +ₖ g x)))) ∣₀ + ≡ (∣ Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , f x))) ∣₀ +ₕ ∣ Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , g x))) ∣₀) + helper2 = + cong ∣_∣₀ + (funExt λ {(x , north) → ((λ i → (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x +ₖ g x) i)) + ∙∙ sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ cong (λ x → ((0ₖ +ₖ 0ₖ) +ₖ x)) (rUnitₖ 0ₖ) + ∙∙ refl) + ; (x , south) → refl + ∙∙ sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ cong (λ x → ((0ₖ +ₖ 0ₖ) +ₖ x)) (rUnitₖ 0ₖ) + ∙∙ (λ i → (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x) i) +ₖ ∣ north ∣ +ₖ Kn→ΩKn+1 1 (g x) i) + ; (x , (merid south i)) j → hcomp (λ k → λ {(j = i0) → ∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x +ₖ g x) (i ∨ (~ k)) + ; (j = i1) → (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x) (i ∧ k)) +ₖ ∣ north ∣ +ₖ Kn→ΩKn+1 1 (g x) (i ∧ k)}) + ((sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ cong (λ x → ((0ₖ +ₖ 0ₖ) +ₖ x)) (rUnitₖ 0ₖ)) j) + ; (x , (merid north u)) → {!!}}) + where + anotherone : (x : _) → cong (0ₖ +ₖ_) (Kn→ΩKn+1 1 (f x +ₖ g x)) ≡ {!!} + anotherone x = {!!} + + helper5 : (x : _) → Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , f x +ₖ g x))) x + ≡ (Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , f x))) x) +ₖ (Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , g x))) x) + helper5 (x , north) = sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i))) + helper5 (x , south) = sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i))) + helper5 (x , merid north i) = sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i))) + helper5 (x , merid south i) j = + {!!} + where + helper13 : sym (sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i)))) + ∙ (λ i → ∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x +ₖ g x) i) ∙ (sym (rUnitₖ (0ₖ +ₖ 0ₖ)) + ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i)))) + ≡ λ i → (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x) i) +ₖ + ∣ north ∣ +ₖ Kn→ΩKn+1 1 (g x) i + helper13 = (λ j → sym ((λ i → (rUnitₖ (lUnitₖ 0ₖ (j ∧ (~ i))) (~ i))) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ lUnitₖ 0ₖ (~ i))) + ∙ lUnit (rUnit ((λ i → lUnitₖ (Kn→ΩKn+1 1 (f x +ₖ g x) i) j)) j) j + ∙ ((λ i → (rUnitₖ (lUnitₖ 0ₖ (j ∧ (~ i))) (~ i)))) ∙ λ i → (0ₖ +ₖ 0ₖ) +ₖ lUnitₖ 0ₖ (~ i)) + ∙ (λ j → sym ((λ i → (rUnitₖ (lUnitₖ 0ₖ (~ i)) (~ i))) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ lUnitₖ 0ₖ (~ i))) ∙ + {!!} ∙ + (λ i → (rUnitₖ (lUnitₖ 0ₖ (~ i)) (~ i))) ∙ λ i → (0ₖ +ₖ 0ₖ) +ₖ lUnitₖ 0ₖ (~ i)) + ∙ {!!} + ∙ {!!} +{- +i = i0 ⊢ (sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i)))) j +i = i1 ⊢ (sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i)))) j +j = i0 ⊢ ∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x +ₖ g x) i +j = i1 ⊢ (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x) i) +ₖ + ∣ north ∣ +ₖ Kn→ΩKn+1 1 (g x) i +-} + + + helper' : Iso Int (Unit × Int) + Iso.inv helper' = proj₂ + Iso.fun helper' x = tt , x + Iso.leftInv helper' x = refl + Iso.rightInv helper' x = ×≡ refl refl + + helper : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} → Iso C Unit → Iso A B → Iso (A × C) B + Iso.fun (helper cUnit iAB) x = Iso.fun iAB (proj₁ x) + Iso.inv (helper cUnit iAB) x = (Iso.inv iAB x) , (Iso.inv cUnit tt ) + Iso.rightInv (helper cUnit iAB) = Iso.rightInv iAB + Iso.leftInv (helper cUnit iAB) b = ×≡ (Iso.leftInv iAB (proj₁ b)) (Iso.leftInv cUnit (proj₂ b)) + + helper2 : Iso (coHom 2 ((S₊ 1) × (S₊ 1))) (coHom 2 ((S₊ 1) × hLevelTrunc 3 (S₊ 1))) + Iso.fun helper2 = sRec setTruncIsSet (λ f → ∣ (λ {(x , y) → f (x , trRec isGroupoidS1 (idfun (S₊ 1)) y)}) ∣₀) + Iso.inv helper2 = sRec setTruncIsSet (λ f → ∣ (λ {(x , y) → f (x , ∣ y ∣)}) ∣₀) + Iso.rightInv helper2 = + sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + (λ f → cong ∣_∣₀ + (funExt λ {(x , y) → + trElim {B = λ y → f (x , ∣ trRec isGroupoidS1 (λ x₁ → x₁) y ∣) ≡ f (x , y)} + (λ _ → isOfHLevelPath' 3 (isOfHLevelTrunc 4) _ _) + (λ a → refl) y})) + Iso.leftInv helper2 = + sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ (funExt λ {(x , y) → refl}) + +-- basechange* : (x y : S¹) → x ≡ y → x ≡ y → ΩS¹ +-- basechange* x y = J (λ y p → (x ≡ y) → ΩS¹) (basechange x) + + +-- test1 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S₊ 1 × Int) +-- Iso.fun test1 f = (trRec isGroupoidS1 (λ a → a) (f north)) +-- , winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (f (loop* i))))) +-- Iso.inv test1 (north , b) x = ∣ x ∣ +-- Iso.inv test1 (south , b) x = ∣ x ∣ +-- Iso.inv test1 (merid a i , b) x = {!!} +-- Iso.rightInv test1 = {!!} +-- Iso.leftInv test1 = {!!} + +-- funRec : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) {C : (A → hLevelTrunc n B) → Type ℓ''} +-- → isOfHLevel n B +-- → ((f : A → B) → C (λ a → ∣ f a ∣)) +-- → (f : A → hLevelTrunc n B) → C f +-- funRec {A = A} {B = B} n {C = C} hLev ind f = subst C (helper f) (ind (λ a → trRec hLev (λ x → x) (f a))) +-- where +-- helper : retract {A = A → hLevelTrunc n B} {B = A → B} (λ f₁ a → trRec hLev (λ x → x) (f₁ a)) (λ f₁ a → ∣ f₁ a ∣) +-- helper f = funExt λ a → helper2 (f a) +-- where +-- helper2 : (x : hLevelTrunc n B) → ∣ trRec hLev (λ x → x) x ∣ ≡ x +-- helper2 = trElim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → refl + +-- invMapSurj : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (ϕ : morph G H) → ((x : Group.type H) → isInIm G H (fst ϕ) x) +-- → morph H G +-- fst (invMapSurj G H (ϕ , pf) surj) a = {!pRec!} +-- snd (invMapSurj G H (ϕ , pf) surj) = {!!} + +{- +ImIso : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (ϕ : morph G H) → ((x : Group.type H) → isInIm G H (fst ϕ) x) + → grIso H (imGroup G H ϕ) +ImIso G H (ϕ , mf) surj = + let idH = isGroup.id (Group.groupStruc H) + idG = isGroup.id (Group.groupStruc G) + _+G_ = isGroup.comp (Group.groupStruc G) + _+H_ = isGroup.comp (Group.groupStruc H) + _+Im_ = isGroup.comp (Group.groupStruc (imGroup G H (ϕ , mf))) + invG = isGroup.inv (Group.groupStruc G) + invH = isGroup.inv (Group.groupStruc H) + lUnit = isGroup.lUnit (Group.groupStruc H) + lCancel = isGroup.rCancel (Group.groupStruc H) + in + Iso''→Iso _ _ + (iso'' ((λ x → x , pRec propTruncIsProp (λ (a , b) → ∣ a , b ∣₋₁) (surj x)) + , λ a b → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) + (λ surja → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) + (λ surjb → + pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) + (λ surja+b → + (λ i → (a +H b) , (pRec (propTruncIsProp) + (λ (a , b) → ∣ a , b ∣₋₁) + (propTruncIsProp (surj (isGroup.comp (Group.groupStruc H) a b)) ∣ surja+b ∣₋₁ i))) ∙ + (λ i → (a +H b) , ∣ (fst surja+b) , (snd surja+b) ∣₋₁) ∙ + ΣProp≡ (λ _ → propTruncIsProp) refl ∙ + λ i → (a +H b) , pRec (propTruncIsProp) + (λ p1 → + pRec (λ x y → squash x y) + (λ p2 → + ∣ + isGroup.comp (Group.groupStruc G) (fst p1) (fst p2) , + mf (fst p1) (fst p2) ∙ + cong₂ (isGroup.comp (Group.groupStruc H)) (snd p1) (snd p2) + ∣₋₁) + (pRec (propTruncIsProp) + ∣_∣₋₁ (propTruncIsProp ∣ surjb ∣₋₁ (surj b) i))) + (pRec (propTruncIsProp) + ∣_∣₋₁ (propTruncIsProp ∣ surja ∣₋₁ (surj a) i ))) + (surj (isGroup.comp (Group.groupStruc H) a b))) + (surj b)) + (surj a)) + (λ x inker → cong fst inker) + λ x → pRec propTruncIsProp (λ inimx → ∣ (ϕ (inimx .fst)) , ΣProp≡ (λ _ → propTruncIsProp) (inimx .snd) ∣₋₁) (snd x)) +-} + + +{- +H¹-S¹≃Int : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +H¹-S¹≃Int = + Iso''→Iso _ _ + (iso'' ((λ x → ∣ theFuns x ∣₀) , λ a b → cong ∣_∣₀ (funExt λ x → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 1) _) ∙ sym (cong (ΩKn+1→Kn) (theFunsId2 a b x)))) + (λ x inker → pRec (isSetInt _ _) (inj x) (Iso.fun PathIdTrunc₀Iso inker)) + finIm) + where + d : _ + d = MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 + + i : _ + i = MV.i _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1 + + Δ : _ + Δ = MV.Δ _ _ (S₊ 1) (λ _ → tt) (λ _ → tt) 0 + + + d-surj : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) + → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x + d-surj = sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ x → MV.Ker-i⊂Im-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ + (sym (isContrHelper .snd _)) + where + isContrHelper : isContr (Group.type (×coHomGr 1 Unit Unit)) + isContrHelper = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) + , λ y → ×≡ (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₁ y)) + (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₂ y)) + + theFuns : (a : Int) → Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 + theFuns a (inl x) = 0ₖ + theFuns a (inr x) = 0ₖ + theFuns a (push north i) = Kn→ΩKn+1 0 a i + theFuns a (push south i) = 0ₖ + + + theFunsId2 : (a b : Int) (x : Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) + → Kn→ΩKn+1 1 (theFuns a x) ∙ Kn→ΩKn+1 1 (theFuns b x) ≡ Kn→ΩKn+1 1 (theFuns (a +ℤ b) x) + theFunsId2 a b (inl x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) + theFunsId2 a b (inr x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) + theFunsId2 a b (push north i) j = + hcomp (λ k → λ {(i = i0) → ((λ i₁ → + (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) + j + ; (i = i1) → ((λ i₁ → + (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) + j + ; (j = i0) → cong₂Funct (λ p q → Kn→ΩKn+1 1 p ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b) (~ k) i + ; (j = i1) → (helper2 a b) k i }) + (hcomp (λ k → λ { (j = i0) → compPath-filler (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) k i + ; (j = i1) → compPath-filler (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) k i + ; (i = i1) → RHS-filler b j k + ; (i = i0) → ((λ i₁ → + (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) + j}) + (bottom-filler a j i)) + + where + + bottom-filler : (a : Int) → + PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) + (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) + ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) + j ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) + (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) + ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) + j) (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) + bottom-filler a j i = + hcomp (λ k → λ {(j = i0) → helper2 (~ k) i ; + (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 a) i}) + ((λ j₂ → ∣ rCancel (merid north) j j₂ ∣) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + + where + helper2 : Path (Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣ ≡ Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣) + (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i) ∙ Kn→ΩKn+1 1 ∣ north ∣) + (λ i → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + helper2 = cong₂Sym1 (Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) (~ i) j ∣) (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + + RHS-filler : (b : Int) → + PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j + ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j) + (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) + (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) + RHS-filler b j i = + hcomp (λ k → λ {(j = i0) → cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i ; + (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 b) i}) + (cong (λ q → (λ i → ∣ rCancel (merid north) j i ∣) ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i) + + helper2 : (a b : Int) → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b) ≡ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ℤ b)) + helper2 a b = + sym (congFunct (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b)) + ∙ (λ i → cong (Kn→ΩKn+1 1) (Iso.rightInv (Iso3-Kn-ΩKn+1 0) (Kn→ΩKn+1 0 a ∙ Kn→ΩKn+1 0 b) (~ i))) + ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ₖ b)) ) + ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (addLemma a b i))) + + theFunsId2 a b (push south i) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ sym (lUnit _) + + inj : (a : Int) → theFuns a ≡ (λ _ → ∣ north ∣) → a ≡ pos 0 + inj a id = + pRec (isSetInt _ _) + (sigmaElim (λ _ → isOfHLevelPath 2 isSetInt _ _) + (λ a p → pRec (isSetInt _ _) + (λ id2 → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 0) _) + ∙ cong (ΩKn+1→Kn) + (PathP→compPathR + (cong (λ f → cong f (push north)) id) + ∙ test)) + (Iso.fun PathIdTrunc₀Iso p))) (d-surj ∣ theFuns a ∣₀) + where + + test : (λ i → id i (inl tt)) ∙ (λ i → ∣ north ∣) ∙ sym (λ i → id i (inr tt)) ≡ refl + test = (λ i → cong (λ f → f (inl tt)) id + ∙ lUnit (sym (cong (λ f → f (inr tt)) id)) (~ i)) + ∙ (λ i → cong (λ f → f (push south i)) id + ∙ sym (cong (λ f → f (inr tt)) id)) + ∙ rCancel (cong (λ f → f (inr tt)) id) + + + consMember : (a : Int) → coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) + consMember a = d ∣ (λ _ → a) ∣₀ + + consMember≡0 : (a : Int) → consMember a ≡ 0ₕ + consMember≡0 a = + (MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ (λ _ → a) ∣₀ ∣ + (∣ (λ _ → a) ∣₀ , ∣ (λ _ → 0) ∣₀) + , cong ∣_∣₀ (λ i x → (rUnitₖ a i)) ∣₋₁) + + f+consMember' : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int × Int ] (f +ₕ (-ₕ (consMember (proj₁ x))) ≡ ∣ theFuns (proj₂ x) ∣₀) + f+consMember' = + sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ f → pRec propTruncIsProp + (sigmaElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) + (λ g id → ∣ ((g south) , ((g north) +ₖ (-ₖ g south))) + , (pRec (setTruncIsSet _ _) + (λ id → (λ i → ∣ id (~ i) ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀) ∙ funId1 g) + (Iso.fun PathIdTrunc₀Iso id)) ∣₋₁)) + (d-surj ∣ f ∣₀) + where + funId1 : (g : S₊ 0 → Int) + → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀ ≡ + ∣ theFuns ((g north) +ₖ (-ₖ (g south))) ∣₀ + funId1 g = (λ i → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ + +ₕ (morphMinus (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) d + (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ∣ (λ _ → g south) ∣₀ (~ i))) + ∙ sym (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ g ∣₀ (-ₕ ∣ (λ _ → g south) ∣₀)) + ∙ (cong (λ x → d ∣ x ∣₀) g'Id) + ∙ cong ∣_∣₀ helper + where + g' : S₊ 0 → Int + g' north = (g north) +ₖ (-ₖ (g south)) + g' south = 0 + + g'Id : (λ x → g x +ₖ (-ₖ (g south))) ≡ g' + g'Id = funExt (λ {north → refl + ; south → rCancelₖ (g south)}) + + helper : MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g' ≡ theFuns (g north +ₖ (-ₖ g south)) + helper = funExt λ {(inl tt) → refl + ; (inr tt) → refl + ; (push north i) → refl + ; (push south i) → refl} + finIm : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int ] (∣ theFuns x ∣₀ ≡ f) + finIm f = + pRec propTruncIsProp + (λ {((a , b) , id) → ∣ b , (sym id ∙ cong (λ x → f +ₕ x) ((λ i → (-ₕ (consMember≡0 a i))) ∙ sym (lUnitₕ (-ₕ 0ₕ)) ∙ rCancelₕ 0ₕ) ∙ (rUnitₕ f)) ∣₋₁}) + (f+consMember' f) +-} + + + +-- Hⁿ-Sⁿ≃Int : (n : ℕ) → grIso intGroup (coHomGr (suc n) (S₊ (suc n))) +-- Hⁿ-Sⁿ≃Int zero = +-- compGrIso {F = intGroup} {G = {!!}} {H = {!coHomGr 1 (S₊ 1)!}} +-- (Iso''→Iso +-- (iso'' ((λ x → ∣ theFuns x ∣₀) , λ a b → cong ∣_∣₀ (funExt λ x → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 1) _) ∙ sym (cong (ΩKn+1→Kn) (theFunsId2 a b x)))) +-- (λ x inker → pRec (isSetInt _ _) (inj x) (Iso.fun PathIdTrunc₀Iso inker)) +-- finIm)) +-- {!invGrIso _ _ (coHomPushout≡coHomSn 0 1)!} +-- where +-- d : _ +-- d = MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 + +-- i : _ +-- i = MV.i _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1 + +-- Δ : _ +-- Δ = MV.Δ _ _ (S₊ 1) (λ _ → tt) (λ _ → tt) 0 + + +-- d-surj : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x +-- d-surj = sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- λ x → MV.Ker-i⊂Im-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ +-- (sym (isContrHelper .snd _)) +-- where +-- isContrHelper : isContr (Group.type (×coHomGr 1 Unit Unit)) +-- isContrHelper = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) +-- , λ y → ×≡ (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₁ y)) +-- (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₂ y)) + +-- theFuns : (a : Int) → Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 +-- theFuns a (inl x) = 0ₖ +-- theFuns a (inr x) = 0ₖ +-- theFuns a (push north i) = Kn→ΩKn+1 0 a i +-- theFuns a (push south i) = 0ₖ + + +-- theFunsId2 : (a b : Int) (x : Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) +-- → Kn→ΩKn+1 1 (theFuns a x) ∙ Kn→ΩKn+1 1 (theFuns b x) ≡ Kn→ΩKn+1 1 (theFuns (a +ℤ b) x) +-- theFunsId2 a b (inl x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) +-- theFunsId2 a b (inr x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) +-- theFunsId2 a b (push north i) j = +-- hcomp (λ k → λ {(i = i0) → ((λ i₁ → +-- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) +-- j +-- ; (i = i1) → ((λ i₁ → +-- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) +-- j +-- ; (j = i0) → cong₂Funct (λ p q → Kn→ΩKn+1 1 p ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b) (~ k) i +-- ; (j = i1) → (helper2 a b) k i }) +-- (hcomp (λ k → λ { (j = i0) → compPath-filler (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) k i +-- ; (j = i1) → compPath-filler (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) k i +-- ; (i = i1) → RHS-filler b j k +-- ; (i = i0) → ((λ i₁ → +-- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) +-- j}) +-- (bottom-filler a j i)) + +-- where + +-- bottom-filler : (a : Int) → +-- PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) +-- ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) +-- j ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) +-- ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) +-- j) (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) +-- bottom-filler a j i = +-- hcomp (λ k → λ {(j = i0) → helper2 (~ k) i ; +-- (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 a) i}) +-- ((λ j₂ → ∣ rCancel (merid north) j j₂ ∣) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + +-- where +-- helper2 : Path (Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣ ≡ Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- (λ i → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) +-- helper2 = cong₂Sym1 (Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) (~ i) j ∣) (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + +-- RHS-filler : (b : Int) → +-- PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j +-- ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j) +-- (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) +-- (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) +-- RHS-filler b j i = +-- hcomp (λ k → λ {(j = i0) → cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i ; +-- (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 b) i}) +-- (cong (λ q → (λ i → ∣ rCancel (merid north) j i ∣) ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i) + +-- helper2 : (a b : Int) → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b) ≡ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ℤ b)) +-- helper2 a b = +-- sym (congFunct (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b)) +-- ∙ (λ i → cong (Kn→ΩKn+1 1) (Iso.rightInv (Iso3-Kn-ΩKn+1 0) (Kn→ΩKn+1 0 a ∙ Kn→ΩKn+1 0 b) (~ i))) +-- ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ₖ b)) ) +-- ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (addLemma a b i))) + +-- theFunsId2 a b (push south i) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- ∙ sym (lUnit _) + +-- inj : (a : Int) → theFuns a ≡ (λ _ → ∣ north ∣) → a ≡ pos 0 +-- inj a id = +-- pRec (isSetInt _ _) +-- (sigmaElim (λ _ → isOfHLevelPath 2 isSetInt _ _) +-- (λ a p → pRec (isSetInt _ _) +-- (λ id2 → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 0) _) +-- ∙ cong (ΩKn+1→Kn) +-- (PathP→compPathR +-- (cong (λ f → cong f (push north)) id) +-- ∙ test)) +-- (Iso.fun PathIdTrunc₀Iso p))) (d-surj ∣ theFuns a ∣₀) +-- where + +-- test : (λ i → id i (inl tt)) ∙ (λ i → ∣ north ∣) ∙ sym (λ i → id i (inr tt)) ≡ refl +-- test = (λ i → cong (λ f → f (inl tt)) id +-- ∙ lUnit (sym (cong (λ f → f (inr tt)) id)) (~ i)) +-- ∙ (λ i → cong (λ f → f (push south i)) id +-- ∙ sym (cong (λ f → f (inr tt)) id)) +-- ∙ rCancel (cong (λ f → f (inr tt)) id) + + +-- consMember : (a : Int) → coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) +-- consMember a = d ∣ (λ _ → a) ∣₀ + +-- consMember≡0 : (a : Int) → consMember a ≡ 0ₕ +-- consMember≡0 a = +-- (MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ (λ _ → a) ∣₀ ∣ +-- (∣ (λ _ → a) ∣₀ , ∣ (λ _ → 0) ∣₀) +-- , cong ∣_∣₀ (λ i x → (rUnitₖ a i)) ∣₋₁) + +-- f+consMember' : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int × Int ] (f +ₕ (-ₕ (consMember (proj₁ x))) ≡ ∣ theFuns (proj₂ x) ∣₀) +-- f+consMember' = +-- sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- λ f → pRec propTruncIsProp +-- (sigmaElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- (λ g id → ∣ ((g south) , ((g north) +ₖ (-ₖ g south))) +-- , (pRec (setTruncIsSet _ _) +-- (λ id → (λ i → ∣ id (~ i) ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀) ∙ funId1 g) +-- (Iso.fun PathIdTrunc₀Iso id)) ∣₋₁)) +-- (d-surj ∣ f ∣₀) +-- where +-- funId1 : (g : S₊ 0 → Int) +-- → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀ ≡ +-- ∣ theFuns ((g north) +ₖ (-ₖ (g south))) ∣₀ +-- funId1 g = (λ i → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +-- +ₕ (morphMinus (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) d +-- (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ∣ (λ _ → g south) ∣₀ (~ i))) +-- ∙ sym (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ g ∣₀ (-ₕ ∣ (λ _ → g south) ∣₀)) +-- ∙ (cong (λ x → d ∣ x ∣₀) g'Id) +-- ∙ cong ∣_∣₀ helper +-- where +-- g' : S₊ 0 → Int +-- g' north = (g north) +ₖ (-ₖ (g south)) +-- g' south = 0 + +-- g'Id : (λ x → g x +ₖ (-ₖ (g south))) ≡ g' +-- g'Id = funExt (λ {north → refl +-- ; south → rCancelₖ (g south)}) + +-- helper : MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g' ≡ theFuns (g north +ₖ (-ₖ g south)) +-- helper = funExt λ {(inl tt) → refl +-- ; (inr tt) → refl +-- ; (push north i) → refl +-- ; (push south i) → refl} +-- finIm : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int ] (∣ theFuns x ∣₀ ≡ f) +-- finIm f = +-- pRec propTruncIsProp +-- (λ {((a , b) , id) → ∣ b , (sym id ∙ cong (λ x → f +ₕ x) ((λ i → (-ₕ (consMember≡0 a i))) ∙ sym (lUnitₕ (-ₕ 0ₕ)) ∙ rCancelₕ 0ₕ) ∙ (rUnitₕ f)) ∣₋₁}) +-- (f+consMember' f) +-- Hⁿ-Sⁿ≃Int (suc n) = +-- compGrIso (Hⁿ-Sⁿ≃Int n) +-- (transport (λ i → grIso {!!} {!coHomGr (suc (suc n)) (Pushout (λ _ → tt) (λ _ → tt))!}) {!!}) + + +-- {- +-- coHom1S1≃ℤ : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- grIso.fun coHom1S1≃ℤ = grIso.fun {!compGrIso coHom1Iso (invGrIso _ _ (ImIso _ _ (d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ?))!} +-- grIso.inv coHom1S1≃ℤ = {!!} +-- grIso.rightInv coHom1S1≃ℤ = {!!} +-- grIso.leftInv coHom1S1≃ℤ = {!!} +-- -} + +-- -- compGrIso coHom1Iso (invGrIso _ _ (ImIso _ _ {!d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0!} {!!})) + + +-- -- coHomGrIm : grIso (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- (imGroup (coHomGr 0 (S₊ 0)) +-- -- (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 +-- -- , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) +-- -- coHomGrIm = +-- -- ImIso _ +-- -- _ +-- -- _ +-- -- {!!} + + +-- -- -- coHom1RedS1 : Iso (coHom 1 (S₊ 1)) (coHomRed 1 (S₊ 1 , north)) +-- -- -- Iso.fun coHom1RedS1 = sRec setTruncIsSet λ f → ∣ f , (pRec {!!} {!!} ((transport (λ i → (b : truncIdempotent 3 {!S₊ 1!} (~ i)) → ∥ (transp (λ j → truncIdempotent {!3!} {!!} (~ i ∨ (~ j))) (~ i) north) ≡ b ∥₋₁) isConnectedS1) (f north)) ) ∣₀ +-- -- -- Iso.inv coHom1RedS1 = {!!} +-- -- -- Iso.rightInv coHom1RedS1 = {!setTruncIdempotent!} +-- -- -- Iso.leftInv coHom1RedS1 = {!!} + +-- -- -- coHom1Red : ∀ {ℓ} (A : Pointed ℓ) → Iso (coHom 1 (typ A)) (coHomRed 1 A) +-- -- -- Iso.fun (coHom1Red A) = sRec setTruncIsSet λ f → ∣ f , {!!} ∣₀ +-- -- -- Iso.inv (coHom1Red A) = {!!} +-- -- -- Iso.rightInv (coHom1Red A) = {!!} +-- -- -- Iso.leftInv (coHom1Red A) = {!!} + +-- -- -- -- morphtest : morph (coHomGr 1 (S₊ 1)) intGroup +-- -- -- -- fst morphtest = sRec isSetInt λ f → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (f (loop* i))))) +-- -- -- -- snd morphtest = sElim2 {!!} +-- -- -- -- (funRec 3 isGroupoidS1 +-- -- -- -- λ f → funRec 3 isGroupoidS1 +-- -- -- -- λ g → pRec (isSetInt _ _) +-- -- -- -- (λ n=fn → +-- -- -- -- pRec (isSetInt _ _) +-- -- -- -- (λ n=gn → (λ j → winding (basechange (SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (∣ f (north) ∣ +ₖ ∣ n=gn (~ j) ∣)))) (λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (∣ f (loop* i) ∣ +ₖ ∣ transp (λ i → n=gn ((~ i) ∨ (~ j)) ≡ n=gn ((~ i) ∨ (~ j))) (~ j) (λ i → g (loop* i)) i ∣)))))) +-- -- -- -- ∙ {!!} +-- -- -- -- ∙ {!!}) +-- -- -- -- (isConnectedS1 (g north))) +-- -- -- -- (isConnectedS1 (f north))) + + +-- -- -- -- {- (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (∣ f (loop* i) ∣ +ₖ ∣ g (loop* i) ∣))))) +-- -- -- -- ∙ (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (Kn→ΩKn+1 1 ∣ f (loop* i) ∣ ∙ Kn→ΩKn+1 1 ∣ g (loop* i) ∣)))))) +-- -- -- -- ∙ (λ j → winding (basechange _ (cong (λ x → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (Kn→ΩKn+1 1 ∣ f x ∣ ∙ Kn→ΩKn+1 1 ∣ g x ∣))))) loop*)) ) +-- -- -- -- ∙ (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn ((cong ∣_∣ (merid (f (loop* i)) ∙ sym (merid north)) ∙ cong ∣_∣ (merid (g (loop* i)) ∙ sym (merid north))))))))) +-- -- -- -- ∙ (λ j → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (congFunct ∣_∣ (merid (f (loop* i)) ∙ sym (merid north)) (merid (g (loop* i)) ∙ sym (merid north)) (~ j))))))) +-- -- -- -- ∙ (λ j → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (cong ∣_∣ (({!!} ∙ {!!}) ∙ {!!}))))))) +-- -- -- -- ∙ {!!} +-- -- -- -- ∙ {!!} +-- -- -- -- ∙ {!!}) -} + +-- -- -- -- where +-- -- -- -- helper : ∀ {ℓ} {A : Type ℓ} (a : A) (f : A → S¹) (p q : a ≡ a) → winding (basechange (f a) (cong f (p ∙ q))) ≡ winding (basechange (f a) (cong f p ∙ cong f q)) +-- -- -- -- helper a f p q i = winding (basechange (f a) (congFunct f p q i)) +-- -- -- -- helper2 : (x : S¹) (p q : x ≡ x) → basechange x (p ∙ q) ≡ basechange x p ∙ basechange x q +-- -- -- -- helper2 base p q = refl +-- -- -- -- helper2 (loop i) p q = {!!} +-- -- -- -- helper4 : (x y z : S¹) (p : x ≡ z) (q r : x ≡ y) (s : y ≡ z) → basechange* x z p (q ∙ s) ≡ basechange* x y {!!} q ∙ {!!} +-- -- -- -- helper4 = {!!} +-- -- -- -- helper3 : (p q : ΩS¹) → winding (p ∙ q) ≡ (winding p +ℤ winding q) +-- -- -- -- helper3 = {!!} + + +-- -- -- -- -- fstmap : morph (dirProd intGroup intGroup) (coHomGr 0 (S₊ 0)) +-- -- -- -- -- fstmap = compMorph {F = dirProd intGroup intGroup} {G = ×coHomGr 0 Unit Unit} {H = coHomGr 0 (S₊ 0)} +-- -- -- -- -- (×morph (grIso.inv coHomGrUnit0) (grIso.inv coHomGrUnit0)) +-- -- -- -- -- (((MV.Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) zero)) , +-- -- -- -- -- {!MV.ΔIsHom _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) zero!}) + +-- -- -- -- -- fstMapId : (a : Int × Int) → fstmap .fst a ≡ ∣ (λ _ → proj₁ a +ℤ (0 - proj₂ a)) ∣₀ +-- -- -- -- -- fstMapId (a , b) = (λ i → ∣ (λ _ → a +ₖ (-ₖ b)) ∣₀) ∙ {!addLemma!} ∙ {!!} ∙ {!!} + +-- -- -- -- -- isoAgain : grIso intGroup (coHomGr 1 (S₊ 1)) +-- -- -- -- -- isoAgain = +-- -- -- -- -- Iso''→Iso _ _ +-- -- -- -- -- (iso'' ((λ a → ∣ (λ {north → 0ₖ ; south → 0ₖ ; (merid north i) → {!a!} ; (merid south i) → {!!}}) ∣₀) , {!!}) +-- -- -- -- -- {!!} +-- -- -- -- -- {!!}) + +-- -- -- -- -- -- -- test2 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → S¹) +-- -- -- -- -- -- -- Iso.fun test2 f = {!!} +-- -- -- -- -- -- -- Iso.inv test2 f north = ∣ transport (sym S¹≡S1) (f base) ∣ +-- -- -- -- -- -- -- Iso.inv test2 f south = ∣ transport (sym S¹≡S1) (f base) ∣ +-- -- -- -- -- -- -- Iso.inv test2 f (merid a i) = cong ∣_∣ {!transport (sym S¹≡S1) (f base)!} i +-- -- -- -- -- -- -- Iso.rightInv test2 = {!!} +-- -- -- -- -- -- -- Iso.leftInv test2 = {!!} + +-- -- -- -- -- -- -- F : winding (basechange base loop) ≡ 1 +-- -- -- -- -- -- -- F = refl + +-- -- -- -- -- -- -- another : (f g : Int) → winding (basechange {!!} {!!}) ≡ {!!} +-- -- -- -- -- -- -- another = {!!} + +-- -- -- -- -- -- -- test : Iso (S¹ → S¹) (S¹ × Int) +-- -- -- -- -- -- -- Iso.fun test f = f base , winding (basechange (f base) (cong f loop)) +-- -- -- -- -- -- -- Iso.inv test (x , int) base = x +-- -- -- -- -- -- -- Iso.inv test (x , int) (loop i) = {!!} +-- -- -- -- -- -- -- Iso.rightInv test = {!!} +-- -- -- -- -- -- -- Iso.leftInv test = {!!} + +-- -- -- -- -- -- -- -- test13 : Iso ∥ (S¹ → S¹) ∥₀ Int +-- -- -- -- -- -- -- -- Iso.fun test13 = sRec isSetInt λ f → winding (basechange (f base) (λ i → f (loop i))) +-- -- -- -- -- -- -- -- Iso.inv test13 a = ∣ (λ {base → {!!} ; (loop i) → {!!}}) ∣₀ +-- -- -- -- -- -- -- -- Iso.rightInv test13 = {!!} +-- -- -- -- -- -- -- -- Iso.leftInv test13 = {!!} + +-- -- -- -- -- -- -- -- test : Iso (S¹ → S¹) (S¹ × Int) +-- -- -- -- -- -- -- -- Iso.fun test f = (f base) , transport (basedΩS¹≡Int (f base)) λ i → f (loop i) +-- -- -- -- -- -- -- -- Iso.inv test (x , int) base = x +-- -- -- -- -- -- -- -- Iso.inv test (x , int) (loop i) = transport (sym (basedΩS¹≡Int x)) int i +-- -- -- -- -- -- -- -- Iso.rightInv test (x , int) i = (x , transportTransport⁻ (basedΩS¹≡Int x) int i) +-- -- -- -- -- -- -- -- Iso.leftInv test f = +-- -- -- -- -- -- -- -- funExt λ { base → refl +-- -- -- -- -- -- -- -- ; (loop i) j → transport⁻Transport (basedΩS¹≡Int (f base)) (λ i → f (loop i)) j i} + + +-- -- -- -- -- -- -- -- lem : S¹ ≡ hLevelTrunc 3 (S₊ 1) +-- -- -- -- -- -- -- -- lem = sym (truncIdempotent 3 isGroupoidS¹) ∙ λ i → hLevelTrunc 3 (S¹≡S1 (~ i)) + +-- -- -- -- -- -- -- -- ×≡ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (a b : A × B) → proj₁ a ≡ proj₁ b → proj₂ a ≡ proj₂ b → a ≡ b +-- -- -- -- -- -- -- -- ×≡ (_ , _) (_ , _) id1 id2 i = (id1 i) , (id2 i) + +-- -- -- -- -- -- -- -- test22 : Iso (S₊ 1 → coHomK 1) (S₊ 1 × Int) +-- -- -- -- -- -- -- -- Iso.fun test22 f = {!f north!} , {!!} +-- -- -- -- -- -- -- -- Iso.inv test22 = {!!} +-- -- -- -- -- -- -- -- Iso.rightInv test22 = {!!} +-- -- -- -- -- -- -- -- Iso.leftInv test22 = {!!} + + + + + + +-- -- -- -- -- -- -- -- coHom1≃∥S1×ℤ∥₀ : Iso (coHom 1 (S₊ 1)) ∥ S₊ 1 × Int ∥₀ +-- -- -- -- -- -- -- -- coHom1≃∥S1×ℤ∥₀ = setTruncIso test2 +-- -- -- -- -- -- -- -- where +-- -- -- -- -- -- -- -- test2 : Iso (S₊ 1 → coHomK 1) (S₊ 1 × Int) +-- -- -- -- -- -- -- -- Iso.fun test2 f = transport (λ i → S¹≡S1 (~ i) × Int) (Iso.fun test (transport (λ i → (S¹≡S1 i → lem (~ i))) f)) +-- -- -- -- -- -- -- -- Iso.inv test2 x = transport (λ i → (S¹≡S1 (~ i) → lem i)) (Iso.inv test (transport (λ i → S¹≡S1 i × Int) x)) +-- -- -- -- -- -- -- -- Iso.rightInv test2 (s , int) = ×≡ _ _ {!!} {!!} +-- -- -- -- -- -- -- -- Iso.leftInv test2 f = {!!} ∙ {!!} ∙ {!!} + +-- -- -- -- -- -- -- -- test2Id : (a b : (S₊ 1 → coHomK 1)) → proj₂ (Iso.fun test2 (λ x → a x +ₖ b x)) ≡ (proj₂ (Iso.fun test2 a) +ₖ proj₂ (Iso.fun test2 a)) +-- -- -- -- -- -- -- -- test2Id a b = {! +-- -- -- -- -- -- -- -- transport +-- -- -- -- -- -- -- -- (basedΩS¹≡Int +-- -- -- -- -- -- -- -- (transport (λ i → S¹≡S1 i → lem (~ i)) (λ x → a x +ₖ b x) base)) +-- -- -- -- -- -- -- -- (λ i → +-- -- -- -- -- -- -- -- transport (λ i₁ → S¹≡S1 i₁ → lem (~ i₁)) (λ x → a x +ₖ b x) +-- -- -- -- -- -- -- -- (loop i))!} ∙ {!transport (λ i → S¹≡S1 i → lem (~ i)) (λ x → a x +ₖ b x) base!} + + +-- -- -- -- -- -- -- -- main : grIso intGroup (coHomGr 1 (S₊ 1)) +-- -- -- -- -- -- -- -- main = Iso'→Iso +-- -- -- -- -- -- -- -- (iso' {!!} +-- -- -- -- -- -- -- -- {!!}) + + +-- -- -- -- -- -- coHom1 : grIso (coHomGr 1 (S₊ 1)) intGroup +-- -- -- -- -- -- coHom1 = +-- -- -- -- -- -- Iso'→Iso +-- -- -- -- -- -- (iso' (iso ({!!} ∘ {!!} ∘ {!!} ∘ {!!}) +-- -- -- -- -- -- {!!} +-- -- -- -- -- -- {!!} +-- -- -- -- -- -- {!!}) +-- -- -- -- -- -- {!!}) + + + +-- -- -- -- -- -- schonf : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : (A × B) → Type ℓ''} +-- -- -- -- -- -- → ((a : A) (b : B) → C (a , b)) +-- -- -- -- -- -- → (x : A × B) → C x +-- -- -- -- -- -- schonf f (a , b) = f a b + +-- -- -- -- -- -- -- -- setTruncProdIso : ∀ {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') → Iso ∥ (A × B) ∥₀ (∥ A ∥₀ × ∥ B ∥₀) +-- -- -- -- -- -- -- -- Iso.fun (setTruncProdIso A B) = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) λ {(a , b) → ∣ a ∣₀ , ∣ b ∣₀} +-- -- -- -- -- -- -- -- Iso.inv (setTruncProdIso A B) (a , b) = sRec setTruncIsSet (λ a → sRec setTruncIsSet (λ b → ∣ a , b ∣₀) b) a +-- -- -- -- -- -- -- -- Iso.rightInv (setTruncProdIso A B) = +-- -- -- -- -- -- -- -- schonf (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- -- -- -- -- -- -- -- λ _ → sElim (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- -- -- -- -- -- -- -- λ _ → refl) +-- -- -- -- -- -- -- -- Iso.leftInv (setTruncProdIso A B) = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ {(a , b) → refl} + +-- -- -- -- -- -- -- -- setTruncProdLemma : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (a1 a2 : A) (b : B) → isHLevelConnected 2 A +-- -- -- -- -- -- -- -- → Path ∥ A × B ∥₀ ∣ a1 , b ∣₀ ∣ a2 , b ∣₀ +-- -- -- -- -- -- -- -- setTruncProdLemma {A = A} {B = B} a1 a2 b conA i = Iso.inv (setTruncProdIso A B) (Iso.inv setTruncTrunc0Iso ((sym (conA .snd ∣ a1 ∣) ∙ (conA .snd ∣ a2 ∣)) i) , ∣ b ∣₀) + +-- -- -- -- -- -- -- -- test3 : Iso ∥ S₊ 1 × Int ∥₀ Int +-- -- -- -- -- -- -- -- Iso.fun test3 = sRec isSetInt proj₂ +-- -- -- -- -- -- -- -- Iso.inv test3 a = ∣ north , a ∣₀ +-- -- -- -- -- -- -- -- Iso.rightInv test3 a = refl +-- -- -- -- -- -- -- -- Iso.leftInv test3 = +-- -- -- -- -- -- -- -- sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- -- -- -- -- -- -- -- λ {(s , int) → setTruncProdLemma north s int (sphereConnected 1)} + +-- -- -- -- -- -- -- -- coHomGr0-S1 : grIso intGroup (coHomGr 1 (S₊ 1)) +-- -- -- -- -- -- -- -- coHomGr0-S1 = +-- -- -- -- -- -- -- -- Iso'→Iso +-- -- -- -- -- -- -- -- (iso' (compIso (symIso test3) (symIso coHom1≃∥S1×ℤ∥₀)) +-- -- -- -- -- -- -- -- (indIntGroup {G = coHomGr 1 (S₊ 1)} +-- -- -- -- -- -- -- -- (Iso.fun (compIso (symIso test3) (symIso coHom1≃∥S1×ℤ∥₀))) +-- -- -- -- -- -- -- -- ((λ i → ∣ transport (λ i → (S¹≡S1 (~ i) → lem i)) (Iso.inv test (base , 0)) ∣₀) +-- -- -- -- -- -- -- -- ∙ (λ i → ∣ transport (λ i → (S¹≡S1 (~ i) → lem i)) (helper2 i) ∣₀) +-- -- -- -- -- -- -- -- ∙ cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → {!!}})) +-- -- -- -- -- -- -- -- {!!} +-- -- -- -- -- -- -- -- {!!})) +-- -- -- -- -- -- -- -- where +-- -- -- -- -- -- -- -- helper : basedΩS¹≡ΩS¹ base ≡ {!basechange!} +-- -- -- -- -- -- -- -- helper = {!substComposite!} + +-- -- -- -- -- -- -- -- substComposite2 : ∀ {ℓ} {A B C : Type ℓ} +-- -- -- -- -- -- -- -- (P : A ≡ B) (Q : B ≡ C) (a : A) +-- -- -- -- -- -- -- -- → transport (P ∙ Q) a ≡ transport Q (transport P a) +-- -- -- -- -- -- -- -- substComposite2 = {!!} + +-- -- -- -- -- -- -- -- helper1 : transport (λ i → S¹≡S1 i × Int) (north , 0) ≡ (base , 0) +-- -- -- -- -- -- -- -- helper1 = refl +-- -- -- -- -- -- -- -- helper3 : transport (sym (basedΩS¹≡Int base)) 0 ≡ refl +-- -- -- -- -- -- -- -- helper3 = (λ i → transport (symDistr (basedΩS¹≡ΩS¹ base) ΩS¹≡Int i) 0) +-- -- -- -- -- -- -- -- ∙ substComposite2 (sym ΩS¹≡Int) (sym (basedΩS¹≡ΩS¹ base)) 0 +-- -- -- -- -- -- -- -- ∙ (λ i → transport (λ i → basedΩS¹≡ΩS¹ base (~ i)) refl) -- +-- -- -- -- -- -- -- -- ∙ transportRefl ((equiv-proof (basechange-isequiv base) refl) .fst .fst) +-- -- -- -- -- -- -- -- ∙ (λ i → equiv-proof (transport (λ j → isEquiv (refl-conjugation j)) (basedΩS¹→ΩS¹-isequiv i0)) refl .fst .fst) +-- -- -- -- -- -- -- -- ∙ (λ i → {!equiv-proof (transport (λ j → isEquiv (refl-conjugation j)) (basedΩS¹→ΩS¹-isequiv i0)) refl .fst!}) +-- -- -- -- -- -- -- -- ∙ {!basedΩS¹→ΩS¹!} +-- -- -- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- -- -- helper4 : (x : S¹) → Iso.inv test (base , 0) x ≡ x +-- -- -- -- -- -- -- -- helper4 = {!!} +-- -- -- -- -- -- -- -- helper2 : Iso.inv test (transport (λ i → S¹≡S1 i × Int) (north , 0)) ≡ λ x → x +-- -- -- -- -- -- -- -- helper2 = (λ i → Iso.inv test (base , 0)) ∙ {!!} ∙ {!!} + +-- -- -- -- -- -- ×≡ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A × B} → proj₁ x ≡ proj₁ y → proj₂ x ≡ proj₂ y → x ≡ y +-- -- -- -- -- -- ×≡ {x = (a , b)} {y = (c , d)} id1 id2 i = (id1 i) , (id2 i) + +-- -- -- -- -- -- fstFunHelper : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- -- → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 _) (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x +-- -- -- -- -- -- fstFunHelper a = MV.Ker-i⊂Im-d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a +-- -- -- -- -- -- (sym (isContrH1Unit×H1Unit .snd _) ∙ (isContrH1Unit×H1Unit .snd _)) +-- -- -- -- -- -- where +-- -- -- -- -- -- isContrH1Unit×H1Unit : isContr (Group.type (×coHomGr 1 Unit Unit)) +-- -- -- -- -- -- isContrH1Unit×H1Unit = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) +-- -- -- -- -- -- , λ {(a , b) → sigmaProdElim {D = λ (x : Σ[ x ∈ Group.type (×coHomGr 1 Unit Unit)] Unit) → (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) ≡ fst x} +-- -- -- -- -- -- (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- -- -- -- -- -- (λ a b _ → ×≡ (grIso.leftInv (coHomGrUnit≥1 0) ∣ a ∣₀) (grIso.leftInv (coHomGrUnit≥1 0) ∣ b ∣₀)) ((a , b) , tt)} + + + +-- -- -- -- -- -- helperMorph : isMorph intGroup (dirProd intGroup intGroup) λ a → a , (0 - a) +-- -- -- -- -- -- helperMorph = +-- -- -- -- -- -- indIntGroup {G = dirProd intGroup intGroup} +-- -- -- -- -- -- (λ a → a , (0 - a)) +-- -- -- -- -- -- refl +-- -- -- -- -- -- (λ a → ×≡ refl (helper2 a)) +-- -- -- -- -- -- λ a → ×≡ refl (helper3 a) +-- -- -- -- -- -- where +-- -- -- -- -- -- helper1 : (a : Int) → predInt (sucInt a) ≡ a +-- -- -- -- -- -- helper1 (pos zero) = refl +-- -- -- -- -- -- helper1 (pos (suc n)) = refl +-- -- -- -- -- -- helper1 (negsuc zero) = refl +-- -- -- -- -- -- helper1 (negsuc (suc n)) = refl + +-- -- -- -- -- -- helper4 : (a : Int) → sucInt (predInt a) ≡ a +-- -- -- -- -- -- helper4 (pos zero) = refl +-- -- -- -- -- -- helper4 (pos (suc n)) = refl +-- -- -- -- -- -- helper4 (negsuc zero) = refl +-- -- -- -- -- -- helper4 (negsuc (suc n)) = refl + +-- -- -- -- -- -- helper2 : (a : Int) → (pos 0 - sucInt a) ≡ predInt (pos 0 - a) +-- -- -- -- -- -- helper2 (pos zero) = refl +-- -- -- -- -- -- helper2 (pos (suc n)) = refl +-- -- -- -- -- -- helper2 (negsuc zero) = refl +-- -- -- -- -- -- helper2 (negsuc (suc n)) = sym (helper1 _) + +-- -- -- -- -- -- helper3 : (a : Int) → (pos 0 - predInt a) ≡ sucInt (pos 0 - a) +-- -- -- -- -- -- helper3 (pos zero) = refl +-- -- -- -- -- -- helper3 (pos (suc zero)) = refl +-- -- -- -- -- -- helper3 (pos (suc (suc n))) = sym (helper4 _) +-- -- -- -- -- -- helper3 (negsuc zero) = refl +-- -- -- -- -- -- helper3 (negsuc (suc n)) = refl + +-- -- -- -- -- -- helper : (a b : Int) → (pos 0 - (a +ℤ b)) ≡ ((pos 0 - a) +ℤ (pos 0 - b)) +-- -- -- -- -- -- helper a (pos zero) = refl +-- -- -- -- -- -- helper (pos zero) (pos (suc n)) = +-- -- -- -- -- -- cong (λ x → pos 0 - sucInt x) (+ℤ-comm (pos zero) (pos n)) +-- -- -- -- -- -- ∙ +ℤ-comm (pos 0 +negsuc n) (pos zero) +-- -- -- -- -- -- helper (pos (suc n₁)) (pos (suc n)) = +-- -- -- -- -- -- {!!} +-- -- -- -- -- -- helper (negsuc n₁) (pos (suc n)) = {!!} +-- -- -- -- -- -- helper a (negsuc n) = {!!} + +-- -- -- -- -- -- fun : morph intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- -- fst fun = MV.d _ _ _ (λ _ → tt) (λ _ → tt) 0 ∘ grIso.inv coHom0-S0 .fst ∘ λ a → a , (0 - a) +-- -- -- -- -- -- snd fun = {!!} +-- -- -- -- -- -- {- compMorph {F = intGroup} {G = dirProd intGroup intGroup} {H = coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))} +-- -- -- -- -- -- ((λ a → a , a) , (λ a b → refl)) +-- -- -- -- -- -- (compMorph {F = dirProd intGroup intGroup} {G = coHomGr 0 (S₊ 0)} {H = coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))} (grIso.inv coHom0-S0) +-- -- -- -- -- -- (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 +-- -- -- -- -- -- , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) .snd -} +-- -- -- -- -- -- {- theIso : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- -- theIso = Iso''→Iso _ _ +-- -- -- -- -- -- (iso'' ((λ x → ∣ (λ {(inl tt) → 0ₖ ; (inr tt) → 0ₖ ; (push a i) → Kn→ΩKn+1 0 x i}) ∣₀) , {!!}) +-- -- -- -- -- -- {!!} +-- -- -- -- -- -- {!MV.d!}) +-- -- -- -- -- -- -} + + + +-- -- -- -- -- -- theIso : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- -- theIso = +-- -- -- -- -- -- Iso''→Iso _ _ +-- -- -- -- -- -- (iso'' fun +-- -- -- -- -- -- (λ x inker → {!MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 +-- -- -- -- -- -- (grIso.inv coHom0-S0 .fst (g , g))!}) +-- -- -- -- -- -- (sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- -- -- -- -- -- λ x → pRec propTruncIsProp +-- -- -- -- -- -- (λ {(a , b) → {!fun!} }) +-- -- -- -- -- -- (fstFunHelper (∣ x ∣₀)))) +-- -- -- -- -- -- where +-- -- -- -- -- -- whoKnows : (a : S₊ 0 → Int) (x : MV.D Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt)) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (λ _ → a north) x +-- -- -- -- -- -- ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 a x +-- -- -- -- -- -- whoKnows a (inl x) = refl +-- -- -- -- -- -- whoKnows a (inr x) = refl +-- -- -- -- -- -- whoKnows a (push north i) = refl +-- -- -- -- -- -- whoKnows a (push south i) j = {!!} + +-- -- -- -- -- -- helper : (a : Int) → (grIso.inv coHom0-S0 .fst (a , a)) ≡ ∣ S0→Int (a , a) ∣₀ +-- -- -- -- -- -- helper a = {!have : + +-- -- -- -- -- -- ∣ +-- -- -- -- -- -- MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 +-- -- -- -- -- -- (S0→Int (x , x)) +-- -- -- -- -- -- ∣₀ +-- -- -- -- -- -- ≡ ∣ (λ _ → ∣ north ∣) ∣₀!} + +-- -- -- -- -- -- helper2 : (a b : Int) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a)) ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (b , b)) +-- -- -- -- -- -- → a ≡ b +-- -- -- -- -- -- helper2 a b id = pRec (isSetInt a b) (λ {(pt , id) → {!!}}) (fstFunHelper ∣ (MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a))) ∣₀) + +-- -- -- -- -- -- idFun : (a : Int) → MV.D Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 +-- -- -- -- -- -- idFun a (inl x) = 0ₖ +-- -- -- -- -- -- idFun a (inr x) = 0ₖ +-- -- -- -- -- -- idFun a (push north i) = Kn→ΩKn+1 zero a i +-- -- -- -- -- -- idFun a (push south i) = Kn→ΩKn+1 zero a i + +-- -- -- -- -- -- helper3 : (a : Int) → (MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a))) ≡ idFun a +-- -- -- -- -- -- helper3 a = funExt λ {(inl x) → refl ; (inr x) → refl ; (push north i) → refl ; (push south i) → refl } + +-- -- -- -- -- -- helper4 : (a b : Int) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a)) ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (b , b)) +-- -- -- -- -- -- → a ≡ b +-- -- -- -- -- -- helper4 a b id = +-- -- -- -- -- -- {!!} +-- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- where +-- -- -- -- -- -- helper5 : {!!} --PathP (λ k → id k (inl tt) ≡ id k (inr tt)) (Kn→ΩKn+1 zero a) (Kn→ΩKn+1 zero a) +-- -- -- -- -- -- helper5 i j = {!id i!} + +-- -- -- -- -- -- -- fun : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) → coHom 0 (S₊ 0) +-- -- -- -- -- -- -- fun a = (pRec {P = Σ[ x ∈ coHom 0 (S₊ 0)] (MV.d _ _ _ (λ _ → tt) (λ _ → tt) 0 x ≡ a) × isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) (MV.Δ _ _ _ (λ _ → tt) (λ _ → tt) 0) x} +-- -- -- -- -- -- -- (λ {(a1 , b) (c , d) → ΣProp≡ (λ x → isOfHLevelProd 1 (setTruncIsSet _ _) propTruncIsProp) +-- -- -- -- -- -- -- (sElim2 {B = λ a1 c → (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a1 ≡ a) +-- -- -- -- -- -- -- → MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 c ≡ a +-- -- -- -- -- -- -- → isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) +-- -- -- -- -- -- -- (λ z → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 z) a1 +-- -- -- -- -- -- -- → isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) +-- -- -- -- -- -- -- (λ z → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 z) c → a1 ≡ c} +-- -- -- -- -- -- -- (λ _ _ → {!!}) +-- -- -- -- -- -- -- (λ a c b1 d1 → pElim (λ _ → isOfHLevelΠ 1 λ _ → setTruncIsSet _ _) +-- -- -- -- -- -- -- λ b2 → pElim (λ _ → setTruncIsSet _ _) +-- -- -- -- -- -- -- λ d2 → {!d2!}) +-- -- -- -- -- -- -- a1 c (proj₁ b) (proj₁ d) (proj₂ b) (proj₂ d))}) +-- -- -- -- -- -- -- (λ {(a , b) → a , b , MV.Ker-d⊂Im-Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a {!!}}) +-- -- -- -- -- -- -- (fstFunHelper a)) .fst -- pRec {!!} {!!} (fstFunHelper a) diff --git a/Cubical/ZCohomology/MayerVietoris.agda b/Cubical/ZCohomology/MayerVietoris.agda deleted file mode 100644 index 08db1c8e7..000000000 --- a/Cubical/ZCohomology/MayerVietoris.agda +++ /dev/null @@ -1,235 +0,0 @@ - -{-# OPTIONS --cubical --safe #-} -module Cubical.ZCohomology.MayerVietoris where - -open import Cubical.ZCohomology.Base -open import Cubical.HITs.S1 -open import Cubical.ZCohomology.Properties -open import Cubical.HITs.Sn -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Function -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Transport -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Univalence -open import Cubical.Data.NatMinusTwo.Base -open import Cubical.Data.Empty -open import Cubical.Data.Sigma hiding (_×_) -open import Cubical.HITs.Susp -open import Cubical.HITs.Wedge -open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) -open import Cubical.HITs.Nullification -open import Cubical.Data.Int hiding (_+_) -open import Cubical.Data.Nat -open import Cubical.Data.Prod -open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) -open import Cubical.Homotopy.Loopspace -open import Cubical.Homotopy.Connected -open import Cubical.Homotopy.Freudenthal -open import Cubical.HITs.SmashProduct.Base -open import Cubical.Data.Group.Base - - -open import Cubical.ZCohomology.KcompPrelims - - -open import Cubical.HITs.Pushout -open import Cubical.Data.Sum.Base -open import Cubical.Data.HomotopyGroup -open import Cubical.Data.Bool hiding (_⊕_) - - --- {- --- record AbGroup {ℓ} (A : Type ℓ) : Type ℓ where --- constructor abgr --- field --- noll : A --- _⊕_ : A → A → A --- associate : (a b c : A) → (a ⊕ b) ⊕ c ≡ a ⊕ (b ⊕ c) --- commute : (a b : A) → a ⊕ b ≡ (b ⊕ a) --- r-unit : (a : A) → a ⊕ noll ≡ a --- negate : (a : A) → Σ[ a⁻ ∈ A ] (a ⊕ a⁻ ≡ noll) --- -} --- {- --- grHom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (ϕ : A → B) --- → AbGroup A → AbGroup B → Type (ℓ-max ℓ ℓ') --- grHom {A = A} {B = B} ϕ (abgr 0A _+A_ as-A comm-A r-A neg-A) (abgr 0B _+B_ as-B comm-B r-B neg-B) = (x y : A) → ϕ (x +A y) ≡ (ϕ x +B ϕ y) --- ImHom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (ϕ : A → B) (A' : AbGroup A) (B' : AbGroup B) --- → grHom ϕ A' B' → Type (ℓ-max ℓ ℓ') --- ImHom {A = A} {B = B} ϕ Agr Bgr hom = Σ[ b ∈ B ] Σ[ a ∈ A ] ϕ a ≡ b --- KerHom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (ϕ : A → B) (A' : AbGroup A) (B' : AbGroup B) --- → grHom ϕ A' B' → Type (ℓ-max ℓ ℓ') --- KerHom {A = A} {B = B} ϕ Agr (abgr 0B _+B_ as-B comm-B r-B neg-B) hom = Σ[ a ∈ A ] ϕ a ≡ 0B --- -} - - --- --- - - -private - variable - ℓ ℓ' ℓ'' : Level - A : Type ℓ - B : Type ℓ' - C : Type ℓ'' - A' : Pointed ℓ - B' : Pointed ℓ' - - --- coHomFun : (n : ℕ) (f : A → B) → coHom n B → coHom n A --- coHomFun n f = sRec setTruncIsSet λ β → ∣ β ∘ f ∣₀ - --- coHomFun∙ : (n : ℕ) (f : A' →∙ B') → coHomRed n B' → coHomRed n A' --- coHomFun∙ n (f , fpt) = sRec setTruncIsSet λ { (β , βpt) → ∣ β ∘ f , (λ i → β (fpt i)) ∙ βpt ∣₀} - --- coHomFun∙2 : (n : ℕ) (f : A → B' .fst) → coHomRed n B' → coHom n A --- coHomFun∙2 zero f = sRec setTruncIsSet λ { (β , βpt) → ∣ β ∘ f ∣₀} --- coHomFun∙2 (suc n) f = sRec setTruncIsSet λ { (β , βpt) → ∣ β ∘ f ∣₀} - --- module MV {ℓ ℓ' ℓ''} ((A , a₀) : Pointed ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) where --- D : Pointed (ℓ-max (ℓ-max ℓ ℓ') ℓ'') --- D = Pushout f g , inl a₀ - --- -- σ : typ A → typ (Ω (∙Susp (typ A))) --- -- σ a = merid a ∙ merid (pt A) ⁻¹ - --- i : (n : ℕ) → coHomRed n D → coHomRed n (A , a₀) × coHom n B --- i zero = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) --- λ { (δ , δpt) → ∣ (λ x → δ (inl x)) , δpt ∣₀ , ∣ (λ x → δ (inr x)) ∣₀ } --- i (suc n) = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) --- λ { (δ , δpt) → ∣ (λ x → δ (inl x)) , δpt ∣₀ , ∣ (λ x → δ (inr x)) ∣₀ } - --- Δ : (n : ℕ) → coHomRed n (A , a₀) × coHom n B → coHom n C --- Δ n (α , β) = coHomFun∙2 n f α +ₕ (-ₕ (coHomFun n g β)) - --- -- - --- d-pre1 : (n : ℕ) → (C → coHomK n) → D .fst → coHomK-ptd (suc n) .fst --- d-pre1 zero γ (inl x) = 0ₖ --- d-pre1 zero γ (inr x) = 0ₖ --- d-pre1 zero γ (push a i) = Kn→ΩKn+1 zero (γ a) i --- d-pre1 (suc n) γ (inl x) = 0ₖ --- d-pre1 (suc n) γ (inr x) = 0ₖ --- d-pre1 (suc n) γ (push a i) = Kn→ΩKn+1 (suc n) (γ a) i - --- d-pre1∙ : (n : ℕ) → (γ : (C → coHomK n)) → d-pre1 n γ (snd D) ≡ ∣ north ∣ --- d-pre1∙ zero γ = refl --- d-pre1∙ (suc n) γ = refl - --- d-pre2 : (n : ℕ) → (C → coHomK n) → D →∙ coHomK-ptd (suc n) --- d-pre2 n γ = d-pre1 n γ , d-pre1∙ n γ - --- d : (n : ℕ) → coHom n C → coHomRed (suc n) D --- d n = sRec setTruncIsSet λ a → ∣ d-pre2 n a ∣₀ - --- iIsHom : (n : ℕ) (x y : coHomRed n D) → i n (+ₕ∙ n x y) ≡ +prod n (i n x) (i n y) --- iIsHom zero = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) --- λ { (f , fpt) (g , gpt) → refl} --- iIsHom (suc n) = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) --- λ { (f , fpt) (g , gpt) → refl} - --- prodElim : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : ∥ A ∥₀ × ∥ B ∥₀ → Type ℓ''} --- → ((x : ∥ A ∥₀ × ∥ B ∥₀) → isOfHLevel 2 (C x)) --- → ((a : A) (b : B) → C (∣ a ∣₀ , ∣ b ∣₀)) --- → (x : ∥ A ∥₀ × ∥ B ∥₀) → C x --- prodElim {A = A} {B = B} {C = C} hlevel ind (a , b) = schonf a b --- where --- schonf : (a : ∥ A ∥₀) (b : ∥ B ∥₀) → C (a , b) --- schonf = sElim (λ x → isOfHLevelΠ 2 λ y → hlevel (_ , _)) λ a → sElim (λ x → hlevel (_ , _)) --- λ b → ind a b - --- ΔIsHom : (n : ℕ) (x y : coHomRed n (A , a₀) × coHom n B) → Δ n (+prod n x y) ≡ (Δ n x +ₕ Δ n y) --- ΔIsHom zero = prodElim (λ x → isOfHLevelΠ 2 λ y → isOfHLevelPath 2 setTruncIsSet _ _ ) --- λ {(f' , fpt) x1 → prodElim (λ x → isOfHLevelPath 2 setTruncIsSet _ _) --- λ {(g' , gpt) x2 → (λ i → ∣ (λ x → (f' (f x) +ₖ g' (f x)) +ₖ -distrₖ (x1 (g x)) (x2 (g x)) i) ∣₀) ∙ --- (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ g' (f x)) (-ₖ x1 (g x)) (-ₖ x2 (g x)) (~ i)) ∣₀) ∙ --- (λ i → ∣ (λ x → assocₖ (f' (f x)) (g' (f x)) (-ₖ x1 (g x)) i +ₖ (-ₖ x2 (g x))) ∣₀) ∙ --- (λ i → ∣ (λ x → (f' (f x) +ₖ commₖ (g' (f x)) (-ₖ x1 (g x)) i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ --- (λ i → ∣ (λ x → assocₖ (f' (f x)) (-ₖ x1 (g x)) (g' (f x)) (~ i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ --- (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ (-ₖ x1 (g x))) (g' (f x)) (-ₖ x2 (g x)) i) ∣₀)}} --- ΔIsHom (suc n) = prodElim (λ x → isOfHLevelΠ 2 λ y → isOfHLevelPath 2 setTruncIsSet _ _ ) --- λ {(f' , fpt) x1 → prodElim (λ x → isOfHLevelPath 2 setTruncIsSet _ _) --- λ {(g' , gpt) x2 → (λ i → ∣ (λ x → (f' (f x) +ₖ g' (f x)) +ₖ -distrₖ (x1 (g x)) (x2 (g x)) i) ∣₀) ∙ --- (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ g' (f x)) (-ₖ x1 (g x)) (-ₖ x2 (g x)) (~ i)) ∣₀) ∙ --- (λ i → ∣ (λ x → assocₖ (f' (f x)) (g' (f x)) (-ₖ x1 (g x)) i +ₖ (-ₖ x2 (g x))) ∣₀) ∙ --- (λ i → ∣ (λ x → (f' (f x) +ₖ commₖ (g' (f x)) (-ₖ x1 (g x)) i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ --- (λ i → ∣ (λ x → assocₖ (f' (f x)) (-ₖ x1 (g x)) (g' (f x)) (~ i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ --- (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ (-ₖ x1 (g x))) (g' (f x)) (-ₖ x2 (g x)) i) ∣₀)}} - --- niceHelper : (n : ℕ) (h l : C → coHomK n) (x : D .fst) → (d-pre1 n h x +ₖ d-pre1 n l x) ≡ d-pre1 n (λ x → h x +ₖ l x) x --- niceHelper zero h l (inl x) = {!!} --- niceHelper zero h l (inr x) = {!!} --- niceHelper zero h l (push a i₁) = {!!} --- niceHelper (suc n) h l (inl x) = {!!} -- lUnitₖ 0ₖ --- niceHelper (suc n) h l (inr x) = {!!} -- lUnitₖ 0ₖ --- niceHelper (suc n) h l (push a i) j = --- hcomp (λ k → λ {(i = i0) → {!lUnitₖ 0ₖ j!} --- ; (i = i1) → {!!} --- ; (j = i0) → {!!} --- ; (j = i1) → {!!}}) --- {!!} --- where --- side : I → I → coHomK n --- side i j = --- hcomp {!!} --- {!!} - --- helper : (h l : C → Int) (x : D .fst) → (d-pre1 zero h x +ₖ d-pre1 zero l x) ≡ d-pre1 zero (λ x → h x +ₖ l x) x --- helper h l (inl x) = lUnitₖ 0ₖ --- helper h l (inr x) = lUnitₖ 0ₖ --- helper h l (push a i) = {!!} --- where --- helper1 : Kn→ΩKn+1 zero (h a +ₖ l a) ≡ cong₂ (λ x y → x +ₖ y) (Kn→ΩKn+1 zero (h a)) (Kn→ΩKn+1 zero (l a)) --- helper1 = (λ i → Kn→ΩKn+1 zero (h a +ₖ l a)) ∙ +ₖ→∙ zero (h a) (l a) ∙ --- {!!} ∙ --- sym (cong₂Funct (λ x y → x +ₖ y) (Kn→ΩKn+1 zero (h a)) (Kn→ΩKn+1 zero (l a))) ∙ --- (λ i j → Kn→ΩKn+1 zero (h a) j +ₖ Kn→ΩKn+1 zero (l a) j) ∙ --- λ i j → ΩKn+1→Kn (Kn→ΩKn+1 1 (Kn→ΩKn+1 zero (h a) j) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 zero (l a) j)) --- helper2 : PathP (λ i → rUnitₖ 0ₖ i ≡ rUnitₖ 0ₖ i ) (cong (λ x → x +ₖ ∣ north ∣) (Kn→ΩKn+1 zero (h a)) ∙ cong (λ y → ∣ north ∣ +ₖ y) (Kn→ΩKn+1 zero (l a))) --- (Kn→ΩKn+1 zero (h a) ∙ Kn→ΩKn+1 zero (l a)) --- helper2 j i = hcomp (λ k → λ { (i = i0) → rUnitₖ ∣ north ∣ j --- ; (i = i1) → {!!}}) --- (rUnitₖ (Kn→ΩKn+1 zero (h a) i) j) -- (cong (λ x → {!rUnitₖ {n = 1} x i!}) (Kn→ΩKn+1 zero (h a)) ∙ cong (λ y → lUnitₖ {n = 1} y i) (Kn→ΩKn+1 zero (l a))) - --- helper3 : PathP (λ i → rUnitₖ 0ₖ i ≡ rUnitₖ 0ₖ i) (cong (λ x → x +ₖ ∣ north ∣) (Kn→ΩKn+1 zero (h a))) (Kn→ΩKn+1 zero (h a)) --- helper3 i = cong (λ x → rUnitₖ x i) (Kn→ΩKn+1 zero (h a)) - --- heler4 : PathP (λ i → lUnitₖ 0ₖ i ≡ lUnitₖ 0ₖ i) (cong (λ x → ∣ north ∣ +ₖ x) (Kn→ΩKn+1 zero (l a))) (Kn→ΩKn+1 zero (l a)) --- heler4 i = cong (λ x → lUnitₖ x i) (Kn→ΩKn+1 zero (l a)) - --- ttest : {!!} --- ttest = congFunct --- -- test2 : (n : ℕ) (a : C) (h l : C → coHomK n) → cong₂ (λ x y → x +ₖ y) (Kn→ΩKn+1 n (h a)) (Kn→ΩKn+1 n (l a)) ≡ {!!} ∙ Kn→ΩKn+1 n (h a +ₖ l a) ∙ {!!} --- -- test2 = {!!} - --- -- pp : PathP (λ i → lUnitₖ (0ₖ {n = 1}) i ≡ lUnitₖ 0ₖ i) (λ i → Kn→ΩKn+1 zero (h a) i +ₖ Kn→ΩKn+1 zero (l a) i) (Kn→ΩKn+1 zero (h a +ₖ l a)) --- -- pp = toPathP ((λ z → transport (λ i → lUnitₖ (0ₖ {n = 1}) i ≡ lUnitₖ 0ₖ i) (lUnit (rUnit (λ i → Kn→ΩKn+1 zero (h a) i +ₖ Kn→ΩKn+1 zero (l a) i) z) z )) --- -- ∙ (λ z → transp (λ i → lUnitₖ (0ₖ {n = 1}) (i ∨ z) ≡ lUnitₖ 0ₖ (i ∨ z)) z --- -- (((λ i → lUnitₖ (0ₖ {n = 1}) ((~ i) ∧ z)) ∙ (λ i → Kn→ΩKn+1 zero (h a) i +ₖ Kn→ΩKn+1 zero (l a) i) --- -- ∙ λ i → lUnitₖ (0ₖ {n = 1}) (i ∧ z)))) --- -- ∙ (λ z → sym ((λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ 1 i ∙ Kn→ΩKn+1 1 0ₖ)) ∙ --- -- (λ i → ΩKn+1→Kn (lUnit (Kn→ΩKn+1 1 0ₖ) (~ i))) ∙ --- -- Iso.leftInv (Iso2-Kn-ΩKn+1 1) 0ₖ) --- -- ∙ (lUnit (rUnit (λ i → Kn→ΩKn+1 zero (h a) i +ₖ Kn→ΩKn+1 zero (l a) i) z) z) --- -- ∙ ((λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ 1 i ∙ Kn→ΩKn+1 1 0ₖ)) ∙ --- -- (λ i → ΩKn+1→Kn (lUnit (Kn→ΩKn+1 1 0ₖ) (~ i))) ∙ --- -- Iso.leftInv (Iso2-Kn-ΩKn+1 1) 0ₖ)) --- -- ∙ (λ z → sym (lUnitₖ (0ₖ {n = 1})) --- -- ∙ (lUnit (rUnit refl z ) z ∙ (λ i → ΩKn+1→Kn (Kn→ΩKn+1 1 (Kn→ΩKn+1 zero (h a) i) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 zero (l a) i))) ∙ lUnit (rUnit refl z) z) --- -- ∙ lUnitₖ (0ₖ {n = 1})) --- -- ∙ ((λ z → sym (lUnitₖ (0ₖ {n = 1})) --- -- ∙ (({!!} ∙ {!ΩKn+1→Kn (Kn→ΩKn+1 1 (Kn→ΩKn+1 zero (h a) i) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 zero (l a) i)!} ∙ λ i → {!l!}) --- -- ∙ (λ i → Iso.leftInv (Iso2-Kn-ΩKn+1 1) {!!} {!!}) ∙ (λ i → (Iso.leftInv (Iso2-Kn-ΩKn+1 1) 0ₖ (~ i ∧ z))) ∙ {!!} ∙ {!!}) --- -- ∙ lUnitₖ (0ₖ {n = 1}))) --- -- ∙ {!!} --- -- ∙ {!!} --- -- ∙ {!!}) -- compPathP (λ j i → {!!}) (compPathP {!!} (compPathP {!!} {!!})) - --- -- dIsHom : (n : ℕ) (x y : coHom n C) → d n (x +ₕ y) ≡ +ₕ∙ (suc n) (d n x) (d n y) --- -- dIsHom zero = sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) --- -- λ h l → (λ i → ∣ (λ x → helper h l x (~ i)) , lUnit (λ j → lUnitₖ 0ₖ (j ∨ (~ i))) i ∣₀) ∙ --- -- (λ i → ∣ (λ x → d-pre1 zero h x +ₖ d-pre1 zero l x) , refl ∙ lUnitₖ 0ₖ ∣₀) --- -- dIsHom (suc n) = {!d-pre1 zero h x +ₖ d-pre1 zero l x!} diff --git a/Cubical/ZCohomology/MayerVietorisReduced.agda b/Cubical/ZCohomology/MayerVietorisUnreduced.agda similarity index 67% rename from Cubical/ZCohomology/MayerVietorisReduced.agda rename to Cubical/ZCohomology/MayerVietorisUnreduced.agda index c813e3898..ff032ba5f 100644 --- a/Cubical/ZCohomology/MayerVietorisReduced.agda +++ b/Cubical/ZCohomology/MayerVietorisUnreduced.agda @@ -1,6 +1,5 @@ - {-# OPTIONS --cubical --safe #-} -module Cubical.ZCohomology.MayerVietorisReduced where +module Cubical.ZCohomology.MayerVietorisUnreduced where open import Cubical.ZCohomology.Base open import Cubical.HITs.S1 @@ -55,27 +54,40 @@ private coHomFun : (n : ℕ) (f : A → B) → coHom n B → coHom n A coHomFun n f = sRec setTruncIsSet λ β → ∣ β ∘ f ∣₀ +prodElim2 : ∀ {ℓ ℓ' ℓ'' ℓ''' ℓ''''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} + {E : (∥ A ∥₀ × ∥ B ∥₀) → (∥ C ∥₀ × ∥ D ∥₀) → Type ℓ''''} + → ((x : ∥ A ∥₀ × ∥ B ∥₀) (y : ∥ C ∥₀ × ∥ D ∥₀) → isOfHLevel 2 (E x y)) + → ((a : A) (b : B) (c : C) (d : D) → E (∣ a ∣₀ , ∣ b ∣₀) (∣ c ∣₀ , ∣ d ∣₀)) + → ((x : ∥ A ∥₀ × ∥ B ∥₀) (y : ∥ C ∥₀ × ∥ D ∥₀) → (E x y)) +prodElim2 isset f = prodElim (λ _ → isOfHLevelΠ 2 λ _ → isset _ _) + λ a b → prodElim (λ _ → isset _ _) + λ c d → f a b c d + module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) where D : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') D = Pushout f g - i : (n : ℕ) → coHom n D → coHom n A × coHom n B + i : (n : ℕ) → morph (coHomGr n D) (×coHomGr n A B) + morph.fun (i zero) = + sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) + λ { δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀} + morph.fun (i (suc n)) = + sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) + λ { δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀} + morph.ismorph (i zero) = + sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ f g → refl + morph.ismorph (i (suc n)) = + sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ f g → refl + +{- i zero = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) λ { δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀} i (suc n) = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) λ { δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀} - - Δ : (n : ℕ) → coHom n A × coHom n B → coHom n C - Δ n (α , β) = coHomFun n f α +ₕ -ₕ (coHomFun n g β) - - d-pre : (n : ℕ) → (C → coHomK n) → D → coHomK (suc n) - d-pre n γ (inl x) = 0ₖ - d-pre n γ (inr x) = 0ₖ - d-pre zero γ (push a i) = Kn→ΩKn+1 zero (γ a) i - d-pre (suc n) γ (push a i) = Kn→ΩKn+1 (suc n) (γ a) i - - d : (n : ℕ) → coHom n C → coHom (suc n) D - d n = sRec setTruncIsSet λ a → ∣ d-pre n a ∣₀ +-} +{- iIsHom : (n : ℕ) → isMorph (coHomGr n D) (×coHomGr n A B) (i n) iIsHom zero = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) @@ -83,27 +95,27 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : iIsHom (suc n) = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) λ f g → refl - - prodElim : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : ∥ A ∥₀ × ∥ B ∥₀ → Type ℓ''} - → ((x : ∥ A ∥₀ × ∥ B ∥₀) → isOfHLevel 2 (C x)) - → ((a : A) (b : B) → C (∣ a ∣₀ , ∣ b ∣₀)) - → (x : ∥ A ∥₀ × ∥ B ∥₀) → C x - prodElim {A = A} {B = B} {C = C} hlevel ind (a , b) = schonf a b - where - schonf : (a : ∥ A ∥₀) (b : ∥ B ∥₀) → C (a , b) - schonf = sElim (λ x → isOfHLevelΠ 2 λ y → hlevel (_ , _)) λ a → sElim (λ x → hlevel (_ , _)) - λ b → ind a b - - - prodElim2 : ∀ {ℓ ℓ' ℓ'' ℓ''' ℓ''''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} - {E : (∥ A ∥₀ × ∥ B ∥₀) → (∥ C ∥₀ × ∥ D ∥₀) → Type ℓ''''} - → ((x : ∥ A ∥₀ × ∥ B ∥₀) (y : ∥ C ∥₀ × ∥ D ∥₀) → isOfHLevel 2 (E x y)) - → ((a : A) (b : B) (c : C) (d : D) → E (∣ a ∣₀ , ∣ b ∣₀) (∣ c ∣₀ , ∣ d ∣₀)) - → ((x : ∥ A ∥₀ × ∥ B ∥₀) (y : ∥ C ∥₀ × ∥ D ∥₀) → (E x y)) - prodElim2 isset f = prodElim (λ _ → isOfHLevelΠ 2 λ _ → isset _ _) - λ a b → prodElim (λ _ → isset _ _) - λ c d → f a b c d - +-} + Δ : (n : ℕ) → morph (×coHomGr n A B) (coHomGr n C) + morph.fun (Δ n) (α , β) = coHomFun n f α +ₕ -ₕ (coHomFun n g β) + morph.ismorph (Δ zero) = + prodElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _ ) + λ f' x1 g' x2 → (λ i → ∣ (λ x → (f' (f x) +ₖ g' (f x)) +ₖ -distrₖ (x1 (g x)) (x2 (g x)) i) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ g' (f x)) (-ₖ x1 (g x)) (-ₖ x2 (g x)) (~ i)) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x)) (g' (f x)) (-ₖ x1 (g x)) i +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → (f' (f x) +ₖ commₖ (g' (f x)) (-ₖ x1 (g x)) i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x)) (-ₖ x1 (g x)) (g' (f x)) (~ i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ (-ₖ x1 (g x))) (g' (f x)) (-ₖ x2 (g x)) i) ∣₀) + morph.ismorph (Δ (suc n)) = + prodElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _ ) + λ f' x1 g' x2 → (λ i → ∣ (λ x → (f' (f x) +ₖ g' (f x)) +ₖ -distrₖ (x1 (g x)) (x2 (g x)) i) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ g' (f x)) (-ₖ x1 (g x)) (-ₖ x2 (g x)) (~ i)) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x)) (g' (f x)) (-ₖ x1 (g x)) i +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → (f' (f x) +ₖ commₖ (g' (f x)) (-ₖ x1 (g x)) i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x)) (-ₖ x1 (g x)) (g' (f x)) (~ i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ (-ₖ x1 (g x))) (g' (f x)) (-ₖ x2 (g x)) i) ∣₀) + +{- ΔIsHom : (n : ℕ) → isMorph (×coHomGr n A B) (coHomGr n C) (Δ n) ΔIsHom zero = prodElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _ ) λ f' x1 g' x2 → (λ i → ∣ (λ x → (f' (f x) +ₖ g' (f x)) +ₖ -distrₖ (x1 (g x)) (x2 (g x)) i) ∣₀) ∙ @@ -120,6 +132,22 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : (λ i → ∣ (λ x → assocₖ (f' (f x)) (-ₖ x1 (g x)) (g' (f x)) (~ i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ (-ₖ x1 (g x))) (g' (f x)) (-ₖ x2 (g x)) i) ∣₀) +-} + + d-pre : (n : ℕ) → (C → coHomK n) → D → coHomK (suc n) + d-pre n γ (inl x) = 0ₖ + d-pre n γ (inr x) = 0ₖ + d-pre zero γ (push a i) = Kn→ΩKn+1 zero (γ a) i + d-pre (suc n) γ (push a i) = Kn→ΩKn+1 (suc n) (γ a) i + + + + + + + + + dHomHelperPath : (n : ℕ) (h l : C → coHomK n) (a : C) → I → I → coHomK (suc n) dHomHelperPath n h l a i j = hcomp (λ k → λ { (i = i0) → lUnitₖ 0ₖ (~ j) @@ -147,17 +175,27 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : dHomHelper zero h l (push a i) j = dHomHelperPath zero h l a i j dHomHelper (suc n) h l (push a i) j = dHomHelperPath (suc n) h l a i j + d : (n : ℕ) → morph (coHomGr n C) (coHomGr (suc n) D) + morph.fun (d n) = sRec setTruncIsSet λ a → ∣ d-pre n a ∣₀ + morph.ismorph (d zero) = sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f g i → ∣ funExt (λ x → dHomHelper zero f g x) i ∣₀ + morph.ismorph (d (suc n)) = sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f g i → ∣ funExt (λ x → dHomHelper (suc n) f g x) i ∣₀ + +{- dIsHom : (n : ℕ) → isMorph (coHomGr n C) (coHomGr (suc n) D) (d n) dIsHom zero = sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ f g i → ∣ funExt (λ x → dHomHelper zero f g x) i ∣₀ dIsHom (suc n) = sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ f g i → ∣ funExt (λ x → dHomHelper (suc n) f g x) i ∣₀ +-} + -- Long Exact Sequence Im-d⊂Ker-i : (n : ℕ) (x : coHom (suc n) D) - → isInIm (coHomGr n C) (coHomGr (suc n) D) (d n) x - → isInKer (coHomGr (suc n) D) (×coHomGr (suc n) A B) (i (suc n)) x + → isInIm (coHomGr n C) (coHomGr (suc n) D) (morph.fun (d n)) x + → isInKer (coHomGr (suc n) D) (×coHomGr (suc n) A B) (morph.fun (i (suc n))) x Im-d⊂Ker-i n = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) λ a → pElim (λ _ → isOfHLevelPath' 1 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) (sigmaElim (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) @@ -165,8 +203,8 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : (λ { δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀ }) (b (~ i)))) Ker-i⊂Im-d : (n : ℕ) (x : coHom (suc n) D) - → isInKer (coHomGr (suc n) D) (×coHomGr (suc n) A B) (i (suc n)) x - → isInIm (coHomGr n C) (coHomGr (suc n) D) (d n) x + → isInKer (coHomGr (suc n) D) (×coHomGr (suc n) A B) (morph.fun (i (suc n))) x + → isInIm (coHomGr n C) (coHomGr (suc n) D) (morph.fun (d n)) x Ker-i⊂Im-d n = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) λ a p → pRec {A = (λ x → a (inl x)) ≡ λ _ → 0ₖ} (isOfHLevelΠ 1 (λ _ → propTruncIsProp)) (λ p1 → pRec propTruncIsProp λ p2 → ∣ ∣ (λ c → ΩKn+1→Kn (sym (cong (λ F → F (f c)) p1) @@ -212,11 +250,11 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : (pushFiller (suc n) F p1 p2 a j i) Im-i⊂Ker-Δ : (n : ℕ) (x : coHom n A × coHom n B) - → isInIm (coHomGr n D) (×coHomGr n A B) (i n) x - → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) x + → isInIm (coHomGr n D) (×coHomGr n A B) (morph.fun (i n)) x + → isInKer (×coHomGr n A B) (coHomGr n C) (morph.fun (Δ n)) x Im-i⊂Ker-Δ n (Fa , Fb) = - sElim {B = λ Fa → (Fb : _) → isInIm (coHomGr n D) (×coHomGr n A B) (i n) (Fa , Fb) - → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) (Fa , Fb)} + sElim {B = λ Fa → (Fb : _) → isInIm (coHomGr n D) (×coHomGr n A B) (morph.fun (i n)) (Fa , Fb) + → isInKer (×coHomGr n A B) (coHomGr n C) (morph.fun (Δ n)) (Fa , Fb)} (λ _ → isOfHLevelΠ 2 λ _ → (isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _)) (λ Fa → sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ Fb → pRec (setTruncIsSet _ _) @@ -226,23 +264,23 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : Fb where helper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) (Fd : D → coHomK n) - → (i n ∣ Fd ∣₀ ≡ (∣ Fa ∣₀ , ∣ Fb ∣₀)) - → (Δ n) (∣ Fa ∣₀ , ∣ Fb ∣₀) ≡ 0ₕ - helper zero Fa Fb Fd p = cong (Δ zero) (sym p) + → (morph.fun (i n) ∣ Fd ∣₀ ≡ (∣ Fa ∣₀ , ∣ Fb ∣₀)) + → (morph.fun (Δ n)) (∣ Fa ∣₀ , ∣ Fb ∣₀) ≡ 0ₕ + helper zero Fa Fb Fd p = cong (morph.fun (Δ zero)) (sym p) ∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₀ +ₕ -ₕ ∣ (λ x → Fd (push x (~ i))) ∣₀) ∙ rCancelₕ ∣ (λ x → Fd (inl (f x))) ∣₀ - helper (suc n) Fa Fb Fd p = cong (Δ (suc n)) (sym p) + helper (suc n) Fa Fb Fd p = cong (morph.fun (Δ (suc n))) (sym p) ∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₀ +ₕ -ₕ ∣ (λ x → Fd (push x (~ i))) ∣₀) ∙ rCancelₕ ∣ (λ x → Fd (inl (f x))) ∣₀ Ker-Δ⊂Im-i : (n : ℕ) (a : coHom n A × coHom n B) - → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) a - → isInIm (coHomGr n D) (×coHomGr n A B) (i n) a + → isInKer (×coHomGr n A B) (coHomGr n C) (morph.fun (Δ n)) a + → isInIm (coHomGr n D) (×coHomGr n A B) (morph.fun (i n)) a Ker-Δ⊂Im-i n (Fa , Fb) = - sElim {B = λ Fa → (Fb : _) → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) (Fa , Fb) - → isInIm (coHomGr n D) (×coHomGr n A B) (i n) (Fa , Fb)} + sElim {B = λ Fa → (Fb : _) → isInKer (×coHomGr n A B) (coHomGr n C) (morph.fun (Δ n)) (Fa , Fb) + → isInIm (coHomGr n D) (×coHomGr n A B) (morph.fun (i n)) (Fa , Fb)} (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) (λ Fa → sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) λ Fb p → pRec propTruncIsProp @@ -254,7 +292,7 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : where helper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) - → (Δ n) (∣ Fa ∣₀ , ∣ Fb ∣₀) ≡ 0ₕ + → (morph.fun (Δ n)) (∣ Fa ∣₀ , ∣ Fb ∣₀) ≡ 0ₕ → ∥ (Path (_ → _) (λ c → Fa (f c)) (λ c → Fb (g c))) ∥₋₁ helper zero Fa Fb p = Iso.fun (PathIdTrunc₀Iso) (sym (rUnitₕ (coHomFun zero f ∣ Fa ∣₀)) @@ -278,14 +316,14 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : anotherHelper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) → (q : Path (C → coHomK n) (λ c → Fa (f c)) (λ c → Fb (g c))) - → i n ∣ helpFun n Fa Fb (λ x i₁ → q i₁ x) ∣₀ ≡ (∣ Fa ∣₀ , ∣ Fb ∣₀) + → morph.fun (i n) ∣ helpFun n Fa Fb (λ x i₁ → q i₁ x) ∣₀ ≡ (∣ Fa ∣₀ , ∣ Fb ∣₀) anotherHelper zero Fa Fb q = refl anotherHelper (suc n) Fa Fb q = refl Ker-d⊂Im-Δ : (n : ℕ) (a : coHom n C) - → isInKer (coHomGr n C) (coHomGr (suc n) D) (d n) a - → isInIm (×coHomGr n A B) (coHomGr n C) (Δ n) a + → isInKer (coHomGr n C) (coHomGr (suc n) D) (morph.fun (d n)) a + → isInIm (×coHomGr n A B) (coHomGr n C) (morph.fun (Δ n)) a Ker-d⊂Im-Δ n = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) λ Fc p → pRec propTruncIsProp (λ p → ∣ (∣ (λ a → ΩKn+1→Kn (cong (λ f → f (inl a)) p)) ∣₀ , @@ -294,32 +332,38 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : (Iso.fun (PathIdTrunc₀Iso) p) where - distrHelper : (n : ℕ) (p q : _) → ΩKn+1→Kn {n = n} p +ₖ (-ₖ ΩKn+1→Kn q) ≡ ΩKn+1→Kn (p ∙ sym q) - distrHelper n p q i = ΩKn+1→Kn (Iso.rightInv (Iso3-Kn-ΩKn+1 n) p i ∙ Iso.rightInv (Iso3-Kn-ΩKn+1 n) (sym (Iso.rightInv (Iso3-Kn-ΩKn+1 n) q i)) i) + distrHelper : (n : ℕ) (p q : _) + → ΩKn+1→Kn {n = n} p +ₖ (-ₖ ΩKn+1→Kn q) ≡ ΩKn+1→Kn (p ∙ sym q) + distrHelper n p q i = + ΩKn+1→Kn (Iso.rightInv (Iso3-Kn-ΩKn+1 n) p i + ∙ Iso.rightInv (Iso3-Kn-ΩKn+1 n) (sym (Iso.rightInv (Iso3-Kn-ΩKn+1 n) q i)) i) - helper2 : (n : ℕ) (Fc : C → coHomK n) (p : d-pre n Fc ≡ (λ _ → ∣ north ∣)) (c : C) + helper2 : (n : ℕ) (Fc : C → coHomK n) + (p : d-pre n Fc ≡ (λ _ → ∣ north ∣)) (c : C) → ΩKn+1→Kn (λ i₁ → p i₁ (inl (f c))) +ₖ (-ₖ ΩKn+1→Kn (λ i₁ → p i₁ (inr (g c)))) ≡ Fc c - helper2 zero Fc p c = distrHelper zero (cong (λ F → F (inl (f c))) p) (cong (λ F → F (inr (g c))) p) - ∙ cong ΩKn+1→Kn (sym ((PathP→compPathR (cong (λ f → cong f (push c)) p)) - ∙ (λ i → (λ i₁ → p i₁ (inl (f c))) - ∙ (lUnit (sym (λ i₁ → p i₁ (inr (g c)))) (~ i))))) - ∙ Iso.leftInv (Iso3-Kn-ΩKn+1 zero) (Fc c) - helper2 (suc n) Fc p c = distrHelper (suc n) (cong (λ F → F (inl (f c))) p) (cong (λ F → F (inr (g c))) p) - ∙ cong ΩKn+1→Kn (sym ((PathP→compPathR (cong (λ f → cong f (push c)) p)) - ∙ (λ i → (λ i₁ → p i₁ (inl (f c))) - ∙ (lUnit (sym (λ i₁ → p i₁ (inr (g c)))) (~ i))))) - ∙ Iso.leftInv (Iso3-Kn-ΩKn+1 (suc n)) (Fc c) + helper2 zero Fc p c = + distrHelper zero (cong (λ F → F (inl (f c))) p) (cong (λ F → F (inr (g c))) p) + ∙ cong ΩKn+1→Kn (sym ((PathP→compPathR (cong (λ f → cong f (push c)) p)) + ∙ (λ i → (λ i₁ → p i₁ (inl (f c))) + ∙ (lUnit (sym (λ i₁ → p i₁ (inr (g c)))) (~ i))))) + ∙ Iso.leftInv (Iso3-Kn-ΩKn+1 zero) (Fc c) + helper2 (suc n) Fc p c = + distrHelper (suc n) (cong (λ F → F (inl (f c))) p) (cong (λ F → F (inr (g c))) p) + ∙ cong ΩKn+1→Kn (sym ((PathP→compPathR (cong (λ f → cong f (push c)) p)) + ∙ (λ i → (λ i₁ → p i₁ (inl (f c))) + ∙ (lUnit (sym (λ i₁ → p i₁ (inr (g c)))) (~ i))))) + ∙ Iso.leftInv (Iso3-Kn-ΩKn+1 (suc n)) (Fc c) Im-Δ⊂Ker-d : (n : ℕ) (a : coHom n C) - → isInIm (×coHomGr n A B) (coHomGr n C) (Δ n) a - → isInKer (coHomGr n C) (coHomGr (suc n) D) (d n) a + → isInIm (×coHomGr n A B) (coHomGr n C) (morph.fun (Δ n)) a + → isInKer (coHomGr n C) (coHomGr (suc n) D) (morph.fun (d n)) a Im-Δ⊂Ker-d n = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ Fc → pRec (isOfHLevelPath' 1 setTruncIsSet _ _) (sigmaProdElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ Fa Fb p → pRec (isOfHLevelPath' 1 setTruncIsSet _ _) - (λ q → ((λ i → d n ∣ (q (~ i)) ∣₀) ∙ dΔ-Id n Fa Fb)) + (λ q → ((λ i → morph.fun (d n) ∣ (q (~ i)) ∣₀) ∙ dΔ-Id n Fa Fb)) (Iso.fun (PathIdTrunc₀Iso) p)) where @@ -339,32 +383,32 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : d-preRightId (suc n) Fb (push a i) j = Kn→ΩKn+1 (suc n) (Fb (g a)) (~ j ∧ i) dΔ-Id : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) - → d n (Δ n (∣ Fa ∣₀ , ∣ Fb ∣₀)) ≡ 0ₕ + → morph.fun (d n) (morph.fun (Δ n) (∣ Fa ∣₀ , ∣ Fb ∣₀)) ≡ 0ₕ dΔ-Id zero Fa Fb = - dIsHom zero ∣ (λ x → Fa (f x)) ∣₀ (-ₕ ∣ (λ x → Fb (g x)) ∣₀) - ∙ (λ i → d zero ∣ (λ x → Fa (f x)) ∣₀ +ₕ morphMinus (coHomGr zero C) (coHomGr 1 D) (d zero) (dIsHom zero) ∣ (λ x → Fb (g x)) ∣₀ i) - ∙ (λ i → d zero ∣ (λ x → Fa (f x)) ∣₀ +ₕ -ₕ (d zero ∣ (λ x → Fb (g x)) ∣₀)) + morph.ismorph (d zero) ∣ (λ x → Fa (f x)) ∣₀ (-ₕ ∣ (λ x → Fb (g x)) ∣₀) + ∙ (λ i → morph.fun (d zero) ∣ (λ x → Fa (f x)) ∣₀ +ₕ morphMinus (coHomGr zero C) (coHomGr 1 D) (d zero) ∣ (λ x → Fb (g x)) ∣₀ i) + ∙ (λ i → morph.fun (d zero) ∣ (λ x → Fa (f x)) ∣₀ +ₕ -ₕ (morph.fun (d zero) ∣ (λ x → Fb (g x)) ∣₀)) ∙ (λ i → ∣ funExt (d-preLeftId zero Fa) i ∣₀ +ₕ -ₕ ∣ funExt (d-preRightId zero Fb) i ∣₀) ∙ rCancelₕ 0ₕ dΔ-Id (suc n) Fa Fb = - dIsHom (suc n) ∣ (λ x → Fa (f x)) ∣₀ (-ₕ ∣ (λ x → Fb (g x)) ∣₀) - ∙ (λ i → d (suc n) ∣ (λ x → Fa (f x)) ∣₀ +ₕ morphMinus (coHomGr (suc n) C) (coHomGr (2 + n) D) (d (suc n)) (dIsHom (suc n)) ∣ (λ x → Fb (g x)) ∣₀ i) - ∙ (λ i → d (suc n) ∣ (λ x → Fa (f x)) ∣₀ +ₕ -ₕ (d (suc n) ∣ (λ x → Fb (g x)) ∣₀)) + morph.ismorph (d (suc n)) ∣ (λ x → Fa (f x)) ∣₀ (-ₕ ∣ (λ x → Fb (g x)) ∣₀) + ∙ (λ i → morph.fun (d (suc n)) ∣ (λ x → Fa (f x)) ∣₀ +ₕ morphMinus (coHomGr (suc n) C) (coHomGr (2 + n) D) (d (suc n)) ∣ (λ x → Fb (g x)) ∣₀ i) + ∙ (λ i → morph.fun (d (suc n)) ∣ (λ x → Fa (f x)) ∣₀ +ₕ -ₕ (morph.fun (d (suc n)) ∣ (λ x → Fb (g x)) ∣₀)) ∙ (λ i → ∣ funExt (d-preLeftId (suc n) Fa) i ∣₀ +ₕ -ₕ ∣ funExt (d-preRightId (suc n) Fb) i ∣₀) ∙ rCancelₕ 0ₕ -d-morph : ∀ {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) (n : ℕ) - → morph (coHomGr n C) (coHomGr (suc n) (Pushout f g)) -morph.fun (d-morph A B C f g n) = MV.d A B C f g n -morph.ismorph (d-morph A B C f g n) = MV.dIsHom A B C f g n +-- d-morph : ∀ {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) (n : ℕ) +-- → morph (coHomGr n C) (coHomGr (suc n) (Pushout f g)) +-- morph.fun (d-morph A B C f g n) = MV.d A B C f g n +-- morph.ismorph (d-morph A B C f g n) = MV.dIsHom A B C f g n -i-morph : ∀ {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) (n : ℕ) - → morph (coHomGr n (Pushout f g)) (×coHomGr n A B) -morph.fun (i-morph A B C f g n) = MV.i A B C f g n -morph.ismorph (i-morph A B C f g n) = MV.iIsHom A B C f g n +-- i-morph : ∀ {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) (n : ℕ) +-- → morph (coHomGr n (Pushout f g)) (×coHomGr n A B) +-- morph.fun (i-morph A B C f g n) = MV.i A B C f g n +-- morph.ismorph (i-morph A B C f g n) = MV.iIsHom A B C f g n -Δ-morph : ∀ {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) (n : ℕ) - → morph (×coHomGr n A B) (coHomGr n C) -morph.fun (Δ-morph A B C f g n) = MV.Δ A B C f g n -morph.ismorph (Δ-morph A B C f g n) = MV.ΔIsHom A B C f g n +-- Δ-morph : ∀ {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) (n : ℕ) +-- → morph (×coHomGr n A B) (coHomGr n C) +-- morph.fun (Δ-morph A B C f g n) = MV.Δ A B C f g n +-- morph.ismorph (Δ-morph A B C f g n) = MV.ΔIsHom A B C f g n diff --git a/Cubical/ZCohomology/altToFreudenthal.agda b/Cubical/ZCohomology/altToFreudenthal.agda deleted file mode 100644 index 9a5d4e492..000000000 --- a/Cubical/ZCohomology/altToFreudenthal.agda +++ /dev/null @@ -1,201 +0,0 @@ -{-# OPTIONS --cubical --safe #-} -module Cubical.ZCohomology.altToFreudenthal where - -open import Cubical.ZCohomology.Base -open import Cubical.HITs.S1 -open import Cubical.HITs.Sn -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Equiv.HalfAdjoint -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Transport -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Univalence -open import Cubical.Data.NatMinusTwo.Base -open import Cubical.Data.Empty -open import Cubical.Data.Sigma hiding (_×_) -open import Cubical.HITs.Susp -open import Cubical.HITs.Wedge -open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) -open import Cubical.HITs.Nullification -open import Cubical.Data.Int hiding (_+_) -open import Cubical.Data.Nat -open import Cubical.Data.Prod -open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) -open import Cubical.Homotopy.Loopspace -open import Cubical.Homotopy.Connected -open import Cubical.Homotopy.Freudenthal -open import Cubical.HITs.SmashProduct.Base -open import Cubical.Data.Group.Base hiding (_≃_ ; Iso) - -private - variable - ℓ ℓ' : Level - A : Type ℓ - B : Type ℓ' - A' : Pointed ℓ - B' : Pointed ℓ' - - -test : ∀ n → hLevelTrunc (3 + n) (S₊ (1 + n)) → hLevelTrunc (3 + n) (Path (S₊ (2 + n)) north south) -test n = trRec (isOfHLevelTrunc (3 + n)) λ x → ∣ merid x ∣ - -isEquivMerid : ∀ n → isHLevelConnectedFun (3 + n) (test n) -isEquivMerid n = trElim {!!} λ a → isHLevelConnectedPoint2 _ (∣ {!!} ∣ , {!a!}) {!!} - where - ou : (b : Path (S₊ (suc (suc n))) north south) → Iso (fiber (test n) ∣ b ∣) (fiber {A = Unit} {B = hLevelTrunc (suc (suc (suc n))) (Path (S₊ (suc (suc n))) north south)} (λ {tt → ∣ b ∣}) ∣ b ∣) - Iso.fun (ou b) (p , q) = tt , refl - Iso.inv (ou b) (p , q) = ∣ {!!} ∣ , {!!} - Iso.rightInv (ou b) = {!!} - Iso.leftInv (ou b) = {!!} - where - king : (b : Path (S₊ (suc (suc n))) north south) → Iso (fiber (test n) ∣ b ∣) ( hLevelTrunc (suc (suc (suc n))) (S₊ (2 + n))) - Iso.fun (king b) (p , q) = trRec {A = (S₊ (suc n))} (isOfHLevelTrunc (3 + n)) (λ p → ∣ fun p ∣) p - where - fun : S₊ (suc n) → S₊ (suc (suc n)) - fun north = north - fun south = south - fun (merid a i) = merid (merid a i) i - Iso.inv (king b) = {!!} - Iso.rightInv (king b) = {!merid ? ?!} - Iso.leftInv (king b) = {!!} - --- id : (n : ℕ) → typ ((Ω^ (1 + n)) A') --- id zero = refl --- id (suc n) = λ _ → refl - --- indEquiv : (n : ℕ) → ((x : S₊ (suc n)) → typ (Ω ((Ω^ n) (S₊ (suc n) , x)))) → typ (Ω (Ω ((Ω^ n) (Type₀ , S₊ (suc n))))) --- indEquiv zero f = {!!} --- indEquiv (suc n) f = {!!} - - --- funcTypeEq : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → Iso (typ ((Ω^ (2 + n)) (Type₀ , S₊ (1 + n)))) ((x : S₊ (1 + n)) → typ ((Ω^ (1 + n)) (S₊ (1 + n) , x))) --- Iso.fun (funcTypeEq zero) f x = {!transport !} --- Iso.fun (funcTypeEq (suc zero)) f x i = {!cong !} --- Iso.fun (funcTypeEq (suc (suc n))) f x i = {!!} --- Iso.inv (funcTypeEq n) f i = {!f north !} --- Iso.rightInv (funcTypeEq n) = {!!} --- Iso.leftInv (funcTypeEq n) = {!!} - --- idToIso : ∀ {ℓ} {A : Type ℓ} (p : A ≡ A) → Iso A A --- Iso.fun (idToIso {A = A} p) = transport p --- Iso.inv (idToIso {A = A} p) = transport (sym p) --- Iso.rightInv (idToIso {A = A} p) = transportTransport⁻ p --- Iso.leftInv (idToIso {A = A} p) = transport⁻Transport p - --- test : ∀ {ℓ} {A : Type ℓ} (p : A ≡ A) → (typ (Ω ((A ≡ A) , p))) ≃ (typ (Ω ((A ≃ A) , pathToEquiv p))) --- test {ℓ = ℓ} {A = A} p = transport (λ i → helper2 (~ i) ≃ typ (Ω ((A ≃ A) , (pathToEquiv p)))) (compEquiv (invEquiv helper3) (pathToEquiv λ i → typ (Ω (_ , transportRefl (pathToEquiv p) i)))) --- where --- helper2 : typ (Ω ((A ≡ A) , p)) ≡ typ (Ω (Lift (A ≃ A) , transport (univalencePath {A = A} {B = A}) p )) --- helper2 i = typ (Ω ((univalencePath {A = A} {B = A} i) , transp (λ j → univalencePath {A = A} {B = A} (j ∧ i)) (~ i) p)) --- helper3 : typ (Ω ((A ≃ A) , _)) --- ≃ typ (Ω (Lift (A ≃ A) , transport (univalencePath {A = A} {B = A}) p)) - --- helper3 = congEquiv LiftEquiv - --- test2 : ∀ {ℓ} {A : Type ℓ} (p : A ≡ A) → typ (Ω ((A ≃ A) , pathToEquiv p)) ≃ (typ (Ω ((A → A) , transport p))) --- test2 {A = A} p = EquivCat --- where --- EquivCat : (pathToEquiv p ≡ pathToEquiv p) ≃ (pathToEquiv p .fst ≡ pathToEquiv p .fst) --- EquivCat = compEquiv (invEquiv Σ≡) helper -- {!sym ua Σ≡!} ∙ {!!} ∙ {!!} --- where --- helper : Σ (fst (pathToEquiv p) ≡ fst (pathToEquiv p)) (λ p₁ → PathP (λ i → isEquiv (p₁ i)) (snd (pathToEquiv p)) (snd (pathToEquiv p))) ≃ (pathToEquiv p .fst ≡ pathToEquiv p .fst) --- helper = helper2 λ x → toPathP (isPropIsEquiv _ _ _) , λ y → {!!} --- where --- helper2 : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} → ((x : A) → isContr (B x)) → Σ A B ≃ A --- helper2 contr = isoToEquiv (iso (λ { (a , p) → a}) (λ a → a , (contr a .fst)) (λ b → refl) λ { (a , p) i → a , (contr a .snd p) i }) - --- withJ : (p q : A ≃ B) → Iso (Path (A ≃ B) p q) (Path (A → B) (p .fst) (q .fst)) --- Iso.fun (withJ {A = A} {B = B} p q) = cong fst --- Iso.inv (withJ {A = A} p q) path = J (λ f r → (x : isEquiv f) → p ≡ (f , x)) (λ x i → (path i) , {!!}) refl (q .snd) --- Iso.rightInv (withJ p q) a = {!a!} --- Iso.leftInv (withJ p q) = {!!} - --- together : ∀ {ℓ} (n : ℕ) {A : Type ℓ} (p : A ≡ A) → (typ ((Ω^ (1 + n)) ((A ≡ A) , p))) ≃ (typ ((Ω^ (1 + n)) ((A → A) , transport p))) --- toghelper : ∀ {ℓ} (n : ℕ) {A : Type ℓ} (p : A ≡ A) → together n p .fst (snd (Ω ((Ω^ n) ((A ≡ A) , p)))) ≡ snd (Ω ((Ω^ n) ((A → A) , transport p))) --- toghelper zero {A = A} p = {!!} - --- toghelper (suc n) {A = A} p = {!refl!} - --- together zero {A = A} p = helper2 -- (compEquiv (test p) (test2 p)) --- where --- helper : ∀ B p → typ (Ω ((A ≃ B) , univalence .fst p)) ≃ typ (Ω ((A → B) , transport p)) --- helper B p = isoToEquiv (iso (cong fst) {!!} {!!} {!!}) -- isoToEquiv (iso (cong fst) {!univalence .fst p!} {!!} {!!}) --- helper2 : typ (Ω ((A ≡ A) , p)) ≃ typ (Ω ((A → A) , transport p)) --- helper2 = compEquiv (congEquiv univalence) {!!} --- together (suc n) {A = A} p = isoToEquiv helper --- where --- helper : Iso (typ (Ω (Ω ((Ω^ n) ((A ≡ A) , p))))) (typ (Ω (Ω ((Ω^ n) ((A → A) , transport p))))) --- Iso.fun helper a = sym (toghelper n p) ∙ cong (together n p .fst) a ∙ toghelper n p --- Iso.inv helper a = {!? ∙ ? ∙ ?!} --- Iso.rightInv helper = {!!} --- Iso.leftInv helper = {!!} - --- test6 : Iso (A ≡ A) (A ≃ A) --- Iso.fun test6 p = transportEquiv p --- Iso.inv test6 p = ua p --- Iso.rightInv test6 b i = funExt (uaβ b) i , hcomp {!!} (transp (λ j → isEquiv (funExt (λ x → transportRefl (b .fst x)) (i ∧ j))) (~ i) (transp (λ i₁ → isEquiv (transp (λ j → ua b (i₁ ∧ j)) (~ i₁))) i0 (idIsEquiv A))) --- Iso.leftInv test6 b = {! .fst!} - - --- -- S∙ : (n : ℕ) → Pointed₀ --- -- S∙ n = (S₊ n) , north - --- -- ΩS : (n : ℕ) → Type₀ --- -- ΩS n = typ (Ω' (S₊ n , north)) - --- -- promote : (n : ℕ) → S₊ n → ΩS (1 + n) --- -- promote n north = merid north ∙ sym (merid north) --- -- promote n south = merid south ∙ sym (merid north) --- -- promote n (merid a i) = merid (merid a i) ∙ sym (merid north) - --- -- decode' : (n : ℕ) → hLevelTrunc (2 + n) (S₊ n) → hLevelTrunc (2 + n) (ΩS (1 + n)) --- -- decode' n = trRec (isOfHLevelTrunc (2 + n)) --- -- λ a → ∣ promote n a ∣ - - --- -- Codes : (n : ℕ) → S₊ (suc n) → Type₀ --- -- Codes n north = hLevelTrunc (3 + n) (S₊ (1 + n)) --- -- Codes n south = hLevelTrunc (3 + n) (S₊ (1 + n)) --- -- Codes n (merid a i) = {!!} - --- -- SEquiv1 : (m : ℕ) → Iso (ℕ → S₊ (suc m)) (S₊ (suc m) → S₊ (suc m)) --- -- Iso.fun (SEquiv1 m) f north = north --- -- Iso.fun (SEquiv1 m) f south = south --- -- Iso.fun (SEquiv1 m) f (merid a i) = {!!} --- -- Iso.inv (SEquiv1 m) f k = {!!} --- -- Iso.rightInv (SEquiv1 m) = {!!} --- -- Iso.leftInv (SEquiv1 m) = {!!} - - - - - - - --- -- -- SEquiv1 : (n : ℕ) → S₊ (suc n) → S₊ (suc n) → Iso (S₊ (suc n)) (S₊ (suc n)) --- -- -- Iso.fun (SEquiv1 n north north) north = north --- -- -- Iso.fun (SEquiv1 n north north) south = north --- -- -- Iso.fun (SEquiv1 n north north) (merid a i) = (merid a ∙ sym (merid north)) i --- -- -- Iso.fun (SEquiv1 n north south) north = north --- -- -- Iso.fun (SEquiv1 n north south) south = south --- -- -- Iso.fun (SEquiv1 n north south) (merid a i) = merid a i --- -- -- Iso.fun (SEquiv1 n north (merid a i)) = {!!} --- -- -- Iso.fun (SEquiv1 n south t) = {!!} --- -- -- Iso.fun (SEquiv1 n (merid a i) t) = {!!} --- -- -- Iso.inv (SEquiv1 n s t) = {!!} --- -- -- Iso.rightInv (SEquiv1 n s t) = {!!} --- -- -- Iso.leftInv (SEquiv1 n s t) = {!!} - --- -- -- SEquiv : (n : ℕ) → S₊ (suc n) ≃ S₊ (suc n) --- -- -- SEquiv n = isoToEquiv {!!} --- -- -- where --- -- -- isot : Iso (S₊ (suc n)) (S₊ (suc n)) --- -- -- Iso.fun isot north = {!!} --- -- -- Iso.fun isot south = {!!} --- -- -- Iso.fun isot (merid a i) = {!!} --- -- -- Iso.inv isot = {!!} --- -- -- Iso.rightInv isot = {!!} --- -- -- Iso.leftInv isot = {!!} From 7a71bd94f4a5765aef8be915a3f9f0b9e96a3be8 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Fri, 12 Jun 2020 16:27:15 +0200 Subject: [PATCH 44/89] coHom-restructuring --- Cubical/Data/Unit/Properties.agda | 6 + Cubical/HITs/S1/Base.agda | 47 +- Cubical/HITs/Sn/Properties.agda | 28 + Cubical/ZCohomology/Groups/BACKUP.agda | 1778 +++++++++++++++++ Cubical/ZCohomology/Groups/Sn.agda | 1628 +++------------ Cubical/ZCohomology/Groups/Torus.agda | 318 +++ Cubical/ZCohomology/Groups/Unit.agda | 78 + .../ZCohomology/Groups/WedgeOfSpheres.agda | 365 ++++ .../ZCohomology/MayerVietorisUnreduced.agda | 49 +- Cubical/ZCohomology/Properties.agda | 1 + 10 files changed, 2890 insertions(+), 1408 deletions(-) create mode 100644 Cubical/ZCohomology/Groups/BACKUP.agda create mode 100644 Cubical/ZCohomology/Groups/Torus.agda create mode 100644 Cubical/ZCohomology/Groups/Unit.agda create mode 100644 Cubical/ZCohomology/Groups/WedgeOfSpheres.agda diff --git a/Cubical/Data/Unit/Properties.agda b/Cubical/Data/Unit/Properties.agda index f2aac251a..0ce1b09f8 100644 --- a/Cubical/Data/Unit/Properties.agda +++ b/Cubical/Data/Unit/Properties.agda @@ -26,6 +26,12 @@ isOfHLevelUnit n = isContr→isOfHLevel n isContrUnit UnitToTypeId : ∀ {ℓ} (A : Type ℓ) → (Unit → A) ≡ A UnitToTypeId A = isoToPath (iso (λ f → f tt) (λ a _ → a) (λ _ → refl) λ _ → refl) +ContrToTypeIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → isContr A → Iso (A → B) B +Iso.fun (ContrToTypeIso iscontr) f = f (fst iscontr) +Iso.inv (ContrToTypeIso iscontr) b _ = b +Iso.rightInv (ContrToTypeIso iscontr) _ = refl +Iso.leftInv (ContrToTypeIso iscontr) f = funExt λ x → cong f (snd iscontr x) + diagonal-unit : Unit ≡ Unit × Unit diagonal-unit = isoToPath (iso (λ x → tt , tt) (λ x → tt) (λ {(tt , tt) i → tt , tt}) λ {tt i → tt}) diff --git a/Cubical/HITs/S1/Base.agda b/Cubical/HITs/S1/Base.agda index 04f6e23f2..cca7a6ca7 100644 --- a/Cubical/HITs/S1/Base.agda +++ b/Cubical/HITs/S1/Base.agda @@ -32,14 +32,7 @@ module _ where comp (λ _ → S¹) u (outS u0) ≡ hcomp u (outS u0) compS1 φ u u0 = refl - - -toPropElim : ∀ {ℓ} {B : S¹ → Type ℓ} → ((s : S¹) → isProp (B s)) → B base → (s : S¹) → B s -toPropElim {B = B} isprop b base = b -toPropElim {B = B} isprop b (loop i) = hcomp (λ k → λ {(i = i0) → b - ; (i = i1) → isprop base (subst B (loop) b) b k}) - (transp (λ j → B (loop (i ∧ j))) (~ i) b) - +-- ΩS¹ ≡ Int helix : S¹ → Type₀ helix base = Int helix (loop i) = sucPathInt i @@ -93,10 +86,7 @@ isSetΩS¹ p q r s j i = (decode base (isSetInt (winding p) (winding q) (cong winding r) (cong winding s) j i)) -isSetΩx : (x : S¹) → isSet (x ≡ x) -isSetΩx = toPropElim - (λ s → isPropIsSet) - isSetΩS¹ + -- This proof does not rely on rewriting hcomp with empty systems in -- Int as ghcomp has been implemented! @@ -269,9 +259,19 @@ basedΩS¹≡Int : (x : S¹) → basedΩS¹ x ≡ Int basedΩS¹≡Int x = (basedΩS¹≡ΩS¹ x) ∙ ΩS¹≡Int +---- A shorter proof of the same thing ----- +toPropElim : ∀ {ℓ} {B : S¹ → Type ℓ} + → ((s : S¹) → isProp (B s)) + → B base + → (s : S¹) → B s +toPropElim isprop b base = b +toPropElim isprop b (loop i) = + isOfHLevel→isOfHLevelDep 1 isprop b b loop i + +isSetΩx : (x : S¹) → isSet (x ≡ x) +isSetΩx = toPropElim (λ _ → isPropIsSet) isSetΩS¹ ----- A shorter proof of the same thing ----- basechange2 : (x : S¹) → ΩS¹ → (x ≡ x) basechange2 base p = p basechange2 (loop i) p = @@ -302,14 +302,29 @@ basechange2-retr = toPropElim (λ s → isOfHLevelΠ 1 λ x → isSetΩx _ _ _) λ _ → refl +basedΩS¹≡ΩS¹' : (x : S¹) → basedΩS¹ x ≡ ΩS¹ +basedΩS¹≡ΩS¹' x = + isoToPath + (iso (basechange2⁻ x) + (basechange2 x) + (basechange2-retr x) + (basechange2-sect x)) + +-- in fact, the two maps are equal + +basechange≡basechange2 : basechange ≡ basechange2⁻ +basechange≡basechange2 = + funExt (toPropElim (λ _ → isOfHLevelPath' 1 (isOfHLevelΠ 2 (λ _ → isSetΩS¹)) _ _) + refl) -- baschange2⁻ is a morphism basechange2⁻-morph : (x : S¹) (p q : x ≡ x) → basechange2⁻ x (p ∙ q) ≡ basechange2⁻ x p ∙ basechange2⁻ x q -basechange2⁻-morph = toPropElim {B = λ x → (p q : x ≡ x) → basechange2⁻ x (p ∙ q) ≡ basechange2⁻ x p ∙ basechange2⁻ x q} - (λ x → isOfHLevelΠ 1 λ p → isOfHLevelΠ 1 λ q → isSetΩS¹ _ _ ) - λ _ _ → refl +basechange2⁻-morph = + toPropElim {B = λ x → (p q : x ≡ x) → basechange2⁻ x (p ∙ q) ≡ basechange2⁻ x p ∙ basechange2⁻ x q} + (λ x → isOfHLevelΠ 1 λ p → isOfHLevelΠ 1 λ q → isSetΩS¹ _ _ ) + λ _ _ → refl -- Some tests module _ where diff --git a/Cubical/HITs/Sn/Properties.agda b/Cubical/HITs/Sn/Properties.agda index 7f10e899b..4cbd033ca 100644 --- a/Cubical/HITs/Sn/Properties.agda +++ b/Cubical/HITs/Sn/Properties.agda @@ -8,6 +8,7 @@ open import Cubical.Foundations.Function open import Cubical.Foundations.Equiv open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Transport open import Cubical.Foundations.Prelude open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Univalence @@ -20,6 +21,7 @@ open import Cubical.Data.Unit open import Cubical.HITs.Join open import Cubical.HITs.Pushout open import Cubical.HITs.SmashProduct +open import Cubical.HITs.PropositionalTruncation private variable @@ -38,6 +40,15 @@ isOfHLevelS1 = transport (λ i → isOfHLevel 3 (S¹≡S1 (~ i))) λ x y → J (λ y p → (q : x ≡ y) → isProp (p ≡ q)) (transport (λ i → isSet (basedΩS¹≡Int x (~ i))) isSetInt refl) +isGroupoidS1 : isGroupoid (S₊ 1) +isGroupoidS1 = transport (λ i → isGroupoid (S¹≡S1 (~ i))) isGroupoidS¹ + +isConnectedS1 : (x : S₊ 1) → ∥ north ≡ x ∥ +isConnectedS1 x = rec propTruncIsProp + (λ p → ∣ cong (transport (sym (S¹≡S1))) p ∙ transport⁻Transport (S¹≡S1) x ∣) + (isConnectedS¹ (transport S¹≡S1 x)) + + SuspBool→S1 : Susp Bool → S₊ 1 SuspBool→S1 north = north SuspBool→S1 south = south @@ -72,6 +83,23 @@ SuspBool≃S1 = isoToEquiv iso1 Iso.leftInv iso1 = SuspBool→S1-retr +S1→S¹ : S₊ 1 → S¹ +S1→S¹ x = SuspBool→S¹ (S1→SuspBool x) + +S¹→S1 : S¹ → S₊ 1 +S¹→S1 x = SuspBool→S1 (S¹→SuspBool x) + +S1→S¹-sect : section S1→S¹ S¹→S1 +S1→S¹-sect x = + cong SuspBool→S¹ (SuspBool→S1-retr (S¹→SuspBool x)) + ∙ S¹→SuspBool→S¹ x + +S1→S¹-retr : retract S1→S¹ S¹→S1 +S1→S¹-retr x = + cong SuspBool→S1 (SuspBool→S¹→SuspBool (S1→SuspBool x)) + ∙ SuspBool→S1-sect x + + -- map between S¹ ∧ A and Susp A. private f' : {a : A} → A → S₊ 1 → Susp A diff --git a/Cubical/ZCohomology/Groups/BACKUP.agda b/Cubical/ZCohomology/Groups/BACKUP.agda new file mode 100644 index 000000000..d20e121a9 --- /dev/null +++ b/Cubical/ZCohomology/Groups/BACKUP.agda @@ -0,0 +1,1778 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.Groups.Sn where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.S1.S1 +open import Cubical.HITs.S1 +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.MayerVietorisUnreduced +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence +open import Cubical.Data.NatMinusTwo.Base +open import Cubical.Data.Empty +open import Cubical.Data.Sigma hiding (_×_) +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₋₁ ; ∣_∣ to ∣_∣₋₁) +open import Cubical.HITs.Nullification +open import Cubical.Data.Int renaming (_+_ to _+ℤ_; +-comm to +ℤ-comm ; +-assoc to +ℤ-assoc) +open import Cubical.Data.Nat +open import Cubical.Data.Prod +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; recElim to trRec ; rec to trRec2 ; elim3 to trElim3) +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Freudenthal +open import Cubical.HITs.SmashProduct.Base +open import Cubical.Data.Unit +open import Cubical.Data.Group.Base renaming (Iso to grIso ; compIso to compGrIso ; invIso to invGrIso ; idIso to idGrIso) +open import Cubical.Data.Group.GroupLibrary +open import Cubical.ZCohomology.coHomZero-map +open import Cubical.HITs.Wedge + +open import Cubical.Foundations.Equiv.HalfAdjoint + + +open import Cubical.ZCohomology.KcompPrelims + + +open import Cubical.HITs.Pushout +open import Cubical.Data.Sum.Base +open import Cubical.Data.HomotopyGroup +open import Cubical.Data.Bool hiding (_⊕_) + +{- +cong₂Sym : ∀ {ℓ} {A : Type ℓ} {x : A} + (p : x ≡ x) → + (refl ≡ p) → + (P : p ≡ p) → + cong₂ (λ x y → x ∙ y) P (sym P) ≡ refl +cong₂Sym {x = x} p = J (λ p _ → (P : p ≡ p) → cong₂ (λ x y → x ∙ y) P (sym P) ≡ refl) + λ P → cong₂Funct (λ x y → x ∙ y) P (sym P) + ∙ PathP→compPathR (λ j → cong (λ x → rUnit x (~ j)) P ∙ cong (λ x → lUnit x (~ j)) (sym P)) + ∙ (λ j → (sym (rUnit refl)) ∙ rCancel P j ∙ lUnit refl) + ∙ (λ j → sym (rUnit refl) ∙ lUnit (lUnit refl) (~ j) ) + ∙ rCancel (sym (rUnit refl)) + +abstract + cong₂Sym1 : ∀ {ℓ} {A : Type ℓ} {x : A} + (p : x ≡ x) → + (refl ≡ p) → + (P : p ≡ p) → + cong (λ x → x ∙ p) P ≡ cong (λ x → p ∙ x) P + cong₂Sym1 {x = x} p id P = rUnit _ + ∙ (λ i → cong (λ x → x ∙ p) P ∙ lCancel (λ i → p ∙ P i) (~ i)) -- cong (λ x → cong (λ x → x ∙ p) P ∙ x) {!!} + ∙ assoc _ _ _ + ∙ (λ j → cong₂Funct (λ x y → x ∙ y) P (sym P) (~ j) ∙ (λ i → p ∙ P i)) + ∙ (λ j → cong₂Sym p id P j ∙ (λ i → p ∙ P i)) + ∙ sym (lUnit _) +-} + + +test13 : Iso (S₊ 1 → hLevelTrunc 4 (S₊ 2)) (hLevelTrunc 4 (S₊ 2) × hLevelTrunc 3 (S₊ 1)) +Iso.fun test13 f = f north , ΩKn+1→Kn (sym (rCancelₖ (f north)) + ∙ (λ i → f (merid south i) +ₖ (-ₖ f (merid north i))) + ∙ rCancelₖ (f south)) +Iso.inv test13 (a , b) north = a +ₖ 0ₖ +Iso.inv test13 (a , b) south = a +ₖ 0ₖ +Iso.inv test13 (a , b) (merid south i) = a +ₖ (Kn→ΩKn+1 1 b i) +Iso.inv test13 (a , b) (merid north i) = a +ₖ 0ₖ +Iso.rightInv test13 (a , b) = + ×≡ (rUnitₖ a) + ((cong ΩKn+1→Kn (congHelper++ (Kn→ΩKn+1 1 b) (λ x → (a +ₖ x) +ₖ (-ₖ (a +ₖ 0ₖ))) (funExt (λ x → sym (cancelHelper a x))) (rCancelₖ (a +ₖ 0ₖ)))) + ∙ Iso.leftInv (Iso3-Kn-ΩKn+1 1) b) + module _ where + cancelHelper : (a b : hLevelTrunc 4 (S₊ 2)) → (a +ₖ b) +ₖ (-ₖ (a +ₖ 0ₖ)) ≡ b + cancelHelper a b = + (a +ₖ b) +ₖ (-ₖ (a +ₖ 0ₖ)) ≡⟨ (λ i → (a +ₖ b) +ₖ (-ₖ (rUnitₖ a i))) ⟩ + (a +ₖ b) +ₖ (-ₖ a) ≡⟨ cong (λ x → x +ₖ (-ₖ a)) (commₖ a b) ⟩ + (b +ₖ a) +ₖ (-ₖ a) ≡⟨ assocₖ b a (-ₖ a) ⟩ + b +ₖ a +ₖ (-ₖ a) ≡⟨ cong (λ x → b +ₖ x) (rCancelₖ a) ⟩ + b +ₖ 0ₖ ≡⟨ rUnitₖ b ⟩ + b ∎ + + abstract + commHelper : (p q : Path (hLevelTrunc 4 (S₊ 2)) 0ₖ 0ₖ) → p ∙ q ≡ q ∙ p + commHelper p q = + cong₂ _∙_ (sym (Iso.rightInv (Iso3-Kn-ΩKn+1 1) p)) + (sym (Iso.rightInv (Iso3-Kn-ΩKn+1 1) q)) + ∙ sym (Iso.rightInv (Iso3-Kn-ΩKn+1 1) (Kn→ΩKn+1 1 (ΩKn+1→Kn p) ∙ Kn→ΩKn+1 1 (ΩKn+1→Kn q))) + ∙ cong (Kn→ΩKn+1 1) (commₖ (ΩKn+1→Kn p) (ΩKn+1→Kn q)) + ∙ Iso.rightInv (Iso3-Kn-ΩKn+1 1) (Kn→ΩKn+1 1 (ΩKn+1→Kn q) ∙ Kn→ΩKn+1 1 (ΩKn+1→Kn p)) + ∙ sym (cong₂ _∙_ (sym (Iso.rightInv (Iso3-Kn-ΩKn+1 1) q)) + (sym (Iso.rightInv (Iso3-Kn-ΩKn+1 1) p))) + + moveabout : {x : hLevelTrunc 4 (S₊ 2)} (p q : x ≡ 0ₖ) (mid : 0ₖ ≡ 0ₖ) + → sym q ∙ (p ∙ mid ∙ sym p) ∙ q ≡ mid + moveabout p q mid = assoc (sym q) _ _ + ∙ cong (_∙ q) (assoc (sym q) p (mid ∙ sym p)) + ∙ sym (assoc (sym q ∙ p) (mid ∙ sym p) q) + ∙ cong ((sym q ∙ p) ∙_) (sym (assoc mid (sym p) q)) + ∙ cong (λ x → (sym q ∙ p) ∙ (mid ∙ x)) (sym (symDistr (sym q) p)) + ∙ cong ((sym q ∙ p)∙_) (commHelper mid _) + ∙ assoc _ _ _ + ∙ cong (_∙ mid) (rCancel (sym q ∙ p)) + ∙ sym (lUnit mid) + + congHelper : ∀ {ℓ} {A : Type ℓ} {a1 : A} (p : a1 ≡ a1) (f : A → A) (id : (λ x → x) ≡ f) + → cong f p ≡ sym (funExt⁻ id a1) ∙ p ∙ funExt⁻ id a1 + congHelper {a1 = a1} p f id = + (λ i → lUnit (rUnit (cong f p) i) i) + ∙ (λ i → (λ j → id ((~ i) ∨ (~ j)) a1) ∙ cong (id (~ i)) p ∙ λ j → id (~ i ∨ j) a1) + + + congHelper++ : (p : 0ₖ ≡ 0ₖ) (f : hLevelTrunc 4 (S₊ 2) → hLevelTrunc 4 (S₊ 2)) (id : (λ x → x) ≡ f) + → (q : (f 0ₖ) ≡ 0ₖ) + → (sym q) ∙ cong f p ∙ q ≡ p + congHelper++ p f id q = + cong (λ x → sym q ∙ x ∙ q) (congHelper p f id) ∙ + moveabout (sym (funExt⁻ id ∣ north ∣)) q p + +Iso.leftInv test13 a = + funExt λ {north → rUnitₖ (a north) + ; south → rUnitₖ (a north) ∙ cong a (merid north) + ; (merid south i) → test i + ; (merid north i) → test2 i} + where + test : PathP (λ i → a north +ₖ Kn→ΩKn+1 1 (ΩKn+1→Kn (sym (rCancelₖ (a north)) + ∙ (λ i → a (merid south i) +ₖ (-ₖ a (merid north i))) + ∙ rCancelₖ (a south))) i ≡ a (merid south i)) (rUnitₖ (a north)) (rUnitₖ (a north) ∙ cong a (merid north)) + test j i = + hcomp (λ k → λ { (i = i0) → (helper ∙ together ∙ congFunct (_+ₖ 0ₖ) (cong a (merid south)) (cong a (sym (merid north)))) (~ k) j + ; (i = i1) → a (merid south j) + ; (j = i0) → rUnitₖ (a north) i + ; (j = i1) → ((λ j → rUnitₖ (a (merid north ((~ k) ∧ j))) (j ∧ k)) ∙ λ j → rUnitₖ (a (merid north ((~ k) ∨ j))) (k ∨ j)) i }) + (hcomp (λ k → λ { (i = i1) → a (merid south j) + ; (i = i0) → compPath-filler ((cong (λ x → a x +ₖ 0ₖ) (merid south))) (cong (λ x → a x +ₖ 0ₖ) (sym (merid north))) k j + ; (j = i0) → rUnitₖ (a north) i + ; (j = i1) → pathFiller2 (cong (_+ₖ 0ₖ) (cong a (merid north))) (rUnitₖ (a south)) k i}) + (rUnitₖ (a (merid south j)) i)) + + + where + pathFiller : (rUnitₖ (a north) ∙ cong a (merid north)) ≡ cong (_+ₖ 0ₖ) (cong a (merid north)) ∙ rUnitₖ (a south) + pathFiller = (λ i → (λ j → rUnitₖ (a (merid north (i ∧ j))) (j ∧ (~ i))) ∙ λ j → rUnitₖ (a (merid north (i ∨ j))) (~ i ∨ j)) + + pathFiller2 : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) → PathP (λ i → p (~ i) ≡ z) q (p ∙ q) + pathFiller2 p q i = + hcomp (λ k → λ {(i = i0) → lUnit q (~ k) + ; (i = i1) → p ∙ q}) + ((λ j → p (~ i ∨ j)) ∙ q) + + helper : Path (_ ≡ _) (λ i → a north +ₖ Kn→ΩKn+1 1 (ΩKn+1→Kn (sym (rCancelₖ (a north)) + ∙ (λ i → a (merid south i) +ₖ (-ₖ a (merid north i))) + ∙ rCancelₖ (a south))) i) + --- + (cong (a north +ₖ_) (sym (rCancelₖ (a north))) + ∙ ((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north))) + ∙ cong (a north +ₖ_) (rCancelₖ (a north))) + helper = + (λ j i → a north +ₖ Iso.rightInv (Iso3-Kn-ΩKn+1 1) ((sym (rCancelₖ (a north)) + ∙ (λ i → a (merid south i) +ₖ (-ₖ a (merid north i))) + ∙ rCancelₖ (a south))) j i) ∙ + congFunct (a north +ₖ_) (sym (rCancelₖ (a north))) ((λ i → a (merid south i) +ₖ (-ₖ a (merid north i))) ∙ rCancelₖ (a south)) ∙ + (λ j → cong (a north +ₖ_) (sym (rCancelₖ (a north))) ∙ congFunct (a north +ₖ_) ((λ i → a (merid south i) +ₖ (-ₖ a (merid north i)))) (rCancelₖ (a south)) j) ∙ + (λ j → cong (a north +ₖ_) (sym (rCancelₖ (a north))) ∙ cong₂Funct (λ x y → a north +ₖ a x +ₖ (-ₖ (a y))) (merid south) (merid north) j ∙ cong (a north +ₖ_) ((rCancelₖ (a south)))) ∙ + (λ i → cong (a north +ₖ_) (sym (rCancelₖ (a north))) ∙ ((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ i ∨ (~ j))) +ₖ (-ₖ (a (merid north (j ∧ (~ i)))))) ∙ cong (a north +ₖ_) ((rCancelₖ (a (merid north (~ i)))))) + + abstract + pathHelper : (a b : hLevelTrunc 4 (S₊ 2)) → a +ₖ b +ₖ (-ₖ a) ≡ b +ₖ 0ₖ + pathHelper a b = + a +ₖ b +ₖ (-ₖ a) ≡⟨ commₖ a (b +ₖ (-ₖ a)) ⟩ + (b +ₖ (-ₖ a)) +ₖ a ≡⟨ assocₖ b (-ₖ a) a ⟩ + b +ₖ (-ₖ a) +ₖ a ≡⟨ cong (b +ₖ_) (lCancelₖ a) ⟩ + b +ₖ 0ₖ ∎ + + + helperFun : ∀ {ℓ} {A : Type ℓ} {a b : A} (p p' : a ≡ b) (q q' : b ≡ b) (r : a ≡ a) + → ((p q : a ≡ a) → p ∙ q ≡ q ∙ p) + → q ≡ q' + → PathP (λ i → p' (~ i) ≡ p' (~ i)) q' r + → p ∙ q ∙ sym p ≡ r + helperFun p p' q q' r comm qid dep = + p ∙ q ∙ sym p ≡⟨ cong (λ x → p ∙ x ∙ sym p) qid ⟩ + p ∙ q' ∙ sym p ≡⟨ cong (λ x → p ∙ x ∙ sym p) (λ i → lUnit (rUnit q' i) i) ⟩ + p ∙ (refl ∙ q' ∙ refl) ∙ sym p ≡⟨ cong (λ x → p ∙ x ∙ sym p) (λ i → (λ j → p' (~ i ∨ ~ j)) ∙ dep i ∙ λ j → p' (~ i ∨ j)) ⟩ + p ∙ (sym p' ∙ r ∙ p') ∙ sym p ≡⟨ assoc p (sym p' ∙ r ∙ p') (sym p) ⟩ + (p ∙ sym p' ∙ r ∙ p') ∙ sym p ≡⟨ cong (_∙ sym p) (assoc p (sym p') (r ∙ p')) ⟩ + ((p ∙ sym p') ∙ r ∙ p') ∙ sym p ≡⟨ sym (assoc (p ∙ sym p') (r ∙ p') (sym p)) ⟩ + (p ∙ sym p') ∙ (r ∙ p') ∙ sym p ≡⟨ cong ((p ∙ sym p') ∙_) (sym (assoc r p' (sym p))) ⟩ + (p ∙ sym p') ∙ r ∙ (p' ∙ sym p) ≡⟨ cong (λ x → (p ∙ sym p') ∙ r ∙ x) (sym (symDistr p (sym p'))) ⟩ + (p ∙ sym p') ∙ r ∙ sym (p ∙ sym p') ≡⟨ cong ((p ∙ sym p') ∙_) (comm r (sym (p ∙ sym p'))) ⟩ + (p ∙ sym p') ∙ sym (p ∙ sym p') ∙ r ≡⟨ assoc (p ∙ sym p') (sym (p ∙ sym p')) r ⟩ + ((p ∙ sym p') ∙ sym (p ∙ sym p')) ∙ r ≡⟨ cong (_∙ r) (rCancel (p ∙ sym p')) ⟩ + refl ∙ r ≡⟨ sym (lUnit r) ⟩ + r ∎ + together : Path (_ ≡ _) (cong (a north +ₖ_) (sym (rCancelₖ (a north))) + ∙ ((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north))) + ∙ cong (a north +ₖ_) (rCancelₖ (a north))) + (cong(_+ₖ 0ₖ) (cong a (merid south) ∙ cong a (sym (merid north)))) + together = + helperFun (cong (a north +ₖ_) (sym (rCancelₖ (a north)))) + (λ i → pathHelper (a north) (a north) (~ i)) + ((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north))) + (((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north)))) + (cong(_+ₖ 0ₖ) (cong a (merid south) ∙ cong a (sym (merid north)))) + (transport (λ i → (p q : rUnitₖ (a north) (~ i) ≡ rUnitₖ (a north) (~ i)) → p ∙ q ≡ q ∙ p) + (λ p q → isCommS2 _ p q)) + refl + λ i → hcomp (λ k → λ {(i = i0) → (λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north)) + ; (i = i1) → congFunct (_+ₖ 0ₖ) (cong a (merid south)) (cong a (sym (merid north))) (~ k)}) + ((λ j → pathHelper (a north) (a (merid south j)) i) ∙ (λ j → pathHelper (a north) (a (merid north (~ j))) i)) + where + isCommS2 : (x : hLevelTrunc 4 (S₊ 2)) (p q : x ≡ x) → p ∙ q ≡ q ∙ p + isCommS2 x p q = sym (transport⁻Transport (typId x) (p ∙ q)) + ∙ cong (transport (λ i → typId x (~ i))) (typIdHelper p q) + ∙ (λ i → transport (λ i → typId x (~ i)) (commHelper ∣ north ∣ ∣ north ∣ (transport (λ i → typId x i) p) (transport (λ i → typId x i) q) i)) + ∙ cong (transport (λ i → typId x (~ i))) (sym (typIdHelper q p)) + ∙ transport⁻Transport (typId x) (q ∙ p) + where + congIsoHelper : (x : hLevelTrunc 4 (S₊ 2)) → Iso (hLevelTrunc 4 (S₊ 2)) (hLevelTrunc 4 (S₊ 2)) + Iso.fun (congIsoHelper x) = _+ₖ x + Iso.inv (congIsoHelper x) = _+ₖ (-ₖ x) + Iso.rightInv (congIsoHelper x) a = assocₖ a (-ₖ x) x ∙ cong (a +ₖ_) (lCancelₖ x) ∙ rUnitₖ a + Iso.leftInv (congIsoHelper x) a = assocₖ a x (-ₖ x) ∙ cong (a +ₖ_) (rCancelₖ x) ∙ rUnitₖ a + typId : (x : hLevelTrunc 4 (S₊ 2)) → (x ≡ x) ≡ Path (hLevelTrunc 4 (S₊ 2)) 0ₖ 0ₖ + typId x = isoToPath ((congIso (isoToEquiv (symIso (congIsoHelper x))))) ∙ λ i → rCancelₖ x i ≡ rCancelₖ x i + where + + typIdHelper : (p q : (x ≡ x)) → transport (λ i → typId x i) (p ∙ q) ≡ transport (λ i → typId x i) p ∙ transport (λ i → typId x i) q + typIdHelper p q = + (substComposite (λ x → x) (isoToPath ((congIso (isoToEquiv (symIso (congIsoHelper x)))))) (λ i → rCancelₖ x i ≡ rCancelₖ x i) (p ∙ q)) + ∙ (λ i → transport (λ i → rCancelₖ x i ≡ rCancelₖ x i) (transport (isoToPath (congIso (isoToEquiv (symIso (congIsoHelper x))))) (p ∙ q))) + ∙ (λ i → transport (λ i → rCancelₖ x i ≡ rCancelₖ x i) (transportRefl (congFunct (_+ₖ (-ₖ x)) p q i) i)) + ∙ (λ i → transport (λ i → rCancelₖ x i ≡ rCancelₖ x i) ((lUnit (rUnit (transportRefl (cong (_+ₖ (-ₖ x)) p) (~ i)) i) i) ∙ (lUnit (rUnit (transportRefl (cong (_+ₖ (-ₖ x)) q) (~ i)) i) i))) + ∙ (λ i → transp (λ j → rCancelₖ x (i ∨ j) ≡ rCancelₖ x (i ∨ j)) i (((λ j → rCancelₖ x (i ∧ (~ j))) ∙ transport refl (cong (_+ₖ (-ₖ x)) p) ∙ λ j → rCancelₖ x (i ∧ j)) ∙ (λ j → rCancelₖ x (i ∧ (~ j))) ∙ transport refl (cong (_+ₖ (-ₖ x)) q) ∙ λ j → rCancelₖ x (i ∧ j))) + ∙ (λ i → transp (λ j → rCancelₖ x ((~ i) ∨ j) ≡ rCancelₖ x ((~ i) ∨ j)) (~ i) ((λ j → rCancelₖ x ((~ i) ∧ (~ j))) ∙ transport refl (cong (_+ₖ (-ₖ x)) p) ∙ λ j → rCancelₖ x ((~ i) ∧ j)) + ∙ transp (λ j → rCancelₖ x ((~ i) ∨ j) ≡ rCancelₖ x ((~ i) ∨ j)) (~ i) ((λ j → rCancelₖ x ((~ i) ∧ (~ j))) ∙ transport refl (cong (_+ₖ (-ₖ x)) q) ∙ λ j → rCancelₖ x ((~ i) ∧ j))) + ∙ (λ i → transport (λ j → rCancelₖ x j ≡ rCancelₖ x j) (lUnit (rUnit (transport refl (cong (_+ₖ (-ₖ x)) p)) (~ i)) (~ i)) + ∙ transport (λ j → rCancelₖ x j ≡ rCancelₖ x j) (lUnit (rUnit (transport refl (cong (_+ₖ (-ₖ x)) q)) (~ i)) (~ i))) + ∙ cong₂ (_∙_) (sym (substComposite (λ x → x) (isoToPath ((congIso (isoToEquiv (symIso (congIsoHelper x)))))) (λ i → rCancelₖ x i ≡ rCancelₖ x i) p)) + (sym (substComposite (λ x → x) (isoToPath ((congIso (isoToEquiv (symIso (congIsoHelper x)))))) (λ i → rCancelₖ x i ≡ rCancelₖ x i) q)) + + + test2 : PathP (λ i → a north +ₖ ∣ north ∣ ≡ a (merid north i)) (rUnitₖ (a north)) (rUnitₖ (a north) ∙ cong a (merid north)) + test2 i j = + hcomp (λ k → λ { (i = i0) → rUnitₖ (a north) j + ; (j = i1) → a (merid north (i ∧ k)) + ; (j = i0) → a north +ₖ ∣ north ∣}) + (rUnitₖ (a north) j) + + +S1→S¹ : S₊ 1 → S¹ +S1→S¹ x = SuspBool→S¹ (S1→SuspBool x) + +S¹→S1 : S¹ → S₊ 1 +S¹→S1 x = SuspBool→S1 (S¹→SuspBool x) + +S1→S¹-sect : section S1→S¹ S¹→S1 +S1→S¹-sect x = + cong SuspBool→S¹ (SuspBool→S1-retr (S¹→SuspBool x)) + ∙ S¹→SuspBool→S¹ x + +S1→S¹-retr : retract S1→S¹ S¹→S1 +S1→S¹-retr x = + cong SuspBool→S1 (SuspBool→S¹→SuspBool (S1→SuspBool x)) + ∙ SuspBool→S1-sect x + + + + + + + +isGroupoidS1 : isGroupoid (S₊ 1) +isGroupoidS1 = transport (λ i → isGroupoid (S¹≡S1 (~ i))) isGroupoidS¹ +isConnectedS1 : (x : S₊ 1) → ∥ north ≡ x ∥₋₁ +isConnectedS1 x = pRec propTruncIsProp + (λ p → ∣ cong (transport (sym (S¹≡S1))) p ∙ transport⁻Transport (S¹≡S1) x ∣₋₁) + (isConnectedS¹ (transport S¹≡S1 x)) + + +open import Cubical.HITs.S2 + +test : (S₊ 2) → S₊ 1 +test north = north +test south = south +test (merid a i) = merid north i + +test2 : (S₊ 1) → (S₊ 2) +test2 north = north +test2 south = south +test2 (merid a i) = merid (merid a i) i + + +testS² : S² → S¹ +testS² base = base +testS² (surf i i₁) = base + +test4 : S₊ 1 → hLevelTrunc 4 (S₊ 2) +test4 north = ∣ north ∣ +test4 south = ∣ north ∣ +test4 (merid a i) = (Kn→ΩKn+1 1 ∣ south ∣) i + + + + +coHom0Torus : grIso (coHomGr 0 (S₊ 1 × S₊ 1)) intGroup +coHom0Torus = + Iso'→Iso + (iso' + (iso (sRec isSetInt (λ f → f (north , north))) + (λ a → ∣ (λ x → a) ∣₀) + (λ a → refl) + (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ + (funExt λ {(x , y) → suspToPropRec2 + {B = λ x y → f (north , north) ≡ f (x , y)} + north + (λ _ _ → isSetInt _ _) + refl + x y}))) + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) λ a b → addLemma (a (north , north)) (b (north , north)))) + +private + S¹map : hLevelTrunc 3 S¹ → S¹ + S¹map = trElim (λ _ → isGroupoidS¹) (idfun S¹) + + S¹map-id : (x : hLevelTrunc 3 S¹) → Path (hLevelTrunc 3 S¹) ∣ S¹map x ∣ x + S¹map-id = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → refl + + S1map : hLevelTrunc 3 (S₊ 1) → (S₊ 1) + S1map = trElim (λ _ → isGroupoidS1) (idfun _) + + +S¹→S¹ : Iso (S¹ → hLevelTrunc 3 S¹) (S¹ × Int) +Iso.fun S¹→S¹ f = S¹map (f base) + , winding (basechange2⁻ (S¹map (f base)) λ i → S¹map (f (loop i))) +Iso.inv S¹→S¹ (s , int) base = ∣ s ∣ +Iso.inv S¹→S¹ (s , int) (loop i) = ∣ basechange2 s (intLoop int) i ∣ +Iso.rightInv S¹→S¹ (s , int) = ×≡ refl ((λ i → winding (basechange2-retr s (λ i → intLoop int i) i)) + ∙ windingIntLoop int) +Iso.leftInv S¹→S¹ f = funExt λ { base → S¹map-id (f base) + ; (loop i) j → helper2 j i} + where + helper2 : PathP (λ i → S¹map-id (f base) i ≡ S¹map-id (f base) i) + (λ i → ∣ basechange2 (S¹map (f base)) (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁)))))) i ∣) + (cong f loop) + helper2 i j = + hcomp (λ k → λ { (i = i0) → cong ∣_∣ (basechange2 (S¹map (f base)) (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁))))))) j + ; (i = i1) → S¹map-id (f (loop j)) k + ; (j = i0) → S¹map-id (f base) (i ∧ k) + ; (j = i1) → S¹map-id (f base) (i ∧ k)}) + (helper4 i j) + where + helper4 : Path (Path (hLevelTrunc 3 _) _ _) + (cong ∣_∣ (basechange2 (S¹map (f base)) + (intLoop + (winding + (basechange2⁻ (S¹map (f base)) + (λ i₁ → S¹map (f (loop i₁)))))))) + λ i → ∣ S¹map (f (loop i)) ∣ + helper4 i = + cong ∣_∣ + ((cong (basechange2 (S¹map (f base))) + (decodeEncode base (basechange2⁻ (S¹map (f base)) + (λ i₁ → S¹map (f (loop i₁))))) + ∙ basechange2-sect (S¹map (f base)) + (λ i → S¹map (f (loop i)))) i) + +S1→S1→S1×Int : Iso ((S₊ 1) → hLevelTrunc 3 (S₊ 1)) ((hLevelTrunc 3 (S₊ 1)) × Int) +S1→S1→S1×Int = compIso helper2 (compIso S¹→S¹ helper) + where + helper : Iso (S¹ × Int) (hLevelTrunc 3 (S₊ 1) × Int) + Iso.fun helper (s , int) = ∣ S¹→S1 s ∣ , int + Iso.inv helper (s , int) = (S1→S¹ (S1map s)) , int + Iso.rightInv helper (s , int) = + trElim {B = λ s → (∣ S¹→S1 (S1→S¹ (S1map s)) ∣ , int) ≡ (s , int)} + (λ _ → isOfHLevelPath 3 (isOfHLevelProd 3 (isOfHLevelTrunc 3) (isOfHLevelSuc 2 isSetInt)) _ _) + (λ a → ×≡ (cong ∣_∣ (S1→S¹-retr a)) refl) + s + Iso.leftInv helper (s , int) = ×≡ (S1→S¹-sect s) refl + + helper2 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → hLevelTrunc 3 S¹) + Iso.fun helper2 f x = trMap S1→S¹ (f (S¹→S1 x)) + Iso.inv helper2 f x = trMap S¹→S1 (f (S1→S¹ x)) + Iso.rightInv helper2 f = funExt λ x i → helper3 (f (S1→S¹-sect x i)) i + where + helper3 : (x : hLevelTrunc 3 S¹) → trMap S1→S¹ (trMap S¹→S1 x) ≡ x + helper3 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → cong ∣_∣ (S1→S¹-sect a) + Iso.leftInv helper2 f = funExt λ x i → helper3 (f (S1→S¹-retr x i)) i + where + helper3 : (x : hLevelTrunc 3 (S₊ 1)) → trMap S¹→S1 (trMap S1→S¹ x) ≡ x + helper3 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → cong ∣_∣ (S1→S¹-retr a) + +S¹→S¹≡S1→S1 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → hLevelTrunc 3 S¹) +Iso.fun S¹→S¹≡S1→S1 f x = trMap S1→S¹ (f (S¹→S1 x)) +Iso.inv S¹→S¹≡S1→S1 f x = trMap S¹→S1 (f (S1→S¹ x)) +Iso.rightInv S¹→S¹≡S1→S1 F = funExt λ x i → helper2 (F (S1→S¹-sect x i)) i + where + helper2 : (x : hLevelTrunc 3 S¹) → trMap S1→S¹ (trMap S¹→S1 x) ≡ x + helper2 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → cong ∣_∣ (S1→S¹-sect a) +Iso.leftInv S¹→S¹≡S1→S1 F = funExt λ x i → helper2 (F (S1→S¹-retr x i)) i + where + helper2 : (x : hLevelTrunc 3 (S₊ 1)) → trMap S¹→S1 (trMap S1→S¹ x) ≡ x + helper2 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → cong ∣_∣ (S1→S¹-retr a) + + + + +basechange-lemma : ∀ {ℓ} {A : Type ℓ} {a : A} (x y : S¹) (F : a ≡ a → S¹) (f : S¹ → a ≡ a) (g : S¹ → a ≡ a) + → (f base ≡ refl) + → (g base ≡ refl) + → basechange2⁻ (F (f base ∙ g base)) (cong₂ {A = S¹} {B = λ x → S¹} (λ x y → F (f x ∙ g y)) loop loop) + ≡ basechange2⁻ (F (f base)) (cong (λ x → F (f x)) loop) ∙ basechange2⁻ (F (g base)) (cong (λ x → F (g x)) loop) +basechange-lemma x y F f g frefl grefl = + (λ i → basechange2⁻ (F (f base ∙ g base)) (cong₂Funct (λ x y → F (f x ∙ g y)) loop loop i)) + ∙ (λ i → basechange2⁻ (F (f base ∙ g base)) (cong (λ x₁ → F (f x₁ ∙ g base)) loop ∙ cong (λ y₁ → F (f base ∙ g y₁)) loop)) + ∙ basechange2⁻-morph (F (f base ∙ g base)) _ _ + ∙ (λ j → basechange2⁻ (F (f base ∙ grefl j)) + (λ i → F (f (loop i) ∙ grefl j)) + ∙ basechange2⁻ (F (frefl j ∙ g base)) + (λ i → F (frefl j ∙ g (loop i)))) + ∙ ((λ j → basechange2⁻ (F (rUnit (f base) (~ j))) + (λ i → F (rUnit (f (loop i)) (~ j))) + ∙ basechange2⁻ (F (lUnit (g base) (~ j))) + (λ i → F (lUnit (g (loop i)) (~ j))))) + + +basechange-lemma2 : (f g : S¹ → hLevelTrunc 3 (S₊ 1)) (F : hLevelTrunc 3 (S₊ 1) → S¹) + → ((basechange2⁻ (F (f base +ₖ g base)) λ i → F ((f (loop i)) +ₖ g (loop i))) + ≡ basechange2⁻ (F (f base)) (cong (F ∘ f) loop) + ∙ basechange2⁻ (F (g base)) (cong (F ∘ g) loop)) +basechange-lemma2 f g F = coInd f g F (f base) (g base) refl refl + where + coInd : (f g : S¹ → hLevelTrunc 3 (S₊ 1)) (F : hLevelTrunc 3 (S₊ 1) → S¹) (x y : hLevelTrunc 3 (S₊ 1)) + → f base ≡ x + → g base ≡ y + → ((basechange2⁻ (F (f base +ₖ g base)) λ i → F ((f (loop i)) +ₖ g (loop i))) + ≡ basechange2⁻ (F (f base)) (cong (F ∘ f) loop) + ∙ basechange2⁻ (F (g base)) (cong (F ∘ g) loop)) + coInd f g F = + elim2 (λ _ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 (isGroupoidS¹ base base)) _ _ ) + (suspToPropRec2 + north + (λ _ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → isGroupoidS¹ _ _ _ _) + λ fid gid → + (λ i → basechange2⁻ (F (f base +ₖ g base)) (cong₂Funct (λ x y → F (f x +ₖ g y)) loop loop i)) + ∙ (basechange2⁻-morph (F (f base +ₖ g base)) + (cong (λ x → F (f x +ₖ g base)) loop) + (cong (λ x → F (f base +ₖ g x)) loop)) + ∙ (λ i → basechange2⁻ (F (f base +ₖ gid i)) (cong (λ x → F (f x +ₖ gid i)) loop) + ∙ basechange2⁻ (F (fid i +ₖ g base)) (cong (λ x → F (fid i +ₖ g x)) loop)) + ∙ (λ i → basechange2⁻ (F (rUnitₖ (f base) i)) (cong (λ x → F (rUnitₖ (f x) i)) loop) + ∙ basechange2⁻ (F (lUnitₖ (g base) i)) (cong (λ x → F (lUnitₖ (g x) i)) loop))) + + + +coHom1S1≃ℤ : grIso (coHomGr 1 (S₊ 1)) intGroup +coHom1S1≃ℤ = + Iso'→Iso + (iso' + (iso + (sRec isSetInt (λ f → proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) + (λ a → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , a)) ∣₀) + (λ a → (λ i → proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 (Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , a)))))) + ∙ (λ i → proj₂ (Iso.fun S¹→S¹ (Iso.rightInv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , a)) i))) + ∙ λ i → proj₂ (Iso.rightInv S¹→S¹ (base , a) i)) + (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → (λ i → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹ (base , proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) ∣₀) + ∙ (λ i → sRec setTruncIsSet + (λ x → ∣ Iso.inv S¹→S¹≡S1→S1 x ∣₀) + (sRec setTruncIsSet + (λ x → ∣ Iso.inv S¹→S¹ (x , (proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) ∣₀) + ∣ base ∣₀)) + ∙ (λ i → sRec setTruncIsSet + (λ x → ∣ Iso.inv S¹→S¹≡S1→S1 x ∣₀) + (sRec setTruncIsSet + (λ x → ∣ Iso.inv S¹→S¹ (x , (proj₂ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) ∣₀) + (Iso.inv PathIdTrunc₀Iso (isConnectedS¹ (proj₁ (Iso.fun S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f)))) i))) + ∙ (λ i → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.leftInv S¹→S¹ (Iso.fun S¹→S¹≡S1→S1 f) i) ∣₀) + ∙ (λ i → ∣ Iso.leftInv S¹→S¹≡S1→S1 f i ∣₀))) + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) + λ f g → (λ i → winding (basechange2⁻ (S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)) ∙ Kn→ΩKn+1 1 (g (S¹→S1 base)))))) + (λ i → S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 (loop i))) ∙ Kn→ΩKn+1 1 (g (S¹→S1 (loop i))))))))) + ∙ cong winding (helper2 (f (S¹→S1 base)) (g (S¹→S1 base)) f g refl refl) + ∙ winding-hom ((basechange2⁻ (S¹map (trMap S1→S¹ (f north))) + (λ i → S¹map (trMap S1→S¹ (f (S¹→S1 (loop i))))))) + ((basechange2⁻ (S¹map (trMap S1→S¹ (g north))) + (λ i → S¹map (trMap S1→S¹ (g (S¹→S1 (loop i))))))))) + + + where + helper2 : (x y : hLevelTrunc 3 (S₊ 1)) (f g : S₊ 1 → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1) + → (f (S¹→S1 base)) ≡ x + → (g (S¹→S1 base)) ≡ y + → (basechange2⁻ (S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)) ∙ Kn→ΩKn+1 1 (g (S¹→S1 base))))))) + (λ i → S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 (loop i))) ∙ Kn→ΩKn+1 1 (g (S¹→S1 (loop i))))))) + ≡ (basechange2⁻ (S¹map (trMap S1→S¹ ((f (S¹→S1 base)))))) + (λ i → S¹map (trMap S1→S¹ (f (S¹→S1 (loop i))))) + ∙ (basechange2⁻ (S¹map (trMap S1→S¹ ((g (S¹→S1 base))))) + (λ i → S¹map (trMap S1→S¹ ((g (S¹→S1 (loop i))))))) + helper2 = elim2 + (λ _ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 + λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 + λ _ → isOfHLevelPath 3 (isOfHLevelSuc 3 (isGroupoidS¹) base base) _ _) + (suspToPropRec2 {A = S₊ 0} north + (λ _ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → (isGroupoidS¹) _ _ _ _ ) + λ f g reflf reflg → + (basechange-lemma + base + base + (λ x → S¹map (trMap S1→S¹ (ΩKn+1→Kn x))) + (λ x → Kn→ΩKn+1 1 (f (S¹→S1 x))) ((λ x → Kn→ΩKn+1 1 (g (S¹→S1 x)))) + (cong (Kn→ΩKn+1 1) reflf ∙ Kn→ΩKn+10ₖ 1) + (cong (Kn→ΩKn+1 1) reflg ∙ Kn→ΩKn+10ₖ 1)) + ∙ λ j → basechange2⁻ (S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 base)) j))) + (λ i → S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 (loop i))) j))) + ∙ basechange2⁻ (S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 base)) j))) + (λ i → S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 (loop i))) j)))) + +pushoutSn : (n : ℕ) → Iso (S₊ (suc n)) (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt) +Iso.fun (pushoutSn n) north = inl tt +Iso.fun (pushoutSn n) south = inr tt +Iso.fun (pushoutSn n) (merid a i) = push a i +Iso.inv (pushoutSn n) (inl x) = north +Iso.inv (pushoutSn n) (inr x) = south +Iso.inv (pushoutSn n) (push a i) = merid a i +Iso.rightInv (pushoutSn n) (inl x) = refl +Iso.rightInv (pushoutSn n) (inr x) = refl +Iso.rightInv (pushoutSn n) (push a i) = refl +Iso.leftInv (pushoutSn n) north = refl +Iso.leftInv (pushoutSn n) south = refl +Iso.leftInv (pushoutSn n) (merid a i) = refl + +Sn≡Pushout : (n : ℕ) → (S₊ (suc n)) ≡ (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt) +Sn≡Pushout n = isoToPath (pushoutSn n) + +coHomPushout≡coHomSn' : (n m : ℕ) → grIso (coHomGr m (S₊ (suc n))) (coHomGr m (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt)) +morph.fun (grIso.fun (coHomPushout≡coHomSn' n m)) = + sRec setTruncIsSet + λ f → ∣ (λ {(inl x) → f north ; (inr x) → f south ; (push a i) → f (merid a i)}) ∣₀ +morph.ismorph (grIso.fun (coHomPushout≡coHomSn' n m)) = + sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ a b → cong ∣_∣₀ (funExt λ {(inl x) → refl ; (inr x) → refl ; (push a i) → refl }) +morph.fun (grIso.inv (coHomPushout≡coHomSn' n m)) = sRec setTruncIsSet (λ f → ∣ (λ {north → f (inl tt) ; south → f (inr tt) ; (merid a i) → f (push a i)}) ∣₀) +morph.ismorph (grIso.inv (coHomPushout≡coHomSn' n m)) = + sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ a b → cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → refl }) +grIso.rightInv (coHomPushout≡coHomSn' n m) = + sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ (funExt λ {(inl x) → refl ; (inr x) → refl ; (push a i) → refl }) +grIso.leftInv (coHomPushout≡coHomSn' n m) = + sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → refl }) + + + + + + +coHomGrUnit0 : grIso (coHomGr 0 Unit) intGroup +grIso.fun coHomGrUnit0 = mph (sRec isSetInt (λ f → f tt)) + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) + (λ a b → addLemma (a tt) (b tt))) +grIso.inv coHomGrUnit0 = mph (λ a → ∣ (λ _ → a) ∣₀) (λ a b i → ∣ (λ _ → addLemma a b (~ i)) ∣₀) +grIso.rightInv coHomGrUnit0 a = refl +grIso.leftInv coHomGrUnit0 = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ a → refl + +isContrCohomUnit : (n : ℕ) → isContr (coHom (suc n) Unit) +isContrCohomUnit n = subst isContr (λ i → ∥ UnitToTypeId (coHomK (suc n)) (~ i) ∥₀) (helper n) + where + helper : (n : ℕ) → isContr (∥ coHomK (suc n) ∥₀) + helper n = subst isContr ((isoToPath (truncOfTruncIso {A = S₊ (1 + n)} 2 (1 + n))) ∙ sym propTrunc≡Trunc0 ∙ λ i → ∥ hLevelTrunc (suc (+-comm n 2 i)) (S₊ (1 + n)) ∥₀) + (isConnectedSubtr 2 (helper2 n .fst) (subst (λ x → isHLevelConnected x (S₊ (suc n))) (sym (helper2 n .snd)) (sphereConnected (suc n))) ) + where + helper2 : (n : ℕ) → Σ[ m ∈ ℕ ] m + 2 ≡ 2 + n + helper2 zero = 0 , refl + helper2 (suc n) = (suc n) , λ i → suc (+-comm n 2 i) + +coHomGrUnit≥1 : (n : ℕ) → grIso (coHomGr (suc n) Unit) trivialGroup +grIso.fun (coHomGrUnit≥1 n) = mph (λ _ → tt) (λ _ _ → refl) +grIso.inv (coHomGrUnit≥1 n) = mph (λ _ → ∣ (λ _ → ∣ north ∣) ∣₀) (λ _ _ → sym (rUnitₕ 0ₕ)) +grIso.rightInv (coHomGrUnit≥1 n) a = refl +grIso.leftInv (coHomGrUnit≥1 n) a = sym (isContrCohomUnit n .snd 0ₕ) ∙ isContrCohomUnit n .snd a + +S0→Int : (a : Int × Int) → S₊ 0 → Int +S0→Int a north = proj₁ a +S0→Int a south = proj₂ a + +coHom0-S0 : grIso (coHomGr 0 (S₊ 0)) (dirProd intGroup intGroup) +coHom0-S0 = + Iso'→Iso + (iso' + (iso (sRec (isOfHLevelProd 2 isSetInt isSetInt) + λ f → (f north) , (f south)) + (λ a → ∣ S0→Int a ∣₀) + (λ { (a , b) → refl}) + (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ f → cong ∣_∣₀ (funExt (λ {north → refl ; south → refl})))) + (sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 isSetInt isSetInt) _ _) + λ a b i → addLemma (a north) (b north) i , addLemma (a south) (b south) i)) + + + + +coHom0S1 : grIso intGroup (coHomGr 0 (S₊ 1)) +coHom0S1 = + Iso'→Iso + (iso' + (iso + (λ a → ∣ (λ x → a) ∣₀) + (sRec isSetInt (λ f → f north)) + (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ (funExt (suspToPropRec north (λ _ → isSetInt _ _) refl))) + (λ a → refl)) + λ a b → cong ∣_∣₀ (funExt λ x → sym (addLemma a b))) + +coHom1S1-helper : Iso (coHom 2 (S₊ 1)) Unit +coHom1S1-helper = + compIso (setTruncIso test13) + (compIso + setTruncOfProdIso + (compIso + (prodIso setTruncTrunc0Iso setTruncTrunc0Iso) + (pathToIso + ((λ i → isContr→≡Unit helper2 i × isContr→≡Unit helper3 i) + ∙ sym diagonal-unit)))) + where + + helper2 : isContr (hLevelTrunc 2 (hLevelTrunc 4 (S₊ 2))) + helper2 = + transport (λ i → isContr (isoToPath (truncOfTruncIso {A = (S₊ 2)} 2 2) i)) + (isConnectedSubtr 2 1 (sphereConnected 2)) + + helper3 : isContr (hLevelTrunc 2 (hLevelTrunc 3 (S₊ 1))) + helper3 = + transport (λ i → isContr (hLevelTrunc 2 (truncIdempotent {A = S₊ 1} 3 isGroupoidS1 (~ i)))) + (sphereConnected 1) + + + +coHom1S1 : grIso intGroup (coHomGr 1 (S₊ 1)) +coHom1S1 = + compGrIso + (diagonalIso1 + _ + (coHomGr 0 (S₊ 0)) + _ + (invGrIso coHom0-S0) + (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) + (λ x → MV.Ker-i⊂Im-d _ _(S₊ 0) (λ _ → tt) (λ _ → tt) 0 x + (×≡ (isOfHLevelSuc 0 (isContrCohomUnit 0) _ _) + (isOfHLevelSuc 0 (isContrCohomUnit 0) _ _))) + (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) + (λ x inker + → pRec propTruncIsProp + (λ {((f , g) , id') → helper x f g id' inker}) + ((MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ inker)))) + (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ F surj + → pRec (setTruncIsSet _ _) (λ { (x , id) → MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ F ∣₀ ∣ (∣ (λ _ → x) ∣₀ , ∣ (λ _ → 0) ∣₀) , + (cong ∣_∣₀ (funExt (surjHelper x))) ∙ sym id ∣₋₁ }) surj) ) + (invGrIso (coHomPushout≡coHomSn' 0 1)) + + where + surjHelper : (x : Int) (x₁ : S₊ 0) → x +ₖ (-ₖ pos 0) ≡ S0→Int (x , x) x₁ + surjHelper x north = cong (x +ₖ_) (-0ₖ) ∙ rUnitₖ x + surjHelper x south = cong (x +ₖ_) (-0ₖ) ∙ rUnitₖ x + + helper : (F : S₊ 0 → Int) (f g : ∥ (Unit → Int) ∥₀) (id : morph.fun (MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) (f , g) ≡ ∣ F ∣₀) + → isInKer (coHomGr 0 (S₊ 0)) + (coHomGr 1 (Pushout (λ _ → tt) (λ _ → tt))) + (morph.fun (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) + ∣ F ∣₀ + → ∃[ x ∈ Int ] ∣ F ∣₀ ≡ morph.fun (grIso.inv coHom0-S0) (x , x) + helper F = + sElim2 (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ f g id inker + → pRec propTruncIsProp + (λ {((a , b) , id2) + → sElim2 {B = λ f g → morph.fun (MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) (f , g) ≡ ∣ F ∣₀ → _ } + (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) + (λ f g id → ∣ (helper2 f g .fst) , (sym id ∙ sym (helper2 f g .snd)) ∣₋₁) + a b id2}) + (MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ F ∣₀ inker) + where + helper2 : (f g : Unit → Int) + → Σ[ x ∈ Int ] morph.fun (grIso.inv coHom0-S0) (x , x) + ≡ morph.fun (MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) (∣ f ∣₀ , ∣ g ∣₀) + helper2 f g = (f _ +ₖ (-ₖ g _) ) , cong ∣_∣₀ (funExt (λ {north → refl ; south → refl})) + + + +coHom-n-Sn : (n : ℕ) → grIso intGroup (coHomGr (suc n) (S₊ (suc n))) +coHom-n-Sn zero = coHom1S1 +coHom-n-Sn (suc n) = + compGrIso + (compGrIso + (coHom-n-Sn n) + theIso) + (invGrIso (coHomPushout≡coHomSn' (suc n) (suc (suc n)))) + where + theIso : grIso (coHomGr (suc n) (S₊ (suc n))) (coHomGr (suc (suc n)) + (Pushout {A = S₊ (suc n)} (λ _ → tt) (λ _ → tt))) + theIso = + SES→Iso + (×coHomGr (suc n) Unit Unit) + (×coHomGr (suc (suc n)) Unit Unit) + (ses (λ p q → ×≡ (isOfHLevelSuc 0 (isContrCohomUnit n) (proj₁ p) (proj₁ q)) + (isOfHLevelSuc 0 (isContrCohomUnit n) (proj₂ p) (proj₂ q))) + (λ p q → ×≡ (isOfHLevelSuc 0 (isContrCohomUnit (suc n)) (proj₁ p) (proj₁ q)) + (isOfHLevelSuc 0 (isContrCohomUnit (suc n)) (proj₂ p) (proj₂ q))) + (MV.Δ _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) + (MV.i _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (2 + n)) + (MV.d _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) + (MV.Ker-d⊂Im-Δ _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) + (MV.Ker-i⊂Im-d _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n))) + + +coHom1-s0 : grIso (coHomGr 1 (S₊ 0)) trivialGroup +coHom1-s0 = + Iso'→Iso + (iso' + (iso (λ _ → tt) + (λ _ → 0ₕ) + (λ _ → refl) + (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → pRec (setTruncIsSet _ _) + (λ id → cong ∣_∣₀ (sym id)) + (helper f (f north) (f south) refl refl))) + λ _ _ → refl) + where + helper : (f : S₊ 0 → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1) → (x y : ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1) → (f north ≡ x) → (f south ≡ y) → ∥ f ≡ (λ _ → 0ₖ) ∥₋₁ + helper f = + elim2 (λ _ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelPlus {n = 1} 2 propTruncIsProp) + (suspToPropRec2 north (λ _ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → propTruncIsProp) + λ id id2 → ∣ funExt (λ {north → id ; south → id2}) ∣₋₁) + +coHom2-S1 : grIso (coHomGr 2 (S₊ 1)) trivialGroup +coHom2-S1 = + Iso'→Iso + (iso' + (iso + (λ _ → tt) + (λ _ → 0ₕ) + (λ _ → refl) + (λ _ → isPropcoHom1S1 _ _)) + λ _ _ → refl) + where + isPropcoHom1S1 : isProp (coHom 2 (S₊ 1)) + isPropcoHom1S1 = + isOfHLevelSuc 0 + (transport (λ i → isContr (isoToPath coHom1S1-helper (~ i))) + isContrUnit) + +coHom1Torus : grIso (coHomGr 1 ((S₊ 1) × (S₊ 1))) (dirProd intGroup intGroup) +coHom1Torus = + compGrIso + (Iso'→Iso + (iso' (compIso + (setTruncIso (compIso + schönfinkelIso + (compIso + (codomainIso S1→S1→S1×Int) + toProdIso))) + (setTruncOfProdIso)) + (sElim2 + (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ f g → ×≡ (cong ∣_∣₀ + (funExt (λ x → helper (f (x , S¹→S1 base) +ₖ g (x , S¹→S1 base)) + ∙ sym (cong₂ (λ x y → x +ₖ y) + (helper (f (x , S¹→S1 base))) + (helper (g (x , S¹→S1 base))))))) + (cong ∣_∣₀ + (funExt + (suspToPropRec + north + (λ _ → isSetInt _ _) + (cong winding + (basechange-lemma2 + (λ x → f (north , S¹→S1 x)) + (λ x → g (north , S¹→S1 x)) + λ x → S¹map (trMap S1→S¹ x)) + ∙ winding-hom + (basechange2⁻ + (S¹map (trMap S1→S¹ (f (north , S¹→S1 base)))) + (λ i → S¹map (trMap S1→S¹ (f (north , S¹→S1 (loop i)))))) + (basechange2⁻ + (S¹map (trMap S1→S¹ (g (north , S¹→S1 base)))) + (λ i → S¹map (trMap S1→S¹ (g (north , S¹→S1 (loop i)))))) + ∙ sym (addLemma + (winding + (basechange2⁻ + (S¹map (trMap S1→S¹ (f (north , S¹→S1 base)))) + (λ i → S¹map (trMap S1→S¹ (f (north , S¹→S1 (loop i))))))) + (winding + (basechange2⁻ + (S¹map (trMap S1→S¹ (g (north , S¹→S1 base)))) + (λ i → S¹map (trMap S1→S¹ (g (north , S¹→S1 (loop i))))))))))))))) + (dirProdIso (invGrIso (coHom-n-Sn 0)) (invGrIso coHom0S1)) + + where + helper : (x : hLevelTrunc 3 (S₊ 1)) → ∣ S¹→S1 (S¹map (trMap S1→S¹ x)) ∣ ≡ x + helper = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ a → cong ∣_∣ (S1→S¹-retr a) + + +coHom2Torus : grIso (coHomGr 2 ((S₊ 1) × (S₊ 1))) intGroup +coHom2Torus = + invGrIso + (Iso'→Iso + (iso' + (compIso + helper' + (compIso + (symIso (prodIso (groupIso→Iso coHom2-S1) + (groupIso→Iso (invGrIso (coHom-n-Sn 0))))) + (compIso + (symIso setTruncOfProdIso) + (symIso + (setTruncIso + (compIso + schönfinkelIso + (compIso + (codomainIso test13) + toProdIso))))))) + λ a b → {!!} {- (λ i → Iso.fun + (symIso + (setTruncIso + (compIso schönfinkelIso (compIso (codomainIso test13) toProdIso)))) + (Iso.fun (symIso setTruncOfProdIso) + (helper'' (Iso.inv (groupIso→Iso coHom2-S1) tt) 0ₕ i , morph.ismorph (grIso.inv (invGrIso coHom1S1)) a b i))) + ∙ (λ i → Iso.fun + (symIso + (setTruncIso + (compIso schönfinkelIso (compIso (codomainIso test13) toProdIso)))) + ∣ (λ _ → ∣ north ∣) , {!!} ∣₀) + ∙ {!toProdIso!} -})) + where + helper'' : isProp ∥ (S₊ 1 → hLevelTrunc 4 (S₊ 2)) ∥₀ + helper'' = {!!} + + helper2' : (f g : (S₊ 1 → hLevelTrunc 3 (S₊ 1))) → + Path (coHom 2 ((S₊ 1) × (S₊ 1))) (Iso.fun + (symIso + (setTruncIso + (compIso schönfinkelIso (compIso (codomainIso test13) toProdIso)))) + (Iso.fun (symIso setTruncOfProdIso) (0ₕ , ∣ f ∣₀ +ₕ ∣ g ∣₀))) + (Iso.fun (symIso + (setTruncIso + (compIso schönfinkelIso (compIso (codomainIso test13) toProdIso)))) + (Iso.fun (symIso setTruncOfProdIso) (0ₕ , ∣ f ∣₀)) +ₕ Iso.fun (symIso + (setTruncIso + (compIso schönfinkelIso (compIso (codomainIso test13) toProdIso)))) + (Iso.fun (symIso setTruncOfProdIso) (0ₕ , ∣ g ∣₀))) + helper2' f g = (λ i → Iso.fun (symIso + (setTruncIso + (compIso schönfinkelIso ((codomainIso test13))))) ∣ (λ x → 0ₖ , (f x +ₖ g x)) ∣₀) + ∙ (λ i → ∣ Iso.fun (symIso (compIso schönfinkelIso (codomainIso test13))) (λ x → 0ₖ , (f x +ₖ g x)) ∣₀) + ∙ (λ i → ∣ Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , (f x +ₖ g x)))) ∣₀) + ∙ (λ i → ∣ {!!} ∣₀) + ∙ {!!} + ∙ {!!} + where + fun : S₊ 1 × S₊ 1 → hLevelTrunc 4 (S₊ 2) + fun (x , north) = 0ₖ +ₖ 0ₖ + fun (x , south) = 0ₖ +ₖ 0ₖ + fun (x , (merid south i)) = 0ₖ +ₖ (Kn→ΩKn+1 1 (f x +ₖ g x) i) + fun (x , (merid north i)) = 0ₖ +ₖ 0ₖ + + helper : Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , (f x +ₖ g x)))) ≡ fun + helper = funExt λ {(x , north) → refl ; (x , south) → refl ; (x , (merid north i)) → refl ; (x , (merid south i)) → refl} + + helper2 : ∣ Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , (f x +ₖ g x)))) ∣₀ + ≡ (∣ Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , f x))) ∣₀ +ₕ ∣ Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , g x))) ∣₀) + helper2 = + cong ∣_∣₀ + (funExt λ {(x , north) → ((λ i → (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x +ₖ g x) i)) + ∙∙ sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ cong (λ x → ((0ₖ +ₖ 0ₖ) +ₖ x)) (rUnitₖ 0ₖ) + ∙∙ refl) + ; (x , south) → refl + ∙∙ sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ cong (λ x → ((0ₖ +ₖ 0ₖ) +ₖ x)) (rUnitₖ 0ₖ) + ∙∙ (λ i → (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x) i) +ₖ ∣ north ∣ +ₖ Kn→ΩKn+1 1 (g x) i) + ; (x , (merid south i)) j → hcomp (λ k → λ {(j = i0) → ∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x +ₖ g x) (i ∨ (~ k)) + ; (j = i1) → (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x) (i ∧ k)) +ₖ ∣ north ∣ +ₖ Kn→ΩKn+1 1 (g x) (i ∧ k)}) + ((sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ cong (λ x → ((0ₖ +ₖ 0ₖ) +ₖ x)) (rUnitₖ 0ₖ)) j) + ; (x , (merid north u)) → {!!}}) + where + anotherone : (x : _) → cong (0ₖ +ₖ_) (Kn→ΩKn+1 1 (f x +ₖ g x)) ≡ {!!} + anotherone x = {!!} + + helper5 : (x : _) → Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , f x +ₖ g x))) x + ≡ (Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , f x))) x) +ₖ (Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , g x))) x) + helper5 (x , north) = sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i))) + helper5 (x , south) = sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i))) + helper5 (x , merid north i) = sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i))) + helper5 (x , merid south i) j = + {!!} + where + helper13 : sym (sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i)))) + ∙ (λ i → ∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x +ₖ g x) i) ∙ (sym (rUnitₖ (0ₖ +ₖ 0ₖ)) + ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i)))) + ≡ λ i → (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x) i) +ₖ + ∣ north ∣ +ₖ Kn→ΩKn+1 1 (g x) i + helper13 = (λ j → sym ((λ i → (rUnitₖ (lUnitₖ 0ₖ (j ∧ (~ i))) (~ i))) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ lUnitₖ 0ₖ (~ i))) + ∙ lUnit (rUnit ((λ i → lUnitₖ (Kn→ΩKn+1 1 (f x +ₖ g x) i) j)) j) j + ∙ ((λ i → (rUnitₖ (lUnitₖ 0ₖ (j ∧ (~ i))) (~ i)))) ∙ λ i → (0ₖ +ₖ 0ₖ) +ₖ lUnitₖ 0ₖ (~ i)) + ∙ (λ j → sym ((λ i → (rUnitₖ (lUnitₖ 0ₖ (~ i)) (~ i))) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ lUnitₖ 0ₖ (~ i))) ∙ + {!!} ∙ + (λ i → (rUnitₖ (lUnitₖ 0ₖ (~ i)) (~ i))) ∙ λ i → (0ₖ +ₖ 0ₖ) +ₖ lUnitₖ 0ₖ (~ i)) + ∙ {!!} + ∙ {!!} +{- +i = i0 ⊢ (sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i)))) j +i = i1 ⊢ (sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i)))) j +j = i0 ⊢ ∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x +ₖ g x) i +j = i1 ⊢ (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x) i) +ₖ + ∣ north ∣ +ₖ Kn→ΩKn+1 1 (g x) i +-} + + + helper' : Iso Int (Unit × Int) + Iso.inv helper' = proj₂ + Iso.fun helper' x = tt , x + Iso.leftInv helper' x = refl + Iso.rightInv helper' x = ×≡ refl refl + + helper : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} → Iso C Unit → Iso A B → Iso (A × C) B + Iso.fun (helper cUnit iAB) x = Iso.fun iAB (proj₁ x) + Iso.inv (helper cUnit iAB) x = (Iso.inv iAB x) , (Iso.inv cUnit tt ) + Iso.rightInv (helper cUnit iAB) = Iso.rightInv iAB + Iso.leftInv (helper cUnit iAB) b = ×≡ (Iso.leftInv iAB (proj₁ b)) (Iso.leftInv cUnit (proj₂ b)) + + helper2 : Iso (coHom 2 ((S₊ 1) × (S₊ 1))) (coHom 2 ((S₊ 1) × hLevelTrunc 3 (S₊ 1))) + Iso.fun helper2 = sRec setTruncIsSet (λ f → ∣ (λ {(x , y) → f (x , trRec isGroupoidS1 (idfun (S₊ 1)) y)}) ∣₀) + Iso.inv helper2 = sRec setTruncIsSet (λ f → ∣ (λ {(x , y) → f (x , ∣ y ∣)}) ∣₀) + Iso.rightInv helper2 = + sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + (λ f → cong ∣_∣₀ + (funExt λ {(x , y) → + trElim {B = λ y → f (x , ∣ trRec isGroupoidS1 (λ x₁ → x₁) y ∣) ≡ f (x , y)} + (λ _ → isOfHLevelPath' 3 (isOfHLevelTrunc 4) _ _) + (λ a → refl) y})) + Iso.leftInv helper2 = + sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ (funExt λ {(x , y) → refl}) + +-- basechange* : (x y : S¹) → x ≡ y → x ≡ y → ΩS¹ +-- basechange* x y = J (λ y p → (x ≡ y) → ΩS¹) (basechange x) + + +-- test1 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S₊ 1 × Int) +-- Iso.fun test1 f = (trRec isGroupoidS1 (λ a → a) (f north)) +-- , winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (f (loop* i))))) +-- Iso.inv test1 (north , b) x = ∣ x ∣ +-- Iso.inv test1 (south , b) x = ∣ x ∣ +-- Iso.inv test1 (merid a i , b) x = {!!} +-- Iso.rightInv test1 = {!!} +-- Iso.leftInv test1 = {!!} + +-- funRec : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) {C : (A → hLevelTrunc n B) → Type ℓ''} +-- → isOfHLevel n B +-- → ((f : A → B) → C (λ a → ∣ f a ∣)) +-- → (f : A → hLevelTrunc n B) → C f +-- funRec {A = A} {B = B} n {C = C} hLev ind f = subst C (helper f) (ind (λ a → trRec hLev (λ x → x) (f a))) +-- where +-- helper : retract {A = A → hLevelTrunc n B} {B = A → B} (λ f₁ a → trRec hLev (λ x → x) (f₁ a)) (λ f₁ a → ∣ f₁ a ∣) +-- helper f = funExt λ a → helper2 (f a) +-- where +-- helper2 : (x : hLevelTrunc n B) → ∣ trRec hLev (λ x → x) x ∣ ≡ x +-- helper2 = trElim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → refl + +-- invMapSurj : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (ϕ : morph G H) → ((x : Group.type H) → isInIm G H (fst ϕ) x) +-- → morph H G +-- fst (invMapSurj G H (ϕ , pf) surj) a = {!pRec!} +-- snd (invMapSurj G H (ϕ , pf) surj) = {!!} + +{- +ImIso : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (ϕ : morph G H) → ((x : Group.type H) → isInIm G H (fst ϕ) x) + → grIso H (imGroup G H ϕ) +ImIso G H (ϕ , mf) surj = + let idH = isGroup.id (Group.groupStruc H) + idG = isGroup.id (Group.groupStruc G) + _+G_ = isGroup.comp (Group.groupStruc G) + _+H_ = isGroup.comp (Group.groupStruc H) + _+Im_ = isGroup.comp (Group.groupStruc (imGroup G H (ϕ , mf))) + invG = isGroup.inv (Group.groupStruc G) + invH = isGroup.inv (Group.groupStruc H) + lUnit = isGroup.lUnit (Group.groupStruc H) + lCancel = isGroup.rCancel (Group.groupStruc H) + in + Iso''→Iso _ _ + (iso'' ((λ x → x , pRec propTruncIsProp (λ (a , b) → ∣ a , b ∣₋₁) (surj x)) + , λ a b → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) + (λ surja → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) + (λ surjb → + pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) + (λ surja+b → + (λ i → (a +H b) , (pRec (propTruncIsProp) + (λ (a , b) → ∣ a , b ∣₋₁) + (propTruncIsProp (surj (isGroup.comp (Group.groupStruc H) a b)) ∣ surja+b ∣₋₁ i))) ∙ + (λ i → (a +H b) , ∣ (fst surja+b) , (snd surja+b) ∣₋₁) ∙ + ΣProp≡ (λ _ → propTruncIsProp) refl ∙ + λ i → (a +H b) , pRec (propTruncIsProp) + (λ p1 → + pRec (λ x y → squash x y) + (λ p2 → + ∣ + isGroup.comp (Group.groupStruc G) (fst p1) (fst p2) , + mf (fst p1) (fst p2) ∙ + cong₂ (isGroup.comp (Group.groupStruc H)) (snd p1) (snd p2) + ∣₋₁) + (pRec (propTruncIsProp) + ∣_∣₋₁ (propTruncIsProp ∣ surjb ∣₋₁ (surj b) i))) + (pRec (propTruncIsProp) + ∣_∣₋₁ (propTruncIsProp ∣ surja ∣₋₁ (surj a) i ))) + (surj (isGroup.comp (Group.groupStruc H) a b))) + (surj b)) + (surj a)) + (λ x inker → cong fst inker) + λ x → pRec propTruncIsProp (λ inimx → ∣ (ϕ (inimx .fst)) , ΣProp≡ (λ _ → propTruncIsProp) (inimx .snd) ∣₋₁) (snd x)) +-} + + +{- +H¹-S¹≃Int : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +H¹-S¹≃Int = + Iso''→Iso _ _ + (iso'' ((λ x → ∣ theFuns x ∣₀) , λ a b → cong ∣_∣₀ (funExt λ x → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 1) _) ∙ sym (cong (ΩKn+1→Kn) (theFunsId2 a b x)))) + (λ x inker → pRec (isSetInt _ _) (inj x) (Iso.fun PathIdTrunc₀Iso inker)) + finIm) + where + d : _ + d = MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 + + i : _ + i = MV.i _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1 + + Δ : _ + Δ = MV.Δ _ _ (S₊ 1) (λ _ → tt) (λ _ → tt) 0 + + + d-surj : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) + → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x + d-surj = sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ x → MV.Ker-i⊂Im-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ + (sym (isContrHelper .snd _)) + where + isContrHelper : isContr (Group.type (×coHomGr 1 Unit Unit)) + isContrHelper = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) + , λ y → ×≡ (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₁ y)) + (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₂ y)) + + theFuns : (a : Int) → Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 + theFuns a (inl x) = 0ₖ + theFuns a (inr x) = 0ₖ + theFuns a (push north i) = Kn→ΩKn+1 0 a i + theFuns a (push south i) = 0ₖ + + + theFunsId2 : (a b : Int) (x : Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) + → Kn→ΩKn+1 1 (theFuns a x) ∙ Kn→ΩKn+1 1 (theFuns b x) ≡ Kn→ΩKn+1 1 (theFuns (a +ℤ b) x) + theFunsId2 a b (inl x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) + theFunsId2 a b (inr x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) + theFunsId2 a b (push north i) j = + hcomp (λ k → λ {(i = i0) → ((λ i₁ → + (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) + j + ; (i = i1) → ((λ i₁ → + (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) + j + ; (j = i0) → cong₂Funct (λ p q → Kn→ΩKn+1 1 p ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b) (~ k) i + ; (j = i1) → (helper2 a b) k i }) + (hcomp (λ k → λ { (j = i0) → compPath-filler (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) k i + ; (j = i1) → compPath-filler (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) k i + ; (i = i1) → RHS-filler b j k + ; (i = i0) → ((λ i₁ → + (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) + j}) + (bottom-filler a j i)) + + where + + bottom-filler : (a : Int) → + PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) + (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) + ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) + j ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) + (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) + ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) + j) (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) + bottom-filler a j i = + hcomp (λ k → λ {(j = i0) → helper2 (~ k) i ; + (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 a) i}) + ((λ j₂ → ∣ rCancel (merid north) j j₂ ∣) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + + where + helper2 : Path (Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣ ≡ Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣) + (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i) ∙ Kn→ΩKn+1 1 ∣ north ∣) + (λ i → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + helper2 = cong₂Sym1 (Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) (~ i) j ∣) (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + + RHS-filler : (b : Int) → + PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j + ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j) + (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) + (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) + RHS-filler b j i = + hcomp (λ k → λ {(j = i0) → cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i ; + (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 b) i}) + (cong (λ q → (λ i → ∣ rCancel (merid north) j i ∣) ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i) + + helper2 : (a b : Int) → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b) ≡ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ℤ b)) + helper2 a b = + sym (congFunct (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b)) + ∙ (λ i → cong (Kn→ΩKn+1 1) (Iso.rightInv (Iso3-Kn-ΩKn+1 0) (Kn→ΩKn+1 0 a ∙ Kn→ΩKn+1 0 b) (~ i))) + ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ₖ b)) ) + ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (addLemma a b i))) + + theFunsId2 a b (push south i) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) + ∙ sym (lUnit _) + + inj : (a : Int) → theFuns a ≡ (λ _ → ∣ north ∣) → a ≡ pos 0 + inj a id = + pRec (isSetInt _ _) + (sigmaElim (λ _ → isOfHLevelPath 2 isSetInt _ _) + (λ a p → pRec (isSetInt _ _) + (λ id2 → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 0) _) + ∙ cong (ΩKn+1→Kn) + (PathP→compPathR + (cong (λ f → cong f (push north)) id) + ∙ test)) + (Iso.fun PathIdTrunc₀Iso p))) (d-surj ∣ theFuns a ∣₀) + where + + test : (λ i → id i (inl tt)) ∙ (λ i → ∣ north ∣) ∙ sym (λ i → id i (inr tt)) ≡ refl + test = (λ i → cong (λ f → f (inl tt)) id + ∙ lUnit (sym (cong (λ f → f (inr tt)) id)) (~ i)) + ∙ (λ i → cong (λ f → f (push south i)) id + ∙ sym (cong (λ f → f (inr tt)) id)) + ∙ rCancel (cong (λ f → f (inr tt)) id) + + + consMember : (a : Int) → coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) + consMember a = d ∣ (λ _ → a) ∣₀ + + consMember≡0 : (a : Int) → consMember a ≡ 0ₕ + consMember≡0 a = + (MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ (λ _ → a) ∣₀ ∣ + (∣ (λ _ → a) ∣₀ , ∣ (λ _ → 0) ∣₀) + , cong ∣_∣₀ (λ i x → (rUnitₖ a i)) ∣₋₁) + + f+consMember' : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int × Int ] (f +ₕ (-ₕ (consMember (proj₁ x))) ≡ ∣ theFuns (proj₂ x) ∣₀) + f+consMember' = + sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ f → pRec propTruncIsProp + (sigmaElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) + (λ g id → ∣ ((g south) , ((g north) +ₖ (-ₖ g south))) + , (pRec (setTruncIsSet _ _) + (λ id → (λ i → ∣ id (~ i) ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀) ∙ funId1 g) + (Iso.fun PathIdTrunc₀Iso id)) ∣₋₁)) + (d-surj ∣ f ∣₀) + where + funId1 : (g : S₊ 0 → Int) + → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀ ≡ + ∣ theFuns ((g north) +ₖ (-ₖ (g south))) ∣₀ + funId1 g = (λ i → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ + +ₕ (morphMinus (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) d + (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ∣ (λ _ → g south) ∣₀ (~ i))) + ∙ sym (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ g ∣₀ (-ₕ ∣ (λ _ → g south) ∣₀)) + ∙ (cong (λ x → d ∣ x ∣₀) g'Id) + ∙ cong ∣_∣₀ helper + where + g' : S₊ 0 → Int + g' north = (g north) +ₖ (-ₖ (g south)) + g' south = 0 + + g'Id : (λ x → g x +ₖ (-ₖ (g south))) ≡ g' + g'Id = funExt (λ {north → refl + ; south → rCancelₖ (g south)}) + + helper : MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g' ≡ theFuns (g north +ₖ (-ₖ g south)) + helper = funExt λ {(inl tt) → refl + ; (inr tt) → refl + ; (push north i) → refl + ; (push south i) → refl} + finIm : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int ] (∣ theFuns x ∣₀ ≡ f) + finIm f = + pRec propTruncIsProp + (λ {((a , b) , id) → ∣ b , (sym id ∙ cong (λ x → f +ₕ x) ((λ i → (-ₕ (consMember≡0 a i))) ∙ sym (lUnitₕ (-ₕ 0ₕ)) ∙ rCancelₕ 0ₕ) ∙ (rUnitₕ f)) ∣₋₁}) + (f+consMember' f) +-} + + + +-- Hⁿ-Sⁿ≃Int : (n : ℕ) → grIso intGroup (coHomGr (suc n) (S₊ (suc n))) +-- Hⁿ-Sⁿ≃Int zero = +-- compGrIso {F = intGroup} {G = {!!}} {H = {!coHomGr 1 (S₊ 1)!}} +-- (Iso''→Iso +-- (iso'' ((λ x → ∣ theFuns x ∣₀) , λ a b → cong ∣_∣₀ (funExt λ x → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 1) _) ∙ sym (cong (ΩKn+1→Kn) (theFunsId2 a b x)))) +-- (λ x inker → pRec (isSetInt _ _) (inj x) (Iso.fun PathIdTrunc₀Iso inker)) +-- finIm)) +-- {!invGrIso _ _ (coHomPushout≡coHomSn 0 1)!} +-- where +-- d : _ +-- d = MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 + +-- i : _ +-- i = MV.i _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1 + +-- Δ : _ +-- Δ = MV.Δ _ _ (S₊ 1) (λ _ → tt) (λ _ → tt) 0 + + +-- d-surj : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x +-- d-surj = sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- λ x → MV.Ker-i⊂Im-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ +-- (sym (isContrHelper .snd _)) +-- where +-- isContrHelper : isContr (Group.type (×coHomGr 1 Unit Unit)) +-- isContrHelper = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) +-- , λ y → ×≡ (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₁ y)) +-- (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₂ y)) + +-- theFuns : (a : Int) → Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 +-- theFuns a (inl x) = 0ₖ +-- theFuns a (inr x) = 0ₖ +-- theFuns a (push north i) = Kn→ΩKn+1 0 a i +-- theFuns a (push south i) = 0ₖ + + +-- theFunsId2 : (a b : Int) (x : Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) +-- → Kn→ΩKn+1 1 (theFuns a x) ∙ Kn→ΩKn+1 1 (theFuns b x) ≡ Kn→ΩKn+1 1 (theFuns (a +ℤ b) x) +-- theFunsId2 a b (inl x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) +-- theFunsId2 a b (inr x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) +-- theFunsId2 a b (push north i) j = +-- hcomp (λ k → λ {(i = i0) → ((λ i₁ → +-- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) +-- j +-- ; (i = i1) → ((λ i₁ → +-- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) +-- j +-- ; (j = i0) → cong₂Funct (λ p q → Kn→ΩKn+1 1 p ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b) (~ k) i +-- ; (j = i1) → (helper2 a b) k i }) +-- (hcomp (λ k → λ { (j = i0) → compPath-filler (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) k i +-- ; (j = i1) → compPath-filler (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) k i +-- ; (i = i1) → RHS-filler b j k +-- ; (i = i0) → ((λ i₁ → +-- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) +-- j}) +-- (bottom-filler a j i)) + +-- where + +-- bottom-filler : (a : Int) → +-- PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) +-- ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) +-- j ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) +-- ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) +-- j) (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) +-- bottom-filler a j i = +-- hcomp (λ k → λ {(j = i0) → helper2 (~ k) i ; +-- (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 a) i}) +-- ((λ j₂ → ∣ rCancel (merid north) j j₂ ∣) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + +-- where +-- helper2 : Path (Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣ ≡ Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- (λ i → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) +-- helper2 = cong₂Sym1 (Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) (~ i) j ∣) (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) + +-- RHS-filler : (b : Int) → +-- PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j +-- ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j) +-- (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) +-- (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) +-- RHS-filler b j i = +-- hcomp (λ k → λ {(j = i0) → cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i ; +-- (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 b) i}) +-- (cong (λ q → (λ i → ∣ rCancel (merid north) j i ∣) ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i) + +-- helper2 : (a b : Int) → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b) ≡ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ℤ b)) +-- helper2 a b = +-- sym (congFunct (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b)) +-- ∙ (λ i → cong (Kn→ΩKn+1 1) (Iso.rightInv (Iso3-Kn-ΩKn+1 0) (Kn→ΩKn+1 0 a ∙ Kn→ΩKn+1 0 b) (~ i))) +-- ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ₖ b)) ) +-- ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (addLemma a b i))) + +-- theFunsId2 a b (push south i) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) +-- ∙ sym (lUnit _) + +-- inj : (a : Int) → theFuns a ≡ (λ _ → ∣ north ∣) → a ≡ pos 0 +-- inj a id = +-- pRec (isSetInt _ _) +-- (sigmaElim (λ _ → isOfHLevelPath 2 isSetInt _ _) +-- (λ a p → pRec (isSetInt _ _) +-- (λ id2 → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 0) _) +-- ∙ cong (ΩKn+1→Kn) +-- (PathP→compPathR +-- (cong (λ f → cong f (push north)) id) +-- ∙ test)) +-- (Iso.fun PathIdTrunc₀Iso p))) (d-surj ∣ theFuns a ∣₀) +-- where + +-- test : (λ i → id i (inl tt)) ∙ (λ i → ∣ north ∣) ∙ sym (λ i → id i (inr tt)) ≡ refl +-- test = (λ i → cong (λ f → f (inl tt)) id +-- ∙ lUnit (sym (cong (λ f → f (inr tt)) id)) (~ i)) +-- ∙ (λ i → cong (λ f → f (push south i)) id +-- ∙ sym (cong (λ f → f (inr tt)) id)) +-- ∙ rCancel (cong (λ f → f (inr tt)) id) + + +-- consMember : (a : Int) → coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) +-- consMember a = d ∣ (λ _ → a) ∣₀ + +-- consMember≡0 : (a : Int) → consMember a ≡ 0ₕ +-- consMember≡0 a = +-- (MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ (λ _ → a) ∣₀ ∣ +-- (∣ (λ _ → a) ∣₀ , ∣ (λ _ → 0) ∣₀) +-- , cong ∣_∣₀ (λ i x → (rUnitₖ a i)) ∣₋₁) + +-- f+consMember' : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int × Int ] (f +ₕ (-ₕ (consMember (proj₁ x))) ≡ ∣ theFuns (proj₂ x) ∣₀) +-- f+consMember' = +-- sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- λ f → pRec propTruncIsProp +-- (sigmaElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- (λ g id → ∣ ((g south) , ((g north) +ₖ (-ₖ g south))) +-- , (pRec (setTruncIsSet _ _) +-- (λ id → (λ i → ∣ id (~ i) ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀) ∙ funId1 g) +-- (Iso.fun PathIdTrunc₀Iso id)) ∣₋₁)) +-- (d-surj ∣ f ∣₀) +-- where +-- funId1 : (g : S₊ 0 → Int) +-- → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀ ≡ +-- ∣ theFuns ((g north) +ₖ (-ₖ (g south))) ∣₀ +-- funId1 g = (λ i → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +-- +ₕ (morphMinus (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) d +-- (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ∣ (λ _ → g south) ∣₀ (~ i))) +-- ∙ sym (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ g ∣₀ (-ₕ ∣ (λ _ → g south) ∣₀)) +-- ∙ (cong (λ x → d ∣ x ∣₀) g'Id) +-- ∙ cong ∣_∣₀ helper +-- where +-- g' : S₊ 0 → Int +-- g' north = (g north) +ₖ (-ₖ (g south)) +-- g' south = 0 + +-- g'Id : (λ x → g x +ₖ (-ₖ (g south))) ≡ g' +-- g'Id = funExt (λ {north → refl +-- ; south → rCancelₖ (g south)}) + +-- helper : MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g' ≡ theFuns (g north +ₖ (-ₖ g south)) +-- helper = funExt λ {(inl tt) → refl +-- ; (inr tt) → refl +-- ; (push north i) → refl +-- ; (push south i) → refl} +-- finIm : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int ] (∣ theFuns x ∣₀ ≡ f) +-- finIm f = +-- pRec propTruncIsProp +-- (λ {((a , b) , id) → ∣ b , (sym id ∙ cong (λ x → f +ₕ x) ((λ i → (-ₕ (consMember≡0 a i))) ∙ sym (lUnitₕ (-ₕ 0ₕ)) ∙ rCancelₕ 0ₕ) ∙ (rUnitₕ f)) ∣₋₁}) +-- (f+consMember' f) +-- Hⁿ-Sⁿ≃Int (suc n) = +-- compGrIso (Hⁿ-Sⁿ≃Int n) +-- (transport (λ i → grIso {!!} {!coHomGr (suc (suc n)) (Pushout (λ _ → tt) (λ _ → tt))!}) {!!}) + + +-- {- +-- coHom1S1≃ℤ : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- grIso.fun coHom1S1≃ℤ = grIso.fun {!compGrIso coHom1Iso (invGrIso _ _ (ImIso _ _ (d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ?))!} +-- grIso.inv coHom1S1≃ℤ = {!!} +-- grIso.rightInv coHom1S1≃ℤ = {!!} +-- grIso.leftInv coHom1S1≃ℤ = {!!} +-- -} + +-- -- compGrIso coHom1Iso (invGrIso _ _ (ImIso _ _ {!d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0!} {!!})) + + +-- -- coHomGrIm : grIso (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- (imGroup (coHomGr 0 (S₊ 0)) +-- -- (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 +-- -- , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) +-- -- coHomGrIm = +-- -- ImIso _ +-- -- _ +-- -- _ +-- -- {!!} + + +-- -- -- coHom1RedS1 : Iso (coHom 1 (S₊ 1)) (coHomRed 1 (S₊ 1 , north)) +-- -- -- Iso.fun coHom1RedS1 = sRec setTruncIsSet λ f → ∣ f , (pRec {!!} {!!} ((transport (λ i → (b : truncIdempotent 3 {!S₊ 1!} (~ i)) → ∥ (transp (λ j → truncIdempotent {!3!} {!!} (~ i ∨ (~ j))) (~ i) north) ≡ b ∥₋₁) isConnectedS1) (f north)) ) ∣₀ +-- -- -- Iso.inv coHom1RedS1 = {!!} +-- -- -- Iso.rightInv coHom1RedS1 = {!setTruncIdempotent!} +-- -- -- Iso.leftInv coHom1RedS1 = {!!} + +-- -- -- coHom1Red : ∀ {ℓ} (A : Pointed ℓ) → Iso (coHom 1 (typ A)) (coHomRed 1 A) +-- -- -- Iso.fun (coHom1Red A) = sRec setTruncIsSet λ f → ∣ f , {!!} ∣₀ +-- -- -- Iso.inv (coHom1Red A) = {!!} +-- -- -- Iso.rightInv (coHom1Red A) = {!!} +-- -- -- Iso.leftInv (coHom1Red A) = {!!} + +-- -- -- -- morphtest : morph (coHomGr 1 (S₊ 1)) intGroup +-- -- -- -- fst morphtest = sRec isSetInt λ f → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (f (loop* i))))) +-- -- -- -- snd morphtest = sElim2 {!!} +-- -- -- -- (funRec 3 isGroupoidS1 +-- -- -- -- λ f → funRec 3 isGroupoidS1 +-- -- -- -- λ g → pRec (isSetInt _ _) +-- -- -- -- (λ n=fn → +-- -- -- -- pRec (isSetInt _ _) +-- -- -- -- (λ n=gn → (λ j → winding (basechange (SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (∣ f (north) ∣ +ₖ ∣ n=gn (~ j) ∣)))) (λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (∣ f (loop* i) ∣ +ₖ ∣ transp (λ i → n=gn ((~ i) ∨ (~ j)) ≡ n=gn ((~ i) ∨ (~ j))) (~ j) (λ i → g (loop* i)) i ∣)))))) +-- -- -- -- ∙ {!!} +-- -- -- -- ∙ {!!}) +-- -- -- -- (isConnectedS1 (g north))) +-- -- -- -- (isConnectedS1 (f north))) + + +-- -- -- -- {- (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (∣ f (loop* i) ∣ +ₖ ∣ g (loop* i) ∣))))) +-- -- -- -- ∙ (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (Kn→ΩKn+1 1 ∣ f (loop* i) ∣ ∙ Kn→ΩKn+1 1 ∣ g (loop* i) ∣)))))) +-- -- -- -- ∙ (λ j → winding (basechange _ (cong (λ x → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (Kn→ΩKn+1 1 ∣ f x ∣ ∙ Kn→ΩKn+1 1 ∣ g x ∣))))) loop*)) ) +-- -- -- -- ∙ (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn ((cong ∣_∣ (merid (f (loop* i)) ∙ sym (merid north)) ∙ cong ∣_∣ (merid (g (loop* i)) ∙ sym (merid north))))))))) +-- -- -- -- ∙ (λ j → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (congFunct ∣_∣ (merid (f (loop* i)) ∙ sym (merid north)) (merid (g (loop* i)) ∙ sym (merid north)) (~ j))))))) +-- -- -- -- ∙ (λ j → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (cong ∣_∣ (({!!} ∙ {!!}) ∙ {!!}))))))) +-- -- -- -- ∙ {!!} +-- -- -- -- ∙ {!!} +-- -- -- -- ∙ {!!}) -} + +-- -- -- -- where +-- -- -- -- helper : ∀ {ℓ} {A : Type ℓ} (a : A) (f : A → S¹) (p q : a ≡ a) → winding (basechange (f a) (cong f (p ∙ q))) ≡ winding (basechange (f a) (cong f p ∙ cong f q)) +-- -- -- -- helper a f p q i = winding (basechange (f a) (congFunct f p q i)) +-- -- -- -- helper2 : (x : S¹) (p q : x ≡ x) → basechange x (p ∙ q) ≡ basechange x p ∙ basechange x q +-- -- -- -- helper2 base p q = refl +-- -- -- -- helper2 (loop i) p q = {!!} +-- -- -- -- helper4 : (x y z : S¹) (p : x ≡ z) (q r : x ≡ y) (s : y ≡ z) → basechange* x z p (q ∙ s) ≡ basechange* x y {!!} q ∙ {!!} +-- -- -- -- helper4 = {!!} +-- -- -- -- helper3 : (p q : ΩS¹) → winding (p ∙ q) ≡ (winding p +ℤ winding q) +-- -- -- -- helper3 = {!!} + + +-- -- -- -- -- fstmap : morph (dirProd intGroup intGroup) (coHomGr 0 (S₊ 0)) +-- -- -- -- -- fstmap = compMorph {F = dirProd intGroup intGroup} {G = ×coHomGr 0 Unit Unit} {H = coHomGr 0 (S₊ 0)} +-- -- -- -- -- (×morph (grIso.inv coHomGrUnit0) (grIso.inv coHomGrUnit0)) +-- -- -- -- -- (((MV.Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) zero)) , +-- -- -- -- -- {!MV.ΔIsHom _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) zero!}) + +-- -- -- -- -- fstMapId : (a : Int × Int) → fstmap .fst a ≡ ∣ (λ _ → proj₁ a +ℤ (0 - proj₂ a)) ∣₀ +-- -- -- -- -- fstMapId (a , b) = (λ i → ∣ (λ _ → a +ₖ (-ₖ b)) ∣₀) ∙ {!addLemma!} ∙ {!!} ∙ {!!} + +-- -- -- -- -- isoAgain : grIso intGroup (coHomGr 1 (S₊ 1)) +-- -- -- -- -- isoAgain = +-- -- -- -- -- Iso''→Iso _ _ +-- -- -- -- -- (iso'' ((λ a → ∣ (λ {north → 0ₖ ; south → 0ₖ ; (merid north i) → {!a!} ; (merid south i) → {!!}}) ∣₀) , {!!}) +-- -- -- -- -- {!!} +-- -- -- -- -- {!!}) + +-- -- -- -- -- -- -- test2 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → S¹) +-- -- -- -- -- -- -- Iso.fun test2 f = {!!} +-- -- -- -- -- -- -- Iso.inv test2 f north = ∣ transport (sym S¹≡S1) (f base) ∣ +-- -- -- -- -- -- -- Iso.inv test2 f south = ∣ transport (sym S¹≡S1) (f base) ∣ +-- -- -- -- -- -- -- Iso.inv test2 f (merid a i) = cong ∣_∣ {!transport (sym S¹≡S1) (f base)!} i +-- -- -- -- -- -- -- Iso.rightInv test2 = {!!} +-- -- -- -- -- -- -- Iso.leftInv test2 = {!!} + +-- -- -- -- -- -- -- F : winding (basechange base loop) ≡ 1 +-- -- -- -- -- -- -- F = refl + +-- -- -- -- -- -- -- another : (f g : Int) → winding (basechange {!!} {!!}) ≡ {!!} +-- -- -- -- -- -- -- another = {!!} + +-- -- -- -- -- -- -- test : Iso (S¹ → S¹) (S¹ × Int) +-- -- -- -- -- -- -- Iso.fun test f = f base , winding (basechange (f base) (cong f loop)) +-- -- -- -- -- -- -- Iso.inv test (x , int) base = x +-- -- -- -- -- -- -- Iso.inv test (x , int) (loop i) = {!!} +-- -- -- -- -- -- -- Iso.rightInv test = {!!} +-- -- -- -- -- -- -- Iso.leftInv test = {!!} + +-- -- -- -- -- -- -- -- test13 : Iso ∥ (S¹ → S¹) ∥₀ Int +-- -- -- -- -- -- -- -- Iso.fun test13 = sRec isSetInt λ f → winding (basechange (f base) (λ i → f (loop i))) +-- -- -- -- -- -- -- -- Iso.inv test13 a = ∣ (λ {base → {!!} ; (loop i) → {!!}}) ∣₀ +-- -- -- -- -- -- -- -- Iso.rightInv test13 = {!!} +-- -- -- -- -- -- -- -- Iso.leftInv test13 = {!!} + +-- -- -- -- -- -- -- -- test : Iso (S¹ → S¹) (S¹ × Int) +-- -- -- -- -- -- -- -- Iso.fun test f = (f base) , transport (basedΩS¹≡Int (f base)) λ i → f (loop i) +-- -- -- -- -- -- -- -- Iso.inv test (x , int) base = x +-- -- -- -- -- -- -- -- Iso.inv test (x , int) (loop i) = transport (sym (basedΩS¹≡Int x)) int i +-- -- -- -- -- -- -- -- Iso.rightInv test (x , int) i = (x , transportTransport⁻ (basedΩS¹≡Int x) int i) +-- -- -- -- -- -- -- -- Iso.leftInv test f = +-- -- -- -- -- -- -- -- funExt λ { base → refl +-- -- -- -- -- -- -- -- ; (loop i) j → transport⁻Transport (basedΩS¹≡Int (f base)) (λ i → f (loop i)) j i} + + +-- -- -- -- -- -- -- -- lem : S¹ ≡ hLevelTrunc 3 (S₊ 1) +-- -- -- -- -- -- -- -- lem = sym (truncIdempotent 3 isGroupoidS¹) ∙ λ i → hLevelTrunc 3 (S¹≡S1 (~ i)) + +-- -- -- -- -- -- -- -- ×≡ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (a b : A × B) → proj₁ a ≡ proj₁ b → proj₂ a ≡ proj₂ b → a ≡ b +-- -- -- -- -- -- -- -- ×≡ (_ , _) (_ , _) id1 id2 i = (id1 i) , (id2 i) + +-- -- -- -- -- -- -- -- test22 : Iso (S₊ 1 → coHomK 1) (S₊ 1 × Int) +-- -- -- -- -- -- -- -- Iso.fun test22 f = {!f north!} , {!!} +-- -- -- -- -- -- -- -- Iso.inv test22 = {!!} +-- -- -- -- -- -- -- -- Iso.rightInv test22 = {!!} +-- -- -- -- -- -- -- -- Iso.leftInv test22 = {!!} + + + + + + +-- -- -- -- -- -- -- -- coHom1≃∥S1×ℤ∥₀ : Iso (coHom 1 (S₊ 1)) ∥ S₊ 1 × Int ∥₀ +-- -- -- -- -- -- -- -- coHom1≃∥S1×ℤ∥₀ = setTruncIso test2 +-- -- -- -- -- -- -- -- where +-- -- -- -- -- -- -- -- test2 : Iso (S₊ 1 → coHomK 1) (S₊ 1 × Int) +-- -- -- -- -- -- -- -- Iso.fun test2 f = transport (λ i → S¹≡S1 (~ i) × Int) (Iso.fun test (transport (λ i → (S¹≡S1 i → lem (~ i))) f)) +-- -- -- -- -- -- -- -- Iso.inv test2 x = transport (λ i → (S¹≡S1 (~ i) → lem i)) (Iso.inv test (transport (λ i → S¹≡S1 i × Int) x)) +-- -- -- -- -- -- -- -- Iso.rightInv test2 (s , int) = ×≡ _ _ {!!} {!!} +-- -- -- -- -- -- -- -- Iso.leftInv test2 f = {!!} ∙ {!!} ∙ {!!} + +-- -- -- -- -- -- -- -- test2Id : (a b : (S₊ 1 → coHomK 1)) → proj₂ (Iso.fun test2 (λ x → a x +ₖ b x)) ≡ (proj₂ (Iso.fun test2 a) +ₖ proj₂ (Iso.fun test2 a)) +-- -- -- -- -- -- -- -- test2Id a b = {! +-- -- -- -- -- -- -- -- transport +-- -- -- -- -- -- -- -- (basedΩS¹≡Int +-- -- -- -- -- -- -- -- (transport (λ i → S¹≡S1 i → lem (~ i)) (λ x → a x +ₖ b x) base)) +-- -- -- -- -- -- -- -- (λ i → +-- -- -- -- -- -- -- -- transport (λ i₁ → S¹≡S1 i₁ → lem (~ i₁)) (λ x → a x +ₖ b x) +-- -- -- -- -- -- -- -- (loop i))!} ∙ {!transport (λ i → S¹≡S1 i → lem (~ i)) (λ x → a x +ₖ b x) base!} + + +-- -- -- -- -- -- -- -- main : grIso intGroup (coHomGr 1 (S₊ 1)) +-- -- -- -- -- -- -- -- main = Iso'→Iso +-- -- -- -- -- -- -- -- (iso' {!!} +-- -- -- -- -- -- -- -- {!!}) + + +-- -- -- -- -- -- coHom1 : grIso (coHomGr 1 (S₊ 1)) intGroup +-- -- -- -- -- -- coHom1 = +-- -- -- -- -- -- Iso'→Iso +-- -- -- -- -- -- (iso' (iso ({!!} ∘ {!!} ∘ {!!} ∘ {!!}) +-- -- -- -- -- -- {!!} +-- -- -- -- -- -- {!!} +-- -- -- -- -- -- {!!}) +-- -- -- -- -- -- {!!}) + + + +-- -- -- -- -- -- schonf : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : (A × B) → Type ℓ''} +-- -- -- -- -- -- → ((a : A) (b : B) → C (a , b)) +-- -- -- -- -- -- → (x : A × B) → C x +-- -- -- -- -- -- schonf f (a , b) = f a b + +-- -- -- -- -- -- -- -- setTruncProdIso : ∀ {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') → Iso ∥ (A × B) ∥₀ (∥ A ∥₀ × ∥ B ∥₀) +-- -- -- -- -- -- -- -- Iso.fun (setTruncProdIso A B) = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) λ {(a , b) → ∣ a ∣₀ , ∣ b ∣₀} +-- -- -- -- -- -- -- -- Iso.inv (setTruncProdIso A B) (a , b) = sRec setTruncIsSet (λ a → sRec setTruncIsSet (λ b → ∣ a , b ∣₀) b) a +-- -- -- -- -- -- -- -- Iso.rightInv (setTruncProdIso A B) = +-- -- -- -- -- -- -- -- schonf (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- -- -- -- -- -- -- -- λ _ → sElim (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- -- -- -- -- -- -- -- λ _ → refl) +-- -- -- -- -- -- -- -- Iso.leftInv (setTruncProdIso A B) = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ {(a , b) → refl} + +-- -- -- -- -- -- -- -- setTruncProdLemma : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (a1 a2 : A) (b : B) → isHLevelConnected 2 A +-- -- -- -- -- -- -- -- → Path ∥ A × B ∥₀ ∣ a1 , b ∣₀ ∣ a2 , b ∣₀ +-- -- -- -- -- -- -- -- setTruncProdLemma {A = A} {B = B} a1 a2 b conA i = Iso.inv (setTruncProdIso A B) (Iso.inv setTruncTrunc0Iso ((sym (conA .snd ∣ a1 ∣) ∙ (conA .snd ∣ a2 ∣)) i) , ∣ b ∣₀) + +-- -- -- -- -- -- -- -- test3 : Iso ∥ S₊ 1 × Int ∥₀ Int +-- -- -- -- -- -- -- -- Iso.fun test3 = sRec isSetInt proj₂ +-- -- -- -- -- -- -- -- Iso.inv test3 a = ∣ north , a ∣₀ +-- -- -- -- -- -- -- -- Iso.rightInv test3 a = refl +-- -- -- -- -- -- -- -- Iso.leftInv test3 = +-- -- -- -- -- -- -- -- sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- -- -- -- -- -- -- -- λ {(s , int) → setTruncProdLemma north s int (sphereConnected 1)} + +-- -- -- -- -- -- -- -- coHomGr0-S1 : grIso intGroup (coHomGr 1 (S₊ 1)) +-- -- -- -- -- -- -- -- coHomGr0-S1 = +-- -- -- -- -- -- -- -- Iso'→Iso +-- -- -- -- -- -- -- -- (iso' (compIso (symIso test3) (symIso coHom1≃∥S1×ℤ∥₀)) +-- -- -- -- -- -- -- -- (indIntGroup {G = coHomGr 1 (S₊ 1)} +-- -- -- -- -- -- -- -- (Iso.fun (compIso (symIso test3) (symIso coHom1≃∥S1×ℤ∥₀))) +-- -- -- -- -- -- -- -- ((λ i → ∣ transport (λ i → (S¹≡S1 (~ i) → lem i)) (Iso.inv test (base , 0)) ∣₀) +-- -- -- -- -- -- -- -- ∙ (λ i → ∣ transport (λ i → (S¹≡S1 (~ i) → lem i)) (helper2 i) ∣₀) +-- -- -- -- -- -- -- -- ∙ cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → {!!}})) +-- -- -- -- -- -- -- -- {!!} +-- -- -- -- -- -- -- -- {!!})) +-- -- -- -- -- -- -- -- where +-- -- -- -- -- -- -- -- helper : basedΩS¹≡ΩS¹ base ≡ {!basechange!} +-- -- -- -- -- -- -- -- helper = {!substComposite!} + +-- -- -- -- -- -- -- -- substComposite2 : ∀ {ℓ} {A B C : Type ℓ} +-- -- -- -- -- -- -- -- (P : A ≡ B) (Q : B ≡ C) (a : A) +-- -- -- -- -- -- -- -- → transport (P ∙ Q) a ≡ transport Q (transport P a) +-- -- -- -- -- -- -- -- substComposite2 = {!!} + +-- -- -- -- -- -- -- -- helper1 : transport (λ i → S¹≡S1 i × Int) (north , 0) ≡ (base , 0) +-- -- -- -- -- -- -- -- helper1 = refl +-- -- -- -- -- -- -- -- helper3 : transport (sym (basedΩS¹≡Int base)) 0 ≡ refl +-- -- -- -- -- -- -- -- helper3 = (λ i → transport (symDistr (basedΩS¹≡ΩS¹ base) ΩS¹≡Int i) 0) +-- -- -- -- -- -- -- -- ∙ substComposite2 (sym ΩS¹≡Int) (sym (basedΩS¹≡ΩS¹ base)) 0 +-- -- -- -- -- -- -- -- ∙ (λ i → transport (λ i → basedΩS¹≡ΩS¹ base (~ i)) refl) -- +-- -- -- -- -- -- -- -- ∙ transportRefl ((equiv-proof (basechange-isequiv base) refl) .fst .fst) +-- -- -- -- -- -- -- -- ∙ (λ i → equiv-proof (transport (λ j → isEquiv (refl-conjugation j)) (basedΩS¹→ΩS¹-isequiv i0)) refl .fst .fst) +-- -- -- -- -- -- -- -- ∙ (λ i → {!equiv-proof (transport (λ j → isEquiv (refl-conjugation j)) (basedΩS¹→ΩS¹-isequiv i0)) refl .fst!}) +-- -- -- -- -- -- -- -- ∙ {!basedΩS¹→ΩS¹!} +-- -- -- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- -- -- helper4 : (x : S¹) → Iso.inv test (base , 0) x ≡ x +-- -- -- -- -- -- -- -- helper4 = {!!} +-- -- -- -- -- -- -- -- helper2 : Iso.inv test (transport (λ i → S¹≡S1 i × Int) (north , 0)) ≡ λ x → x +-- -- -- -- -- -- -- -- helper2 = (λ i → Iso.inv test (base , 0)) ∙ {!!} ∙ {!!} + +-- -- -- -- -- -- ×≡ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A × B} → proj₁ x ≡ proj₁ y → proj₂ x ≡ proj₂ y → x ≡ y +-- -- -- -- -- -- ×≡ {x = (a , b)} {y = (c , d)} id1 id2 i = (id1 i) , (id2 i) + +-- -- -- -- -- -- fstFunHelper : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- -- → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 _) (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x +-- -- -- -- -- -- fstFunHelper a = MV.Ker-i⊂Im-d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a +-- -- -- -- -- -- (sym (isContrH1Unit×H1Unit .snd _) ∙ (isContrH1Unit×H1Unit .snd _)) +-- -- -- -- -- -- where +-- -- -- -- -- -- isContrH1Unit×H1Unit : isContr (Group.type (×coHomGr 1 Unit Unit)) +-- -- -- -- -- -- isContrH1Unit×H1Unit = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) +-- -- -- -- -- -- , λ {(a , b) → sigmaProdElim {D = λ (x : Σ[ x ∈ Group.type (×coHomGr 1 Unit Unit)] Unit) → (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) ≡ fst x} +-- -- -- -- -- -- (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- -- -- -- -- -- (λ a b _ → ×≡ (grIso.leftInv (coHomGrUnit≥1 0) ∣ a ∣₀) (grIso.leftInv (coHomGrUnit≥1 0) ∣ b ∣₀)) ((a , b) , tt)} + + + +-- -- -- -- -- -- helperMorph : isMorph intGroup (dirProd intGroup intGroup) λ a → a , (0 - a) +-- -- -- -- -- -- helperMorph = +-- -- -- -- -- -- indIntGroup {G = dirProd intGroup intGroup} +-- -- -- -- -- -- (λ a → a , (0 - a)) +-- -- -- -- -- -- refl +-- -- -- -- -- -- (λ a → ×≡ refl (helper2 a)) +-- -- -- -- -- -- λ a → ×≡ refl (helper3 a) +-- -- -- -- -- -- where +-- -- -- -- -- -- helper1 : (a : Int) → predInt (sucInt a) ≡ a +-- -- -- -- -- -- helper1 (pos zero) = refl +-- -- -- -- -- -- helper1 (pos (suc n)) = refl +-- -- -- -- -- -- helper1 (negsuc zero) = refl +-- -- -- -- -- -- helper1 (negsuc (suc n)) = refl + +-- -- -- -- -- -- helper4 : (a : Int) → sucInt (predInt a) ≡ a +-- -- -- -- -- -- helper4 (pos zero) = refl +-- -- -- -- -- -- helper4 (pos (suc n)) = refl +-- -- -- -- -- -- helper4 (negsuc zero) = refl +-- -- -- -- -- -- helper4 (negsuc (suc n)) = refl + +-- -- -- -- -- -- helper2 : (a : Int) → (pos 0 - sucInt a) ≡ predInt (pos 0 - a) +-- -- -- -- -- -- helper2 (pos zero) = refl +-- -- -- -- -- -- helper2 (pos (suc n)) = refl +-- -- -- -- -- -- helper2 (negsuc zero) = refl +-- -- -- -- -- -- helper2 (negsuc (suc n)) = sym (helper1 _) + +-- -- -- -- -- -- helper3 : (a : Int) → (pos 0 - predInt a) ≡ sucInt (pos 0 - a) +-- -- -- -- -- -- helper3 (pos zero) = refl +-- -- -- -- -- -- helper3 (pos (suc zero)) = refl +-- -- -- -- -- -- helper3 (pos (suc (suc n))) = sym (helper4 _) +-- -- -- -- -- -- helper3 (negsuc zero) = refl +-- -- -- -- -- -- helper3 (negsuc (suc n)) = refl + +-- -- -- -- -- -- helper : (a b : Int) → (pos 0 - (a +ℤ b)) ≡ ((pos 0 - a) +ℤ (pos 0 - b)) +-- -- -- -- -- -- helper a (pos zero) = refl +-- -- -- -- -- -- helper (pos zero) (pos (suc n)) = +-- -- -- -- -- -- cong (λ x → pos 0 - sucInt x) (+ℤ-comm (pos zero) (pos n)) +-- -- -- -- -- -- ∙ +ℤ-comm (pos 0 +negsuc n) (pos zero) +-- -- -- -- -- -- helper (pos (suc n₁)) (pos (suc n)) = +-- -- -- -- -- -- {!!} +-- -- -- -- -- -- helper (negsuc n₁) (pos (suc n)) = {!!} +-- -- -- -- -- -- helper a (negsuc n) = {!!} + +-- -- -- -- -- -- fun : morph intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- -- fst fun = MV.d _ _ _ (λ _ → tt) (λ _ → tt) 0 ∘ grIso.inv coHom0-S0 .fst ∘ λ a → a , (0 - a) +-- -- -- -- -- -- snd fun = {!!} +-- -- -- -- -- -- {- compMorph {F = intGroup} {G = dirProd intGroup intGroup} {H = coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))} +-- -- -- -- -- -- ((λ a → a , a) , (λ a b → refl)) +-- -- -- -- -- -- (compMorph {F = dirProd intGroup intGroup} {G = coHomGr 0 (S₊ 0)} {H = coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))} (grIso.inv coHom0-S0) +-- -- -- -- -- -- (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 +-- -- -- -- -- -- , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) .snd -} +-- -- -- -- -- -- {- theIso : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- -- theIso = Iso''→Iso _ _ +-- -- -- -- -- -- (iso'' ((λ x → ∣ (λ {(inl tt) → 0ₖ ; (inr tt) → 0ₖ ; (push a i) → Kn→ΩKn+1 0 x i}) ∣₀) , {!!}) +-- -- -- -- -- -- {!!} +-- -- -- -- -- -- {!MV.d!}) +-- -- -- -- -- -- -} + + + +-- -- -- -- -- -- theIso : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) +-- -- -- -- -- -- theIso = +-- -- -- -- -- -- Iso''→Iso _ _ +-- -- -- -- -- -- (iso'' fun +-- -- -- -- -- -- (λ x inker → {!MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 +-- -- -- -- -- -- (grIso.inv coHom0-S0 .fst (g , g))!}) +-- -- -- -- -- -- (sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- -- -- -- -- -- λ x → pRec propTruncIsProp +-- -- -- -- -- -- (λ {(a , b) → {!fun!} }) +-- -- -- -- -- -- (fstFunHelper (∣ x ∣₀)))) +-- -- -- -- -- -- where +-- -- -- -- -- -- whoKnows : (a : S₊ 0 → Int) (x : MV.D Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt)) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (λ _ → a north) x +-- -- -- -- -- -- ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 a x +-- -- -- -- -- -- whoKnows a (inl x) = refl +-- -- -- -- -- -- whoKnows a (inr x) = refl +-- -- -- -- -- -- whoKnows a (push north i) = refl +-- -- -- -- -- -- whoKnows a (push south i) j = {!!} + +-- -- -- -- -- -- helper : (a : Int) → (grIso.inv coHom0-S0 .fst (a , a)) ≡ ∣ S0→Int (a , a) ∣₀ +-- -- -- -- -- -- helper a = {!have : + +-- -- -- -- -- -- ∣ +-- -- -- -- -- -- MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 +-- -- -- -- -- -- (S0→Int (x , x)) +-- -- -- -- -- -- ∣₀ +-- -- -- -- -- -- ≡ ∣ (λ _ → ∣ north ∣) ∣₀!} + +-- -- -- -- -- -- helper2 : (a b : Int) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a)) ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (b , b)) +-- -- -- -- -- -- → a ≡ b +-- -- -- -- -- -- helper2 a b id = pRec (isSetInt a b) (λ {(pt , id) → {!!}}) (fstFunHelper ∣ (MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a))) ∣₀) + +-- -- -- -- -- -- idFun : (a : Int) → MV.D Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 +-- -- -- -- -- -- idFun a (inl x) = 0ₖ +-- -- -- -- -- -- idFun a (inr x) = 0ₖ +-- -- -- -- -- -- idFun a (push north i) = Kn→ΩKn+1 zero a i +-- -- -- -- -- -- idFun a (push south i) = Kn→ΩKn+1 zero a i + +-- -- -- -- -- -- helper3 : (a : Int) → (MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a))) ≡ idFun a +-- -- -- -- -- -- helper3 a = funExt λ {(inl x) → refl ; (inr x) → refl ; (push north i) → refl ; (push south i) → refl } + +-- -- -- -- -- -- helper4 : (a b : Int) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a)) ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (b , b)) +-- -- -- -- -- -- → a ≡ b +-- -- -- -- -- -- helper4 a b id = +-- -- -- -- -- -- {!!} +-- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- ∙ {!!} +-- -- -- -- -- -- where +-- -- -- -- -- -- helper5 : {!!} --PathP (λ k → id k (inl tt) ≡ id k (inr tt)) (Kn→ΩKn+1 zero a) (Kn→ΩKn+1 zero a) +-- -- -- -- -- -- helper5 i j = {!id i!} + +-- -- -- -- -- -- -- fun : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) → coHom 0 (S₊ 0) +-- -- -- -- -- -- -- fun a = (pRec {P = Σ[ x ∈ coHom 0 (S₊ 0)] (MV.d _ _ _ (λ _ → tt) (λ _ → tt) 0 x ≡ a) × isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) (MV.Δ _ _ _ (λ _ → tt) (λ _ → tt) 0) x} +-- -- -- -- -- -- -- (λ {(a1 , b) (c , d) → ΣProp≡ (λ x → isOfHLevelProd 1 (setTruncIsSet _ _) propTruncIsProp) +-- -- -- -- -- -- -- (sElim2 {B = λ a1 c → (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a1 ≡ a) +-- -- -- -- -- -- -- → MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 c ≡ a +-- -- -- -- -- -- -- → isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) +-- -- -- -- -- -- -- (λ z → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 z) a1 +-- -- -- -- -- -- -- → isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) +-- -- -- -- -- -- -- (λ z → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 z) c → a1 ≡ c} +-- -- -- -- -- -- -- (λ _ _ → {!!}) +-- -- -- -- -- -- -- (λ a c b1 d1 → pElim (λ _ → isOfHLevelΠ 1 λ _ → setTruncIsSet _ _) +-- -- -- -- -- -- -- λ b2 → pElim (λ _ → setTruncIsSet _ _) +-- -- -- -- -- -- -- λ d2 → {!d2!}) +-- -- -- -- -- -- -- a1 c (proj₁ b) (proj₁ d) (proj₂ b) (proj₂ d))}) +-- -- -- -- -- -- -- (λ {(a , b) → a , b , MV.Ker-d⊂Im-Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a {!!}}) +-- -- -- -- -- -- -- (fstFunHelper a)) .fst -- pRec {!!} {!!} (fstFunHelper a) diff --git a/Cubical/ZCohomology/Groups/Sn.agda b/Cubical/ZCohomology/Groups/Sn.agda index ac11a8579..245eeb952 100644 --- a/Cubical/ZCohomology/Groups/Sn.agda +++ b/Cubical/ZCohomology/Groups/Sn.agda @@ -40,6 +40,8 @@ open import Cubical.HITs.Wedge open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.ZCohomology.Groups.Unit + open import Cubical.ZCohomology.KcompPrelims @@ -49,43 +51,19 @@ open import Cubical.Data.Sum.Base open import Cubical.Data.HomotopyGroup open import Cubical.Data.Bool hiding (_⊕_) -{- -cong₂Sym : ∀ {ℓ} {A : Type ℓ} {x : A} - (p : x ≡ x) → - (refl ≡ p) → - (P : p ≡ p) → - cong₂ (λ x y → x ∙ y) P (sym P) ≡ refl -cong₂Sym {x = x} p = J (λ p _ → (P : p ≡ p) → cong₂ (λ x y → x ∙ y) P (sym P) ≡ refl) - λ P → cong₂Funct (λ x y → x ∙ y) P (sym P) - ∙ PathP→compPathR (λ j → cong (λ x → rUnit x (~ j)) P ∙ cong (λ x → lUnit x (~ j)) (sym P)) - ∙ (λ j → (sym (rUnit refl)) ∙ rCancel P j ∙ lUnit refl) - ∙ (λ j → sym (rUnit refl) ∙ lUnit (lUnit refl) (~ j) ) - ∙ rCancel (sym (rUnit refl)) - -abstract - cong₂Sym1 : ∀ {ℓ} {A : Type ℓ} {x : A} - (p : x ≡ x) → - (refl ≡ p) → - (P : p ≡ p) → - cong (λ x → x ∙ p) P ≡ cong (λ x → p ∙ x) P - cong₂Sym1 {x = x} p id P = rUnit _ - ∙ (λ i → cong (λ x → x ∙ p) P ∙ lCancel (λ i → p ∙ P i) (~ i)) -- cong (λ x → cong (λ x → x ∙ p) P ∙ x) {!!} - ∙ assoc _ _ _ - ∙ (λ j → cong₂Funct (λ x y → x ∙ y) P (sym P) (~ j) ∙ (λ i → p ∙ P i)) - ∙ (λ j → cong₂Sym p id P j ∙ (λ i → p ∙ P i)) - ∙ sym (lUnit _) --} - -test13 : Iso (S₊ 1 → hLevelTrunc 4 (S₊ 2)) (hLevelTrunc 4 (S₊ 2) × hLevelTrunc 3 (S₊ 1)) -Iso.fun test13 f = f north , ΩKn+1→Kn (sym (rCancelₖ (f north)) - ∙ (λ i → f (merid south i) +ₖ (-ₖ f (merid north i))) - ∙ rCancelₖ (f south)) -Iso.inv test13 (a , b) north = a +ₖ 0ₖ -Iso.inv test13 (a , b) south = a +ₖ 0ₖ -Iso.inv test13 (a , b) (merid south i) = a +ₖ (Kn→ΩKn+1 1 b i) -Iso.inv test13 (a , b) (merid north i) = a +ₖ 0ₖ -Iso.rightInv test13 (a , b) = +{- Proof that S¹→K2 is isomorphic to K2×K1 +Gives a better computing proof that Used for H¹(S²)≡0 which does not rely om Mayer Vietoris +Also needed for H²(T²) -} +S1→K2≡K2×K1 : Iso (S₊ 1 → hLevelTrunc 4 (S₊ 2)) (hLevelTrunc 4 (S₊ 2) × hLevelTrunc 3 (S₊ 1)) +Iso.fun S1→K2≡K2×K1 f = f north , ΩKn+1→Kn (sym (rCancelₖ (f north)) + ∙ (λ i → f (merid south i) +ₖ (-ₖ f (merid north i))) + ∙ rCancelₖ (f south)) +Iso.inv S1→K2≡K2×K1 (a , b) north = a +ₖ 0ₖ +Iso.inv S1→K2≡K2×K1 (a , b) south = a +ₖ 0ₖ +Iso.inv S1→K2≡K2×K1 (a , b) (merid south i) = a +ₖ (Kn→ΩKn+1 1 b i) +Iso.inv S1→K2≡K2×K1 (a , b) (merid north i) = a +ₖ 0ₖ +Iso.rightInv S1→K2≡K2×K1 (a , b) = ×≡ (rUnitₖ a) ((cong ΩKn+1→Kn (congHelper++ (Kn→ΩKn+1 1 b) (λ x → (a +ₖ x) +ₖ (-ₖ (a +ₖ 0ₖ))) (funExt (λ x → sym (cancelHelper a x))) (rCancelₖ (a +ₖ 0ₖ)))) ∙ Iso.leftInv (Iso3-Kn-ΩKn+1 1) b) @@ -136,7 +114,7 @@ Iso.rightInv test13 (a , b) = cong (λ x → sym q ∙ x ∙ q) (congHelper p f id) ∙ moveabout (sym (funExt⁻ id ∣ north ∣)) q p -Iso.leftInv test13 a = +Iso.leftInv S1→K2≡K2×K1 a = funExt λ {north → rUnitₖ (a north) ; south → rUnitₖ (a north) ∙ cong a (merid north) ; (merid south i) → test i @@ -198,17 +176,17 @@ Iso.leftInv test13 a = → PathP (λ i → p' (~ i) ≡ p' (~ i)) q' r → p ∙ q ∙ sym p ≡ r helperFun p p' q q' r comm qid dep = - p ∙ q ∙ sym p ≡⟨ cong (λ x → p ∙ x ∙ sym p) qid ⟩ - p ∙ q' ∙ sym p ≡⟨ cong (λ x → p ∙ x ∙ sym p) (λ i → lUnit (rUnit q' i) i) ⟩ - p ∙ (refl ∙ q' ∙ refl) ∙ sym p ≡⟨ cong (λ x → p ∙ x ∙ sym p) (λ i → (λ j → p' (~ i ∨ ~ j)) ∙ dep i ∙ λ j → p' (~ i ∨ j)) ⟩ - p ∙ (sym p' ∙ r ∙ p') ∙ sym p ≡⟨ assoc p (sym p' ∙ r ∙ p') (sym p) ⟩ - (p ∙ sym p' ∙ r ∙ p') ∙ sym p ≡⟨ cong (_∙ sym p) (assoc p (sym p') (r ∙ p')) ⟩ - ((p ∙ sym p') ∙ r ∙ p') ∙ sym p ≡⟨ sym (assoc (p ∙ sym p') (r ∙ p') (sym p)) ⟩ - (p ∙ sym p') ∙ (r ∙ p') ∙ sym p ≡⟨ cong ((p ∙ sym p') ∙_) (sym (assoc r p' (sym p))) ⟩ - (p ∙ sym p') ∙ r ∙ (p' ∙ sym p) ≡⟨ cong (λ x → (p ∙ sym p') ∙ r ∙ x) (sym (symDistr p (sym p'))) ⟩ - (p ∙ sym p') ∙ r ∙ sym (p ∙ sym p') ≡⟨ cong ((p ∙ sym p') ∙_) (comm r (sym (p ∙ sym p'))) ⟩ - (p ∙ sym p') ∙ sym (p ∙ sym p') ∙ r ≡⟨ assoc (p ∙ sym p') (sym (p ∙ sym p')) r ⟩ - ((p ∙ sym p') ∙ sym (p ∙ sym p')) ∙ r ≡⟨ cong (_∙ r) (rCancel (p ∙ sym p')) ⟩ + p ∙ q ∙ sym p ≡⟨ cong (λ x → p ∙ x ∙ sym p) qid ⟩ + p ∙ q' ∙ sym p ≡⟨ cong (λ x → p ∙ x ∙ sym p) (λ i → lUnit (rUnit q' i) i) ⟩ + p ∙ (refl ∙ q' ∙ refl) ∙ sym p ≡⟨ cong (λ x → p ∙ x ∙ sym p) (λ i → (λ j → p' (~ i ∨ ~ j)) ∙ dep i ∙ λ j → p' (~ i ∨ j)) ⟩ + p ∙ (sym p' ∙ r ∙ p') ∙ sym p ≡⟨ assoc p (sym p' ∙ r ∙ p') (sym p) ⟩ + (p ∙ sym p' ∙ r ∙ p') ∙ sym p ≡⟨ cong (_∙ sym p) (assoc p (sym p') (r ∙ p')) ⟩ + ((p ∙ sym p') ∙ r ∙ p') ∙ sym p ≡⟨ sym (assoc (p ∙ sym p') (r ∙ p') (sym p)) ⟩ + (p ∙ sym p') ∙ (r ∙ p') ∙ sym p ≡⟨ cong ((p ∙ sym p') ∙_) (sym (assoc r p' (sym p))) ⟩ + (p ∙ sym p') ∙ r ∙ (p' ∙ sym p) ≡⟨ cong (λ x → (p ∙ sym p') ∙ r ∙ x) (sym (symDistr p (sym p'))) ⟩ + (p ∙ sym p') ∙ r ∙ sym (p ∙ sym p') ≡⟨ cong ((p ∙ sym p') ∙_) (comm r (sym (p ∙ sym p'))) ⟩ + (p ∙ sym p') ∙ sym (p ∙ sym p') ∙ r ≡⟨ assoc (p ∙ sym p') (sym (p ∙ sym p')) r ⟩ + ((p ∙ sym p') ∙ sym (p ∙ sym p')) ∙ r ≡⟨ cong (_∙ r) (rCancel (p ∙ sym p')) ⟩ refl ∙ r ≡⟨ sym (lUnit r) ⟩ r ∎ together : Path (_ ≡ _) (cong (a north +ₖ_) (sym (rCancelₖ (a north))) @@ -240,115 +218,273 @@ Iso.leftInv test13 a = Iso.inv (congIsoHelper x) = _+ₖ (-ₖ x) Iso.rightInv (congIsoHelper x) a = assocₖ a (-ₖ x) x ∙ cong (a +ₖ_) (lCancelₖ x) ∙ rUnitₖ a Iso.leftInv (congIsoHelper x) a = assocₖ a x (-ₖ x) ∙ cong (a +ₖ_) (rCancelₖ x) ∙ rUnitₖ a + typId : (x : hLevelTrunc 4 (S₊ 2)) → (x ≡ x) ≡ Path (hLevelTrunc 4 (S₊ 2)) 0ₖ 0ₖ typId x = isoToPath ((congIso (isoToEquiv (symIso (congIsoHelper x))))) ∙ λ i → rCancelₖ x i ≡ rCancelₖ x i - where typIdHelper : (p q : (x ≡ x)) → transport (λ i → typId x i) (p ∙ q) ≡ transport (λ i → typId x i) p ∙ transport (λ i → typId x i) q typIdHelper p q = - (substComposite (λ x → x) (isoToPath ((congIso (isoToEquiv (symIso (congIsoHelper x)))))) (λ i → rCancelₖ x i ≡ rCancelₖ x i) (p ∙ q)) - ∙ (λ i → transport (λ i → rCancelₖ x i ≡ rCancelₖ x i) (transport (isoToPath (congIso (isoToEquiv (symIso (congIsoHelper x))))) (p ∙ q))) - ∙ (λ i → transport (λ i → rCancelₖ x i ≡ rCancelₖ x i) (transportRefl (congFunct (_+ₖ (-ₖ x)) p q i) i)) - ∙ (λ i → transport (λ i → rCancelₖ x i ≡ rCancelₖ x i) ((lUnit (rUnit (transportRefl (cong (_+ₖ (-ₖ x)) p) (~ i)) i) i) ∙ (lUnit (rUnit (transportRefl (cong (_+ₖ (-ₖ x)) q) (~ i)) i) i))) - ∙ (λ i → transp (λ j → rCancelₖ x (i ∨ j) ≡ rCancelₖ x (i ∨ j)) i (((λ j → rCancelₖ x (i ∧ (~ j))) ∙ transport refl (cong (_+ₖ (-ₖ x)) p) ∙ λ j → rCancelₖ x (i ∧ j)) ∙ (λ j → rCancelₖ x (i ∧ (~ j))) ∙ transport refl (cong (_+ₖ (-ₖ x)) q) ∙ λ j → rCancelₖ x (i ∧ j))) - ∙ (λ i → transp (λ j → rCancelₖ x ((~ i) ∨ j) ≡ rCancelₖ x ((~ i) ∨ j)) (~ i) ((λ j → rCancelₖ x ((~ i) ∧ (~ j))) ∙ transport refl (cong (_+ₖ (-ₖ x)) p) ∙ λ j → rCancelₖ x ((~ i) ∧ j)) - ∙ transp (λ j → rCancelₖ x ((~ i) ∨ j) ≡ rCancelₖ x ((~ i) ∨ j)) (~ i) ((λ j → rCancelₖ x ((~ i) ∧ (~ j))) ∙ transport refl (cong (_+ₖ (-ₖ x)) q) ∙ λ j → rCancelₖ x ((~ i) ∧ j))) - ∙ (λ i → transport (λ j → rCancelₖ x j ≡ rCancelₖ x j) (lUnit (rUnit (transport refl (cong (_+ₖ (-ₖ x)) p)) (~ i)) (~ i)) - ∙ transport (λ j → rCancelₖ x j ≡ rCancelₖ x j) (lUnit (rUnit (transport refl (cong (_+ₖ (-ₖ x)) q)) (~ i)) (~ i))) - ∙ cong₂ (_∙_) (sym (substComposite (λ x → x) (isoToPath ((congIso (isoToEquiv (symIso (congIsoHelper x)))))) (λ i → rCancelₖ x i ≡ rCancelₖ x i) p)) - (sym (substComposite (λ x → x) (isoToPath ((congIso (isoToEquiv (symIso (congIsoHelper x)))))) (λ i → rCancelₖ x i ≡ rCancelₖ x i) q)) - - - test2 : PathP (λ i → a north +ₖ ∣ north ∣ ≡ a (merid north i)) (rUnitₖ (a north)) (rUnitₖ (a north) ∙ cong a (merid north)) + (substComposite (λ x → x) (isoToPath ((congIso (isoToEquiv (symIso (congIsoHelper x)))))) + (λ i → rCancelₖ x i ≡ rCancelₖ x i) (p ∙ q)) + ∙ (λ i → transport (λ i → rCancelₖ x i ≡ rCancelₖ x i) + (transportRefl (congFunct (_+ₖ (-ₖ x)) p q i) i)) + ∙ (λ i → transport (λ i → rCancelₖ x i ≡ rCancelₖ x i) ( + (lUnit (rUnit (transportRefl (cong (_+ₖ (-ₖ x)) p) (~ i)) i) i) + ∙ (lUnit (rUnit (transportRefl (cong (_+ₖ (-ₖ x)) q) (~ i)) i) i))) + ∙ (λ i → transp (λ j → rCancelₖ x (i ∨ j) ≡ rCancelₖ x (i ∨ j)) i + (((λ j → rCancelₖ x (i ∧ (~ j))) ∙ transport refl (cong (_+ₖ (-ₖ x)) p) ∙ λ j → rCancelₖ x (i ∧ j)) + ∙ (λ j → rCancelₖ x (i ∧ (~ j))) ∙ transport refl (cong (_+ₖ (-ₖ x)) q) ∙ λ j → rCancelₖ x (i ∧ j))) + ∙ (λ i → transp (λ j → rCancelₖ x ((~ i) ∨ j) ≡ rCancelₖ x ((~ i) ∨ j)) (~ i) + ((λ j → rCancelₖ x ((~ i) ∧ (~ j))) ∙ transport refl (cong (_+ₖ (-ₖ x)) p) ∙ λ j → rCancelₖ x ((~ i) ∧ j)) + ∙ transp (λ j → rCancelₖ x ((~ i) ∨ j) ≡ rCancelₖ x ((~ i) ∨ j)) (~ i) + ((λ j → rCancelₖ x ((~ i) ∧ (~ j))) ∙ transport refl (cong (_+ₖ (-ₖ x)) q) ∙ λ j → rCancelₖ x ((~ i) ∧ j))) + ∙ (λ i → transport (λ j → rCancelₖ x j ≡ rCancelₖ x j) + (lUnit (rUnit (transport refl (cong (_+ₖ (-ₖ x)) p)) (~ i)) (~ i)) + ∙ transport (λ j → rCancelₖ x j ≡ rCancelₖ x j) + (lUnit (rUnit (transport refl (cong (_+ₖ (-ₖ x)) q)) (~ i)) (~ i))) + ∙ cong₂ (_∙_) (sym (substComposite (λ x → x) (isoToPath ((congIso (isoToEquiv (symIso (congIsoHelper x)))))) + (λ i → rCancelₖ x i ≡ rCancelₖ x i) p)) + (sym (substComposite (λ x → x) (isoToPath ((congIso (isoToEquiv (symIso (congIsoHelper x)))))) + (λ i → rCancelₖ x i ≡ rCancelₖ x i) q)) + + + test2 : PathP (λ i → a north +ₖ ∣ north ∣ ≡ a (merid north i)) + (rUnitₖ (a north)) (rUnitₖ (a north) ∙ cong a (merid north)) test2 i j = hcomp (λ k → λ { (i = i0) → rUnitₖ (a north) j ; (j = i1) → a (merid north (i ∧ k)) ; (j = i0) → a north +ₖ ∣ north ∣}) (rUnitₖ (a north) j) +---------------------------------------------------------------------- +--- We will need to switch between Sⁿ defined using suspensions and using pushouts +--- in order to apply Mayer Vietoris. +coHomPushout≅coHomSn : (n m : ℕ) → grIso (coHomGr m (S₊ (suc n))) (coHomGr m (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt)) +morph.fun (grIso.fun (coHomPushout≅coHomSn n m)) = + sRec setTruncIsSet + λ f → ∣ (λ {(inl x) → f north + ; (inr x) → f south + ; (push a i) → f (merid a i)}) ∣₀ +morph.ismorph (grIso.fun (coHomPushout≅coHomSn n m)) = + sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ a b → cong ∣_∣₀ (funExt λ {(inl x) → refl + ; (inr x) → refl + ; (push a i) → refl }) +morph.fun (grIso.inv (coHomPushout≅coHomSn n m)) = + sRec setTruncIsSet + (λ f → ∣ (λ {north → f (inl tt) + ; south → f (inr tt) + ; (merid a i) → f (push a i)}) ∣₀) +morph.ismorph (grIso.inv (coHomPushout≅coHomSn n m)) = + sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ a b → cong ∣_∣₀ (funExt λ {north → refl + ; south → refl + ; (merid a i) → refl }) +grIso.rightInv (coHomPushout≅coHomSn n m) = + sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ (funExt λ {(inl x) → refl + ; (inr x) → refl + ; (push a i) → refl }) +grIso.leftInv (coHomPushout≅coHomSn n m) = + sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ (funExt λ {north → refl + ; south → refl + ; (merid a i) → refl }) -S1→S¹ : S₊ 1 → S¹ -S1→S¹ x = SuspBool→S¹ (S1→SuspBool x) - -S¹→S1 : S¹ → S₊ 1 -S¹→S1 x = SuspBool→S1 (S¹→SuspBool x) - -S1→S¹-sect : section S1→S¹ S¹→S1 -S1→S¹-sect x = - cong SuspBool→S¹ (SuspBool→S1-retr (S¹→SuspBool x)) - ∙ S¹→SuspBool→S¹ x - -S1→S¹-retr : retract S1→S¹ S¹→S1 -S1→S¹-retr x = - cong SuspBool→S1 (SuspBool→S¹→SuspBool (S1→SuspBool x)) - ∙ SuspBool→S1-sect x +-------------------------- H⁰(S⁰) ----------------------------- +S0→Int : (a : Int × Int) → S₊ 0 → Int +S0→Int a north = proj₁ a +S0→Int a south = proj₂ a +H⁰-S⁰≅ℤ×ℤ : grIso (coHomGr 0 (S₊ 0)) (dirProd intGroup intGroup) +H⁰-S⁰≅ℤ×ℤ = + Iso'→Iso + (iso' + (iso (sRec (isOfHLevelProd 2 isSetInt isSetInt) + λ f → (f north) , (f south)) + (λ a → ∣ S0→Int a ∣₀) + (λ { (a , b) → refl}) + (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ (funExt (λ {north → refl ; south → refl})))) + (sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 isSetInt isSetInt) _ _) + λ a b i → addLemma (a north) (b north) i , addLemma (a south) (b south) i)) +------------------------ H⁰(S¹) ------------------------------ +H⁰-S¹≅ℤ : grIso intGroup (coHomGr 0 (S₊ 1)) +H⁰-S¹≅ℤ = + Iso'→Iso + (iso' + (iso + (λ a → ∣ (λ x → a) ∣₀) + (sRec isSetInt (λ f → f north)) + (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ (funExt (suspToPropRec north (λ _ → isSetInt _ _) refl))) + (λ a → refl)) + λ a b → cong ∣_∣₀ (funExt λ x → sym (addLemma a b))) +----------------------------------------------------- + + +H¹-S¹-helper : Iso (coHom 2 (S₊ 1)) Unit +H¹-S¹-helper = + compIso (setTruncIso S1→K2≡K2×K1) + (compIso + setTruncOfProdIso + (compIso + (prodIso setTruncTrunc0Iso setTruncTrunc0Iso) + (pathToIso + ((λ i → isContr→≡Unit helper2 i × isContr→≡Unit helper3 i) + ∙ sym diagonal-unit)))) + where + helper2 : isContr (hLevelTrunc 2 (hLevelTrunc 4 (S₊ 2))) + helper2 = + transport (λ i → isContr (isoToPath (truncOfTruncIso {A = (S₊ 2)} 2 2) i)) + (isConnectedSubtr 2 1 (sphereConnected 2)) + helper3 : isContr (hLevelTrunc 2 (hLevelTrunc 3 (S₊ 1))) + helper3 = + transport (λ i → isContr (hLevelTrunc 2 (truncIdempotent {A = S₊ 1} 3 isGroupoidS1 (~ i)))) + (sphereConnected 1) -isGroupoidS1 : isGroupoid (S₊ 1) -isGroupoidS1 = transport (λ i → isGroupoid (S¹≡S1 (~ i))) isGroupoidS¹ -isConnectedS1 : (x : S₊ 1) → ∥ north ≡ x ∥₋₁ -isConnectedS1 x = pRec propTruncIsProp - (λ p → ∣ cong (transport (sym (S¹≡S1))) p ∙ transport⁻Transport (S¹≡S1) x ∣₋₁) - (isConnectedS¹ (transport S¹≡S1 x)) +-------------------------------------------------------------------- +--------------------------H¹(S¹) ----------------------------------- +H¹-S¹≅ℤ : grIso intGroup (coHomGr 1 (S₊ 1)) +H¹-S¹≅ℤ = + compGrIso + (diagonalIso1 + _ + (coHomGr 0 (S₊ 0)) + _ + (invGrIso H⁰-S⁰≅ℤ×ℤ) + (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) + (λ x → MV.Ker-i⊂Im-d _ _(S₊ 0) (λ _ → tt) (λ _ → tt) 0 x + (×≡ (isOfHLevelSuc 0 (isContrHⁿ-Unit 0) _ _) + (isOfHLevelSuc 0 (isContrHⁿ-Unit 0) _ _))) + (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) + (λ x inker + → pRec propTruncIsProp + (λ {((f , g) , id') → helper x f g id' inker}) + ((MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ inker)))) + (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ F surj + → pRec (setTruncIsSet _ _) + (λ { (x , id) → MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ F ∣₀ ∣ + (∣ (λ _ → x) ∣₀ , ∣ (λ _ → 0) ∣₀) , + (cong ∣_∣₀ (funExt (surjHelper x))) ∙ sym id ∣₋₁ }) surj) ) + (invGrIso (coHomPushout≅coHomSn 0 1)) + + where + surjHelper : (x : Int) (x₁ : S₊ 0) → x +ₖ (-ₖ pos 0) ≡ S0→Int (x , x) x₁ + surjHelper x north = cong (x +ₖ_) (-0ₖ) ∙ rUnitₖ x + surjHelper x south = cong (x +ₖ_) (-0ₖ) ∙ rUnitₖ x -open import Cubical.HITs.S2 + helper : (F : S₊ 0 → Int) (f g : ∥ (Unit → Int) ∥₀) (id : morph.fun (MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) (f , g) ≡ ∣ F ∣₀) + → isInKer (coHomGr 0 (S₊ 0)) + (coHomGr 1 (Pushout (λ _ → tt) (λ _ → tt))) + (morph.fun (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) + ∣ F ∣₀ + → ∃[ x ∈ Int ] ∣ F ∣₀ ≡ morph.fun (grIso.inv H⁰-S⁰≅ℤ×ℤ) (x , x) + helper F = + sElim2 (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ f g id inker + → pRec propTruncIsProp + (λ {((a , b) , id2) + → sElim2 {B = λ f g → morph.fun (MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) (f , g) ≡ ∣ F ∣₀ → _ } + (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) + (λ f g id → ∣ (helper2 f g .fst) , (sym id ∙ sym (helper2 f g .snd)) ∣₋₁) + a b id2}) + (MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ F ∣₀ inker) + where + helper2 : (f g : Unit → Int) + → Σ[ x ∈ Int ] morph.fun (grIso.inv H⁰-S⁰≅ℤ×ℤ) (x , x) + ≡ morph.fun (MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) (∣ f ∣₀ , ∣ g ∣₀) + helper2 f g = (f _ +ₖ (-ₖ g _) ) , cong ∣_∣₀ (funExt (λ {north → refl ; south → refl})) -test : (S₊ 2) → S₊ 1 -test north = north -test south = south -test (merid a i) = merid north i -test2 : (S₊ 1) → (S₊ 2) -test2 north = north -test2 south = south -test2 (merid a i) = merid (merid a i) i +---------------------------- Hⁿ(Sⁿ) ≅ ℤ , n ≥ 1 ------------------- +coHom-n-Sn : (n : ℕ) → grIso intGroup (coHomGr (suc n) (S₊ (suc n))) +coHom-n-Sn zero = H¹-S¹≅ℤ +coHom-n-Sn (suc n) = + compGrIso + (compGrIso + (coHom-n-Sn n) + theIso) + (invGrIso (coHomPushout≅coHomSn (suc n) (suc (suc n)))) + where + theIso : grIso (coHomGr (suc n) (S₊ (suc n))) (coHomGr (suc (suc n)) + (Pushout {A = S₊ (suc n)} (λ _ → tt) (λ _ → tt))) + theIso = + SES→Iso + (×coHomGr (suc n) Unit Unit) + (×coHomGr (suc (suc n)) Unit Unit) + (ses (λ p q → ×≡ (isOfHLevelSuc 0 (isContrHⁿ-Unit n) (proj₁ p) (proj₁ q)) + (isOfHLevelSuc 0 (isContrHⁿ-Unit n) (proj₂ p) (proj₂ q))) + (λ p q → ×≡ (isOfHLevelSuc 0 (isContrHⁿ-Unit (suc n)) (proj₁ p) (proj₁ q)) + (isOfHLevelSuc 0 (isContrHⁿ-Unit (suc n)) (proj₂ p) (proj₂ q))) + (MV.Δ _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) + (MV.i _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (2 + n)) + (MV.d _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) + (MV.Ker-d⊂Im-Δ _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) + (MV.Ker-i⊂Im-d _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n))) +-------------------------------------------------------------------- -testS² : S² → S¹ -testS² base = base -testS² (surf i i₁) = base +------------------------- H¹(S⁰) ≅ 0 ------------------------------- +H¹-S⁰≅0 : grIso (coHomGr 1 (S₊ 0)) trivialGroup +H¹-S⁰≅0 = + Iso'→Iso + (iso' + (iso (λ _ → tt) + (λ _ → 0ₕ) + (λ _ → refl) + (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → pRec (setTruncIsSet _ _) + (λ id → cong ∣_∣₀ (sym id)) + (helper f (f north) (f south) refl refl))) + λ _ _ → refl) + where + helper : (f : S₊ 0 → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1) → (x y : ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1) → (f north ≡ x) → (f south ≡ y) → ∥ f ≡ (λ _ → 0ₖ) ∥₋₁ + helper f = + elim2 (λ _ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelPlus {n = 1} 2 propTruncIsProp) + (suspToPropRec2 north (λ _ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → propTruncIsProp) + λ id id2 → ∣ funExt (λ {north → id ; south → id2}) ∣₋₁) -test4 : S₊ 1 → hLevelTrunc 4 (S₊ 2) -test4 north = ∣ north ∣ -test4 south = ∣ north ∣ -test4 (merid a i) = (Kn→ΩKn+1 1 ∣ south ∣) i +------------------------- H²(S¹) ≅ 0 ------------------------------- +H²-S¹≅0 : grIso (coHomGr 2 (S₊ 1)) trivialGroup +H²-S¹≅0 = + Iso'→Iso + (iso' + (iso + (λ _ → tt) + (λ _ → 0ₕ) + (λ _ → refl) + (λ _ → isPropH¹-S¹≅ℤ _ _)) + λ _ _ → refl) + where + isPropH¹-S¹≅ℤ : isProp (coHom 2 (S₊ 1)) + isPropH¹-S¹≅ℤ = + isOfHLevelSuc 0 + (transport (λ i → isContr (isoToPath H¹-S¹-helper (~ i))) + isContrUnit) +------------------------------------------------------------------- +--------- Direct proof of H¹(S¹) ≅ ℤ without Mayer-Vietoris ------- +S¹map : hLevelTrunc 3 S¹ → S¹ +S¹map = trElim (λ _ → isGroupoidS¹) (idfun S¹) -coHom0Torus : grIso (coHomGr 0 (S₊ 1 × S₊ 1)) intGroup -coHom0Torus = - Iso'→Iso - (iso' - (iso (sRec isSetInt (λ f → f (north , north))) - (λ a → ∣ (λ x → a) ∣₀) - (λ a → refl) - (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) - λ f → cong ∣_∣₀ - (funExt λ {(x , y) → suspToPropRec2 - {B = λ x y → f (north , north) ≡ f (x , y)} - north - (λ _ _ → isSetInt _ _) - refl - x y}))) - (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) λ a b → addLemma (a (north , north)) (b (north , north)))) +S¹map-id : (x : hLevelTrunc 3 S¹) → Path (hLevelTrunc 3 S¹) ∣ S¹map x ∣ x +S¹map-id = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → refl -private - S¹map : hLevelTrunc 3 S¹ → S¹ - S¹map = trElim (λ _ → isGroupoidS¹) (idfun S¹) +S1map : hLevelTrunc 3 (S₊ 1) → (S₊ 1) +S1map = trElim (λ _ → isGroupoidS1) (idfun _) - S¹map-id : (x : hLevelTrunc 3 S¹) → Path (hLevelTrunc 3 S¹) ∣ S¹map x ∣ x - S¹map-id = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) - λ a → refl - S1map : hLevelTrunc 3 (S₊ 1) → (S₊ 1) - S1map = trElim (λ _ → isGroupoidS1) (idfun _) S¹→S¹ : Iso (S¹ → hLevelTrunc 3 S¹) (S¹ × Int) @@ -449,34 +585,6 @@ basechange-lemma x y F f g frefl grefl = (λ i → F (lUnit (g (loop i)) (~ j))))) -basechange-lemma2 : (f g : S¹ → hLevelTrunc 3 (S₊ 1)) (F : hLevelTrunc 3 (S₊ 1) → S¹) - → ((basechange2⁻ (F (f base +ₖ g base)) λ i → F ((f (loop i)) +ₖ g (loop i))) - ≡ basechange2⁻ (F (f base)) (cong (F ∘ f) loop) - ∙ basechange2⁻ (F (g base)) (cong (F ∘ g) loop)) -basechange-lemma2 f g F = coInd f g F (f base) (g base) refl refl - where - coInd : (f g : S¹ → hLevelTrunc 3 (S₊ 1)) (F : hLevelTrunc 3 (S₊ 1) → S¹) (x y : hLevelTrunc 3 (S₊ 1)) - → f base ≡ x - → g base ≡ y - → ((basechange2⁻ (F (f base +ₖ g base)) λ i → F ((f (loop i)) +ₖ g (loop i))) - ≡ basechange2⁻ (F (f base)) (cong (F ∘ f) loop) - ∙ basechange2⁻ (F (g base)) (cong (F ∘ g) loop)) - coInd f g F = - elim2 (λ _ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 (isGroupoidS¹ base base)) _ _ ) - (suspToPropRec2 - north - (λ _ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → isGroupoidS¹ _ _ _ _) - λ fid gid → - (λ i → basechange2⁻ (F (f base +ₖ g base)) (cong₂Funct (λ x y → F (f x +ₖ g y)) loop loop i)) - ∙ (basechange2⁻-morph (F (f base +ₖ g base)) - (cong (λ x → F (f x +ₖ g base)) loop) - (cong (λ x → F (f base +ₖ g x)) loop)) - ∙ (λ i → basechange2⁻ (F (f base +ₖ gid i)) (cong (λ x → F (f x +ₖ gid i)) loop) - ∙ basechange2⁻ (F (fid i +ₖ g base)) (cong (λ x → F (fid i +ₖ g x)) loop)) - ∙ (λ i → basechange2⁻ (F (rUnitₖ (f base) i)) (cong (λ x → F (rUnitₖ (f x) i)) loop) - ∙ basechange2⁻ (F (lUnitₖ (g base) i)) (cong (λ x → F (lUnitₖ (g x) i)) loop))) - - coHom1S1≃ℤ : grIso (coHomGr 1 (S₊ 1)) intGroup coHom1S1≃ℤ = @@ -540,1219 +648,3 @@ coHom1S1≃ℤ = (λ i → S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 (loop i))) j))) ∙ basechange2⁻ (S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 base)) j))) (λ i → S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 (loop i))) j)))) - -pushoutSn : (n : ℕ) → Iso (S₊ (suc n)) (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt) -Iso.fun (pushoutSn n) north = inl tt -Iso.fun (pushoutSn n) south = inr tt -Iso.fun (pushoutSn n) (merid a i) = push a i -Iso.inv (pushoutSn n) (inl x) = north -Iso.inv (pushoutSn n) (inr x) = south -Iso.inv (pushoutSn n) (push a i) = merid a i -Iso.rightInv (pushoutSn n) (inl x) = refl -Iso.rightInv (pushoutSn n) (inr x) = refl -Iso.rightInv (pushoutSn n) (push a i) = refl -Iso.leftInv (pushoutSn n) north = refl -Iso.leftInv (pushoutSn n) south = refl -Iso.leftInv (pushoutSn n) (merid a i) = refl - -Sn≡Pushout : (n : ℕ) → (S₊ (suc n)) ≡ (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt) -Sn≡Pushout n = isoToPath (pushoutSn n) - -coHomPushout≡coHomSn' : (n m : ℕ) → grIso (coHomGr m (S₊ (suc n))) (coHomGr m (Pushout {A = S₊ n} (λ _ → tt) λ _ → tt)) -morph.fun (grIso.fun (coHomPushout≡coHomSn' n m)) = - sRec setTruncIsSet - λ f → ∣ (λ {(inl x) → f north ; (inr x) → f south ; (push a i) → f (merid a i)}) ∣₀ -morph.ismorph (grIso.fun (coHomPushout≡coHomSn' n m)) = - sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) - λ a b → cong ∣_∣₀ (funExt λ {(inl x) → refl ; (inr x) → refl ; (push a i) → refl }) -morph.fun (grIso.inv (coHomPushout≡coHomSn' n m)) = sRec setTruncIsSet (λ f → ∣ (λ {north → f (inl tt) ; south → f (inr tt) ; (merid a i) → f (push a i)}) ∣₀) -morph.ismorph (grIso.inv (coHomPushout≡coHomSn' n m)) = - sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) - λ a b → cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → refl }) -grIso.rightInv (coHomPushout≡coHomSn' n m) = - sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) - λ f → cong ∣_∣₀ (funExt λ {(inl x) → refl ; (inr x) → refl ; (push a i) → refl }) -grIso.leftInv (coHomPushout≡coHomSn' n m) = - sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) - λ f → cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → refl }) - - - - - - -coHomGrUnit0 : grIso (coHomGr 0 Unit) intGroup -grIso.fun coHomGrUnit0 = mph (sRec isSetInt (λ f → f tt)) - (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) - (λ a b → addLemma (a tt) (b tt))) -grIso.inv coHomGrUnit0 = mph (λ a → ∣ (λ _ → a) ∣₀) (λ a b i → ∣ (λ _ → addLemma a b (~ i)) ∣₀) -grIso.rightInv coHomGrUnit0 a = refl -grIso.leftInv coHomGrUnit0 = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ a → refl - -isContrCohomUnit : (n : ℕ) → isContr (coHom (suc n) Unit) -isContrCohomUnit n = subst isContr (λ i → ∥ UnitToTypeId (coHomK (suc n)) (~ i) ∥₀) (helper n) - where - helper : (n : ℕ) → isContr (∥ coHomK (suc n) ∥₀) - helper n = subst isContr ((isoToPath (truncOfTruncIso {A = S₊ (1 + n)} 2 (1 + n))) ∙ sym propTrunc≡Trunc0 ∙ λ i → ∥ hLevelTrunc (suc (+-comm n 2 i)) (S₊ (1 + n)) ∥₀) - (isConnectedSubtr 2 (helper2 n .fst) (subst (λ x → isHLevelConnected x (S₊ (suc n))) (sym (helper2 n .snd)) (sphereConnected (suc n))) ) - where - helper2 : (n : ℕ) → Σ[ m ∈ ℕ ] m + 2 ≡ 2 + n - helper2 zero = 0 , refl - helper2 (suc n) = (suc n) , λ i → suc (+-comm n 2 i) - -coHomGrUnit≥1 : (n : ℕ) → grIso (coHomGr (suc n) Unit) trivialGroup -grIso.fun (coHomGrUnit≥1 n) = mph (λ _ → tt) (λ _ _ → refl) -grIso.inv (coHomGrUnit≥1 n) = mph (λ _ → ∣ (λ _ → ∣ north ∣) ∣₀) (λ _ _ → sym (rUnitₕ 0ₕ)) -grIso.rightInv (coHomGrUnit≥1 n) a = refl -grIso.leftInv (coHomGrUnit≥1 n) a = sym (isContrCohomUnit n .snd 0ₕ) ∙ isContrCohomUnit n .snd a - -S0→Int : (a : Int × Int) → S₊ 0 → Int -S0→Int a north = proj₁ a -S0→Int a south = proj₂ a - -coHom0-S0 : grIso (coHomGr 0 (S₊ 0)) (dirProd intGroup intGroup) -coHom0-S0 = - Iso'→Iso - (iso' - (iso (sRec (isOfHLevelProd 2 isSetInt isSetInt) - λ f → (f north) , (f south)) - (λ a → ∣ S0→Int a ∣₀) - (λ { (a , b) → refl}) - (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ f → cong ∣_∣₀ (funExt (λ {north → refl ; south → refl})))) - (sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 isSetInt isSetInt) _ _) - λ a b i → addLemma (a north) (b north) i , addLemma (a south) (b south) i)) - - - - -coHom0S1 : grIso intGroup (coHomGr 0 (S₊ 1)) -coHom0S1 = - Iso'→Iso - (iso' - (iso - (λ a → ∣ (λ x → a) ∣₀) - (sRec isSetInt (λ f → f north)) - (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) - λ f → cong ∣_∣₀ (funExt (suspToPropRec north (λ _ → isSetInt _ _) refl))) - (λ a → refl)) - λ a b → cong ∣_∣₀ (funExt λ x → sym (addLemma a b))) - -coHom1S1 : grIso intGroup (coHomGr 1 (S₊ 1)) -coHom1S1 = - compGrIso - (diagonalIso1 - _ - (coHomGr 0 (S₊ 0)) - _ - (invGrIso coHom0-S0) - (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) - (λ x → MV.Ker-i⊂Im-d _ _(S₊ 0) (λ _ → tt) (λ _ → tt) 0 x - (×≡ (isOfHLevelSuc 0 (isContrCohomUnit 0) _ _) - (isOfHLevelSuc 0 (isContrCohomUnit 0) _ _))) - (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) - (λ x inker - → pRec propTruncIsProp - (λ {((f , g) , id') → helper x f g id' inker}) - ((MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ inker)))) - (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) - λ F surj - → pRec (setTruncIsSet _ _) (λ { (x , id) → MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ F ∣₀ ∣ (∣ (λ _ → x) ∣₀ , ∣ (λ _ → 0) ∣₀) , - (cong ∣_∣₀ (funExt (surjHelper x))) ∙ sym id ∣₋₁ }) surj) ) - (invGrIso (coHomPushout≡coHomSn' 0 1)) - - where - surjHelper : (x : Int) (x₁ : S₊ 0) → x +ₖ (-ₖ pos 0) ≡ S0→Int (x , x) x₁ - surjHelper x north = cong (x +ₖ_) (-0ₖ) ∙ rUnitₖ x - surjHelper x south = cong (x +ₖ_) (-0ₖ) ∙ rUnitₖ x - - helper : (F : S₊ 0 → Int) (f g : ∥ (Unit → Int) ∥₀) (id : morph.fun (MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) (f , g) ≡ ∣ F ∣₀) - → isInKer (coHomGr 0 (S₊ 0)) - (coHomGr 1 (Pushout (λ _ → tt) (λ _ → tt))) - (morph.fun (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) - ∣ F ∣₀ - → ∃[ x ∈ Int ] ∣ F ∣₀ ≡ morph.fun (grIso.inv coHom0-S0) (x , x) - helper F = - sElim2 (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) - λ f g id inker - → pRec propTruncIsProp - (λ {((a , b) , id2) - → sElim2 {B = λ f g → morph.fun (MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) (f , g) ≡ ∣ F ∣₀ → _ } - (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) - (λ f g id → ∣ (helper2 f g .fst) , (sym id ∙ sym (helper2 f g .snd)) ∣₋₁) - a b id2}) - (MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ F ∣₀ inker) - where - helper2 : (f g : Unit → Int) - → Σ[ x ∈ Int ] morph.fun (grIso.inv coHom0-S0) (x , x) - ≡ morph.fun (MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) (∣ f ∣₀ , ∣ g ∣₀) - helper2 f g = (f _ +ₖ (-ₖ g _) ) , cong ∣_∣₀ (funExt (λ {north → refl ; south → refl})) - - - -coHom-n-Sn : (n : ℕ) → grIso intGroup (coHomGr (suc n) (S₊ (suc n))) -coHom-n-Sn zero = coHom1S1 -coHom-n-Sn (suc n) = - compGrIso - (compGrIso - (coHom-n-Sn n) - theIso) - (invGrIso (coHomPushout≡coHomSn' (suc n) (suc (suc n)))) - where - theIso : grIso (coHomGr (suc n) (S₊ (suc n))) (coHomGr (suc (suc n)) - (Pushout {A = S₊ (suc n)} (λ _ → tt) (λ _ → tt))) - theIso = - SES→Iso - (×coHomGr (suc n) Unit Unit) - (×coHomGr (suc (suc n)) Unit Unit) - (ses (λ p q → ×≡ (isOfHLevelSuc 0 (isContrCohomUnit n) (proj₁ p) (proj₁ q)) - (isOfHLevelSuc 0 (isContrCohomUnit n) (proj₂ p) (proj₂ q))) - (λ p q → ×≡ (isOfHLevelSuc 0 (isContrCohomUnit (suc n)) (proj₁ p) (proj₁ q)) - (isOfHLevelSuc 0 (isContrCohomUnit (suc n)) (proj₂ p) (proj₂ q))) - (MV.Δ _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) - (MV.i _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (2 + n)) - (MV.d _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) - (MV.Ker-d⊂Im-Δ _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n)) - (MV.Ker-i⊂Im-d _ _ (S₊ (suc n)) (λ _ → tt) (λ _ → tt) (suc n))) - - -coHom1-s0 : grIso (coHomGr 1 (S₊ 0)) trivialGroup -coHom1-s0 = - Iso'→Iso - (iso' - (iso (λ _ → tt) - (λ _ → 0ₕ) - (λ _ → refl) - (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) - λ f → pRec (setTruncIsSet _ _) - (λ id → cong ∣_∣₀ (sym id)) - (helper f (f north) (f south) refl refl))) - λ _ _ → refl) - where - helper : (f : S₊ 0 → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1) → (x y : ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1) → (f north ≡ x) → (f south ≡ y) → ∥ f ≡ (λ _ → 0ₖ) ∥₋₁ - helper f = - elim2 (λ _ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelPlus {n = 1} 2 propTruncIsProp) - (suspToPropRec2 north (λ _ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → propTruncIsProp) - λ id id2 → ∣ funExt (λ {north → id ; south → id2}) ∣₋₁) - -coHom2-S1 : grIso (coHomGr 2 (S₊ 1)) trivialGroup -coHom2-S1 = - compGrIso - (coHomPushout≡coHomSn' 0 2) - (compGrIso - (invGrIso - (SES→Iso - (×coHomGr 1 Unit Unit) - (×coHomGr 2 Unit Unit) - (ses (isOfHLevelProd 1 (isOfHLevelSuc 0 (isContrCohomUnit 0)) (isOfHLevelSuc 0 (isContrCohomUnit 0))) - (isOfHLevelProd 1 (isOfHLevelSuc 0 (isContrCohomUnit 1)) (isOfHLevelSuc 0 (isContrCohomUnit 1))) - (MV.Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1) - (MV.i _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 2) - (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1) - (MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1) - ({! MV.Ker-i⊂Im-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1 !})))) - coHom1-s0) - - - - -coHom1Torus : grIso (coHomGr 1 ((S₊ 1) × (S₊ 1))) (dirProd intGroup intGroup) -coHom1Torus = - compGrIso - (Iso'→Iso - (iso' (compIso - (setTruncIso (compIso - schönfinkelIso - (compIso - (codomainIso S1→S1→S1×Int) - toProdIso))) - (setTruncOfProdIso)) - (sElim2 - (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) - λ f g → ×≡ (cong ∣_∣₀ - (funExt (λ x → helper (f (x , S¹→S1 base) +ₖ g (x , S¹→S1 base)) - ∙ sym (cong₂ (λ x y → x +ₖ y) - (helper (f (x , S¹→S1 base))) - (helper (g (x , S¹→S1 base))))))) - (cong ∣_∣₀ - (funExt - (suspToPropRec - north - (λ _ → isSetInt _ _) - (cong winding - (basechange-lemma2 - (λ x → f (north , S¹→S1 x)) - (λ x → g (north , S¹→S1 x)) - λ x → S¹map (trMap S1→S¹ x)) - ∙ winding-hom - (basechange2⁻ - (S¹map (trMap S1→S¹ (f (north , S¹→S1 base)))) - (λ i → S¹map (trMap S1→S¹ (f (north , S¹→S1 (loop i)))))) - (basechange2⁻ - (S¹map (trMap S1→S¹ (g (north , S¹→S1 base)))) - (λ i → S¹map (trMap S1→S¹ (g (north , S¹→S1 (loop i)))))) - ∙ sym (addLemma - (winding - (basechange2⁻ - (S¹map (trMap S1→S¹ (f (north , S¹→S1 base)))) - (λ i → S¹map (trMap S1→S¹ (f (north , S¹→S1 (loop i))))))) - (winding - (basechange2⁻ - (S¹map (trMap S1→S¹ (g (north , S¹→S1 base)))) - (λ i → S¹map (trMap S1→S¹ (g (north , S¹→S1 (loop i))))))))))))))) - (dirProdIso (invGrIso (coHom-n-Sn 0)) (invGrIso coHom0S1)) - - where - helper : (x : hLevelTrunc 3 (S₊ 1)) → ∣ S¹→S1 (S¹map (trMap S1→S¹ x)) ∣ ≡ x - helper = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ a → cong ∣_∣ (S1→S¹-retr a) - - -coHom2Torus : grIso (coHomGr 2 ((S₊ 1) × (S₊ 1))) intGroup -coHom2Torus = - invGrIso - (Iso'→Iso - (iso' - (compIso - helper' - (compIso - (symIso (prodIso (groupIso→Iso coHom2-S1) - (groupIso→Iso (invGrIso (coHom-n-Sn 0))))) - (compIso - (symIso setTruncOfProdIso) - (symIso - (setTruncIso - (compIso - schönfinkelIso - (compIso - (codomainIso test13) - toProdIso))))))) - λ a b → {!!} {- (λ i → Iso.fun - (symIso - (setTruncIso - (compIso schönfinkelIso (compIso (codomainIso test13) toProdIso)))) - (Iso.fun (symIso setTruncOfProdIso) - (helper'' (Iso.inv (groupIso→Iso coHom2-S1) tt) 0ₕ i , morph.ismorph (grIso.inv (invGrIso coHom1S1)) a b i))) - ∙ (λ i → Iso.fun - (symIso - (setTruncIso - (compIso schönfinkelIso (compIso (codomainIso test13) toProdIso)))) - ∣ (λ _ → ∣ north ∣) , {!!} ∣₀) - ∙ {!toProdIso!} -})) - where - helper'' : isProp ∥ (S₊ 1 → hLevelTrunc 4 (S₊ 2)) ∥₀ - helper'' = {!!} - - helper2' : (f g : (S₊ 1 → hLevelTrunc 3 (S₊ 1))) → - Path (coHom 2 ((S₊ 1) × (S₊ 1))) (Iso.fun - (symIso - (setTruncIso - (compIso schönfinkelIso (compIso (codomainIso test13) toProdIso)))) - (Iso.fun (symIso setTruncOfProdIso) (0ₕ , ∣ f ∣₀ +ₕ ∣ g ∣₀))) - (Iso.fun (symIso - (setTruncIso - (compIso schönfinkelIso (compIso (codomainIso test13) toProdIso)))) - (Iso.fun (symIso setTruncOfProdIso) (0ₕ , ∣ f ∣₀)) +ₕ Iso.fun (symIso - (setTruncIso - (compIso schönfinkelIso (compIso (codomainIso test13) toProdIso)))) - (Iso.fun (symIso setTruncOfProdIso) (0ₕ , ∣ g ∣₀))) - helper2' f g = (λ i → Iso.fun (symIso - (setTruncIso - (compIso schönfinkelIso ((codomainIso test13))))) ∣ (λ x → 0ₖ , (f x +ₖ g x)) ∣₀) - ∙ (λ i → ∣ Iso.fun (symIso (compIso schönfinkelIso (codomainIso test13))) (λ x → 0ₖ , (f x +ₖ g x)) ∣₀) - ∙ (λ i → ∣ Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , (f x +ₖ g x)))) ∣₀) - ∙ (λ i → ∣ {!!} ∣₀) - ∙ {!!} - ∙ {!!} - where - fun : S₊ 1 × S₊ 1 → hLevelTrunc 4 (S₊ 2) - fun (x , north) = 0ₖ +ₖ 0ₖ - fun (x , south) = 0ₖ +ₖ 0ₖ - fun (x , (merid south i)) = 0ₖ +ₖ (Kn→ΩKn+1 1 (f x +ₖ g x) i) - fun (x , (merid north i)) = 0ₖ +ₖ 0ₖ - - helper : Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , (f x +ₖ g x)))) ≡ fun - helper = funExt λ {(x , north) → refl ; (x , south) → refl ; (x , (merid north i)) → refl ; (x , (merid south i)) → refl} - - helper2 : ∣ Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , (f x +ₖ g x)))) ∣₀ - ≡ (∣ Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , f x))) ∣₀ +ₕ ∣ Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , g x))) ∣₀) - helper2 = - cong ∣_∣₀ - (funExt λ {(x , north) → ((λ i → (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x +ₖ g x) i)) - ∙∙ sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ cong (λ x → ((0ₖ +ₖ 0ₖ) +ₖ x)) (rUnitₖ 0ₖ) - ∙∙ refl) - ; (x , south) → refl - ∙∙ sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ cong (λ x → ((0ₖ +ₖ 0ₖ) +ₖ x)) (rUnitₖ 0ₖ) - ∙∙ (λ i → (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x) i) +ₖ ∣ north ∣ +ₖ Kn→ΩKn+1 1 (g x) i) - ; (x , (merid south i)) j → hcomp (λ k → λ {(j = i0) → ∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x +ₖ g x) (i ∨ (~ k)) - ; (j = i1) → (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x) (i ∧ k)) +ₖ ∣ north ∣ +ₖ Kn→ΩKn+1 1 (g x) (i ∧ k)}) - ((sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ cong (λ x → ((0ₖ +ₖ 0ₖ) +ₖ x)) (rUnitₖ 0ₖ)) j) - ; (x , (merid north u)) → {!!}}) - where - anotherone : (x : _) → cong (0ₖ +ₖ_) (Kn→ΩKn+1 1 (f x +ₖ g x)) ≡ {!!} - anotherone x = {!!} - - helper5 : (x : _) → Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , f x +ₖ g x))) x - ≡ (Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , f x))) x) +ₖ (Iso.inv schönfinkelIso (Iso.inv (codomainIso test13) ((λ x → 0ₖ , g x))) x) - helper5 (x , north) = sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i))) - helper5 (x , south) = sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i))) - helper5 (x , merid north i) = sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i))) - helper5 (x , merid south i) j = - {!!} - where - helper13 : sym (sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i)))) - ∙ (λ i → ∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x +ₖ g x) i) ∙ (sym (rUnitₖ (0ₖ +ₖ 0ₖ)) - ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i)))) - ≡ λ i → (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x) i) +ₖ - ∣ north ∣ +ₖ Kn→ΩKn+1 1 (g x) i - helper13 = (λ j → sym ((λ i → (rUnitₖ (lUnitₖ 0ₖ (j ∧ (~ i))) (~ i))) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ lUnitₖ 0ₖ (~ i))) - ∙ lUnit (rUnit ((λ i → lUnitₖ (Kn→ΩKn+1 1 (f x +ₖ g x) i) j)) j) j - ∙ ((λ i → (rUnitₖ (lUnitₖ 0ₖ (j ∧ (~ i))) (~ i)))) ∙ λ i → (0ₖ +ₖ 0ₖ) +ₖ lUnitₖ 0ₖ (~ i)) - ∙ (λ j → sym ((λ i → (rUnitₖ (lUnitₖ 0ₖ (~ i)) (~ i))) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ lUnitₖ 0ₖ (~ i))) ∙ - {!!} ∙ - (λ i → (rUnitₖ (lUnitₖ 0ₖ (~ i)) (~ i))) ∙ λ i → (0ₖ +ₖ 0ₖ) +ₖ lUnitₖ 0ₖ (~ i)) - ∙ {!!} - ∙ {!!} -{- -i = i0 ⊢ (sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i)))) j -i = i1 ⊢ (sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i)))) j -j = i0 ⊢ ∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x +ₖ g x) i -j = i1 ⊢ (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x) i) +ₖ - ∣ north ∣ +ₖ Kn→ΩKn+1 1 (g x) i --} - - - helper' : Iso Int (Unit × Int) - Iso.inv helper' = proj₂ - Iso.fun helper' x = tt , x - Iso.leftInv helper' x = refl - Iso.rightInv helper' x = ×≡ refl refl - - helper : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} → Iso C Unit → Iso A B → Iso (A × C) B - Iso.fun (helper cUnit iAB) x = Iso.fun iAB (proj₁ x) - Iso.inv (helper cUnit iAB) x = (Iso.inv iAB x) , (Iso.inv cUnit tt ) - Iso.rightInv (helper cUnit iAB) = Iso.rightInv iAB - Iso.leftInv (helper cUnit iAB) b = ×≡ (Iso.leftInv iAB (proj₁ b)) (Iso.leftInv cUnit (proj₂ b)) - - helper2 : Iso (coHom 2 ((S₊ 1) × (S₊ 1))) (coHom 2 ((S₊ 1) × hLevelTrunc 3 (S₊ 1))) - Iso.fun helper2 = sRec setTruncIsSet (λ f → ∣ (λ {(x , y) → f (x , trRec isGroupoidS1 (idfun (S₊ 1)) y)}) ∣₀) - Iso.inv helper2 = sRec setTruncIsSet (λ f → ∣ (λ {(x , y) → f (x , ∣ y ∣)}) ∣₀) - Iso.rightInv helper2 = - sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) - (λ f → cong ∣_∣₀ - (funExt λ {(x , y) → - trElim {B = λ y → f (x , ∣ trRec isGroupoidS1 (λ x₁ → x₁) y ∣) ≡ f (x , y)} - (λ _ → isOfHLevelPath' 3 (isOfHLevelTrunc 4) _ _) - (λ a → refl) y})) - Iso.leftInv helper2 = - sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) - λ f → cong ∣_∣₀ (funExt λ {(x , y) → refl}) - --- basechange* : (x y : S¹) → x ≡ y → x ≡ y → ΩS¹ --- basechange* x y = J (λ y p → (x ≡ y) → ΩS¹) (basechange x) - - --- test1 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S₊ 1 × Int) --- Iso.fun test1 f = (trRec isGroupoidS1 (λ a → a) (f north)) --- , winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (f (loop* i))))) --- Iso.inv test1 (north , b) x = ∣ x ∣ --- Iso.inv test1 (south , b) x = ∣ x ∣ --- Iso.inv test1 (merid a i , b) x = {!!} --- Iso.rightInv test1 = {!!} --- Iso.leftInv test1 = {!!} - --- funRec : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) {C : (A → hLevelTrunc n B) → Type ℓ''} --- → isOfHLevel n B --- → ((f : A → B) → C (λ a → ∣ f a ∣)) --- → (f : A → hLevelTrunc n B) → C f --- funRec {A = A} {B = B} n {C = C} hLev ind f = subst C (helper f) (ind (λ a → trRec hLev (λ x → x) (f a))) --- where --- helper : retract {A = A → hLevelTrunc n B} {B = A → B} (λ f₁ a → trRec hLev (λ x → x) (f₁ a)) (λ f₁ a → ∣ f₁ a ∣) --- helper f = funExt λ a → helper2 (f a) --- where --- helper2 : (x : hLevelTrunc n B) → ∣ trRec hLev (λ x → x) x ∣ ≡ x --- helper2 = trElim (λ _ → isOfHLevelPath n (isOfHLevelTrunc n) _ _) λ a → refl - --- invMapSurj : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (ϕ : morph G H) → ((x : Group.type H) → isInIm G H (fst ϕ) x) --- → morph H G --- fst (invMapSurj G H (ϕ , pf) surj) a = {!pRec!} --- snd (invMapSurj G H (ϕ , pf) surj) = {!!} - -{- -ImIso : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (ϕ : morph G H) → ((x : Group.type H) → isInIm G H (fst ϕ) x) - → grIso H (imGroup G H ϕ) -ImIso G H (ϕ , mf) surj = - let idH = isGroup.id (Group.groupStruc H) - idG = isGroup.id (Group.groupStruc G) - _+G_ = isGroup.comp (Group.groupStruc G) - _+H_ = isGroup.comp (Group.groupStruc H) - _+Im_ = isGroup.comp (Group.groupStruc (imGroup G H (ϕ , mf))) - invG = isGroup.inv (Group.groupStruc G) - invH = isGroup.inv (Group.groupStruc H) - lUnit = isGroup.lUnit (Group.groupStruc H) - lCancel = isGroup.rCancel (Group.groupStruc H) - in - Iso''→Iso _ _ - (iso'' ((λ x → x , pRec propTruncIsProp (λ (a , b) → ∣ a , b ∣₋₁) (surj x)) - , λ a b → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) - (λ surja → pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) - (λ surjb → - pRec (Group.setStruc (imGroup G H (ϕ , mf)) _ _) - (λ surja+b → - (λ i → (a +H b) , (pRec (propTruncIsProp) - (λ (a , b) → ∣ a , b ∣₋₁) - (propTruncIsProp (surj (isGroup.comp (Group.groupStruc H) a b)) ∣ surja+b ∣₋₁ i))) ∙ - (λ i → (a +H b) , ∣ (fst surja+b) , (snd surja+b) ∣₋₁) ∙ - ΣProp≡ (λ _ → propTruncIsProp) refl ∙ - λ i → (a +H b) , pRec (propTruncIsProp) - (λ p1 → - pRec (λ x y → squash x y) - (λ p2 → - ∣ - isGroup.comp (Group.groupStruc G) (fst p1) (fst p2) , - mf (fst p1) (fst p2) ∙ - cong₂ (isGroup.comp (Group.groupStruc H)) (snd p1) (snd p2) - ∣₋₁) - (pRec (propTruncIsProp) - ∣_∣₋₁ (propTruncIsProp ∣ surjb ∣₋₁ (surj b) i))) - (pRec (propTruncIsProp) - ∣_∣₋₁ (propTruncIsProp ∣ surja ∣₋₁ (surj a) i ))) - (surj (isGroup.comp (Group.groupStruc H) a b))) - (surj b)) - (surj a)) - (λ x inker → cong fst inker) - λ x → pRec propTruncIsProp (λ inimx → ∣ (ϕ (inimx .fst)) , ΣProp≡ (λ _ → propTruncIsProp) (inimx .snd) ∣₋₁) (snd x)) --} - - -{- -H¹-S¹≃Int : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) -H¹-S¹≃Int = - Iso''→Iso _ _ - (iso'' ((λ x → ∣ theFuns x ∣₀) , λ a b → cong ∣_∣₀ (funExt λ x → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 1) _) ∙ sym (cong (ΩKn+1→Kn) (theFunsId2 a b x)))) - (λ x inker → pRec (isSetInt _ _) (inj x) (Iso.fun PathIdTrunc₀Iso inker)) - finIm) - where - d : _ - d = MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 - - i : _ - i = MV.i _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1 - - Δ : _ - Δ = MV.Δ _ _ (S₊ 1) (λ _ → tt) (λ _ → tt) 0 - - - d-surj : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) - → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x - d-surj = sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) - λ x → MV.Ker-i⊂Im-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ - (sym (isContrHelper .snd _)) - where - isContrHelper : isContr (Group.type (×coHomGr 1 Unit Unit)) - isContrHelper = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) - , λ y → ×≡ (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₁ y)) - (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₂ y)) - - theFuns : (a : Int) → Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 - theFuns a (inl x) = 0ₖ - theFuns a (inr x) = 0ₖ - theFuns a (push north i) = Kn→ΩKn+1 0 a i - theFuns a (push south i) = 0ₖ - - - theFunsId2 : (a b : Int) (x : Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) - → Kn→ΩKn+1 1 (theFuns a x) ∙ Kn→ΩKn+1 1 (theFuns b x) ≡ Kn→ΩKn+1 1 (theFuns (a +ℤ b) x) - theFunsId2 a b (inl x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) - theFunsId2 a b (inr x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) - theFunsId2 a b (push north i) j = - hcomp (λ k → λ {(i = i0) → ((λ i₁ → - (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) - ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) - j - ; (i = i1) → ((λ i₁ → - (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) - ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) - j - ; (j = i0) → cong₂Funct (λ p q → Kn→ΩKn+1 1 p ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b) (~ k) i - ; (j = i1) → (helper2 a b) k i }) - (hcomp (λ k → λ { (j = i0) → compPath-filler (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) k i - ; (j = i1) → compPath-filler (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) k i - ; (i = i1) → RHS-filler b j k - ; (i = i0) → ((λ i₁ → - (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) - ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) - j}) - (bottom-filler a j i)) - - where - - bottom-filler : (a : Int) → - PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) - (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) - ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) - j ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) - (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) - ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) - j) (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) - bottom-filler a j i = - hcomp (λ k → λ {(j = i0) → helper2 (~ k) i ; - (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 a) i}) - ((λ j₂ → ∣ rCancel (merid north) j j₂ ∣) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) - - where - helper2 : Path (Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣ ≡ Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣) - (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i) ∙ Kn→ΩKn+1 1 ∣ north ∣) - (λ i → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) - helper2 = cong₂Sym1 (Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) (~ i) j ∣) (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) - - RHS-filler : (b : Int) → - PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j - ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j) - (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) - (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) - RHS-filler b j i = - hcomp (λ k → λ {(j = i0) → cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i ; - (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 b) i}) - (cong (λ q → (λ i → ∣ rCancel (merid north) j i ∣) ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i) - - helper2 : (a b : Int) → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b) ≡ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ℤ b)) - helper2 a b = - sym (congFunct (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b)) - ∙ (λ i → cong (Kn→ΩKn+1 1) (Iso.rightInv (Iso3-Kn-ΩKn+1 0) (Kn→ΩKn+1 0 a ∙ Kn→ΩKn+1 0 b) (~ i))) - ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ₖ b)) ) - ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (addLemma a b i))) - - theFunsId2 a b (push south i) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) - ∙ sym (lUnit _) - - inj : (a : Int) → theFuns a ≡ (λ _ → ∣ north ∣) → a ≡ pos 0 - inj a id = - pRec (isSetInt _ _) - (sigmaElim (λ _ → isOfHLevelPath 2 isSetInt _ _) - (λ a p → pRec (isSetInt _ _) - (λ id2 → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 0) _) - ∙ cong (ΩKn+1→Kn) - (PathP→compPathR - (cong (λ f → cong f (push north)) id) - ∙ test)) - (Iso.fun PathIdTrunc₀Iso p))) (d-surj ∣ theFuns a ∣₀) - where - - test : (λ i → id i (inl tt)) ∙ (λ i → ∣ north ∣) ∙ sym (λ i → id i (inr tt)) ≡ refl - test = (λ i → cong (λ f → f (inl tt)) id - ∙ lUnit (sym (cong (λ f → f (inr tt)) id)) (~ i)) - ∙ (λ i → cong (λ f → f (push south i)) id - ∙ sym (cong (λ f → f (inr tt)) id)) - ∙ rCancel (cong (λ f → f (inr tt)) id) - - - consMember : (a : Int) → coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) - consMember a = d ∣ (λ _ → a) ∣₀ - - consMember≡0 : (a : Int) → consMember a ≡ 0ₕ - consMember≡0 a = - (MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ (λ _ → a) ∣₀ ∣ - (∣ (λ _ → a) ∣₀ , ∣ (λ _ → 0) ∣₀) - , cong ∣_∣₀ (λ i x → (rUnitₖ a i)) ∣₋₁) - - f+consMember' : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int × Int ] (f +ₕ (-ₕ (consMember (proj₁ x))) ≡ ∣ theFuns (proj₂ x) ∣₀) - f+consMember' = - sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) - λ f → pRec propTruncIsProp - (sigmaElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) - (λ g id → ∣ ((g south) , ((g north) +ₖ (-ₖ g south))) - , (pRec (setTruncIsSet _ _) - (λ id → (λ i → ∣ id (~ i) ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀) ∙ funId1 g) - (Iso.fun PathIdTrunc₀Iso id)) ∣₋₁)) - (d-surj ∣ f ∣₀) - where - funId1 : (g : S₊ 0 → Int) - → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀ ≡ - ∣ theFuns ((g north) +ₖ (-ₖ (g south))) ∣₀ - funId1 g = (λ i → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ - +ₕ (morphMinus (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) d - (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ∣ (λ _ → g south) ∣₀ (~ i))) - ∙ sym (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ g ∣₀ (-ₕ ∣ (λ _ → g south) ∣₀)) - ∙ (cong (λ x → d ∣ x ∣₀) g'Id) - ∙ cong ∣_∣₀ helper - where - g' : S₊ 0 → Int - g' north = (g north) +ₖ (-ₖ (g south)) - g' south = 0 - - g'Id : (λ x → g x +ₖ (-ₖ (g south))) ≡ g' - g'Id = funExt (λ {north → refl - ; south → rCancelₖ (g south)}) - - helper : MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g' ≡ theFuns (g north +ₖ (-ₖ g south)) - helper = funExt λ {(inl tt) → refl - ; (inr tt) → refl - ; (push north i) → refl - ; (push south i) → refl} - finIm : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int ] (∣ theFuns x ∣₀ ≡ f) - finIm f = - pRec propTruncIsProp - (λ {((a , b) , id) → ∣ b , (sym id ∙ cong (λ x → f +ₕ x) ((λ i → (-ₕ (consMember≡0 a i))) ∙ sym (lUnitₕ (-ₕ 0ₕ)) ∙ rCancelₕ 0ₕ) ∙ (rUnitₕ f)) ∣₋₁}) - (f+consMember' f) --} - - - --- Hⁿ-Sⁿ≃Int : (n : ℕ) → grIso intGroup (coHomGr (suc n) (S₊ (suc n))) --- Hⁿ-Sⁿ≃Int zero = --- compGrIso {F = intGroup} {G = {!!}} {H = {!coHomGr 1 (S₊ 1)!}} --- (Iso''→Iso --- (iso'' ((λ x → ∣ theFuns x ∣₀) , λ a b → cong ∣_∣₀ (funExt λ x → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 1) _) ∙ sym (cong (ΩKn+1→Kn) (theFunsId2 a b x)))) --- (λ x inker → pRec (isSetInt _ _) (inj x) (Iso.fun PathIdTrunc₀Iso inker)) --- finIm)) --- {!invGrIso _ _ (coHomPushout≡coHomSn 0 1)!} --- where --- d : _ --- d = MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 - --- i : _ --- i = MV.i _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 1 - --- Δ : _ --- Δ = MV.Δ _ _ (S₊ 1) (λ _ → tt) (λ _ → tt) 0 - - --- d-surj : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x --- d-surj = sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) --- λ x → MV.Ker-i⊂Im-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₀ --- (sym (isContrHelper .snd _)) --- where --- isContrHelper : isContr (Group.type (×coHomGr 1 Unit Unit)) --- isContrHelper = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) --- , λ y → ×≡ (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₁ y)) --- (cong ∣_∣₀ (λ i _ → ∣ merid north i ∣) ∙ isContrCohomUnit 0 .snd (proj₂ y)) - --- theFuns : (a : Int) → Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 --- theFuns a (inl x) = 0ₖ --- theFuns a (inr x) = 0ₖ --- theFuns a (push north i) = Kn→ΩKn+1 0 a i --- theFuns a (push south i) = 0ₖ - - --- theFunsId2 : (a b : Int) (x : Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) --- → Kn→ΩKn+1 1 (theFuns a x) ∙ Kn→ΩKn+1 1 (theFuns b x) ≡ Kn→ΩKn+1 1 (theFuns (a +ℤ b) x) --- theFunsId2 a b (inl x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) --- theFunsId2 a b (inr x) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)) --- theFunsId2 a b (push north i) j = --- hcomp (λ k → λ {(i = i0) → ((λ i₁ → --- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) --- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) --- j --- ; (i = i1) → ((λ i₁ → --- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) --- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) --- j --- ; (j = i0) → cong₂Funct (λ p q → Kn→ΩKn+1 1 p ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b) (~ k) i --- ; (j = i1) → (helper2 a b) k i }) --- (hcomp (λ k → λ { (j = i0) → compPath-filler (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) k i --- ; (j = i1) → compPath-filler (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) k i --- ; (i = i1) → RHS-filler b j k --- ; (i = i0) → ((λ i₁ → --- (λ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) --- ∙ (λ i₁ → lUnit (Kn→ΩKn+1 1 ∣ north ∣) (~ i₁))) --- j}) --- (bottom-filler a j i)) - --- where - --- bottom-filler : (a : Int) → --- PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) --- (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) --- ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) --- j ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) --- (λ i₁ j₁ → ∣ rCancel (merid north) i₁ j₁ ∣) --- ∙ sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣))) --- j) (cong (λ x₁ → Kn→ΩKn+1 1 x₁ ∙ Kn→ΩKn+1 1 ∣ north ∣) (Kn→ΩKn+1 0 a)) (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a)) --- bottom-filler a j i = --- hcomp (λ k → λ {(j = i0) → helper2 (~ k) i ; --- (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 a) i}) --- ((λ j₂ → ∣ rCancel (merid north) j j₂ ∣) ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) - --- where --- helper2 : Path (Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣ ≡ Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 ∣ north ∣) --- (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i) ∙ Kn→ΩKn+1 1 ∣ north ∣) --- (λ i → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) --- helper2 = cong₂Sym1 (Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) (~ i) j ∣) (λ i → Kn→ΩKn+1 1 (Kn→ΩKn+1 0 a i)) - --- RHS-filler : (b : Int) → --- PathP (λ j → (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j --- ≡ (cong (λ x → x ∙ Kn→ΩKn+1 1 ∣ north ∣) (λ i j → ∣ rCancel (merid north) i j ∣) ∙ (sym (lUnit (Kn→ΩKn+1 1 ∣ north ∣)))) j) --- (cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b)) --- (cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b)) --- RHS-filler b j i = --- hcomp (λ k → λ {(j = i0) → cong (λ q → Kn→ΩKn+1 1 ∣ north ∣ ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i ; --- (j = i1) → cong (λ x → lUnit (Kn→ΩKn+1 1 x) (~ k)) (Kn→ΩKn+1 0 b) i}) --- (cong (λ q → (λ i → ∣ rCancel (merid north) j i ∣) ∙ Kn→ΩKn+1 1 q) (Kn→ΩKn+1 0 b) i) - --- helper2 : (a b : Int) → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) ∙ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 b) ≡ cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ℤ b)) --- helper2 a b = --- sym (congFunct (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 a) (Kn→ΩKn+1 0 b)) --- ∙ (λ i → cong (Kn→ΩKn+1 1) (Iso.rightInv (Iso3-Kn-ΩKn+1 0) (Kn→ΩKn+1 0 a ∙ Kn→ΩKn+1 0 b) (~ i))) --- ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (a +ₖ b)) ) --- ∙ (λ i → cong (Kn→ΩKn+1 1) (Kn→ΩKn+1 0 (addLemma a b i))) - --- theFunsId2 a b (push south i) = (λ i → (λ j → ∣ rCancel (merid north) i j ∣) ∙ Kn→ΩKn+1 1 ∣ north ∣) --- ∙ sym (lUnit _) - --- inj : (a : Int) → theFuns a ≡ (λ _ → ∣ north ∣) → a ≡ pos 0 --- inj a id = --- pRec (isSetInt _ _) --- (sigmaElim (λ _ → isOfHLevelPath 2 isSetInt _ _) --- (λ a p → pRec (isSetInt _ _) --- (λ id2 → sym (Iso.leftInv (Iso3-Kn-ΩKn+1 0) _) --- ∙ cong (ΩKn+1→Kn) --- (PathP→compPathR --- (cong (λ f → cong f (push north)) id) --- ∙ test)) --- (Iso.fun PathIdTrunc₀Iso p))) (d-surj ∣ theFuns a ∣₀) --- where - --- test : (λ i → id i (inl tt)) ∙ (λ i → ∣ north ∣) ∙ sym (λ i → id i (inr tt)) ≡ refl --- test = (λ i → cong (λ f → f (inl tt)) id --- ∙ lUnit (sym (cong (λ f → f (inr tt)) id)) (~ i)) --- ∙ (λ i → cong (λ f → f (push south i)) id --- ∙ sym (cong (λ f → f (inr tt)) id)) --- ∙ rCancel (cong (λ f → f (inr tt)) id) - - --- consMember : (a : Int) → coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) --- consMember a = d ∣ (λ _ → a) ∣₀ - --- consMember≡0 : (a : Int) → consMember a ≡ 0ₕ --- consMember≡0 a = --- (MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ (λ _ → a) ∣₀ ∣ --- (∣ (λ _ → a) ∣₀ , ∣ (λ _ → 0) ∣₀) --- , cong ∣_∣₀ (λ i x → (rUnitₖ a i)) ∣₋₁) - --- f+consMember' : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int × Int ] (f +ₕ (-ₕ (consMember (proj₁ x))) ≡ ∣ theFuns (proj₂ x) ∣₀) --- f+consMember' = --- sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) --- λ f → pRec propTruncIsProp --- (sigmaElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) --- (λ g id → ∣ ((g south) , ((g north) +ₖ (-ₖ g south))) --- , (pRec (setTruncIsSet _ _) --- (λ id → (λ i → ∣ id (~ i) ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀) ∙ funId1 g) --- (Iso.fun PathIdTrunc₀Iso id)) ∣₋₁)) --- (d-surj ∣ f ∣₀) --- where --- funId1 : (g : S₊ 0 → Int) --- → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ +ₕ -ₕ ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 (λ _ → g south) ∣₀ ≡ --- ∣ theFuns ((g north) +ₖ (-ₖ (g south))) ∣₀ --- funId1 g = (λ i → ∣ MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g ∣₀ --- +ₕ (morphMinus (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) d --- (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ∣ (λ _ → g south) ∣₀ (~ i))) --- ∙ sym (MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ g ∣₀ (-ₕ ∣ (λ _ → g south) ∣₀)) --- ∙ (cong (λ x → d ∣ x ∣₀) g'Id) --- ∙ cong ∣_∣₀ helper --- where --- g' : S₊ 0 → Int --- g' north = (g north) +ₖ (-ₖ (g south)) --- g' south = 0 - --- g'Id : (λ x → g x +ₖ (-ₖ (g south))) ≡ g' --- g'Id = funExt (λ {north → refl --- ; south → rCancelₖ (g south)}) - --- helper : MV.d-pre Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 g' ≡ theFuns (g north +ₖ (-ₖ g south)) --- helper = funExt λ {(inl tt) → refl --- ; (inr tt) → refl --- ; (push north i) → refl --- ; (push south i) → refl} --- finIm : (f : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) → ∃[ x ∈ Int ] (∣ theFuns x ∣₀ ≡ f) --- finIm f = --- pRec propTruncIsProp --- (λ {((a , b) , id) → ∣ b , (sym id ∙ cong (λ x → f +ₕ x) ((λ i → (-ₕ (consMember≡0 a i))) ∙ sym (lUnitₕ (-ₕ 0ₕ)) ∙ rCancelₕ 0ₕ) ∙ (rUnitₕ f)) ∣₋₁}) --- (f+consMember' f) --- Hⁿ-Sⁿ≃Int (suc n) = --- compGrIso (Hⁿ-Sⁿ≃Int n) --- (transport (λ i → grIso {!!} {!coHomGr (suc (suc n)) (Pushout (λ _ → tt) (λ _ → tt))!}) {!!}) - - --- {- --- coHom1S1≃ℤ : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- grIso.fun coHom1S1≃ℤ = grIso.fun {!compGrIso coHom1Iso (invGrIso _ _ (ImIso _ _ (d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ?))!} --- grIso.inv coHom1S1≃ℤ = {!!} --- grIso.rightInv coHom1S1≃ℤ = {!!} --- grIso.leftInv coHom1S1≃ℤ = {!!} --- -} - --- -- compGrIso coHom1Iso (invGrIso _ _ (ImIso _ _ {!d-morph _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0!} {!!})) - - --- -- coHomGrIm : grIso (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- (imGroup (coHomGr 0 (S₊ 0)) --- -- (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 --- -- , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) --- -- coHomGrIm = --- -- ImIso _ --- -- _ --- -- _ --- -- {!!} - - --- -- -- coHom1RedS1 : Iso (coHom 1 (S₊ 1)) (coHomRed 1 (S₊ 1 , north)) --- -- -- Iso.fun coHom1RedS1 = sRec setTruncIsSet λ f → ∣ f , (pRec {!!} {!!} ((transport (λ i → (b : truncIdempotent 3 {!S₊ 1!} (~ i)) → ∥ (transp (λ j → truncIdempotent {!3!} {!!} (~ i ∨ (~ j))) (~ i) north) ≡ b ∥₋₁) isConnectedS1) (f north)) ) ∣₀ --- -- -- Iso.inv coHom1RedS1 = {!!} --- -- -- Iso.rightInv coHom1RedS1 = {!setTruncIdempotent!} --- -- -- Iso.leftInv coHom1RedS1 = {!!} - --- -- -- coHom1Red : ∀ {ℓ} (A : Pointed ℓ) → Iso (coHom 1 (typ A)) (coHomRed 1 A) --- -- -- Iso.fun (coHom1Red A) = sRec setTruncIsSet λ f → ∣ f , {!!} ∣₀ --- -- -- Iso.inv (coHom1Red A) = {!!} --- -- -- Iso.rightInv (coHom1Red A) = {!!} --- -- -- Iso.leftInv (coHom1Red A) = {!!} - --- -- -- -- morphtest : morph (coHomGr 1 (S₊ 1)) intGroup --- -- -- -- fst morphtest = sRec isSetInt λ f → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (f (loop* i))))) --- -- -- -- snd morphtest = sElim2 {!!} --- -- -- -- (funRec 3 isGroupoidS1 --- -- -- -- λ f → funRec 3 isGroupoidS1 --- -- -- -- λ g → pRec (isSetInt _ _) --- -- -- -- (λ n=fn → --- -- -- -- pRec (isSetInt _ _) --- -- -- -- (λ n=gn → (λ j → winding (basechange (SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (∣ f (north) ∣ +ₖ ∣ n=gn (~ j) ∣)))) (λ i → SuspBool→S¹ (S1→SuspBool (trRec2 isGroupoidS1 (λ x → x) (∣ f (loop* i) ∣ +ₖ ∣ transp (λ i → n=gn ((~ i) ∨ (~ j)) ≡ n=gn ((~ i) ∨ (~ j))) (~ j) (λ i → g (loop* i)) i ∣)))))) --- -- -- -- ∙ {!!} --- -- -- -- ∙ {!!}) --- -- -- -- (isConnectedS1 (g north))) --- -- -- -- (isConnectedS1 (f north))) - - --- -- -- -- {- (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (∣ f (loop* i) ∣ +ₖ ∣ g (loop* i) ∣))))) --- -- -- -- ∙ (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (Kn→ΩKn+1 1 ∣ f (loop* i) ∣ ∙ Kn→ΩKn+1 1 ∣ g (loop* i) ∣)))))) --- -- -- -- ∙ (λ j → winding (basechange _ (cong (λ x → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (Kn→ΩKn+1 1 ∣ f x ∣ ∙ Kn→ΩKn+1 1 ∣ g x ∣))))) loop*)) ) --- -- -- -- ∙ (λ i → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn ((cong ∣_∣ (merid (f (loop* i)) ∙ sym (merid north)) ∙ cong ∣_∣ (merid (g (loop* i)) ∙ sym (merid north))))))))) --- -- -- -- ∙ (λ j → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (congFunct ∣_∣ (merid (f (loop* i)) ∙ sym (merid north)) (merid (g (loop* i)) ∙ sym (merid north)) (~ j))))))) --- -- -- -- ∙ (λ j → winding (basechange _ λ i → SuspBool→S¹ (S1→SuspBool (trRec isGroupoidS1 (λ x → x) (ΩKn+1→Kn (cong ∣_∣ (({!!} ∙ {!!}) ∙ {!!}))))))) --- -- -- -- ∙ {!!} --- -- -- -- ∙ {!!} --- -- -- -- ∙ {!!}) -} - --- -- -- -- where --- -- -- -- helper : ∀ {ℓ} {A : Type ℓ} (a : A) (f : A → S¹) (p q : a ≡ a) → winding (basechange (f a) (cong f (p ∙ q))) ≡ winding (basechange (f a) (cong f p ∙ cong f q)) --- -- -- -- helper a f p q i = winding (basechange (f a) (congFunct f p q i)) --- -- -- -- helper2 : (x : S¹) (p q : x ≡ x) → basechange x (p ∙ q) ≡ basechange x p ∙ basechange x q --- -- -- -- helper2 base p q = refl --- -- -- -- helper2 (loop i) p q = {!!} --- -- -- -- helper4 : (x y z : S¹) (p : x ≡ z) (q r : x ≡ y) (s : y ≡ z) → basechange* x z p (q ∙ s) ≡ basechange* x y {!!} q ∙ {!!} --- -- -- -- helper4 = {!!} --- -- -- -- helper3 : (p q : ΩS¹) → winding (p ∙ q) ≡ (winding p +ℤ winding q) --- -- -- -- helper3 = {!!} - - --- -- -- -- -- fstmap : morph (dirProd intGroup intGroup) (coHomGr 0 (S₊ 0)) --- -- -- -- -- fstmap = compMorph {F = dirProd intGroup intGroup} {G = ×coHomGr 0 Unit Unit} {H = coHomGr 0 (S₊ 0)} --- -- -- -- -- (×morph (grIso.inv coHomGrUnit0) (grIso.inv coHomGrUnit0)) --- -- -- -- -- (((MV.Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) zero)) , --- -- -- -- -- {!MV.ΔIsHom _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) zero!}) - --- -- -- -- -- fstMapId : (a : Int × Int) → fstmap .fst a ≡ ∣ (λ _ → proj₁ a +ℤ (0 - proj₂ a)) ∣₀ --- -- -- -- -- fstMapId (a , b) = (λ i → ∣ (λ _ → a +ₖ (-ₖ b)) ∣₀) ∙ {!addLemma!} ∙ {!!} ∙ {!!} - --- -- -- -- -- isoAgain : grIso intGroup (coHomGr 1 (S₊ 1)) --- -- -- -- -- isoAgain = --- -- -- -- -- Iso''→Iso _ _ --- -- -- -- -- (iso'' ((λ a → ∣ (λ {north → 0ₖ ; south → 0ₖ ; (merid north i) → {!a!} ; (merid south i) → {!!}}) ∣₀) , {!!}) --- -- -- -- -- {!!} --- -- -- -- -- {!!}) - --- -- -- -- -- -- -- test2 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → S¹) --- -- -- -- -- -- -- Iso.fun test2 f = {!!} --- -- -- -- -- -- -- Iso.inv test2 f north = ∣ transport (sym S¹≡S1) (f base) ∣ --- -- -- -- -- -- -- Iso.inv test2 f south = ∣ transport (sym S¹≡S1) (f base) ∣ --- -- -- -- -- -- -- Iso.inv test2 f (merid a i) = cong ∣_∣ {!transport (sym S¹≡S1) (f base)!} i --- -- -- -- -- -- -- Iso.rightInv test2 = {!!} --- -- -- -- -- -- -- Iso.leftInv test2 = {!!} - --- -- -- -- -- -- -- F : winding (basechange base loop) ≡ 1 --- -- -- -- -- -- -- F = refl - --- -- -- -- -- -- -- another : (f g : Int) → winding (basechange {!!} {!!}) ≡ {!!} --- -- -- -- -- -- -- another = {!!} - --- -- -- -- -- -- -- test : Iso (S¹ → S¹) (S¹ × Int) --- -- -- -- -- -- -- Iso.fun test f = f base , winding (basechange (f base) (cong f loop)) --- -- -- -- -- -- -- Iso.inv test (x , int) base = x --- -- -- -- -- -- -- Iso.inv test (x , int) (loop i) = {!!} --- -- -- -- -- -- -- Iso.rightInv test = {!!} --- -- -- -- -- -- -- Iso.leftInv test = {!!} - --- -- -- -- -- -- -- -- test13 : Iso ∥ (S¹ → S¹) ∥₀ Int --- -- -- -- -- -- -- -- Iso.fun test13 = sRec isSetInt λ f → winding (basechange (f base) (λ i → f (loop i))) --- -- -- -- -- -- -- -- Iso.inv test13 a = ∣ (λ {base → {!!} ; (loop i) → {!!}}) ∣₀ --- -- -- -- -- -- -- -- Iso.rightInv test13 = {!!} --- -- -- -- -- -- -- -- Iso.leftInv test13 = {!!} - --- -- -- -- -- -- -- -- test : Iso (S¹ → S¹) (S¹ × Int) --- -- -- -- -- -- -- -- Iso.fun test f = (f base) , transport (basedΩS¹≡Int (f base)) λ i → f (loop i) --- -- -- -- -- -- -- -- Iso.inv test (x , int) base = x --- -- -- -- -- -- -- -- Iso.inv test (x , int) (loop i) = transport (sym (basedΩS¹≡Int x)) int i --- -- -- -- -- -- -- -- Iso.rightInv test (x , int) i = (x , transportTransport⁻ (basedΩS¹≡Int x) int i) --- -- -- -- -- -- -- -- Iso.leftInv test f = --- -- -- -- -- -- -- -- funExt λ { base → refl --- -- -- -- -- -- -- -- ; (loop i) j → transport⁻Transport (basedΩS¹≡Int (f base)) (λ i → f (loop i)) j i} - - --- -- -- -- -- -- -- -- lem : S¹ ≡ hLevelTrunc 3 (S₊ 1) --- -- -- -- -- -- -- -- lem = sym (truncIdempotent 3 isGroupoidS¹) ∙ λ i → hLevelTrunc 3 (S¹≡S1 (~ i)) - --- -- -- -- -- -- -- -- ×≡ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (a b : A × B) → proj₁ a ≡ proj₁ b → proj₂ a ≡ proj₂ b → a ≡ b --- -- -- -- -- -- -- -- ×≡ (_ , _) (_ , _) id1 id2 i = (id1 i) , (id2 i) - --- -- -- -- -- -- -- -- test22 : Iso (S₊ 1 → coHomK 1) (S₊ 1 × Int) --- -- -- -- -- -- -- -- Iso.fun test22 f = {!f north!} , {!!} --- -- -- -- -- -- -- -- Iso.inv test22 = {!!} --- -- -- -- -- -- -- -- Iso.rightInv test22 = {!!} --- -- -- -- -- -- -- -- Iso.leftInv test22 = {!!} - - - - - - --- -- -- -- -- -- -- -- coHom1≃∥S1×ℤ∥₀ : Iso (coHom 1 (S₊ 1)) ∥ S₊ 1 × Int ∥₀ --- -- -- -- -- -- -- -- coHom1≃∥S1×ℤ∥₀ = setTruncIso test2 --- -- -- -- -- -- -- -- where --- -- -- -- -- -- -- -- test2 : Iso (S₊ 1 → coHomK 1) (S₊ 1 × Int) --- -- -- -- -- -- -- -- Iso.fun test2 f = transport (λ i → S¹≡S1 (~ i) × Int) (Iso.fun test (transport (λ i → (S¹≡S1 i → lem (~ i))) f)) --- -- -- -- -- -- -- -- Iso.inv test2 x = transport (λ i → (S¹≡S1 (~ i) → lem i)) (Iso.inv test (transport (λ i → S¹≡S1 i × Int) x)) --- -- -- -- -- -- -- -- Iso.rightInv test2 (s , int) = ×≡ _ _ {!!} {!!} --- -- -- -- -- -- -- -- Iso.leftInv test2 f = {!!} ∙ {!!} ∙ {!!} - --- -- -- -- -- -- -- -- test2Id : (a b : (S₊ 1 → coHomK 1)) → proj₂ (Iso.fun test2 (λ x → a x +ₖ b x)) ≡ (proj₂ (Iso.fun test2 a) +ₖ proj₂ (Iso.fun test2 a)) --- -- -- -- -- -- -- -- test2Id a b = {! --- -- -- -- -- -- -- -- transport --- -- -- -- -- -- -- -- (basedΩS¹≡Int --- -- -- -- -- -- -- -- (transport (λ i → S¹≡S1 i → lem (~ i)) (λ x → a x +ₖ b x) base)) --- -- -- -- -- -- -- -- (λ i → --- -- -- -- -- -- -- -- transport (λ i₁ → S¹≡S1 i₁ → lem (~ i₁)) (λ x → a x +ₖ b x) --- -- -- -- -- -- -- -- (loop i))!} ∙ {!transport (λ i → S¹≡S1 i → lem (~ i)) (λ x → a x +ₖ b x) base!} - - --- -- -- -- -- -- -- -- main : grIso intGroup (coHomGr 1 (S₊ 1)) --- -- -- -- -- -- -- -- main = Iso'→Iso --- -- -- -- -- -- -- -- (iso' {!!} --- -- -- -- -- -- -- -- {!!}) - - --- -- -- -- -- -- coHom1 : grIso (coHomGr 1 (S₊ 1)) intGroup --- -- -- -- -- -- coHom1 = --- -- -- -- -- -- Iso'→Iso --- -- -- -- -- -- (iso' (iso ({!!} ∘ {!!} ∘ {!!} ∘ {!!}) --- -- -- -- -- -- {!!} --- -- -- -- -- -- {!!} --- -- -- -- -- -- {!!}) --- -- -- -- -- -- {!!}) - - - --- -- -- -- -- -- schonf : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : (A × B) → Type ℓ''} --- -- -- -- -- -- → ((a : A) (b : B) → C (a , b)) --- -- -- -- -- -- → (x : A × B) → C x --- -- -- -- -- -- schonf f (a , b) = f a b - --- -- -- -- -- -- -- -- setTruncProdIso : ∀ {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') → Iso ∥ (A × B) ∥₀ (∥ A ∥₀ × ∥ B ∥₀) --- -- -- -- -- -- -- -- Iso.fun (setTruncProdIso A B) = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) λ {(a , b) → ∣ a ∣₀ , ∣ b ∣₀} --- -- -- -- -- -- -- -- Iso.inv (setTruncProdIso A B) (a , b) = sRec setTruncIsSet (λ a → sRec setTruncIsSet (λ b → ∣ a , b ∣₀) b) a --- -- -- -- -- -- -- -- Iso.rightInv (setTruncProdIso A B) = --- -- -- -- -- -- -- -- schonf (sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) --- -- -- -- -- -- -- -- λ _ → sElim (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) --- -- -- -- -- -- -- -- λ _ → refl) --- -- -- -- -- -- -- -- Iso.leftInv (setTruncProdIso A B) = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ {(a , b) → refl} - --- -- -- -- -- -- -- -- setTruncProdLemma : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (a1 a2 : A) (b : B) → isHLevelConnected 2 A --- -- -- -- -- -- -- -- → Path ∥ A × B ∥₀ ∣ a1 , b ∣₀ ∣ a2 , b ∣₀ --- -- -- -- -- -- -- -- setTruncProdLemma {A = A} {B = B} a1 a2 b conA i = Iso.inv (setTruncProdIso A B) (Iso.inv setTruncTrunc0Iso ((sym (conA .snd ∣ a1 ∣) ∙ (conA .snd ∣ a2 ∣)) i) , ∣ b ∣₀) - --- -- -- -- -- -- -- -- test3 : Iso ∥ S₊ 1 × Int ∥₀ Int --- -- -- -- -- -- -- -- Iso.fun test3 = sRec isSetInt proj₂ --- -- -- -- -- -- -- -- Iso.inv test3 a = ∣ north , a ∣₀ --- -- -- -- -- -- -- -- Iso.rightInv test3 a = refl --- -- -- -- -- -- -- -- Iso.leftInv test3 = --- -- -- -- -- -- -- -- sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) --- -- -- -- -- -- -- -- λ {(s , int) → setTruncProdLemma north s int (sphereConnected 1)} - --- -- -- -- -- -- -- -- coHomGr0-S1 : grIso intGroup (coHomGr 1 (S₊ 1)) --- -- -- -- -- -- -- -- coHomGr0-S1 = --- -- -- -- -- -- -- -- Iso'→Iso --- -- -- -- -- -- -- -- (iso' (compIso (symIso test3) (symIso coHom1≃∥S1×ℤ∥₀)) --- -- -- -- -- -- -- -- (indIntGroup {G = coHomGr 1 (S₊ 1)} --- -- -- -- -- -- -- -- (Iso.fun (compIso (symIso test3) (symIso coHom1≃∥S1×ℤ∥₀))) --- -- -- -- -- -- -- -- ((λ i → ∣ transport (λ i → (S¹≡S1 (~ i) → lem i)) (Iso.inv test (base , 0)) ∣₀) --- -- -- -- -- -- -- -- ∙ (λ i → ∣ transport (λ i → (S¹≡S1 (~ i) → lem i)) (helper2 i) ∣₀) --- -- -- -- -- -- -- -- ∙ cong ∣_∣₀ (funExt λ {north → refl ; south → refl ; (merid a i) → {!!}})) --- -- -- -- -- -- -- -- {!!} --- -- -- -- -- -- -- -- {!!})) --- -- -- -- -- -- -- -- where --- -- -- -- -- -- -- -- helper : basedΩS¹≡ΩS¹ base ≡ {!basechange!} --- -- -- -- -- -- -- -- helper = {!substComposite!} - --- -- -- -- -- -- -- -- substComposite2 : ∀ {ℓ} {A B C : Type ℓ} --- -- -- -- -- -- -- -- (P : A ≡ B) (Q : B ≡ C) (a : A) --- -- -- -- -- -- -- -- → transport (P ∙ Q) a ≡ transport Q (transport P a) --- -- -- -- -- -- -- -- substComposite2 = {!!} - --- -- -- -- -- -- -- -- helper1 : transport (λ i → S¹≡S1 i × Int) (north , 0) ≡ (base , 0) --- -- -- -- -- -- -- -- helper1 = refl --- -- -- -- -- -- -- -- helper3 : transport (sym (basedΩS¹≡Int base)) 0 ≡ refl --- -- -- -- -- -- -- -- helper3 = (λ i → transport (symDistr (basedΩS¹≡ΩS¹ base) ΩS¹≡Int i) 0) --- -- -- -- -- -- -- -- ∙ substComposite2 (sym ΩS¹≡Int) (sym (basedΩS¹≡ΩS¹ base)) 0 --- -- -- -- -- -- -- -- ∙ (λ i → transport (λ i → basedΩS¹≡ΩS¹ base (~ i)) refl) -- --- -- -- -- -- -- -- -- ∙ transportRefl ((equiv-proof (basechange-isequiv base) refl) .fst .fst) --- -- -- -- -- -- -- -- ∙ (λ i → equiv-proof (transport (λ j → isEquiv (refl-conjugation j)) (basedΩS¹→ΩS¹-isequiv i0)) refl .fst .fst) --- -- -- -- -- -- -- -- ∙ (λ i → {!equiv-proof (transport (λ j → isEquiv (refl-conjugation j)) (basedΩS¹→ΩS¹-isequiv i0)) refl .fst!}) --- -- -- -- -- -- -- -- ∙ {!basedΩS¹→ΩS¹!} --- -- -- -- -- -- -- -- ∙ {!!} --- -- -- -- -- -- -- -- ∙ {!!} --- -- -- -- -- -- -- -- helper4 : (x : S¹) → Iso.inv test (base , 0) x ≡ x --- -- -- -- -- -- -- -- helper4 = {!!} --- -- -- -- -- -- -- -- helper2 : Iso.inv test (transport (λ i → S¹≡S1 i × Int) (north , 0)) ≡ λ x → x --- -- -- -- -- -- -- -- helper2 = (λ i → Iso.inv test (base , 0)) ∙ {!!} ∙ {!!} - --- -- -- -- -- -- ×≡ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A × B} → proj₁ x ≡ proj₁ y → proj₂ x ≡ proj₂ y → x ≡ y --- -- -- -- -- -- ×≡ {x = (a , b)} {y = (c , d)} id1 id2 i = (id1 i) , (id2 i) - --- -- -- -- -- -- fstFunHelper : (x : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- -- -- -- -- → isInIm (coHomGr 0 (S₊ 0)) (coHomGr 1 _) (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) x --- -- -- -- -- -- fstFunHelper a = MV.Ker-i⊂Im-d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a --- -- -- -- -- -- (sym (isContrH1Unit×H1Unit .snd _) ∙ (isContrH1Unit×H1Unit .snd _)) --- -- -- -- -- -- where --- -- -- -- -- -- isContrH1Unit×H1Unit : isContr (Group.type (×coHomGr 1 Unit Unit)) --- -- -- -- -- -- isContrH1Unit×H1Unit = (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) --- -- -- -- -- -- , λ {(a , b) → sigmaProdElim {D = λ (x : Σ[ x ∈ Group.type (×coHomGr 1 Unit Unit)] Unit) → (∣ (λ _ → ∣ north ∣) ∣₀ , ∣ (λ _ → ∣ north ∣) ∣₀) ≡ fst x} --- -- -- -- -- -- (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) --- -- -- -- -- -- (λ a b _ → ×≡ (grIso.leftInv (coHomGrUnit≥1 0) ∣ a ∣₀) (grIso.leftInv (coHomGrUnit≥1 0) ∣ b ∣₀)) ((a , b) , tt)} - - - --- -- -- -- -- -- helperMorph : isMorph intGroup (dirProd intGroup intGroup) λ a → a , (0 - a) --- -- -- -- -- -- helperMorph = --- -- -- -- -- -- indIntGroup {G = dirProd intGroup intGroup} --- -- -- -- -- -- (λ a → a , (0 - a)) --- -- -- -- -- -- refl --- -- -- -- -- -- (λ a → ×≡ refl (helper2 a)) --- -- -- -- -- -- λ a → ×≡ refl (helper3 a) --- -- -- -- -- -- where --- -- -- -- -- -- helper1 : (a : Int) → predInt (sucInt a) ≡ a --- -- -- -- -- -- helper1 (pos zero) = refl --- -- -- -- -- -- helper1 (pos (suc n)) = refl --- -- -- -- -- -- helper1 (negsuc zero) = refl --- -- -- -- -- -- helper1 (negsuc (suc n)) = refl - --- -- -- -- -- -- helper4 : (a : Int) → sucInt (predInt a) ≡ a --- -- -- -- -- -- helper4 (pos zero) = refl --- -- -- -- -- -- helper4 (pos (suc n)) = refl --- -- -- -- -- -- helper4 (negsuc zero) = refl --- -- -- -- -- -- helper4 (negsuc (suc n)) = refl - --- -- -- -- -- -- helper2 : (a : Int) → (pos 0 - sucInt a) ≡ predInt (pos 0 - a) --- -- -- -- -- -- helper2 (pos zero) = refl --- -- -- -- -- -- helper2 (pos (suc n)) = refl --- -- -- -- -- -- helper2 (negsuc zero) = refl --- -- -- -- -- -- helper2 (negsuc (suc n)) = sym (helper1 _) - --- -- -- -- -- -- helper3 : (a : Int) → (pos 0 - predInt a) ≡ sucInt (pos 0 - a) --- -- -- -- -- -- helper3 (pos zero) = refl --- -- -- -- -- -- helper3 (pos (suc zero)) = refl --- -- -- -- -- -- helper3 (pos (suc (suc n))) = sym (helper4 _) --- -- -- -- -- -- helper3 (negsuc zero) = refl --- -- -- -- -- -- helper3 (negsuc (suc n)) = refl - --- -- -- -- -- -- helper : (a b : Int) → (pos 0 - (a +ℤ b)) ≡ ((pos 0 - a) +ℤ (pos 0 - b)) --- -- -- -- -- -- helper a (pos zero) = refl --- -- -- -- -- -- helper (pos zero) (pos (suc n)) = --- -- -- -- -- -- cong (λ x → pos 0 - sucInt x) (+ℤ-comm (pos zero) (pos n)) --- -- -- -- -- -- ∙ +ℤ-comm (pos 0 +negsuc n) (pos zero) --- -- -- -- -- -- helper (pos (suc n₁)) (pos (suc n)) = --- -- -- -- -- -- {!!} --- -- -- -- -- -- helper (negsuc n₁) (pos (suc n)) = {!!} --- -- -- -- -- -- helper a (negsuc n) = {!!} - --- -- -- -- -- -- fun : morph intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- -- -- -- -- fst fun = MV.d _ _ _ (λ _ → tt) (λ _ → tt) 0 ∘ grIso.inv coHom0-S0 .fst ∘ λ a → a , (0 - a) --- -- -- -- -- -- snd fun = {!!} --- -- -- -- -- -- {- compMorph {F = intGroup} {G = dirProd intGroup intGroup} {H = coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))} --- -- -- -- -- -- ((λ a → a , a) , (λ a b → refl)) --- -- -- -- -- -- (compMorph {F = dirProd intGroup intGroup} {G = coHomGr 0 (S₊ 0)} {H = coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))} (grIso.inv coHom0-S0) --- -- -- -- -- -- (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 --- -- -- -- -- -- , MV.dIsHom Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) .snd -} --- -- -- -- -- -- {- theIso : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- -- -- -- -- theIso = Iso''→Iso _ _ --- -- -- -- -- -- (iso'' ((λ x → ∣ (λ {(inl tt) → 0ₖ ; (inr tt) → 0ₖ ; (push a i) → Kn→ΩKn+1 0 x i}) ∣₀) , {!!}) --- -- -- -- -- -- {!!} --- -- -- -- -- -- {!MV.d!}) --- -- -- -- -- -- -} - - - --- -- -- -- -- -- theIso : grIso intGroup (coHomGr 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt))) --- -- -- -- -- -- theIso = --- -- -- -- -- -- Iso''→Iso _ _ --- -- -- -- -- -- (iso'' fun --- -- -- -- -- -- (λ x inker → {!MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 --- -- -- -- -- -- (grIso.inv coHom0-S0 .fst (g , g))!}) --- -- -- -- -- -- (sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) --- -- -- -- -- -- λ x → pRec propTruncIsProp --- -- -- -- -- -- (λ {(a , b) → {!fun!} }) --- -- -- -- -- -- (fstFunHelper (∣ x ∣₀)))) --- -- -- -- -- -- where --- -- -- -- -- -- whoKnows : (a : S₊ 0 → Int) (x : MV.D Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt)) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (λ _ → a north) x --- -- -- -- -- -- ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 a x --- -- -- -- -- -- whoKnows a (inl x) = refl --- -- -- -- -- -- whoKnows a (inr x) = refl --- -- -- -- -- -- whoKnows a (push north i) = refl --- -- -- -- -- -- whoKnows a (push south i) j = {!!} - --- -- -- -- -- -- helper : (a : Int) → (grIso.inv coHom0-S0 .fst (a , a)) ≡ ∣ S0→Int (a , a) ∣₀ --- -- -- -- -- -- helper a = {!have : - --- -- -- -- -- -- ∣ --- -- -- -- -- -- MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 --- -- -- -- -- -- (S0→Int (x , x)) --- -- -- -- -- -- ∣₀ --- -- -- -- -- -- ≡ ∣ (λ _ → ∣ north ∣) ∣₀!} - --- -- -- -- -- -- helper2 : (a b : Int) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a)) ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (b , b)) --- -- -- -- -- -- → a ≡ b --- -- -- -- -- -- helper2 a b id = pRec (isSetInt a b) (λ {(pt , id) → {!!}}) (fstFunHelper ∣ (MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a))) ∣₀) - --- -- -- -- -- -- idFun : (a : Int) → MV.D Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1 --- -- -- -- -- -- idFun a (inl x) = 0ₖ --- -- -- -- -- -- idFun a (inr x) = 0ₖ --- -- -- -- -- -- idFun a (push north i) = Kn→ΩKn+1 zero a i --- -- -- -- -- -- idFun a (push south i) = Kn→ΩKn+1 zero a i - --- -- -- -- -- -- helper3 : (a : Int) → (MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a))) ≡ idFun a --- -- -- -- -- -- helper3 a = funExt λ {(inl x) → refl ; (inr x) → refl ; (push north i) → refl ; (push south i) → refl } - --- -- -- -- -- -- helper4 : (a b : Int) → MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (a , a)) ≡ MV.d-pre Unit Unit (Susp ⊥) (λ _ → tt) (λ _ → tt) 0 (S0→Int (b , b)) --- -- -- -- -- -- → a ≡ b --- -- -- -- -- -- helper4 a b id = --- -- -- -- -- -- {!!} --- -- -- -- -- -- ∙ {!!} --- -- -- -- -- -- ∙ {!!} --- -- -- -- -- -- where --- -- -- -- -- -- helper5 : {!!} --PathP (λ k → id k (inl tt) ≡ id k (inr tt)) (Kn→ΩKn+1 zero a) (Kn→ΩKn+1 zero a) --- -- -- -- -- -- helper5 i j = {!id i!} - --- -- -- -- -- -- -- fun : coHom 1 (Pushout {A = S₊ 0} (λ _ → tt) (λ _ → tt)) → coHom 0 (S₊ 0) --- -- -- -- -- -- -- fun a = (pRec {P = Σ[ x ∈ coHom 0 (S₊ 0)] (MV.d _ _ _ (λ _ → tt) (λ _ → tt) 0 x ≡ a) × isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) (MV.Δ _ _ _ (λ _ → tt) (λ _ → tt) 0) x} --- -- -- -- -- -- -- (λ {(a1 , b) (c , d) → ΣProp≡ (λ x → isOfHLevelProd 1 (setTruncIsSet _ _) propTruncIsProp) --- -- -- -- -- -- -- (sElim2 {B = λ a1 c → (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a1 ≡ a) --- -- -- -- -- -- -- → MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 c ≡ a --- -- -- -- -- -- -- → isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) --- -- -- -- -- -- -- (λ z → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 z) a1 --- -- -- -- -- -- -- → isInIm (×coHomGr 0 Unit Unit) (coHomGr 0 (S₊ 0)) --- -- -- -- -- -- -- (λ z → MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 z) c → a1 ≡ c} --- -- -- -- -- -- -- (λ _ _ → {!!}) --- -- -- -- -- -- -- (λ a c b1 d1 → pElim (λ _ → isOfHLevelΠ 1 λ _ → setTruncIsSet _ _) --- -- -- -- -- -- -- λ b2 → pElim (λ _ → setTruncIsSet _ _) --- -- -- -- -- -- -- λ d2 → {!d2!}) --- -- -- -- -- -- -- a1 c (proj₁ b) (proj₁ d) (proj₂ b) (proj₂ d))}) --- -- -- -- -- -- -- (λ {(a , b) → a , b , MV.Ker-d⊂Im-Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0 a {!!}}) --- -- -- -- -- -- -- (fstFunHelper a)) .fst -- pRec {!!} {!!} (fstFunHelper a) diff --git a/Cubical/ZCohomology/Groups/Torus.agda b/Cubical/ZCohomology/Groups/Torus.agda new file mode 100644 index 000000000..9822e824b --- /dev/null +++ b/Cubical/ZCohomology/Groups/Torus.agda @@ -0,0 +1,318 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.Groups.Torus where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.S1.S1 +open import Cubical.HITs.S1 +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.MayerVietorisUnreduced +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence +open import Cubical.Data.NatMinusTwo.Base +open import Cubical.Data.Empty +open import Cubical.Data.Sigma hiding (_×_) +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₋₁ ; ∣_∣ to ∣_∣₋₁) +open import Cubical.HITs.Nullification +open import Cubical.Data.Int renaming (_+_ to _+ℤ_; +-comm to +ℤ-comm ; +-assoc to +ℤ-assoc) +open import Cubical.Data.Nat +open import Cubical.Data.Prod +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; recElim to trRec ; rec to trRec2 ; elim3 to trElim3) +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Freudenthal +open import Cubical.HITs.SmashProduct.Base +open import Cubical.Data.Unit +open import Cubical.Data.Group.Base renaming (Iso to grIso ; compIso to compGrIso ; invIso to invGrIso ; idIso to idGrIso) +open import Cubical.Data.Group.GroupLibrary +open import Cubical.ZCohomology.coHomZero-map +open import Cubical.HITs.Wedge + +open import Cubical.Foundations.Equiv.HalfAdjoint + +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Sn + +open import Cubical.ZCohomology.KcompPrelims + + +open import Cubical.HITs.Pushout +open import Cubical.Data.Sum.Base +open import Cubical.Data.HomotopyGroup +open import Cubical.Data.Bool hiding (_⊕_) + + +-- H⁰(T²) +coHom0Torus : grIso (coHomGr 0 (S₊ 1 × S₊ 1)) intGroup +coHom0Torus = + Iso'→Iso + (iso' + (iso (sRec isSetInt (λ f → f (north , north))) + (λ a → ∣ (λ x → a) ∣₀) + (λ a → refl) + (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ + (funExt λ {(x , y) → suspToPropRec2 + {B = λ x y → f (north , north) ≡ f (x , y)} + north + (λ _ _ → isSetInt _ _) + refl + x y}))) + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) + λ a b → addLemma (a (north , north)) (b (north , north)))) + + + +------------------ H¹(T²) ------------------------------- + +-- We first need the following technical lemma +basechange-lemma2 : (f g : S¹ → hLevelTrunc 3 (S₊ 1)) (F : hLevelTrunc 3 (S₊ 1) → S¹) + → ((basechange2⁻ (F (f base +ₖ g base)) λ i → F ((f (loop i)) +ₖ g (loop i))) + ≡ basechange2⁻ (F (f base)) (cong (F ∘ f) loop) + ∙ basechange2⁻ (F (g base)) (cong (F ∘ g) loop)) +basechange-lemma2 f g F = coInd f g F (f base) (g base) refl refl + where + coInd : (f g : S¹ → hLevelTrunc 3 (S₊ 1)) (F : hLevelTrunc 3 (S₊ 1) → S¹) (x y : hLevelTrunc 3 (S₊ 1)) + → f base ≡ x + → g base ≡ y + → ((basechange2⁻ (F (f base +ₖ g base)) λ i → F ((f (loop i)) +ₖ g (loop i))) + ≡ basechange2⁻ (F (f base)) (cong (F ∘ f) loop) + ∙ basechange2⁻ (F (g base)) (cong (F ∘ g) loop)) + coInd f g F = + elim2 (λ _ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 (isGroupoidS¹ base base)) _ _ ) + (suspToPropRec2 + north + (λ _ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → isGroupoidS¹ _ _ _ _) + λ fid gid → + (λ i → basechange2⁻ (F (f base +ₖ g base)) (cong₂Funct (λ x y → F (f x +ₖ g y)) loop loop i)) + ∙ (basechange2⁻-morph (F (f base +ₖ g base)) + (cong (λ x → F (f x +ₖ g base)) loop) + (cong (λ x → F (f base +ₖ g x)) loop)) + ∙ (λ i → basechange2⁻ (F (f base +ₖ gid i)) (cong (λ x → F (f x +ₖ gid i)) loop) + ∙ basechange2⁻ (F (fid i +ₖ g base)) (cong (λ x → F (fid i +ₖ g x)) loop)) + ∙ (λ i → basechange2⁻ (F (rUnitₖ (f base) i)) (cong (λ x → F (rUnitₖ (f x) i)) loop) + ∙ basechange2⁻ (F (lUnitₖ (g base) i)) (cong (λ x → F (lUnitₖ (g x) i)) loop))) + + + +coHom1Torus : grIso (coHomGr 1 ((S₊ 1) × (S₊ 1))) (dirProd intGroup intGroup) +coHom1Torus = + compGrIso + (Iso'→Iso + (iso' (compIso + (setTruncIso (compIso + schönfinkelIso + (compIso + (codomainIso S1→S1→S1×Int) + toProdIso))) + (setTruncOfProdIso)) + (sElim2 + (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ f g → ×≡ (cong ∣_∣₀ + (funExt (λ x → helper (f (x , S¹→S1 base) +ₖ g (x , S¹→S1 base)) + ∙ sym (cong₂ (λ x y → x +ₖ y) + (helper (f (x , S¹→S1 base))) + (helper (g (x , S¹→S1 base))))))) + (cong ∣_∣₀ + (funExt + (suspToPropRec + north + (λ _ → isSetInt _ _) + (cong winding + (basechange-lemma2 + (λ x → f (north , S¹→S1 x)) + (λ x → g (north , S¹→S1 x)) + λ x → S¹map (trMap S1→S¹ x)) + ∙ winding-hom + (basechange2⁻ + (S¹map (trMap S1→S¹ (f (north , S¹→S1 base)))) + (λ i → S¹map (trMap S1→S¹ (f (north , S¹→S1 (loop i)))))) + (basechange2⁻ + (S¹map (trMap S1→S¹ (g (north , S¹→S1 base)))) + (λ i → S¹map (trMap S1→S¹ (g (north , S¹→S1 (loop i)))))) + ∙ sym (addLemma + (winding + (basechange2⁻ + (S¹map (trMap S1→S¹ (f (north , S¹→S1 base)))) + (λ i → S¹map (trMap S1→S¹ (f (north , S¹→S1 (loop i))))))) + (winding + (basechange2⁻ + (S¹map (trMap S1→S¹ (g (north , S¹→S1 base)))) + (λ i → S¹map (trMap S1→S¹ (g (north , S¹→S1 (loop i))))))))))))))) + (dirProdIso (invGrIso (coHom-n-Sn 0)) (invGrIso H⁰-S¹≅ℤ)) + + where + helper : (x : hLevelTrunc 3 (S₊ 1)) → ∣ S¹→S1 (S¹map (trMap S1→S¹ x)) ∣ ≡ x + helper = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ a → cong ∣_∣ (S1→S¹-retr a) + +------------------------------------------------------------- + + +----------------------- H²(T²) ------------------------------ + +H²-T²≡ℤ : Iso Int (coHom 2 (S₊ 1 × S₊ 1)) +H²-T²≡ℤ = + compIso + helper' + (compIso + (symIso (prodIso (groupIso→Iso H²-S¹≅0) + (groupIso→Iso (invGrIso (coHom-n-Sn 0))))) + (compIso + (symIso setTruncOfProdIso) + (symIso + (setTruncIso + (compIso + schönfinkelIso + (compIso + (codomainIso S1→K2≡K2×K1) + toProdIso)))))) + where + helper' : Iso Int (Unit × Int) + Iso.inv helper' = proj₂ + Iso.fun helper' x = tt , x + Iso.leftInv helper' x = refl + Iso.rightInv helper' x = ×≡ refl refl + +{- + +coHom2Torus : grIso (coHomGr 2 ((S₊ 1) × (S₊ 1))) intGroup +coHom2Torus = + invGrIso + (Iso'→Iso + (iso' + H²-T²≡ℤ + λ a b → {!!} {- (λ i → Iso.fun + (symIso + (setTruncIso + (compIso schönfinkelIso (compIso (codomainIso test13) toProdIso)))) + (Iso.fun (symIso setTruncOfProdIso) + (helper'' (Iso.inv (groupIso→Iso coHom2-S1) tt) 0ₕ i , morph.ismorph (grIso.inv (invGrIso coHom1S1)) a b i))) + ∙ (λ i → Iso.fun + (symIso + (setTruncIso + (compIso schönfinkelIso (compIso (codomainIso test13) toProdIso)))) + ∣ (λ _ → ∣ north ∣) , {!!} ∣₀) + ∙ {!toProdIso!} -})) + +-} + +{- + where + helper'' : isProp ∥ (S₊ 1 → hLevelTrunc 4 (S₊ 2)) ∥₀ + helper'' = {!!} + + helper2' : (f g : (S₊ 1 → hLevelTrunc 3 (S₊ 1))) → + Path (coHom 2 ((S₊ 1) × (S₊ 1))) (Iso.fun + (symIso + (setTruncIso + (compIso schönfinkelIso (compIso (codomainIso S1→K2≡K2×K1) toProdIso)))) + (Iso.fun (symIso setTruncOfProdIso) (0ₕ , ∣ f ∣₀ +ₕ ∣ g ∣₀))) + (Iso.fun (symIso + (setTruncIso + (compIso schönfinkelIso (compIso (codomainIso S1→K2≡K2×K1) toProdIso)))) + (Iso.fun (symIso setTruncOfProdIso) (0ₕ , ∣ f ∣₀)) +ₕ Iso.fun (symIso + (setTruncIso + (compIso schönfinkelIso (compIso (codomainIso S1→K2≡K2×K1) toProdIso)))) + (Iso.fun (symIso setTruncOfProdIso) (0ₕ , ∣ g ∣₀))) + helper2' f g = (λ i → Iso.fun (symIso + (setTruncIso + (compIso schönfinkelIso ((codomainIso S1→K2≡K2×K1))))) ∣ (λ x → 0ₖ , (f x +ₖ g x)) ∣₀) + ∙ (λ i → ∣ Iso.fun (symIso (compIso schönfinkelIso (codomainIso S1→K2≡K2×K1))) (λ x → 0ₖ , (f x +ₖ g x)) ∣₀) + ∙ (λ i → ∣ Iso.inv schönfinkelIso (Iso.inv (codomainIso S1→K2≡K2×K1) ((λ x → 0ₖ , (f x +ₖ g x)))) ∣₀) + ∙ (λ i → ∣ {!!} ∣₀) + ∙ {!!} + ∙ {!!} + where + fun : S₊ 1 × S₊ 1 → hLevelTrunc 4 (S₊ 2) + fun (x , north) = 0ₖ +ₖ 0ₖ + fun (x , south) = 0ₖ +ₖ 0ₖ + fun (x , (merid south i)) = 0ₖ +ₖ (Kn→ΩKn+1 1 (f x +ₖ g x) i) + fun (x , (merid north i)) = 0ₖ +ₖ 0ₖ + + helper : Iso.inv schönfinkelIso (Iso.inv (codomainIso S1→K2≡K2×K1) ((λ x → 0ₖ , (f x +ₖ g x)))) ≡ fun + helper = funExt λ {(x , north) → refl ; (x , south) → refl ; (x , (merid north i)) → refl ; (x , (merid south i)) → refl} + + helper2 : ∣ Iso.inv schönfinkelIso (Iso.inv (codomainIso S1→K2≡K2×K1) ((λ x → 0ₖ , (f x +ₖ g x)))) ∣₀ + ≡ (∣ Iso.inv schönfinkelIso (Iso.inv (codomainIso S1→K2≡K2×K1) ((λ x → 0ₖ , f x))) ∣₀ +ₕ ∣ Iso.inv schönfinkelIso (Iso.inv (codomainIso S1→K2≡K2×K1) ((λ x → 0ₖ , g x))) ∣₀) + helper2 = + cong ∣_∣₀ + (funExt λ {(x , north) → ((λ i → (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x +ₖ g x) i)) + ∙∙ sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ cong (λ x → ((0ₖ +ₖ 0ₖ) +ₖ x)) (rUnitₖ 0ₖ) + ∙∙ refl) + ; (x , south) → refl + ∙∙ sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ cong (λ x → ((0ₖ +ₖ 0ₖ) +ₖ x)) (rUnitₖ 0ₖ) + ∙∙ (λ i → (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x) i) +ₖ ∣ north ∣ +ₖ Kn→ΩKn+1 1 (g x) i) + ; (x , (merid south i)) j → hcomp (λ k → λ {(j = i0) → ∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x +ₖ g x) (i ∨ (~ k)) + ; (j = i1) → (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x) (i ∧ k)) +ₖ ∣ north ∣ +ₖ Kn→ΩKn+1 1 (g x) (i ∧ k)}) + ((sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ cong (λ x → ((0ₖ +ₖ 0ₖ) +ₖ x)) (rUnitₖ 0ₖ)) j) + ; (x , (merid north u)) → {!!}}) + where + anotherone : (x : _) → cong (0ₖ +ₖ_) (Kn→ΩKn+1 1 (f x +ₖ g x)) ≡ {!!} + anotherone x = {!!} + + helper5 : (x : _) → Iso.inv schönfinkelIso (Iso.inv (codomainIso S1→K2≡K2×K1) ((λ x → 0ₖ , f x +ₖ g x))) x + ≡ (Iso.inv schönfinkelIso (Iso.inv (codomainIso S1→K2≡K2×K1) ((λ x → 0ₖ , f x))) x) +ₖ (Iso.inv schönfinkelIso (Iso.inv (codomainIso S1→K2≡K2×K1) ((λ x → 0ₖ , g x))) x) + helper5 (x , north) = sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i))) + helper5 (x , south) = sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i))) + helper5 (x , merid north i) = sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i))) + helper5 (x , merid south i) j = + {!!} + where + helper13 : sym (sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i)))) + ∙ (λ i → ∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x +ₖ g x) i) ∙ (sym (rUnitₖ (0ₖ +ₖ 0ₖ)) + ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i)))) + ≡ λ i → (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x) i) +ₖ + ∣ north ∣ +ₖ Kn→ΩKn+1 1 (g x) i + helper13 = (λ j → sym ((λ i → (rUnitₖ (lUnitₖ 0ₖ (j ∧ (~ i))) (~ i))) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ lUnitₖ 0ₖ (~ i))) + ∙ lUnit (rUnit ((λ i → lUnitₖ (Kn→ΩKn+1 1 (f x +ₖ g x) i) j)) j) j + ∙ ((λ i → (rUnitₖ (lUnitₖ 0ₖ (j ∧ (~ i))) (~ i)))) ∙ λ i → (0ₖ +ₖ 0ₖ) +ₖ lUnitₖ 0ₖ (~ i)) + ∙ (λ j → sym ((λ i → (rUnitₖ (lUnitₖ 0ₖ (~ i)) (~ i))) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ lUnitₖ 0ₖ (~ i))) ∙ + {!!} ∙ + (λ i → (rUnitₖ (lUnitₖ 0ₖ (~ i)) (~ i))) ∙ λ i → (0ₖ +ₖ 0ₖ) +ₖ lUnitₖ 0ₖ (~ i)) + ∙ {!!} + ∙ {!!} +{- +i = i0 ⊢ (sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i)))) j +i = i1 ⊢ (sym (rUnitₖ (0ₖ +ₖ 0ₖ)) ∙ (λ i → (0ₖ +ₖ 0ₖ) +ₖ (lUnitₖ 0ₖ (~ i)))) j +j = i0 ⊢ ∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x +ₖ g x) i +j = i1 ⊢ (∣ north ∣ +ₖ Kn→ΩKn+1 1 (f x) i) +ₖ + ∣ north ∣ +ₖ Kn→ΩKn+1 1 (g x) i +-} + + + helper' : Iso Int (Unit × Int) + Iso.inv helper' = proj₂ + Iso.fun helper' x = tt , x + Iso.leftInv helper' x = refl + Iso.rightInv helper' x = ×≡ refl refl + + helper : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} → Iso C Unit → Iso A B → Iso (A × C) B + Iso.fun (helper cUnit iAB) x = Iso.fun iAB (proj₁ x) + Iso.inv (helper cUnit iAB) x = (Iso.inv iAB x) , (Iso.inv cUnit tt ) + Iso.rightInv (helper cUnit iAB) = Iso.rightInv iAB + Iso.leftInv (helper cUnit iAB) b = ×≡ (Iso.leftInv iAB (proj₁ b)) (Iso.leftInv cUnit (proj₂ b)) + + helper2 : Iso (coHom 2 ((S₊ 1) × (S₊ 1))) (coHom 2 ((S₊ 1) × hLevelTrunc 3 (S₊ 1))) + Iso.fun helper2 = sRec setTruncIsSet (λ f → ∣ (λ {(x , y) → f (x , trRec isGroupoidS1 (idfun (S₊ 1)) y)}) ∣₀) + Iso.inv helper2 = sRec setTruncIsSet (λ f → ∣ (λ {(x , y) → f (x , ∣ y ∣)}) ∣₀) + Iso.rightInv helper2 = + sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + (λ f → cong ∣_∣₀ + (funExt λ {(x , y) → + trElim {B = λ y → f (x , ∣ trRec isGroupoidS1 (λ x₁ → x₁) y ∣) ≡ f (x , y)} + (λ _ → isOfHLevelPath' 3 (isOfHLevelTrunc 4) _ _) + (λ a → refl) y})) + Iso.leftInv helper2 = + sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₀ (funExt λ {(x , y) → refl}) +-} diff --git a/Cubical/ZCohomology/Groups/Unit.agda b/Cubical/ZCohomology/Groups/Unit.agda new file mode 100644 index 000000000..5f5b7d250 --- /dev/null +++ b/Cubical/ZCohomology/Groups/Unit.agda @@ -0,0 +1,78 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.Groups.Unit where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.HITs.Susp +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₋₁ ; ∣_∣ to ∣_∣₋₁) +open import Cubical.HITs.Nullification +open import Cubical.Data.Int hiding (_+_ ; +-comm) +open import Cubical.Data.Nat +open import Cubical.HITs.Truncation +open import Cubical.Homotopy.Connected +open import Cubical.Data.Unit +open import Cubical.Data.Group.Base renaming (Iso to grIso ; compIso to compGrIso ; invIso to invGrIso ; idIso to idGrIso) +open import Cubical.Data.Group.GroupLibrary +open import Cubical.ZCohomology.coHomZero-map + +H⁰-Unit≃ℤ : grIso (coHomGr 0 Unit) intGroup +grIso.fun H⁰-Unit≃ℤ = mph (sRec isSetInt (λ f → f tt)) + (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) + (λ a b → addLemma (a tt) (b tt))) +grIso.inv H⁰-Unit≃ℤ = mph (λ a → ∣ (λ _ → a) ∣₀) (λ a b i → ∣ (λ _ → addLemma a b (~ i)) ∣₀) +grIso.rightInv H⁰-Unit≃ℤ a = refl +grIso.leftInv H⁰-Unit≃ℤ = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ a → refl + + + +isContrHⁿ-Unit : (n : ℕ) → isContr (coHom (suc n) Unit) +isContrHⁿ-Unit n = subst isContr (λ i → ∥ UnitToTypeId (coHomK (suc n)) (~ i) ∥₀) (helper' n) + where + helper' : (n : ℕ) → isContr (∥ coHomK (suc n) ∥₀) + helper' n = + subst isContr + ((isoToPath (truncOfTruncIso {A = S₊ (1 + n)} 2 (1 + n))) + ∙ sym propTrunc≡Trunc0 ∙ λ i → ∥ hLevelTrunc (suc (+-comm n 2 i)) (S₊ (1 + n)) ∥₀) + (isConnectedSubtr 2 (helper2 n .fst) + (subst (λ x → isHLevelConnected x (S₊ (suc n))) (sym (helper2 n .snd)) (sphereConnected (suc n))) ) + where + helper2 : (n : ℕ) → Σ[ m ∈ ℕ ] m + 2 ≡ 2 + n + helper2 zero = 0 , refl + helper2 (suc n) = (suc n) , λ i → suc (+-comm n 2 i) + + +{- Hⁿ(Unit) for n ≥ 1 -} +Hⁿ-Unit≃0 : (n : ℕ) → grIso (coHomGr (suc n) Unit) trivialGroup +grIso.fun (Hⁿ-Unit≃0 n) = mph (λ _ → tt) + (λ _ _ → refl) +grIso.inv (Hⁿ-Unit≃0 n) = mph (λ _ → ∣ (λ _ → ∣ north ∣) ∣₀) + (λ _ _ → sym (rUnitₕ 0ₕ)) +grIso.rightInv (Hⁿ-Unit≃0 n) _ = refl +grIso.leftInv (Hⁿ-Unit≃0 n) _ = isOfHLevelSuc 0 (isContrHⁿ-Unit n) _ _ + + +{- Hⁿ for arbitrary contractible types -} +Hⁿ-contrType≃0 : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → isContr A + → grIso (coHomGr (suc n) A) trivialGroup +Hⁿ-contrType≃0 {A = A} n contr = + Iso'→Iso + (iso' + (iso (λ _ → tt) + (λ _ → 0ₕ) + (λ _ → refl) + λ a → isOfHLevelSuc 0 helper _ _) + λ _ _ → refl) + where + helper1 : Iso (coHom (suc n) A) (coHom (suc n) Unit) + helper1 = compIso (setTruncIso (ContrToTypeIso contr)) + (setTruncIso (symIso (ContrToTypeIso isContrUnit))) + + helper : isContr (coHom (suc n) A) + helper = (Iso.inv helper1 0ₕ) + , λ y → cong (Iso.inv helper1) (isOfHLevelSuc 0 (isContrHⁿ-Unit n) 0ₕ (Iso.fun helper1 y) ) + ∙ Iso.leftInv helper1 y diff --git a/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda b/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda new file mode 100644 index 000000000..a4ad82d3d --- /dev/null +++ b/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda @@ -0,0 +1,365 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.Groups.WedgeOfSpheres where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.S1.S1 +open import Cubical.HITs.S1 +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.MayerVietorisUnreduced +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence +open import Cubical.Data.NatMinusTwo.Base +open import Cubical.Data.Empty +open import Cubical.Data.Sigma hiding (_×_) +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₋₁ ; ∣_∣ to ∣_∣₋₁) +open import Cubical.HITs.Nullification +open import Cubical.Data.Int renaming (_+_ to _+ℤ_; +-comm to +ℤ-comm ; +-assoc to +ℤ-assoc) +open import Cubical.Data.Nat +open import Cubical.Data.Prod +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; recElim to trRec ; rec to trRec2 ; elim3 to trElim3) +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Freudenthal +open import Cubical.HITs.SmashProduct.Base +open import Cubical.Data.Unit +open import Cubical.Data.Group.Base renaming (Iso to grIso ; compIso to compGrIso ; invIso to invGrIso ; idIso to idGrIso) +open import Cubical.Data.Group.GroupLibrary +open import Cubical.ZCohomology.coHomZero-map +open import Cubical.HITs.Wedge + +open import Cubical.Foundations.Equiv.HalfAdjoint + +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.Groups.Torus + +open import Cubical.ZCohomology.KcompPrelims + + +open import Cubical.HITs.Pushout +open import Cubical.Data.Sum.Base +open import Cubical.Data.HomotopyGroup +open import Cubical.Data.Bool hiding (_⊕_) + +S₊∙ : (n : ℕ) → Pointed₀ +S₊∙ n = (S₊ n) , north + + +isOfHLevelPushout : ∀ {ℓ} {A : Type ℓ} (a : A) (n : ℕ) → isOfHLevel n A → isOfHLevel n ((A , a) ⋁ (A , a)) +isOfHLevelPushout {A = A} a zero hlev = (inr a) , (λ {(inr a2) → helper2 a2 + ; (inl a2) → helper a2 + ; (push t i) → pathP i}) + where + helper : (a2 : A) → Path ((A , a) ⋁ (A , a)) (inr a) (inl a2) + helper a2 = sym (push tt) ∙ λ i → inl (isOfHLevelSuc 0 hlev a a2 i) + helper2 : (a2 : A) → Path ((A , a) ⋁ (A , a)) (inr a) (inr a2) + helper2 a2 i = inr (isOfHLevelSuc 0 hlev a a2 i) + + pathP : PathP (λ i → Path ((A , a) ⋁ (A , a)) (inr a) (push tt i)) + (helper a) (helper2 a) + pathP j i = + hcomp (λ k → λ { (i = i0) → inr a + ; (i = i1) → push tt j + ; (j = i0) → ((rUnit (sym (push tt))) + ∙ (λ i → sym (push tt) ∙ (λ j → inl (helper3 (~ i) j)))) k i + ; (j = i1) → inr (helper3 (~ k) i)}) + (push tt (~ i ∨ j)) + where + helper3 : (isOfHLevelSuc 0 hlev a a) ≡ refl + helper3 = isOfHLevelPlus 2 hlev a a (isOfHLevelSuc 0 hlev a a) refl + +isOfHLevelPushout {A = A} a (suc zero) hlev x y = {!!} + where + helper : (a2 : A) → Path ((A , a) ⋁ (A , a)) (inr a) (inl a2) + helper a2 = sym (push tt) ∙ λ i → inl (hlev a a2 i) + helper2 : (a2 : A) → Path ((A , a) ⋁ (A , a)) (inr a) (inr a2) + helper2 a2 i = inr (hlev a a2 i) + + pathP : PathP (λ i → Path ((A , a) ⋁ (A , a)) (inr a) (push tt i)) + (helper a) (helper2 a) + pathP j i = + hcomp (λ k → λ { (i = i0) → inr a + ; (i = i1) → push tt j + ; (j = i0) → ((rUnit (sym (push tt))) + ∙ (λ i → sym (push tt) ∙ (λ j → inl (helper3 (~ i) j)))) k i + ; (j = i1) → inr (helper3 (~ k) i)}) + (push tt (~ i ∨ j)) + where + helper3 : hlev a a ≡ refl + helper3 = isOfHLevelSuc 1 hlev a a _ _ + +isOfHLevelPushout a (suc (suc n)) hlev (inl x) (inl y) = {!!} +isOfHLevelPushout a (suc (suc n)) hlev (inl x) (inr x₁) = {!!} +isOfHLevelPushout a (suc (suc n)) hlev (inl x) (push a₁ i) = {!!} +isOfHLevelPushout a (suc (suc n)) hlev (inr x) y = {!!} +isOfHLevelPushout a (suc (suc n)) hlev (push a₁ i) y = {!!} + +PushoutToProp : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {f : A → B} {g : A → C} + {D : Pushout f g → Type ℓ'''} + → ((x : Pushout f g) → isProp (D x)) + → ((a : B) → D (inl a)) + → ((c : C) → D (inr c)) + → (x : Pushout f g) → D x +PushoutToProp isset baseB baseC (inl x) = baseB x +PushoutToProp isset baseB baseC (inr x) = baseC x +PushoutToProp {f = f} {g = g} isset baseB baseC (push a i) = + isOfHLevel→isOfHLevelDep 1 isset (baseB (f a)) (baseC (g a)) (push a) i + +H⁰-S¹⋁S¹ : grIso intGroup (coHomGr 0 ((S₊∙ 1) ⋁ S₊∙ 1)) +H⁰-S¹⋁S¹ = + Iso'→Iso + (iso' + (iso + (λ a → ∣ (λ _ → a) ∣₀) + (sRec isSetInt (λ f → f (inl north))) + (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + (λ f → cong ∣_∣₀ + (funExt (PushoutToProp + (λ _ → isSetInt _ _) + (suspToPropRec north (λ _ → isSetInt _ _) refl) + (suspToPropRec north (λ _ → isSetInt _ _) λ i → f (push tt i)))))) + λ _ → refl) + λ a b → cong ∣_∣₀ (funExt λ _ → sym (addLemma a b))) + +test0 : Iso (((S₊ 1 , north) ⋁ ((S₊ 1) , north) → Int)) Int +Iso.fun test0 f = f (inl north) +Iso.inv test0 a _ = a +Iso.rightInv test0 = λ a → refl +Iso.leftInv test0 a = + funExt (PushoutToProp (λ _ → isSetInt _ _) + (suspToPropRec north (λ _ → isSetInt _ _) refl) + (suspToPropRec north (λ _ → isSetInt _ _) λ i → a (push tt i))) + + + +test : Iso (((S₊ 1 , north) ⋁ ((S₊ 1) , north) → hLevelTrunc 3 (S₊ 1))) ((S₊ 1 → hLevelTrunc 3 (S₊ 1)) × (S₊ 1 → hLevelTrunc 3 (S₊ 1))) +Iso.fun test f = (λ x → f (inl x)) , (λ x → f (inr x)) -- (λ x → f (inl x)) , λ x → f (inr x) +Iso.inv test (f , g) (inl x) = f x +ₖ (-ₖ f north) -- f x +Iso.inv test (f , g) (inr x) = g x +ₖ (-ₖ g north) -- g x +Iso.inv test (f , g) (push a i) = (rCancelₖ (f north) ∙ (sym (rCancelₖ (g north)))) i +Iso.rightInv test (f , g) = ×≡ (funExt {!!}) {!!} +Iso.leftInv test a = funExt λ {(inl x) → {!!} ; (inr x) → {!!} ; (push a i) → {!!}} + + + +mapsTo0 : (f : S₊ 1 → hLevelTrunc 3 (S₊ 1)) + → ∥ f north ≡ ∣ north ∣ ∥₋₁ +mapsTo0 f = coInd f (f north) refl + where + coInd : (f : S₊ 1 → hLevelTrunc 3 (S₊ 1)) (x : hLevelTrunc 3 (S₊ 1)) → f north ≡ x → ∥ f north ≡ ∣ north ∣ ∥₋₁ + coInd f = + trElim (λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelPlus {n = 1} 2 propTruncIsProp) + (suspToPropRec north + (λ _ → isOfHLevelΠ 1 λ _ → propTruncIsProp) + ∣_∣₋₁) + +mapsTo0' : (f : S₊ 1 → hLevelTrunc 3 (S₊ 1)) (x : _) + → ∥ f x ≡ ∣ north ∣ ∥₋₁ +mapsTo0' f = suspToPropRec north (λ _ → propTruncIsProp) (mapsTo0 f) + +testID : Iso ∥ (((S¹ , base) ⋁ (S¹ , base)) → S¹) ∥₀ ∥ ((S¹ × S¹) → S¹) ∥₀ +Iso.fun testID = sRec setTruncIsSet λ f → ∣ (λ {(a , b) → f {!!}}) ∣₀ +Iso.inv testID a = {!!} +Iso.rightInv testID a = {!!} +Iso.leftInv testID = {!!} + + +maybe13 : Iso (Σ[ f ∈ coHom 1 (S₊∙ 1 ⋁ S₊∙ 1) ] isInIm (coHomGr 0 Unit) (coHomGr 1 _) (morph.fun (MV.d (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 0)) f) + ∥ S₊ 1 ∥₀ +Iso.fun maybe13 (f , inim) = + pRec helper (sigmaElim (λ _ → setTruncIsSet) (λ g id → {!!})) inim + where + helper : isProp ∥ S₊ 1 ∥₀ + helper = {!!} +Iso.inv maybe13 = {!!} +{- + sRec (isOfHLevelΣ 2 setTruncIsSet (λ _ → isOfHLevelSuc 1 (propTruncIsProp))) + (λ p → ∣ (λ {(inl x) → 0ₖ ; (inr x) → 0ₖ ; (push tt i) → p i}) ∣₀ + , ∣ ∣ (λ _ → ΩKn+1→Kn p) ∣₀ , (cong ∣_∣₀ (funExt λ {(inl x) → refl + ; (inr x) → refl + ; (push tt i) → cong (λ x → x i) + (Iso.rightInv (Iso3-Kn-ΩKn+1 0) p)})) ∣₋₁) -} +Iso.rightInv maybe13 = {!!} +Iso.leftInv maybe13 (f , inim) = + pRec {!!} {!!} inim + + +maybe : (a b : Int) (x : MV.D (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north)) → MV.d-pre (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 0 (λ _ → a) x ≡ MV.d-pre (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 0 (λ _ → b) x +maybe a b (inl x) = refl +maybe a b (inr x) = refl +maybe a b (push a₁ i) j = {!Kn→ΩKn+1 zero ?!} + +testID2 : Iso (coHom 1 ((S₊∙ 1) ⋁ S₊∙ 1)) (coHom 1 (S₊ 1 × S₊ 1)) +Iso.fun testID2 = sRec setTruncIsSet λ f → ∣ (λ {(a , b) → f (inl a) +ₖ f (inl b)}) ∣₀ +Iso.inv testID2 = sRec setTruncIsSet λ f → ∣ (λ {(inl x) → f (x , north) ; (inr x) → f (north , x) ; (push tt i) → f (north , snd (S₊∙ 1))}) ∣₀ +Iso.rightInv testID2 = sElim {!!} λ f → cong ∣_∣₀ (funExt (λ {(a , b) → {!!}})) + where + helper : (f : S₊ 1 × S₊ 1 → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1) → (f (north , north) ≡ 0ₖ ) → (a b : S₊ 1) → f (a , north) +ₖ f (b , north) ≡ f (a , b) + helper f id north north = (cong ((f (north , north)) +ₖ_) id) ∙ rUnitₖ (f (north , north)) + helper f id north south = cong (f (north , north) +ₖ_) (λ i → f (merid north (~ i) , north)) ∙ cong ((f (north , north)) +ₖ_) id ∙ rUnitₖ (f (north , north)) ∙ λ i → f (north , merid north i) + helper f id north (merid a i) = {!!} + where + testhelper : {!!} ∙ (cong ((f (north , north)) +ₖ_) id) ∙ rUnitₖ (f (north , north)) ∙ {!λ i → f (north , merid a i)!} ≡ cong (f (north , north) +ₖ_) (λ i → f (merid north (~ i) , north)) ∙ cong ((f (north , north)) +ₖ_) id ∙ rUnitₖ (f (north , north)) ∙ λ i → f (north , merid north i) + testhelper = {!!} + helper f id south north = {!!} + helper f id south south = {!!} + helper f id south (merid a i) = {!!} + helper f id (merid a i) north = {!!} + helper f id (merid a i) south = {!!} + helper f id (merid a i) (merid a₁ i₁) j = {!!} +Iso.leftInv testID2 = sElim {!!} λ f → cong ∣_∣₀ (funExt λ {(inl x) → {!!} ; (inr x) → {!!} ; (push tt i) → {!!}} ) + +testId3 : Iso (coHom 1 ((S₊∙ 1) ⋁ S₊∙ 1)) (coHom 1 (S₊ 1) × coHom 1 (S₊ 1)) +Iso.fun testId3 = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) λ f → ∣ (λ x → {!!}) ∣₀ , {!!} +Iso.inv testId3 = {!!} +Iso.rightInv testId3 = {!!} +Iso.leftInv testId3 = {!!} + + + + + +test1' : Iso (((S¹ , base) ⋁ (S¹ , base)) → S¹) Int +Iso.fun test1' f = winding (basechange2⁻ _ ((λ i → f (inl (loop i))) ∙ (λ i → f (push tt i)) ∙ (λ i → f (inr (loop i))) ∙ λ i → f (push tt (~ i)))) +Iso.inv test1' int (inl x) = {!!} +Iso.inv test1' int (inr x) = x +Iso.inv test1' int (push a i) = basechange2 {!!} {!!} i +Iso.rightInv test1' = {!!} +Iso.leftInv test1' = {!!} + +test1 : Iso (coHom 1 (S₊∙ 1 ⋁ S₊∙ 1)) (coHom 1 (S₊ 1) × coHom 1 (S₊ 1)) +Iso.fun test1 = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) + λ f → ∣ (λ x → f (inl x)) ∣₀ , ∣ (λ x → f (inr x)) ∣₀ +Iso.inv test1 (f , g) = + sElim2 (λ _ _ → setTruncIsSet) + (λ f g → ∣ (λ {(inl x) → f x +ₖ (-ₖ f north) + ; (inr x) → g x +ₖ (-ₖ (g north)) + ; (push a i) → (rCancelₖ (f north) ∙ (sym (rCancelₖ (g north)))) i}) ∣₀) + f g + where + +Iso.rightInv test1 = + prodElim (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ f g → pRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet _ _) + (λ f0 → pRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet _ _) + (λ g0 → ×≡ (cong ∣_∣₀ (funExt λ {north → cong (λ x → f north +ₖ (-ₖ x)) f0 ∙ cong (f north +ₖ_) -0ₖ ∙ rUnitₖ (f north) + ; south → cong (λ x → f south +ₖ (-ₖ x)) f0 ∙ cong (f south +ₖ_) -0ₖ ∙ rUnitₖ (f south) + ; (merid a i) → cong (λ x → f (merid a i) +ₖ (-ₖ x)) f0 ∙ cong (f (merid a i) +ₖ_) -0ₖ ∙ rUnitₖ (f (merid a i))})) + ((cong ∣_∣₀ (funExt λ {north → cong (λ x → g north +ₖ (-ₖ x)) g0 ∙ cong (g north +ₖ_) -0ₖ ∙ rUnitₖ (g north) + ; south → cong (λ x → g south +ₖ (-ₖ x)) g0 ∙ cong (g south +ₖ_) -0ₖ ∙ rUnitₖ (g south) + ; (merid a i) → cong (λ x → g (merid a i) +ₖ (-ₖ x)) g0 ∙ cong (g (merid a i) +ₖ_) -0ₖ ∙ rUnitₖ (g (merid a i))})))) + (mapsTo0 g)) + (mapsTo0 f) + +{- d i + H⁰(Unit) → H¹(S¹∨S¹) → H¹(S¹) × H¹(S¹) → H¹(Unit) + ℤ ℤ×ℤ 0 +-} + + + + +Iso.leftInv test1 = + sElim {!!} + λ f → pRec {!f!} + (λ fl → pRec {!!} + (λ fr → cong ∣_∣₀ (funExt λ {(inl x) → cong (λ y → f (inl x) +ₖ (-ₖ y)) fl ∙ cong (f (inl x) +ₖ_) -0ₖ ∙ rUnitₖ (f (inl x)) + ; (inr x) → cong (λ y → f (inr x) +ₖ (-ₖ y)) fr ∙ cong (f (inr x) +ₖ_) -0ₖ ∙ rUnitₖ (f (inr x)) + ; (push a i) → {!fl -- cong (λ y → f (inl ?) +ₖ (-ₖ y)) fl ∙ cong (f (inl ?) +ₖ_) -0ₖ ∙ rUnitₖ (f (inl ?))!}})) + (mapsTo0 λ x → f (inr x))) + (mapsTo0 λ x → f (inl x)) + where + test2 : Iso.inv test1 ≡ {!!} + test2 = {!!} + + helper4 : (f : S₊∙ 1 ⋁ S₊∙ 1 → ∥ S₊ 1 ∥ ℕ→ℕ₋₂ 1) + → (x y : hLevelTrunc 3 (S₊ 1)) + → (p : (x ≡ ∣ north ∣)) + → (q : (y ≡ ∣ north ∣)) + → (P : x ≡ y) + → ∥ PathP (λ i → P i ≡ 0ₖ) p q ∥₋₁ + helper4 f x y p q = + {!!} + where + helper5 : (x y : hLevelTrunc 3 (S₊ 1)) → ∥ x ≡ y ∥₋₁ + helper5 = elim2 (λ _ _ → isOfHLevelPlus {n = 1} 2 propTruncIsProp) + (suspToPropRec2 north (λ _ _ → propTruncIsProp) + ∣ refl ∣₋₁) + + +{- d i + H⁰(Unit) → H¹(S¹∨S¹) → H¹(S¹) × H¹(S¹) → H¹(Unit) + ℤ ℤ×ℤ 0 +-} + +anotherhelper : (x : coHom 0 Unit) + → isInIm (×coHomGr 0 (S₊ 1) (S₊ 1)) (coHomGr 0 Unit) (morph.fun (MV.Δ (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 0)) x +anotherhelper = + sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ f → ∣ (∣ (λ _ → f tt) ∣₀ , 0ₕ) , cong ∣_∣₀ (funExt (λ _ → cong ((f tt) +ₖ_) -0ₖ ∙ rUnitₖ (f tt))) ∣₋₁ + +athirdhelper : (x : coHom 0 Unit) + → isInKer (coHomGr 0 Unit) (coHomGr 1 _) (morph.fun (MV.d (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 0)) x +athirdhelper x = MV.Im-Δ⊂Ker-d (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 0 x + (anotherhelper x) + + +afourthhelper : (x : coHom 1 (S₊∙ 1 ⋁ S₊∙ 1)) → isInIm (coHomGr 0 Unit) (coHomGr 1 _) (morph.fun (MV.d (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 0)) x + → x ≡ 0ₕ +afourthhelper x inim = pRec (setTruncIsSet _ _) (λ {(y , id) → sym id ∙ athirdhelper y}) inim + +H¹-S¹∨S¹ : grIso (coHomGr 1 (S₊∙ 1 ⋁ S₊∙ 1)) (×coHomGr 1 (S₊ 1) (S₊ 1)) +H¹-S¹∨S¹ = + Iso''→Iso + (iso'' + (MV.i (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 1) + (λ x inker → afourthhelper x {!!}) + λ x → {!MV.Ker-Δ⊂Im-i (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 1 x ?!}) + + + + + +-- H¹-S¹∨S¹ : grIso (coHomGr 1 (S₊∙ 1 ⋁ S₊∙ 1)) (×coHomGr 1 (S₊ 1) (S₊ 1)) +-- H¹-S¹∨S¹ = +-- Iso''→Iso +-- (iso'' +-- (MV.i (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 1) +-- (λ a inker → {!!}) +-- λ f → {!!}) +-- where +-- test3 : (g : Unit → Int) → {!!} +-- test3 = {!!} + + +-- coHom1-S1∨S1 : (x : coHom 1 (S₊∙ 1 ⋁ S₊∙ 1)) +-- → isInIm (coHomGr 0 Unit) (coHomGr 1 (S₊∙ 1 ⋁ S₊∙ 1)) +-- (morph.fun (MV.d (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 0)) x +-- → x ≡ 0ₕ +-- coHom1-S1∨S1 = +-- sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- λ f → pRec (setTruncIsSet _ _) +-- (sigmaElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- λ g id → pRec (setTruncIsSet _ _) +-- (λ id → {!!}) +-- (Iso.fun PathIdTrunc₀Iso id)) +-- where +-- id : (g : (Unit → Int)) +-- → (x : MV.D (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north)) +-- → MV.d-pre (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 0 g x ≡ 0ₖ +-- id g (inl x) = refl +-- id g (inr x) = refl +-- id g (push a i) j = {!!} + + diff --git a/Cubical/ZCohomology/MayerVietorisUnreduced.agda b/Cubical/ZCohomology/MayerVietorisUnreduced.agda index ff032ba5f..a842c32c7 100644 --- a/Cubical/ZCohomology/MayerVietorisUnreduced.agda +++ b/Cubical/ZCohomology/MayerVietorisUnreduced.agda @@ -249,31 +249,32 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : ; (j = i1) → F (push a i)}) (pushFiller (suc n) F p1 p2 a j i) - Im-i⊂Ker-Δ : (n : ℕ) (x : coHom n A × coHom n B) - → isInIm (coHomGr n D) (×coHomGr n A B) (morph.fun (i n)) x - → isInKer (×coHomGr n A B) (coHomGr n C) (morph.fun (Δ n)) x - Im-i⊂Ker-Δ n (Fa , Fb) = - sElim {B = λ Fa → (Fb : _) → isInIm (coHomGr n D) (×coHomGr n A B) (morph.fun (i n)) (Fa , Fb) - → isInKer (×coHomGr n A B) (coHomGr n C) (morph.fun (Δ n)) (Fa , Fb)} - (λ _ → isOfHLevelΠ 2 λ _ → (isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _)) - (λ Fa → sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) - λ Fb → pRec (setTruncIsSet _ _) - (sigmaElim (λ x → isOfHLevelSuc 1 (setTruncIsSet _ _)) - λ Fd p → helper n Fa Fb Fd p)) - Fa - Fb - where - helper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) (Fd : D → coHomK n) - → (morph.fun (i n) ∣ Fd ∣₀ ≡ (∣ Fa ∣₀ , ∣ Fb ∣₀)) - → (morph.fun (Δ n)) (∣ Fa ∣₀ , ∣ Fb ∣₀) ≡ 0ₕ - helper zero Fa Fb Fd p = cong (morph.fun (Δ zero)) (sym p) - ∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₀ +ₕ -ₕ ∣ (λ x → Fd (push x (~ i))) ∣₀) - ∙ rCancelₕ ∣ (λ x → Fd (inl (f x))) ∣₀ - helper (suc n) Fa Fb Fd p = cong (morph.fun (Δ (suc n))) (sym p) - ∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₀ +ₕ -ₕ ∣ (λ x → Fd (push x (~ i))) ∣₀) - ∙ rCancelₕ ∣ (λ x → Fd (inl (f x))) ∣₀ + abstract + Im-i⊂Ker-Δ : (n : ℕ) (x : coHom n A × coHom n B) + → isInIm (coHomGr n D) (×coHomGr n A B) (morph.fun (i n)) x + → isInKer (×coHomGr n A B) (coHomGr n C) (morph.fun (Δ n)) x + Im-i⊂Ker-Δ n (Fa , Fb) = + sElim {B = λ Fa → (Fb : _) → isInIm (coHomGr n D) (×coHomGr n A B) (morph.fun (i n)) (Fa , Fb) + → isInKer (×coHomGr n A B) (coHomGr n C) (morph.fun (Δ n)) (Fa , Fb)} + (λ _ → isOfHLevelΠ 2 λ _ → (isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _)) + (λ Fa → sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ Fb → pRec (setTruncIsSet _ _) + (sigmaElim (λ x → isOfHLevelSuc 1 (setTruncIsSet _ _)) + λ Fd p → helper n Fa Fb Fd p)) + Fa + Fb + where + helper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) (Fd : D → coHomK n) + → (morph.fun (i n) ∣ Fd ∣₀ ≡ (∣ Fa ∣₀ , ∣ Fb ∣₀)) + → (morph.fun (Δ n)) (∣ Fa ∣₀ , ∣ Fb ∣₀) ≡ 0ₕ + helper zero Fa Fb Fd p = cong (morph.fun (Δ zero)) (sym p) + ∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₀ +ₕ -ₕ ∣ (λ x → Fd (push x (~ i))) ∣₀) + ∙ rCancelₕ ∣ (λ x → Fd (inl (f x))) ∣₀ + helper (suc n) Fa Fb Fd p = cong (morph.fun (Δ (suc n))) (sym p) + ∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₀ +ₕ -ₕ ∣ (λ x → Fd (push x (~ i))) ∣₀) + ∙ rCancelₕ ∣ (λ x → Fd (inl (f x))) ∣₀ + - Ker-Δ⊂Im-i : (n : ℕ) (a : coHom n A × coHom n B) → isInKer (×coHomGr n A B) (coHomGr n C) (morph.fun (Δ n)) a diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index 220dd93fd..d79507903 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -129,6 +129,7 @@ Kn→ΩKn+10ₖ (suc n) = (λ i → cong ∣_∣ (rCancel (merid north) i)) -- c -- i = i1 ⊢ -0ₖ j -- -} + +ₖ→∙ : (n : ℕ) (a b : coHomK n) → Kn→ΩKn+1 n (a +ₖ b) ≡ Kn→ΩKn+1 n a ∙ Kn→ΩKn+1 n b +ₖ→∙ n a b = Iso.rightInv (Iso3-Kn-ΩKn+1 n) (Kn→ΩKn+1 n a ∙ Kn→ΩKn+1 n b) From b7bc88c97404b2c8e1b3363ff92a0820bd3d296a Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Sun, 14 Jun 2020 12:03:25 +0200 Subject: [PATCH 45/89] backup_before_merge --- Cubical/Data/Group/Base.agda | 145 ++++-- Cubical/Data/Group/GroupLibrary.agda | 18 +- Cubical/HITs/SetTruncation/Properties.agda | 2 +- Cubical/ZCohomology/Groups/Sn.agda | 2 +- Cubical/ZCohomology/Groups/Wedge.agda | 98 ++++ .../ZCohomology/Groups/WedgeOfSpheres.agda | 34 +- .../ZCohomology/MayerVietorisUnreduced.agda | 316 +++++++------ .../ZCohomology/MayerVietorisUnreduced2.agda | 426 ++++++++++++++++++ 8 files changed, 827 insertions(+), 214 deletions(-) create mode 100644 Cubical/ZCohomology/Groups/Wedge.agda create mode 100644 Cubical/ZCohomology/MayerVietorisUnreduced2.agda diff --git a/Cubical/Data/Group/Base.agda b/Cubical/Data/Group/Base.agda index ccf7397f0..001292d7c 100644 --- a/Cubical/Data/Group/Base.agda +++ b/Cubical/Data/Group/Base.agda @@ -38,7 +38,6 @@ isMorph : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (f : (Group.type G isMorph G H f = (g0 g1 : Group.type G) → f (isGroup.comp (Group.groupStruc G) g0 g1) ≡ isGroup.comp (Group.groupStruc H) (f g0) (f g1) record morph {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') : Type (ℓ-max ℓ ℓ') where - no-eta-equality constructor mph field fun : Group.type G → Group.type H @@ -48,20 +47,20 @@ record morph {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') : Type (ℓ-max ℓ ℓ open import Cubical.Data.Sigma hiding (_×_ ; comp) -isInIm : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (Group.type G → Group.type H) +isInIm : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (morph G H) → Group.type H → Type _ -isInIm G H ϕ h = ∃[ g ∈ Group.type G ] ϕ g ≡ h +isInIm G H ϕ h = ∃[ g ∈ Group.type G ] (morph.fun ϕ g) ≡ h -isInKer : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (Group.type G → Group.type H) +isInKer : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → morph G H → Group.type G → Type _ -isInKer G H ϕ g = ϕ g ≡ isGroup.id (Group.groupStruc H) +isInKer G H ϕ g = (morph.fun ϕ g) ≡ isGroup.id (Group.groupStruc H) isSurjective : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → morph G H → Type _ -isSurjective G H ϕ = (x : Group.type H) → isInIm G H (morph.fun ϕ) x +isSurjective G H ϕ = (x : Group.type H) → isInIm G H ϕ x isInjective : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → morph G H → Type _ -isInjective G H ϕ = (x : Group.type G) → isInKer G H (morph.fun ϕ) x → x ≡ isGroup.id (Group.groupStruc G) +isInjective G H ϕ = (x : Group.type G) → isInKer G H ϕ x → x ≡ isGroup.id (Group.groupStruc G) -0≡0 : ∀ {ℓ} {G : Group ℓ} → isGroup.inv (Group.groupStruc G) (isGroup.id (Group.groupStruc G)) ≡ isGroup.id (Group.groupStruc G) @@ -251,15 +250,37 @@ record Iso' {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') : Type (ℓ-max ℓ ℓ' isoSetMorph : isMorph G H (I.Iso.fun isoSet) record Iso'' {ℓ ℓ'} (A : Group ℓ) (B : Group ℓ') : Type (ℓ-max ℓ ℓ') where + no-eta-equality constructor iso'' field ϕ : morph A B inj : isInjective A B ϕ surj : isSurjective A B ϕ +record Iso''' {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') : Type (ℓ-max ℓ ℓ') where + constructor iso''' + field + fun : morph G H + inv : Group.type H → Group.type G + rightInv : I.section (morph.fun fun) inv + leftInv : I.retract (morph.fun fun) inv + _≃_ : ∀ {ℓ ℓ'} (A : Group ℓ) (B : Group ℓ') → Type (ℓ-max ℓ ℓ') A ≃ B = Σ (morph A B) \ f → (E.isEquiv (morph.fun f)) +open isGroup +Iso'''→Iso : ∀ {ℓ ℓ'} {G : Group ℓ} {H : Group ℓ'} → Iso''' G H → Iso G H +Iso.fun (Iso'''→Iso is) = Iso'''.fun is +morph.fun (Iso.inv (Iso'''→Iso is)) = Iso'''.inv is +morph.ismorph (Iso.inv (Iso'''→Iso {G = G} {H = H} is)) a b = + cong₂ (λ x y → (Iso'''.inv is) (comp (Group.groupStruc H) x y)) + (sym (Iso'''.rightInv is _)) + (sym (Iso'''.rightInv is _)) + ∙ cong (Iso'''.inv is) (sym (morph.ismorph (Iso'''.fun is) _ _)) + ∙ Iso'''.leftInv is _ +Iso.rightInv (Iso'''→Iso is) = Iso'''.rightInv is +Iso.leftInv (Iso'''→Iso is) = Iso'''.leftInv is + Iso'→Iso : ∀ {ℓ ℓ'} {G : Group ℓ} {H : Group ℓ'} → Iso' G H → Iso G H morph.fun (Iso.fun (Iso'→Iso {G = G} {H = H} iso1)) = I.Iso.fun (Iso'.isoSet iso1) morph.ismorph (Iso.fun (Iso'→Iso {G = G} {H = H} iso1)) = Iso'.isoSetMorph iso1 @@ -326,39 +347,82 @@ Iso''→Iso {A = A} {B = B} is = λ b i → rec (helper (morph.fun (Iso''.ϕ is) b)) (λ a → a) (propTruncIsProp (Iso''.surj is (morph.fun (Iso''.ϕ is) b)) ∣ b , refl ∣ i) .fst) (morph.ismorph (Iso''.ϕ is))) where - helper : (b : _) → isProp (Σ (Group.type A) (λ a → (morph.fun (Iso''.ϕ is)) a ≡ b)) - helper b (a1 , pf1) (a2 , pf2) = - ΣProp≡ (λ _ → isOfHLevelPath' 1 (Group.setStruc B) _ _) - fstId - where - open Group - open isGroup - open morph - A' = groupStruc A - B' = groupStruc B - ϕ = morph.fun (Iso''.ϕ is) - ϕmf = morph.ismorph (Iso''.ϕ is) - - fstIdHelper : isGroup.comp (Group.groupStruc A) a1 (isGroup.inv (Group.groupStruc A) a2) - ≡ isGroup.id (Group.groupStruc A) - fstIdHelper = - let -A = isGroup.inv (Group.groupStruc A) - -B = isGroup.inv (Group.groupStruc B) - rCancelB = isGroup.rCancel (Group.groupStruc B) - _+B_ = isGroup.comp (Group.groupStruc B) - in Iso''.inj is _ - (ϕmf a1 (inv A' a2) - ∙ cong (λ x → comp B' (ϕ a1) x) (morphMinus A B (Iso''.ϕ is) a2) - ∙ cong (λ x → comp B' x (inv B' (ϕ a2))) (pf1 ∙ sym pf2) - ∙ rCancel B' (ϕ a2)) - fstId : a1 ≡ a2 - fstId = - a1 ≡⟨ sym (rUnit A' a1) ⟩ - comp A' a1 (id A') ≡⟨ cong (λ x → comp A' a1 x) (sym (lCancel A' a2)) ⟩ - comp A' a1 (comp A' (inv A' a2) a2) ≡⟨ sym (assoc A' a1 (inv A' a2) a2) ⟩ - comp A' (comp A' a1 (inv A' a2)) a2 ≡⟨ cong (λ x → comp A' x a2) fstIdHelper ⟩ - comp A' (id A') a2 ≡⟨ lUnit A' a2 ⟩ - a2 ∎ + abstract + helper : (b : _) → isProp (Σ (Group.type A) (λ a → (morph.fun (Iso''.ϕ is)) a ≡ b)) + helper _ a b = + ΣProp≡ (λ _ → isOfHLevelPath' 1 (Group.setStruc B) _ _) + fstId + where + open Group + open isGroup + open morph + A' = groupStruc A + B' = groupStruc B + + fstIdHelper : isGroup.comp (Group.groupStruc A) (fst a) (isGroup.inv (Group.groupStruc A) (fst b)) + ≡ isGroup.id (Group.groupStruc A) + fstIdHelper = + let -A = isGroup.inv (Group.groupStruc A) + -B = isGroup.inv (Group.groupStruc B) + rCancelB = isGroup.rCancel (Group.groupStruc B) + _+B_ = isGroup.comp (Group.groupStruc B) + in Iso''.inj is _ + (morph.ismorph (Iso''.ϕ is) (fst a) (inv A' (fst b)) + ∙ cong (λ x → comp B' (morph.fun (Iso''.ϕ is) (fst a)) x) (morphMinus A B (Iso''.ϕ is) (fst b)) + ∙ cong (λ x → comp B' x (inv B' (morph.fun (Iso''.ϕ is) (fst b)))) ((snd a) ∙ sym (snd b)) + ∙ rCancel B' (morph.fun (Iso''.ϕ is) (fst b))) + fstId : fst a ≡ fst b + fstId = + (fst a) ≡⟨ sym (rUnit A' (fst a)) ⟩ + comp A' (fst a) (id A') ≡⟨ cong (λ x → comp A' (fst a) x) (sym (lCancel A' (fst b))) ⟩ + comp A' (fst a) (comp A' (inv A' (fst b)) (fst b)) ≡⟨ sym (assoc A' (fst a) (inv A' (fst b)) (fst b)) ⟩ + comp A' (comp A' (fst a) (inv A' (fst b))) (fst b) ≡⟨ cong (λ x → comp A' x (fst b)) fstIdHelper ⟩ + comp A' (id A') (fst b) ≡⟨ lUnit A' (fst b) ⟩ + (fst b) ∎ + + + +Iso''→Iso2 : ∀ {ℓ ℓ'} {A : Group ℓ} {B : Group ℓ'} → Iso'' A B → Iso A B +Iso''→Iso2 {A = A} {B = B} is = Iso'''→Iso theIso + where + abstract + helper : (b : _) → isProp (Σ (Group.type A) (λ a → (morph.fun (Iso''.ϕ is)) a ≡ b)) + helper _ a b = + ΣProp≡ (λ _ → isOfHLevelPath' 1 (Group.setStruc B) _ _) + fstId + where + open Group + open isGroup + open morph + A' = groupStruc A + B' = groupStruc B + + fstIdHelper : isGroup.comp (Group.groupStruc A) (fst a) (isGroup.inv (Group.groupStruc A) (fst b)) + ≡ isGroup.id (Group.groupStruc A) + fstIdHelper = + let -A = isGroup.inv (Group.groupStruc A) + -B = isGroup.inv (Group.groupStruc B) + rCancelB = isGroup.rCancel (Group.groupStruc B) + _+B_ = isGroup.comp (Group.groupStruc B) + in Iso''.inj is _ + (morph.ismorph (Iso''.ϕ is) (fst a) (inv A' (fst b)) + ∙ cong (λ x → comp B' (morph.fun (Iso''.ϕ is) (fst a)) x) (morphMinus A B (Iso''.ϕ is) (fst b)) + ∙ cong (λ x → comp B' x (inv B' (morph.fun (Iso''.ϕ is) (fst b)))) ((snd a) ∙ sym (snd b)) + ∙ rCancel B' (morph.fun (Iso''.ϕ is) (fst b))) + fstId : fst a ≡ fst b + fstId = + (fst a) ≡⟨ sym (rUnit A' (fst a)) ⟩ + comp A' (fst a) (id A') ≡⟨ cong (λ x → comp A' (fst a) x) (sym (lCancel A' (fst b))) ⟩ + comp A' (fst a) (comp A' (inv A' (fst b)) (fst b)) ≡⟨ sym (assoc A' (fst a) (inv A' (fst b)) (fst b)) ⟩ + comp A' (comp A' (fst a) (inv A' (fst b))) (fst b) ≡⟨ cong (λ x → comp A' x (fst b)) fstIdHelper ⟩ + comp A' (id A') (fst b) ≡⟨ lUnit A' (fst b) ⟩ + (fst b) ∎ + + theIso : Iso''' A B + Iso'''.fun theIso = Iso''.ϕ is + Iso'''.inv theIso b = rec (helper b) (λ a → a) (Iso''.surj is b) .fst + Iso'''.rightInv theIso b = rec (helper b) (λ a → a) (Iso''.surj is b) .snd + Iso'''.leftInv theIso b i = rec (helper (morph.fun (Iso''.ϕ is) b)) (λ a → a) (propTruncIsProp (Iso''.surj is (morph.fun (Iso''.ϕ is) b)) ∣ b , refl ∣ i) .fst groupIso→Iso : ∀ {ℓ ℓ'} {G : Group ℓ} {H : Group ℓ'} → Iso G H → I.Iso (Group.type G) (Group.type H) I.Iso.fun (groupIso→Iso i) = morph.fun (Iso.fun i) @@ -366,7 +430,6 @@ I.Iso.inv (groupIso→Iso i) = morph.fun (Iso.inv i) I.Iso.rightInv (groupIso→Iso i) = Iso.rightInv i I.Iso.leftInv (groupIso→Iso i) = Iso.leftInv i - -- -- Injectivity and surjectivity in terms of exact sequences -- {- diff --git a/Cubical/Data/Group/GroupLibrary.agda b/Cubical/Data/Group/GroupLibrary.agda index 73efef0f0..09394396f 100644 --- a/Cubical/Data/Group/GroupLibrary.agda +++ b/Cubical/Data/Group/GroupLibrary.agda @@ -157,11 +157,11 @@ record SES {ℓ ℓ' ℓ'' ℓ'''} (A : Group ℓ) (B : Group ℓ') (leftGr : Gr ϕ : morph A B Ker-ϕ⊂Im-left : (x : Group.type A) -- - → isInKer A B (morph.fun ϕ) x - → isInIm leftGr A (morph.fun left) x + → isInKer A B ϕ x + → isInIm leftGr A left x Ker-right⊂Im-ϕ : (x : Group.type B) -- - → isInKer B rightGr (morph.fun right) x - → isInIm A B (morph.fun ϕ) x + → isInKer B rightGr right x + → isInIm A B ϕ x SES→Iso : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group ℓ} {B : Group ℓ'} (leftGr : Group ℓ'') (rightGr : Group ℓ''') → SES A B leftGr rightGr → Iso A B @@ -223,10 +223,10 @@ If ψ is an isomorphism and ϕ is surjective with ker ϕ ≡ {ψ (a , a) ∣ a diagonalIso1 : ∀ {ℓ ℓ' ℓ''} (A : Group ℓ) (B : Group ℓ') (C : Group ℓ'') (ψ : Iso (dirProd A A) B) (ϕ : morph B C) → isSurjective _ _ ϕ - → ((x : Group.type B) → isInKer B C (morph.fun ϕ) x + → ((x : Group.type B) → isInKer B C ϕ x → ∃[ y ∈ Group.type A ] x ≡ morph.fun (Iso.fun ψ) (y , y)) → ((x : Group.type B) → (∃[ y ∈ Group.type A ] x ≡ morph.fun (Iso.fun ψ) (y , y)) - → isInKer B C (morph.fun ϕ) x) + → isInKer B C ϕ x) → Iso A C diagonalIso1 A' B' C' ψ' ϕ' issurj ker→diag diag→ker = Iso''→Iso @@ -295,8 +295,8 @@ If ϕ is surjective with ker ϕ ≡ {(a , a) ∣ a ∈ A}, then C ≃ B diagonalIso : ∀ {ℓ ℓ' ℓ''} (A : Group ℓ) (B : Group ℓ') (C : Group ℓ'') → (ϕ : morph (dirProd A A) B) - → ((x : Group.type (dirProd A A)) → isInKer (dirProd A A) B (morph.fun ϕ) x → ∃[ y ∈ Group.type A ] x ≡ (y , y)) - → ((x : Group.type (dirProd A A)) → ∃[ y ∈ Group.type A ] x ≡ (y , y) → isInKer (dirProd A A) B (morph.fun ϕ) x) + → ((x : Group.type (dirProd A A)) → isInKer (dirProd A A) B ϕ x → ∃[ y ∈ Group.type A ] x ≡ (y , y)) + → ((x : Group.type (dirProd A A)) → ∃[ y ∈ Group.type A ] x ≡ (y , y) → isInKer (dirProd A A) B ϕ x) → isSurjective (dirProd A A) B ϕ → Iso C A → Iso C B @@ -320,7 +320,7 @@ diagonalIso A' B' C' ϕ' diagKer1 diagKer2 issurj I = morph.fun fstProj = λ a → a , (id A) morph.ismorph fstProj = λ g0 g1 i → comp A g0 g1 , lUnit A (id A) (~ i) - inj : (a : Group.type C') → isInKer C' B' (λ x → ϕ (morph.fun fstProj (C→A x))) a + inj : (a : Group.type C') → isInKer C' B' (compMorph (compMorph (Iso.fun I) fstProj) ϕ') a → a ≡ id C inj a inker = rec (Group.setStruc C' _ _) (λ {(b , id') → sym (Iso.leftInv I a) diff --git a/Cubical/HITs/SetTruncation/Properties.agda b/Cubical/HITs/SetTruncation/Properties.agda index 51c6173e8..6db517637 100644 --- a/Cubical/HITs/SetTruncation/Properties.agda +++ b/Cubical/HITs/SetTruncation/Properties.agda @@ -5,7 +5,7 @@ This file contains: - Properties of set truncations -} -{-# OPTIONS --cubical --safe #-} +{-# OPTIONS --cubical --no-import-sorts --safe #-} module Cubical.HITs.SetTruncation.Properties where open import Cubical.HITs.SetTruncation.Base diff --git a/Cubical/ZCohomology/Groups/Sn.agda b/Cubical/ZCohomology/Groups/Sn.agda index 245eeb952..71214b63f 100644 --- a/Cubical/ZCohomology/Groups/Sn.agda +++ b/Cubical/ZCohomology/Groups/Sn.agda @@ -384,7 +384,7 @@ H¹-S¹≅ℤ = helper : (F : S₊ 0 → Int) (f g : ∥ (Unit → Int) ∥₀) (id : morph.fun (MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) (f , g) ≡ ∣ F ∣₀) → isInKer (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout (λ _ → tt) (λ _ → tt))) - (morph.fun (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0)) + (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) ∣ F ∣₀ → ∃[ x ∈ Int ] ∣ F ∣₀ ≡ morph.fun (grIso.inv H⁰-S⁰≅ℤ×ℤ) (x , x) helper F = diff --git a/Cubical/ZCohomology/Groups/Wedge.agda b/Cubical/ZCohomology/Groups/Wedge.agda new file mode 100644 index 000000000..631d66ef5 --- /dev/null +++ b/Cubical/ZCohomology/Groups/Wedge.agda @@ -0,0 +1,98 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.Groups.Wedge where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.S1.S1 +open import Cubical.HITs.S1 +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.MayerVietorisUnreduced +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence +open import Cubical.Data.NatMinusTwo.Base +open import Cubical.Data.Empty +open import Cubical.Data.Sigma hiding (_×_) +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₋₁ ; ∣_∣ to ∣_∣₋₁) +open import Cubical.HITs.Nullification +open import Cubical.Data.Int renaming (_+_ to _+ℤ_; +-comm to +ℤ-comm ; +-assoc to +ℤ-assoc) +open import Cubical.Data.Nat +open import Cubical.Data.Prod +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; recElim to trRec ; rec to trRec2 ; elim3 to trElim3) +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Freudenthal +open import Cubical.HITs.SmashProduct.Base +open import Cubical.Data.Unit +open import Cubical.Data.Group.Base renaming (Iso to grIso ; compIso to compGrIso ; invIso to invGrIso ; idIso to idGrIso) +open import Cubical.Data.Group.GroupLibrary +open import Cubical.ZCohomology.coHomZero-map +open import Cubical.HITs.Wedge + +open import Cubical.Foundations.Equiv.HalfAdjoint + +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.Groups.Torus + +open import Cubical.ZCohomology.KcompPrelims + + +open import Cubical.HITs.Pushout +open import Cubical.Data.Sum.Base +open import Cubical.Data.HomotopyGroup +open import Cubical.Data.Bool hiding (_⊕_) + +private + variable + ℓ : Level + A B : Pointed ℓ + + +private + Δ : (n : ℕ) → morph (×coHomGr n (typ A) (typ B)) (coHomGr n Unit) + Δ {A = A} {B = B} = MV.Δ (typ A) (typ B) Unit (λ _ → pt A) (λ _ → pt B) + d : (n : ℕ) → morph (coHomGr n Unit) (coHomGr (suc n) (A ⋁ B)) + d {A = A} {B = B} = MV.d (typ A) (typ B) Unit (λ _ → pt A) (λ _ → pt B) + i : (n : ℕ) → morph (coHomGr n (A ⋁ B)) (×coHomGr n (typ A) (typ B)) + i {A = A} {B = B} = MV.i (typ A) (typ B) Unit (λ _ → pt A) (λ _ → pt B) + +abstract + anotherhelper : (x : coHom 0 Unit) + → isInIm (×coHomGr 0 (typ A) (typ B)) (coHomGr 0 Unit) (Δ {A = A} {B = B} 0) x + anotherhelper = + sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ f → ∣ (∣ (λ _ → f tt) ∣₀ , 0ₕ) , cong ∣_∣₀ (funExt (λ _ → cong ((f tt) +ₖ_) -0ₖ ∙ rUnitₖ (f tt))) ∣₋₁ + + athirdhelper : (x : coHom 0 Unit) + → isInKer (coHomGr 0 Unit) (coHomGr 1 _) (d {A = A} {B = B} 0) x + athirdhelper {A = A} {B = B} x = MV.Im-Δ⊂Ker-d (typ A) (typ B) Unit (λ _ → pt A) (λ _ → pt B) 0 x + (anotherhelper x) + + afourthhelper : (x : coHom 1 (A ⋁ B)) → isInIm (coHomGr 0 Unit) (coHomGr 1 _) (d {A = A} {B = B} 0) x + → x ≡ 0ₕ + afourthhelper x inim = pRec (setTruncIsSet _ _) (λ p → sym (snd p) ∙ athirdhelper (fst p)) inim + +H¹-⋁ : grIso (coHomGr 1 (Pushout (λ _ → snd A) (λ _ → snd B))) (dirProd (coHomGr 1 (typ A)) (coHomGr 1 (typ B))) +H¹-⋁ {A = A} {B = B} = + Iso''→Iso + (iso'' + (i 1) + {!!} -- (λ x ker → afourthhelper x (MV.Ker-i⊂Im-d (typ A) (typ B) Unit (λ _ → pt A) (λ _ → pt B) 0 x ker)) + λ x → MV.Ker-Δ⊂Im-i (typ A) (typ B) Unit (λ _ → pt A) (λ _ → pt B) 1 x + (isOfHLevelSuc 0 (isContrHⁿ-Unit 0) _ _)) + +H¹-⋁' : Iso'' (coHomGr 1 (Pushout {A = Unit} (λ _ → snd A) (λ _ → snd B))) (×coHomGr 1 (typ A) (typ B)) +Iso''.ϕ (H¹-⋁' {A = A} {B = B}) = MV.i-n+1 (typ A) (typ B) Unit (λ _ → pt A) (λ _ → pt B) 0 -- ? sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) λ _ _ → refl +Iso''.inj (H¹-⋁' {A = A} {B = B}) x ker = {!!} -- afourthhelper x (MV.Ker-i⊂Im-d (typ A) (typ B) Unit (λ _ → pt A) (λ _ → pt B) 0 x ker) -- afourthhelper x (MV.Ker-i⊂Im-d (typ A) (typ B) Unit (λ _ → pt A) (λ _ → pt B) 0 x ker) +Iso''.surj (H¹-⋁' {A = A} {B = B}) x = {!MV.Ker-Δ⊂Im-i (typ A) (typ B) Unit (λ _ → pt A) (λ _ → pt B) 1 x + (isOfHLevelSuc 0 (isContrHⁿ-Unit 0) _ _)!} diff --git a/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda b/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda index a4ad82d3d..1738fcfb6 100644 --- a/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda +++ b/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda @@ -175,7 +175,7 @@ Iso.rightInv testID a = {!!} Iso.leftInv testID = {!!} -maybe13 : Iso (Σ[ f ∈ coHom 1 (S₊∙ 1 ⋁ S₊∙ 1) ] isInIm (coHomGr 0 Unit) (coHomGr 1 _) (morph.fun (MV.d (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 0)) f) +maybe13 : Iso (Σ[ f ∈ coHom 1 (S₊∙ 1 ⋁ S₊∙ 1) ] isInIm (coHomGr 0 Unit) (coHomGr 1 _) (MV.d (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 0) f) ∥ S₊ 1 ∥₀ Iso.fun maybe13 (f , inim) = pRec helper (sigmaElim (λ _ → setTruncIsSet) (λ g id → {!!})) inim @@ -195,10 +195,6 @@ Iso.leftInv maybe13 (f , inim) = pRec {!!} {!!} inim -maybe : (a b : Int) (x : MV.D (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north)) → MV.d-pre (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 0 (λ _ → a) x ≡ MV.d-pre (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 0 (λ _ → b) x -maybe a b (inl x) = refl -maybe a b (inr x) = refl -maybe a b (push a₁ i) j = {!Kn→ΩKn+1 zero ?!} testID2 : Iso (coHom 1 ((S₊∙ 1) ⋁ S₊∙ 1)) (coHom 1 (S₊ 1 × S₊ 1)) Iso.fun testID2 = sRec setTruncIsSet λ f → ∣ (λ {(a , b) → f (inl a) +ₖ f (inl b)}) ∣₀ @@ -304,30 +300,36 @@ Iso.leftInv test1 = -} anotherhelper : (x : coHom 0 Unit) - → isInIm (×coHomGr 0 (S₊ 1) (S₊ 1)) (coHomGr 0 Unit) (morph.fun (MV.Δ (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 0)) x + → isInIm (×coHomGr 0 (S₊ 1) (S₊ 1)) (coHomGr 0 Unit) (MV.Δ (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 0) x anotherhelper = sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) λ f → ∣ (∣ (λ _ → f tt) ∣₀ , 0ₕ) , cong ∣_∣₀ (funExt (λ _ → cong ((f tt) +ₖ_) -0ₖ ∙ rUnitₖ (f tt))) ∣₋₁ athirdhelper : (x : coHom 0 Unit) - → isInKer (coHomGr 0 Unit) (coHomGr 1 _) (morph.fun (MV.d (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 0)) x + → isInKer (coHomGr 0 Unit) (coHomGr 1 _) (MV.d (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 0) x athirdhelper x = MV.Im-Δ⊂Ker-d (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 0 x (anotherhelper x) -afourthhelper : (x : coHom 1 (S₊∙ 1 ⋁ S₊∙ 1)) → isInIm (coHomGr 0 Unit) (coHomGr 1 _) (morph.fun (MV.d (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 0)) x +afourthhelper : (x : coHom 1 (S₊∙ 1 ⋁ S₊∙ 1)) → isInIm (coHomGr 0 Unit) (coHomGr 1 _) (MV.d (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 0) x → x ≡ 0ₕ afourthhelper x inim = pRec (setTruncIsSet _ _) (λ {(y , id) → sym id ∙ athirdhelper y}) inim -H¹-S¹∨S¹ : grIso (coHomGr 1 (S₊∙ 1 ⋁ S₊∙ 1)) (×coHomGr 1 (S₊ 1) (S₊ 1)) -H¹-S¹∨S¹ = - Iso''→Iso - (iso'' - (MV.i (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 1) - (λ x inker → afourthhelper x {!!}) - λ x → {!MV.Ker-Δ⊂Im-i (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 1 x ?!}) - +H¹-S¹∨S¹ : Iso'' (coHomGr 1 (Pushout (λ _ → north) (λ _ → north))) (×coHomGr 1 (S₊ 1) (S₊ 1)) +Iso''.ϕ H¹-S¹∨S¹ = MV.i (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 1 +Iso''.inj H¹-S¹∨S¹ x inker = {!afourthhelper x (MV.Ker-i⊂Im-d (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 0 x inker)!} +Iso''.surj H¹-S¹∨S¹ x = {!MV.Ker-Δ⊂Im-i (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 1 x ?!} +{- + (iso'' + {!!} + ? -- (λ x inker → afourthhelper x (MV.Ker-i⊂Im-d (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 0 x inker)) + λ x → MV.Ker-Δ⊂Im-i (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 1 x (helper x)) + where + helper : (x : Group.type (×coHomGr 1 (S₊ 1) (S₊ 1))) → isInKer (×coHomGr 1 (S₊ 1) (S₊ 1)) (coHomGr 1 Unit) + (MV.Δ (S₊ 1) (S₊ 1) Unit (λ _ → north) (λ _ → north) 1) x + helper = {!!} +-} diff --git a/Cubical/ZCohomology/MayerVietorisUnreduced.agda b/Cubical/ZCohomology/MayerVietorisUnreduced.agda index a842c32c7..b0b63ce22 100644 --- a/Cubical/ZCohomology/MayerVietorisUnreduced.agda +++ b/Cubical/ZCohomology/MayerVietorisUnreduced.agda @@ -64,22 +64,35 @@ prodElim2 isset f = prodElim (λ _ → isOfHLevelΠ 2 λ _ → isset _ _) λ c d → f a b c d module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) where - D : Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') - D = Pushout f g - - i : (n : ℕ) → morph (coHomGr n D) (×coHomGr n A B) - morph.fun (i zero) = - sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) - λ { δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀} - morph.fun (i (suc n)) = - sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) - λ { δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀} - morph.ismorph (i zero) = - sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) - λ f g → refl - morph.ismorph (i (suc n)) = - sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) - λ f g → refl + + i-0 : morph (coHomGr 0 (Pushout f g)) (×coHomGr 0 A B) + morph.fun i-0 = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) + λ δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀ + morph.ismorph i-0 = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ f g → refl + i-n+1 : (n : ℕ) → morph (coHomGr (suc n) (Pushout f g)) (×coHomGr (suc n) A B) + morph.fun (i-n+1 n) = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) + λ δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀ + morph.ismorph (i-n+1 n) = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ f g → refl + + + + i* : (n : ℕ) → coHom n (Pushout f g) → coHom n A × coHom n B + i* zero = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) + λ δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀ + i* (suc n) = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) + λ δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀ + + iIsHom : (n : ℕ) → isMorph (coHomGr n (Pushout f g)) (×coHomGr n A B) (i* n) + iIsHom zero = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ f g → refl + iIsHom (suc n) = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ f g → refl + + i : (n : ℕ) → morph (coHomGr n (Pushout f g)) (×coHomGr n A B) + morph.fun (i n) = i* n + morph.ismorph (i n) = iIsHom n {- i zero = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) @@ -134,7 +147,7 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : -} - d-pre : (n : ℕ) → (C → coHomK n) → D → coHomK (suc n) + d-pre : (n : ℕ) → (C → coHomK n) → (Pushout f g) → coHomK (suc n) d-pre n γ (inl x) = 0ₖ d-pre n γ (inr x) = 0ₖ d-pre zero γ (push a i) = Kn→ΩKn+1 zero (γ a) i @@ -169,13 +182,13 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : (cong (λ x → rUnitₖ x (~ j)) (Kn→ΩKn+1 n (h a)) i) - dHomHelper : (n : ℕ) (h l : C → coHomK n) (x : D) → d-pre n (λ x → h x +ₖ l x) x ≡ d-pre n h x +ₖ d-pre n l x + dHomHelper : (n : ℕ) (h l : C → coHomK n) (x : Pushout f g) → d-pre n (λ x → h x +ₖ l x) x ≡ d-pre n h x +ₖ d-pre n l x dHomHelper n h l (inl x) = sym (lUnitₖ 0ₖ) dHomHelper n h l (inr x) = sym (lUnitₖ 0ₖ) dHomHelper zero h l (push a i) j = dHomHelperPath zero h l a i j dHomHelper (suc n) h l (push a i) j = dHomHelperPath (suc n) h l a i j - d : (n : ℕ) → morph (coHomGr n C) (coHomGr (suc n) D) + d : (n : ℕ) → morph (coHomGr n C) (coHomGr (suc n) (Pushout f g)) morph.fun (d n) = sRec setTruncIsSet λ a → ∣ d-pre n a ∣₀ morph.ismorph (d zero) = sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ f g i → ∣ funExt (λ x → dHomHelper zero f g x) i ∣₀ @@ -193,138 +206,149 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : -- Long Exact Sequence - Im-d⊂Ker-i : (n : ℕ) (x : coHom (suc n) D) - → isInIm (coHomGr n C) (coHomGr (suc n) D) (morph.fun (d n)) x - → isInKer (coHomGr (suc n) D) (×coHomGr (suc n) A B) (morph.fun (i (suc n))) x + Im-d⊂Ker-i : (n : ℕ) (x : coHom (suc n) (Pushout f g)) + → isInIm (coHomGr n C) (coHomGr (suc n) (Pushout f g)) (d n) x + → isInKer (coHomGr (suc n) (Pushout f g)) (×coHomGr (suc n) A B) (i (suc n)) x Im-d⊂Ker-i n = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) λ a → pElim (λ _ → isOfHLevelPath' 1 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) (sigmaElim (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) λ δ b → (λ i → sElim (λ _ → isOfHLevelProd 2 setTruncIsSet setTruncIsSet) - (λ { δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀ }) (b (~ i)))) - - Ker-i⊂Im-d : (n : ℕ) (x : coHom (suc n) D) - → isInKer (coHomGr (suc n) D) (×coHomGr (suc n) A B) (morph.fun (i (suc n))) x - → isInIm (coHomGr n C) (coHomGr (suc n) D) (morph.fun (d n)) x - Ker-i⊂Im-d n = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) - λ a p → pRec {A = (λ x → a (inl x)) ≡ λ _ → 0ₖ} (isOfHLevelΠ 1 (λ _ → propTruncIsProp)) - (λ p1 → pRec propTruncIsProp λ p2 → ∣ ∣ (λ c → ΩKn+1→Kn (sym (cong (λ F → F (f c)) p1) - ∙∙ cong a (push c) - ∙∙ cong (λ F → F (g c)) p2)) ∣₀ - , cong ∣_∣₀ (funExt (λ δ → helper n a p1 p2 δ)) ∣₋₁) - (Iso.fun (PathIdTrunc₀Iso) (cong proj₁ p)) - (Iso.fun (PathIdTrunc₀Iso) (cong proj₂ p)) - where - pushFiller : (n : ℕ) (F : D → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) - (p1 : Path (_ → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) (λ a₁ → F (inl a₁)) (λ _ → ∣ north ∣)) - (p2 : Path (_ → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) (λ a₁ → F (inr a₁)) (λ _ → ∣ north ∣)) → - (c : C) → I → I → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n) - pushFiller n F p1 p2 c i j = - hcomp (λ k → λ { (i = i1) → F (push c j) - ; (j = i0) → p1 (~ i ∧ k) (f c) - ; (j = i1) → p2 (~ i ∧ k) (g c)}) - (F (push c j)) - helper : (n : ℕ) (F : D → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) - (p1 : Path (_ → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) (λ a₁ → F (inl a₁)) (λ _ → ∣ north ∣)) - (p2 : Path (_ → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) (λ a₁ → F (inr a₁)) (λ _ → ∣ north ∣)) → - (δ : D) → d-pre n (λ c → ΩKn+1→Kn ((λ i₁ → p1 (~ i₁) (f c)) - ∙∙ cong F (push c) - ∙∙ cong (λ F → F (g c)) p2)) δ - ≡ F δ - helper n F p1 p2 (inl x) = sym (cong (λ f → f x) p1) - helper n F p1 p2 (inr x) = sym (cong (λ f → f x) p2) - helper zero F p1 p2 (push a i) j = - hcomp (λ k → λ { (i = i0) → p1 (~ j) (f a) - ; (i = i1) → p2 (~ j) (g a) - ; (j = i0) → Iso.rightInv (Iso3-Kn-ΩKn+1 zero) ((λ i₁ → p1 (~ i₁) (f a)) - ∙∙ cong F (push a) - ∙∙ cong (λ F₁ → F₁ (g a)) p2) (~ k) i - ; (j = i1) → F (push a i)}) - (pushFiller zero F p1 p2 a j i) - helper (suc n) F p1 p2 (push a i) j = - hcomp (λ k → λ { (i = i0) → p1 (~ j) (f a) - ; (i = i1) → p2 (~ j) (g a) - ; (j = i0) → Iso.rightInv (Iso3-Kn-ΩKn+1 (suc n)) ((λ i₁ → p1 (~ i₁) (f a)) - ∙∙ cong F (push a) - ∙∙ cong (λ F₁ → F₁ (g a)) p2) (~ k) i - ; (j = i1) → F (push a i)}) - (pushFiller (suc n) F p1 p2 a j i) + (λ δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀ ) (b (~ i)))) + +--------------------------------- + Im-d⊂Ker-i1 : (x : coHom 1 (Pushout f g)) + → isInIm (coHomGr 0 C) (coHomGr 1 (Pushout f g)) (d 0) x + → isInKer (coHomGr 1 (Pushout f g)) (×coHomGr 1 A B) (i 1) x + Im-d⊂Ker-i1 = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ a → pElim (λ _ → isOfHLevelPath' 1 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + (sigmaElim (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ δ b → (λ i → sElim (λ _ → isOfHLevelProd 2 setTruncIsSet setTruncIsSet) + (λ δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀ ) (b (~ i)))) +--------------------------------- abstract - Im-i⊂Ker-Δ : (n : ℕ) (x : coHom n A × coHom n B) - → isInIm (coHomGr n D) (×coHomGr n A B) (morph.fun (i n)) x - → isInKer (×coHomGr n A B) (coHomGr n C) (morph.fun (Δ n)) x - Im-i⊂Ker-Δ n (Fa , Fb) = - sElim {B = λ Fa → (Fb : _) → isInIm (coHomGr n D) (×coHomGr n A B) (morph.fun (i n)) (Fa , Fb) - → isInKer (×coHomGr n A B) (coHomGr n C) (morph.fun (Δ n)) (Fa , Fb)} - (λ _ → isOfHLevelΠ 2 λ _ → (isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _)) - (λ Fa → sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) - λ Fb → pRec (setTruncIsSet _ _) - (sigmaElim (λ x → isOfHLevelSuc 1 (setTruncIsSet _ _)) - λ Fd p → helper n Fa Fb Fd p)) - Fa - Fb + Ker-i⊂Im-d : (n : ℕ) (x : coHom (suc n) (Pushout f g)) + → isInKer (coHomGr (suc n) (Pushout f g)) (×coHomGr (suc n) A B) (i (suc n)) x + → isInIm (coHomGr n C) (coHomGr (suc n) (Pushout f g)) (d n) x + Ker-i⊂Im-d n = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ a p → pRec {A = (λ x → a (inl x)) ≡ λ _ → 0ₖ} (isOfHLevelΠ 1 (λ _ → propTruncIsProp)) + (λ p1 → pRec propTruncIsProp λ p2 → ∣ ∣ (λ c → ΩKn+1→Kn (sym (cong (λ F → F (f c)) p1) + ∙∙ cong a (push c) + ∙∙ cong (λ F → F (g c)) p2)) ∣₀ + , cong ∣_∣₀ (funExt (λ δ → helper n a p1 p2 δ)) ∣₋₁) + (Iso.fun (PathIdTrunc₀Iso) (cong proj₁ p)) + (Iso.fun (PathIdTrunc₀Iso) (cong proj₂ p)) where - helper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) (Fd : D → coHomK n) - → (morph.fun (i n) ∣ Fd ∣₀ ≡ (∣ Fa ∣₀ , ∣ Fb ∣₀)) - → (morph.fun (Δ n)) (∣ Fa ∣₀ , ∣ Fb ∣₀) ≡ 0ₕ - helper zero Fa Fb Fd p = cong (morph.fun (Δ zero)) (sym p) - ∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₀ +ₕ -ₕ ∣ (λ x → Fd (push x (~ i))) ∣₀) - ∙ rCancelₕ ∣ (λ x → Fd (inl (f x))) ∣₀ - helper (suc n) Fa Fb Fd p = cong (morph.fun (Δ (suc n))) (sym p) - ∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₀ +ₕ -ₕ ∣ (λ x → Fd (push x (~ i))) ∣₀) - ∙ rCancelₕ ∣ (λ x → Fd (inl (f x))) ∣₀ - - - - Ker-Δ⊂Im-i : (n : ℕ) (a : coHom n A × coHom n B) - → isInKer (×coHomGr n A B) (coHomGr n C) (morph.fun (Δ n)) a - → isInIm (coHomGr n D) (×coHomGr n A B) (morph.fun (i n)) a - Ker-Δ⊂Im-i n (Fa , Fb) = - sElim {B = λ Fa → (Fb : _) → isInKer (×coHomGr n A B) (coHomGr n C) (morph.fun (Δ n)) (Fa , Fb) - → isInIm (coHomGr n D) (×coHomGr n A B) (morph.fun (i n)) (Fa , Fb)} - (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) - (λ Fa → sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) - λ Fb p → pRec propTruncIsProp - (λ q → ∣ ∣ helpFun n Fa Fb (funExt⁻ q) ∣₀ - , anotherHelper n Fa Fb q ∣₋₁) - (helper n Fa Fb p)) - Fa - Fb + pushFiller : (n : ℕ) (F : (Pushout f g) → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) + (p1 : Path (_ → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) (λ a₁ → F (inl a₁)) (λ _ → ∣ north ∣)) + (p2 : Path (_ → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) (λ a₁ → F (inr a₁)) (λ _ → ∣ north ∣)) → + (c : C) → I → I → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n) + pushFiller n F p1 p2 c i j = + hcomp (λ k → λ { (i = i1) → F (push c j) + ; (j = i0) → p1 (~ i ∧ k) (f c) + ; (j = i1) → p2 (~ i ∧ k) (g c)}) + (F (push c j)) + helper : (n : ℕ) (F : (Pushout f g) → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) + (p1 : Path (_ → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) (λ a₁ → F (inl a₁)) (λ _ → ∣ north ∣)) + (p2 : Path (_ → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) (λ a₁ → F (inr a₁)) (λ _ → ∣ north ∣)) → + (δ : (Pushout f g)) → d-pre n (λ c → ΩKn+1→Kn ((λ i₁ → p1 (~ i₁) (f c)) + ∙∙ cong F (push c) + ∙∙ cong (λ F → F (g c)) p2)) δ + ≡ F δ + helper n F p1 p2 (inl x) = sym (cong (λ f → f x) p1) + helper n F p1 p2 (inr x) = sym (cong (λ f → f x) p2) + helper zero F p1 p2 (push a i) j = + hcomp (λ k → λ { (i = i0) → p1 (~ j) (f a) + ; (i = i1) → p2 (~ j) (g a) + ; (j = i0) → Iso.rightInv (Iso3-Kn-ΩKn+1 zero) ((λ i₁ → p1 (~ i₁) (f a)) + ∙∙ cong F (push a) + ∙∙ cong (λ F₁ → F₁ (g a)) p2) (~ k) i + ; (j = i1) → F (push a i)}) + (pushFiller zero F p1 p2 a j i) + helper (suc n) F p1 p2 (push a i) j = + hcomp (λ k → λ { (i = i0) → p1 (~ j) (f a) + ; (i = i1) → p2 (~ j) (g a) + ; (j = i0) → Iso.rightInv (Iso3-Kn-ΩKn+1 (suc n)) ((λ i₁ → p1 (~ i₁) (f a)) + ∙∙ cong F (push a) + ∙∙ cong (λ F₁ → F₁ (g a)) p2) (~ k) i + ; (j = i1) → F (push a i)}) + (pushFiller (suc n) F p1 p2 a j i) + + abstract + Im-i⊂Ker-Δ : (n : ℕ) (x : coHom n A × coHom n B) + → isInIm (coHomGr n (Pushout f g)) (×coHomGr n A B) (i n) x + → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) x + Im-i⊂Ker-Δ n (Fa , Fb) = + sElim {B = λ Fa → (Fb : _) → isInIm (coHomGr n (Pushout f g)) (×coHomGr n A B) (i n) (Fa , Fb) + → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) (Fa , Fb)} + (λ _ → isOfHLevelΠ 2 λ _ → (isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _)) + (λ Fa → sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ Fb → pRec (setTruncIsSet _ _) + (sigmaElim (λ x → isOfHLevelSuc 1 (setTruncIsSet _ _)) + λ Fd p → helper n Fa Fb Fd p)) + Fa + Fb + where + helper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) (Fd : (Pushout f g) → coHomK n) + → (morph.fun (i n) ∣ Fd ∣₀ ≡ (∣ Fa ∣₀ , ∣ Fb ∣₀)) + → (morph.fun (Δ n)) (∣ Fa ∣₀ , ∣ Fb ∣₀) ≡ 0ₕ + helper zero Fa Fb Fd p = cong (morph.fun (Δ zero)) (sym p) + ∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₀ +ₕ -ₕ ∣ (λ x → Fd (push x (~ i))) ∣₀) + ∙ rCancelₕ ∣ (λ x → Fd (inl (f x))) ∣₀ + helper (suc n) Fa Fb Fd p = cong (morph.fun (Δ (suc n))) (sym p) + ∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₀ +ₕ -ₕ ∣ (λ x → Fd (push x (~ i))) ∣₀) + ∙ rCancelₕ ∣ (λ x → Fd (inl (f x))) ∣₀ + + + Ker-Δ⊂Im-i : (n : ℕ) (a : coHom n A × coHom n B) + → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) a + → isInIm (coHomGr n (Pushout f g)) (×coHomGr n A B) (i n) a + Ker-Δ⊂Im-i n (Fa , Fb) = + sElim {B = λ Fa → (Fb : _) → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) (Fa , Fb) + → isInIm (coHomGr n (Pushout f g)) (×coHomGr n A B) (i n) (Fa , Fb)} + (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) + (λ Fa → sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ Fb p → pRec propTruncIsProp + (λ q → ∣ ∣ helpFun n Fa Fb (funExt⁻ q) ∣₀ + , anotherHelper n Fa Fb q ∣₋₁) + (helper n Fa Fb p)) + Fa + Fb - where - helper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) - → (morph.fun (Δ n)) (∣ Fa ∣₀ , ∣ Fb ∣₀) ≡ 0ₕ - → ∥ (Path (_ → _) (λ c → Fa (f c)) (λ c → Fb (g c))) ∥₋₁ - helper zero Fa Fb p = Iso.fun (PathIdTrunc₀Iso) - (sym (rUnitₕ (coHomFun zero f ∣ Fa ∣₀)) - ∙ (λ i → coHomFun zero f ∣ Fa ∣₀ +ₕ (lCancelₕ (coHomFun zero g ∣ Fb ∣₀) (~ i))) - ∙ sym (assocₕ (coHomFun zero f ∣ Fa ∣₀) (-ₕ (coHomFun zero g ∣ Fb ∣₀)) (coHomFun zero g ∣ Fb ∣₀)) - ∙ cong (λ x → x +ₕ (coHomFun zero g ∣ Fb ∣₀)) p - ∙ lUnitₕ (coHomFun zero g ∣ Fb ∣₀)) - helper (suc n) Fa Fb p = Iso.fun (PathIdTrunc₀Iso) - (sym (rUnitₕ (coHomFun (suc n) f ∣ Fa ∣₀)) - ∙ (λ i → coHomFun (suc n) f ∣ Fa ∣₀ +ₕ (lCancelₕ (coHomFun (suc n) g ∣ Fb ∣₀) (~ i))) - ∙ sym (assocₕ (coHomFun (suc n) f ∣ Fa ∣₀) (-ₕ (coHomFun (suc n) g ∣ Fb ∣₀)) (coHomFun (suc n) g ∣ Fb ∣₀)) - ∙ cong (λ x → x +ₕ (coHomFun (suc n) g ∣ Fb ∣₀)) p - ∙ lUnitₕ (coHomFun (suc n) g ∣ Fb ∣₀)) - - helpFun : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) - → ((c : C) → Fa (f c) ≡ Fb (g c)) - → D → coHomK n - helpFun n Fa Fb p (inl x) = Fa x - helpFun n Fa Fb p (inr x) = Fb x - helpFun n Fa Fb p (push a i) = p a i - - anotherHelper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) - → (q : Path (C → coHomK n) (λ c → Fa (f c)) (λ c → Fb (g c))) - → morph.fun (i n) ∣ helpFun n Fa Fb (λ x i₁ → q i₁ x) ∣₀ ≡ (∣ Fa ∣₀ , ∣ Fb ∣₀) - anotherHelper zero Fa Fb q = refl - anotherHelper (suc n) Fa Fb q = refl + where + helper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) + → (morph.fun (Δ n)) (∣ Fa ∣₀ , ∣ Fb ∣₀) ≡ 0ₕ + → ∥ (Path (_ → _) (λ c → Fa (f c)) (λ c → Fb (g c))) ∥₋₁ + helper zero Fa Fb p = Iso.fun (PathIdTrunc₀Iso) + (sym (rUnitₕ (coHomFun zero f ∣ Fa ∣₀)) + ∙ (λ i → coHomFun zero f ∣ Fa ∣₀ +ₕ (lCancelₕ (coHomFun zero g ∣ Fb ∣₀) (~ i))) + ∙ sym (assocₕ (coHomFun zero f ∣ Fa ∣₀) (-ₕ (coHomFun zero g ∣ Fb ∣₀)) (coHomFun zero g ∣ Fb ∣₀)) + ∙ cong (λ x → x +ₕ (coHomFun zero g ∣ Fb ∣₀)) p + ∙ lUnitₕ (coHomFun zero g ∣ Fb ∣₀)) + helper (suc n) Fa Fb p = Iso.fun (PathIdTrunc₀Iso) + (sym (rUnitₕ (coHomFun (suc n) f ∣ Fa ∣₀)) + ∙ (λ i → coHomFun (suc n) f ∣ Fa ∣₀ +ₕ (lCancelₕ (coHomFun (suc n) g ∣ Fb ∣₀) (~ i))) + ∙ sym (assocₕ (coHomFun (suc n) f ∣ Fa ∣₀) (-ₕ (coHomFun (suc n) g ∣ Fb ∣₀)) (coHomFun (suc n) g ∣ Fb ∣₀)) + ∙ cong (λ x → x +ₕ (coHomFun (suc n) g ∣ Fb ∣₀)) p + ∙ lUnitₕ (coHomFun (suc n) g ∣ Fb ∣₀)) + + helpFun : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) + → ((c : C) → Fa (f c) ≡ Fb (g c)) + → (Pushout f g) → coHomK n + helpFun n Fa Fb p (inl x) = Fa x + helpFun n Fa Fb p (inr x) = Fb x + helpFun n Fa Fb p (push a i) = p a i + + anotherHelper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) + → (q : Path (C → coHomK n) (λ c → Fa (f c)) (λ c → Fb (g c))) + → morph.fun (i n) ∣ helpFun n Fa Fb (λ x i₁ → q i₁ x) ∣₀ ≡ (∣ Fa ∣₀ , ∣ Fb ∣₀) + anotherHelper zero Fa Fb q = refl + anotherHelper (suc n) Fa Fb q = refl Ker-d⊂Im-Δ : (n : ℕ) (a : coHom n C) - → isInKer (coHomGr n C) (coHomGr (suc n) D) (morph.fun (d n)) a - → isInIm (×coHomGr n A B) (coHomGr n C) (morph.fun (Δ n)) a + → isInKer (coHomGr n C) (coHomGr (suc n) (Pushout f g)) (d n) a + → isInIm (×coHomGr n A B) (coHomGr n C) (Δ n) a Ker-d⊂Im-Δ n = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) λ Fc p → pRec propTruncIsProp (λ p → ∣ (∣ (λ a → ΩKn+1→Kn (cong (λ f → f (inl a)) p)) ∣₀ , @@ -357,8 +381,8 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : ∙ Iso.leftInv (Iso3-Kn-ΩKn+1 (suc n)) (Fc c) Im-Δ⊂Ker-d : (n : ℕ) (a : coHom n C) - → isInIm (×coHomGr n A B) (coHomGr n C) (morph.fun (Δ n)) a - → isInKer (coHomGr n C) (coHomGr (suc n) D) (morph.fun (d n)) a + → isInIm (×coHomGr n A B) (coHomGr n C) (Δ n) a + → isInKer (coHomGr n C) (coHomGr (suc n) (Pushout f g)) (d n) a Im-Δ⊂Ker-d n = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ Fc → pRec (isOfHLevelPath' 1 setTruncIsSet _ _) @@ -369,14 +393,14 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : where - d-preLeftId : (n : ℕ) (Fa : A → coHomK n)(d : D) + d-preLeftId : (n : ℕ) (Fa : A → coHomK n)(d : (Pushout f g)) → d-pre n (Fa ∘ f) d ≡ 0ₖ d-preLeftId n Fa (inl x) = Kn→ΩKn+1 n (Fa x) d-preLeftId n Fa (inr x) = refl d-preLeftId zero Fa (push a i) j = Kn→ΩKn+1 zero (Fa (f a)) (j ∨ i) d-preLeftId (suc n) Fa (push a i) j = Kn→ΩKn+1 (suc n) (Fa (f a)) (j ∨ i) - d-preRightId : (n : ℕ) (Fb : B → coHomK n) (d : D) + d-preRightId : (n : ℕ) (Fb : B → coHomK n) (d : (Pushout f g)) → d-pre n (Fb ∘ g) d ≡ 0ₖ d-preRightId n Fb (inl x) = refl d-preRightId n Fb (inr x) = sym (Kn→ΩKn+1 n (Fb x)) @@ -387,13 +411,13 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : → morph.fun (d n) (morph.fun (Δ n) (∣ Fa ∣₀ , ∣ Fb ∣₀)) ≡ 0ₕ dΔ-Id zero Fa Fb = morph.ismorph (d zero) ∣ (λ x → Fa (f x)) ∣₀ (-ₕ ∣ (λ x → Fb (g x)) ∣₀) - ∙ (λ i → morph.fun (d zero) ∣ (λ x → Fa (f x)) ∣₀ +ₕ morphMinus (coHomGr zero C) (coHomGr 1 D) (d zero) ∣ (λ x → Fb (g x)) ∣₀ i) + ∙ (λ i → morph.fun (d zero) ∣ (λ x → Fa (f x)) ∣₀ +ₕ morphMinus (coHomGr zero C) (coHomGr 1 (Pushout f g)) (d zero) ∣ (λ x → Fb (g x)) ∣₀ i) ∙ (λ i → morph.fun (d zero) ∣ (λ x → Fa (f x)) ∣₀ +ₕ -ₕ (morph.fun (d zero) ∣ (λ x → Fb (g x)) ∣₀)) ∙ (λ i → ∣ funExt (d-preLeftId zero Fa) i ∣₀ +ₕ -ₕ ∣ funExt (d-preRightId zero Fb) i ∣₀) ∙ rCancelₕ 0ₕ dΔ-Id (suc n) Fa Fb = morph.ismorph (d (suc n)) ∣ (λ x → Fa (f x)) ∣₀ (-ₕ ∣ (λ x → Fb (g x)) ∣₀) - ∙ (λ i → morph.fun (d (suc n)) ∣ (λ x → Fa (f x)) ∣₀ +ₕ morphMinus (coHomGr (suc n) C) (coHomGr (2 + n) D) (d (suc n)) ∣ (λ x → Fb (g x)) ∣₀ i) + ∙ (λ i → morph.fun (d (suc n)) ∣ (λ x → Fa (f x)) ∣₀ +ₕ morphMinus (coHomGr (suc n) C) (coHomGr (2 + n) (Pushout f g)) (d (suc n)) ∣ (λ x → Fb (g x)) ∣₀ i) ∙ (λ i → morph.fun (d (suc n)) ∣ (λ x → Fa (f x)) ∣₀ +ₕ -ₕ (morph.fun (d (suc n)) ∣ (λ x → Fb (g x)) ∣₀)) ∙ (λ i → ∣ funExt (d-preLeftId (suc n) Fa) i ∣₀ +ₕ -ₕ ∣ funExt (d-preRightId (suc n) Fb) i ∣₀) ∙ rCancelₕ 0ₕ diff --git a/Cubical/ZCohomology/MayerVietorisUnreduced2.agda b/Cubical/ZCohomology/MayerVietorisUnreduced2.agda new file mode 100644 index 000000000..876a0c22f --- /dev/null +++ b/Cubical/ZCohomology/MayerVietorisUnreduced2.agda @@ -0,0 +1,426 @@ +{-# OPTIONS --cubical --safe #-} +module Cubical.ZCohomology.MayerVietorisUnreduced2 where + +open import Cubical.ZCohomology.Base +open import Cubical.HITs.S1 +open import Cubical.ZCohomology.Properties +open import Cubical.HITs.Sn +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence +open import Cubical.Data.NatMinusTwo.Base +open import Cubical.Data.Empty +open import Cubical.Data.Sigma hiding (_×_) +open import Cubical.HITs.Susp +open import Cubical.HITs.Wedge +open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₋₁ ; ∣_∣ to ∣_∣₋₁) +open import Cubical.HITs.Nullification +open import Cubical.Data.Int hiding (_+_) +open import Cubical.Data.Nat +open import Cubical.Data.Prod +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Freudenthal +open import Cubical.HITs.SmashProduct.Base +open import Cubical.Data.Group.Base + + +open import Cubical.ZCohomology.KcompPrelims + + +open import Cubical.HITs.Pushout +open import Cubical.Data.Sum.Base +open import Cubical.Data.HomotopyGroup +open import Cubical.Data.Bool hiding (_⊕_) + +private + variable + ℓ : Level + A B C : Type ℓ + f : C → A + g : C → B + + +coHomFun : (n : ℕ) (f : A → B) → coHom n B → coHom n A +coHomFun n f = sRec setTruncIsSet λ β → ∣ β ∘ f ∣₀ + +prodElim2 : ∀ {ℓ ℓ' ℓ'' ℓ''' ℓ''''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} + {E : (∥ A ∥₀ × ∥ B ∥₀) → (∥ C ∥₀ × ∥ D ∥₀) → Type ℓ''''} + → ((x : ∥ A ∥₀ × ∥ B ∥₀) (y : ∥ C ∥₀ × ∥ D ∥₀) → isOfHLevel 2 (E x y)) + → ((a : A) (b : B) (c : C) (d : D) → E (∣ a ∣₀ , ∣ b ∣₀) (∣ c ∣₀ , ∣ d ∣₀)) + → ((x : ∥ A ∥₀ × ∥ B ∥₀) (y : ∥ C ∥₀ × ∥ D ∥₀) → (E x y)) +prodElim2 isset f = prodElim (λ _ → isOfHLevelΠ 2 λ _ → isset _ _) + λ a b → prodElim (λ _ → isset _ _) + λ c d → f a b c d + + + + +i-0 : morph (coHomGr 0 (Pushout f g)) (×coHomGr 0 A B) +morph.fun i-0 = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) + λ δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀ +morph.ismorph i-0 = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ f g → refl +i-n+1 : (n : ℕ) → morph (coHomGr (suc n) (Pushout f g)) (×coHomGr (suc n) A B) +morph.fun (i-n+1 n) = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) + λ δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀ +morph.ismorph (i-n+1 n) = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ f g → refl + + + +i* : (n : ℕ) → coHom n (Pushout f g) → coHom n A × coHom n B +i* zero = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) + λ δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀ +i* (suc n) = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) + λ δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀ + +iIsHom : (n : ℕ) → isMorph (coHomGr n (Pushout f g)) (×coHomGr n A B) (i* n) +iIsHom zero = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ f g → refl +iIsHom (suc n) = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ f g → refl + +i : (n : ℕ) → morph (coHomGr n (Pushout f g)) (×coHomGr n A B) +i n = mph (i* n) (iIsHom n) + +{- +i zero = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) + λ { δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀} +i (suc n) = sRec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) + λ { δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀} +-} +{- + +iIsHom : (n : ℕ) → isMorph (coHomGr n D) (×coHomGr n A B) (i n) +iIsHom zero = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ f g → refl +iIsHom (suc n) = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) + λ f g → refl + +-} +Δ : (n : ℕ) → morph (×coHomGr n A B) (coHomGr n C) +morph.fun (Δ {f = f} {g = g} n) (α , β) = coHomFun n f α +ₕ -ₕ (coHomFun n g β) +morph.ismorph (Δ {f = f} {g = g} zero) = + prodElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _ ) + λ f' x1 g' x2 → (λ i → ∣ (λ x → (f' (f x) +ₖ g' (f x)) +ₖ -distrₖ (x1 (g x)) (x2 (g x)) i) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ g' (f x)) (-ₖ x1 (g x)) (-ₖ x2 (g x)) (~ i)) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x)) (g' (f x)) (-ₖ x1 (g x)) i +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → (f' (f x) +ₖ commₖ (g' (f x)) (-ₖ x1 (g x)) i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x)) (-ₖ x1 (g x)) (g' (f x)) (~ i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ (-ₖ x1 (g x))) (g' (f x)) (-ₖ x2 (g x)) i) ∣₀) +morph.ismorph (Δ {f = f} {g = g} (suc n)) = + prodElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _ ) + λ f' x1 g' x2 → (λ i → ∣ (λ x → (f' (f x) +ₖ g' (f x)) +ₖ -distrₖ (x1 (g x)) (x2 (g x)) i) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ g' (f x)) (-ₖ x1 (g x)) (-ₖ x2 (g x)) (~ i)) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x)) (g' (f x)) (-ₖ x1 (g x)) i +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → (f' (f x) +ₖ commₖ (g' (f x)) (-ₖ x1 (g x)) i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x)) (-ₖ x1 (g x)) (g' (f x)) (~ i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ (-ₖ x1 (g x))) (g' (f x)) (-ₖ x2 (g x)) i) ∣₀) + +{- +ΔIsHom : (n : ℕ) → isMorph (×coHomGr n A B) (coHomGr n C) (Δ n) +ΔIsHom zero = prodElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _ ) + λ f' x1 g' x2 → (λ i → ∣ (λ x → (f' (f x) +ₖ g' (f x)) +ₖ -distrₖ (x1 (g x)) (x2 (g x)) i) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ g' (f x)) (-ₖ x1 (g x)) (-ₖ x2 (g x)) (~ i)) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x)) (g' (f x)) (-ₖ x1 (g x)) i +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → (f' (f x) +ₖ commₖ (g' (f x)) (-ₖ x1 (g x)) i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x)) (-ₖ x1 (g x)) (g' (f x)) (~ i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ (-ₖ x1 (g x))) (g' (f x)) (-ₖ x2 (g x)) i) ∣₀) +ΔIsHom (suc n) = prodElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _ ) + λ f' x1 g' x2 → (λ i → ∣ (λ x → (f' (f x) +ₖ g' (f x)) +ₖ -distrₖ (x1 (g x)) (x2 (g x)) i) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ g' (f x)) (-ₖ x1 (g x)) (-ₖ x2 (g x)) (~ i)) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x)) (g' (f x)) (-ₖ x1 (g x)) i +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → (f' (f x) +ₖ commₖ (g' (f x)) (-ₖ x1 (g x)) i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x)) (-ₖ x1 (g x)) (g' (f x)) (~ i) +ₖ (-ₖ x2 (g x))) ∣₀) ∙ + (λ i → ∣ (λ x → assocₖ (f' (f x) +ₖ (-ₖ x1 (g x))) (g' (f x)) (-ₖ x2 (g x)) i) ∣₀) + +-} + +d-pre : (n : ℕ) → (C → coHomK n) → (Pushout f g) → coHomK (suc n) +d-pre n γ (inl x) = 0ₖ +d-pre n γ (inr x) = 0ₖ +d-pre zero γ (push a i) = Kn→ΩKn+1 zero (γ a) i +d-pre (suc n) γ (push a i) = Kn→ΩKn+1 (suc n) (γ a) i + + + + + + + + + +-- dHomHelperPath : (n : ℕ) (h l : C → coHomK n) (a : C) → I → I → coHomK (suc n) +-- dHomHelperPath n h l a i j = +-- hcomp (λ k → λ { (i = i0) → lUnitₖ 0ₖ (~ j) +-- ; (i = i1) → lUnitₖ 0ₖ (~ j) +-- ; (j = i0) → +ₖ→∙ n (h a) (l a) (~ k) i +-- ; (j = i1) → cong₂Funct (λ x y → x +ₖ y) (Kn→ΩKn+1 n (h a)) (Kn→ΩKn+1 n (l a)) (~ k) i}) +-- (bottom i j) +-- where +-- bottom : I → I → coHomK (suc n) +-- bottom i j = hcomp (λ k → λ { (i = i0) → lUnitₖ 0ₖ (~ j) +-- ; (i = i1) → cong (λ x → lUnitₖ x (~ j)) (Kn→ΩKn+1 n (l a)) k}) +-- (anotherbottom i j) +-- where +-- anotherbottom : I → I → coHomK (suc n) +-- anotherbottom i j = hcomp (λ k → λ { (i = i0) → rUnitlUnit0 k (~ j) +-- ; (i = i1) → rUnitlUnit0 k (~ j) +-- ; (j = i0) → Kn→ΩKn+1 n (h a) i +-- ; (j = i1) → cong (λ x → x +ₖ 0ₖ) (Kn→ΩKn+1 n (h a)) i}) +-- (cong (λ x → rUnitₖ x (~ j)) (Kn→ΩKn+1 n (h a)) i) + + +-- dHomHelper : (n : ℕ) (h l : C → coHomK n) (x : Pushout f g) → d-pre n (λ x → h x +ₖ l x) x ≡ d-pre n h x +ₖ d-pre n l x +-- dHomHelper n h l (inl x) = sym (lUnitₖ 0ₖ) +-- dHomHelper n h l (inr x) = sym (lUnitₖ 0ₖ) +-- dHomHelper zero h l (push a i) j = dHomHelperPath zero h l a i j +-- dHomHelper (suc n) h l (push a i) j = dHomHelperPath (suc n) h l a i j + +-- d : (n : ℕ) → morph (coHomGr n C) (coHomGr (suc n) (Pushout f g)) +-- morph.fun (d n) = sRec setTruncIsSet λ a → ∣ d-pre n a ∣₀ +-- morph.ismorph (d zero) = sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- λ f g i → ∣ funExt (λ x → dHomHelper zero f g x) i ∣₀ +-- morph.ismorph (d (suc n)) = sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- λ f g i → ∣ funExt (λ x → dHomHelper (suc n) f g x) i ∣₀ + +-- {- +-- dIsHom : (n : ℕ) → isMorph (coHomGr n C) (coHomGr (suc n) D) (d n) +-- dIsHom zero = sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- λ f g i → ∣ funExt (λ x → dHomHelper zero f g x) i ∣₀ +-- dIsHom (suc n) = sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- λ f g i → ∣ funExt (λ x → dHomHelper (suc n) f g x) i ∣₀ +-- -} + + +-- -- Long Exact Sequence + +-- Im-d⊂Ker-i : (n : ℕ) (x : coHom (suc n) (Pushout f g)) +-- → isInIm (coHomGr n C) (coHomGr (suc n) (Pushout f g)) (d n) x +-- → isInKer (coHomGr (suc n) (Pushout f g)) (×coHomGr (suc n) A B) (i (suc n)) x +-- Im-d⊂Ker-i n = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- λ a → pElim (λ _ → isOfHLevelPath' 1 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- (sigmaElim (λ _ → isOfHLevelPath 2 (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) _ _) +-- λ δ b → (λ i → sElim (λ _ → isOfHLevelProd 2 setTruncIsSet setTruncIsSet) +-- (λ δ → ∣ (λ x → δ (inl x)) ∣₀ , ∣ (λ x → δ (inr x)) ∣₀ ) (b (~ i)))) + +-- Ker-i⊂Im-d : (n : ℕ) (x : coHom (suc n) (Pushout f g)) +-- → isInKer (coHomGr (suc n) (Pushout f g)) (×coHomGr (suc n) A B) (i (suc n)) x +-- → isInIm (coHomGr n C) (coHomGr (suc n) (Pushout f g)) (d n) x +-- Ker-i⊂Im-d n = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- λ a p → pRec {A = (λ x → a (inl x)) ≡ λ _ → 0ₖ} (isOfHLevelΠ 1 (λ _ → propTruncIsProp)) +-- (λ p1 → pRec propTruncIsProp λ p2 → ∣ ∣ (λ c → ΩKn+1→Kn (sym (cong (λ F → F (f c)) p1) +-- ∙∙ cong a (push c) +-- ∙∙ cong (λ F → F (g c)) p2)) ∣₀ +-- , cong ∣_∣₀ (funExt (λ δ → helper n a p1 p2 δ)) ∣₋₁) +-- (Iso.fun (PathIdTrunc₀Iso) (cong proj₁ p)) +-- (Iso.fun (PathIdTrunc₀Iso) (cong proj₂ p)) +-- where +-- pushFiller : (n : ℕ) (F : (Pushout f g) → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) +-- (p1 : Path (_ → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) (λ a₁ → F (inl a₁)) (λ _ → ∣ north ∣)) +-- (p2 : Path (_ → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) (λ a₁ → F (inr a₁)) (λ _ → ∣ north ∣)) → +-- (c : C) → I → I → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n) +-- pushFiller n F p1 p2 c i j = +-- hcomp (λ k → λ { (i = i1) → F (push c j) +-- ; (j = i0) → p1 (~ i ∧ k) (f c) +-- ; (j = i1) → p2 (~ i ∧ k) (g c)}) +-- (F (push c j)) +-- helper : (n : ℕ) (F : (Pushout f g) → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) +-- (p1 : Path (_ → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) (λ a₁ → F (inl a₁)) (λ _ → ∣ north ∣)) +-- (p2 : Path (_ → ∥ S₊ (suc n) ∥ ℕ→ℕ₋₂ (suc n)) (λ a₁ → F (inr a₁)) (λ _ → ∣ north ∣)) → +-- (δ : (Pushout f g)) → d-pre n (λ c → ΩKn+1→Kn ((λ i₁ → p1 (~ i₁) (f c)) +-- ∙∙ cong F (push c) +-- ∙∙ cong (λ F → F (g c)) p2)) δ +-- ≡ F δ +-- helper n F p1 p2 (inl x) = sym (cong (λ f → f x) p1) +-- helper n F p1 p2 (inr x) = sym (cong (λ f → f x) p2) +-- helper zero F p1 p2 (push a i) j = +-- hcomp (λ k → λ { (i = i0) → p1 (~ j) (f a) +-- ; (i = i1) → p2 (~ j) (g a) +-- ; (j = i0) → Iso.rightInv (Iso3-Kn-ΩKn+1 zero) ((λ i₁ → p1 (~ i₁) (f a)) +-- ∙∙ cong F (push a) +-- ∙∙ cong (λ F₁ → F₁ (g a)) p2) (~ k) i +-- ; (j = i1) → F (push a i)}) +-- (pushFiller zero F p1 p2 a j i) +-- helper (suc n) F p1 p2 (push a i) j = +-- hcomp (λ k → λ { (i = i0) → p1 (~ j) (f a) +-- ; (i = i1) → p2 (~ j) (g a) +-- ; (j = i0) → Iso.rightInv (Iso3-Kn-ΩKn+1 (suc n)) ((λ i₁ → p1 (~ i₁) (f a)) +-- ∙∙ cong F (push a) +-- ∙∙ cong (λ F₁ → F₁ (g a)) p2) (~ k) i +-- ; (j = i1) → F (push a i)}) +-- (pushFiller (suc n) F p1 p2 a j i) + +-- abstract +-- Im-i⊂Ker-Δ : (n : ℕ) (x : coHom n A × coHom n B) +-- → isInIm (coHomGr n (Pushout f g)) (×coHomGr n A B) (i n) x +-- → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) x +-- Im-i⊂Ker-Δ n (Fa , Fb) = +-- sElim {B = λ Fa → (Fb : _) → isInIm (coHomGr n (Pushout f g)) (×coHomGr n A B) (i n) (Fa , Fb) +-- → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) (Fa , Fb)} +-- (λ _ → isOfHLevelΠ 2 λ _ → (isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _)) +-- (λ Fa → sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- λ Fb → pRec (setTruncIsSet _ _) +-- (sigmaElim (λ x → isOfHLevelSuc 1 (setTruncIsSet _ _)) +-- λ Fd p → helper n Fa Fb Fd p)) +-- Fa +-- Fb +-- where +-- helper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) (Fd : (Pushout f g) → coHomK n) +-- → (morph.fun (i n) ∣ Fd ∣₀ ≡ (∣ Fa ∣₀ , ∣ Fb ∣₀)) +-- → (morph.fun (Δ n)) (∣ Fa ∣₀ , ∣ Fb ∣₀) ≡ 0ₕ +-- helper zero Fa Fb Fd p = cong (morph.fun (Δ zero)) (sym p) +-- ∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₀ +ₕ -ₕ ∣ (λ x → Fd (push x (~ i))) ∣₀) +-- ∙ rCancelₕ ∣ (λ x → Fd (inl (f x))) ∣₀ +-- helper (suc n) Fa Fb Fd p = cong (morph.fun (Δ (suc n))) (sym p) +-- ∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₀ +ₕ -ₕ ∣ (λ x → Fd (push x (~ i))) ∣₀) +-- ∙ rCancelₕ ∣ (λ x → Fd (inl (f x))) ∣₀ + + + +-- Ker-Δ⊂Im-i : (n : ℕ) (a : coHom n A × coHom n B) +-- → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) a +-- → isInIm (coHomGr n (Pushout f g)) (×coHomGr n A B) (i n) a +-- Ker-Δ⊂Im-i n (Fa , Fb) = +-- sElim {B = λ Fa → (Fb : _) → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) (Fa , Fb) +-- → isInIm (coHomGr n (Pushout f g)) (×coHomGr n A B) (i n) (Fa , Fb)} +-- (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- (λ Fa → sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- λ Fb p → pRec propTruncIsProp +-- (λ q → ∣ ∣ helpFun n Fa Fb (funExt⁻ q) ∣₀ +-- , anotherHelper n Fa Fb q ∣₋₁) +-- (helper n Fa Fb p)) +-- Fa +-- Fb + +-- where +-- helper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) +-- → (morph.fun (Δ n)) (∣ Fa ∣₀ , ∣ Fb ∣₀) ≡ 0ₕ +-- → ∥ (Path (_ → _) (λ c → Fa (f c)) (λ c → Fb (g c))) ∥₋₁ +-- helper zero Fa Fb p = Iso.fun (PathIdTrunc₀Iso) +-- (sym (rUnitₕ (coHomFun zero f ∣ Fa ∣₀)) +-- ∙ (λ i → coHomFun zero f ∣ Fa ∣₀ +ₕ (lCancelₕ (coHomFun zero g ∣ Fb ∣₀) (~ i))) +-- ∙ sym (assocₕ (coHomFun zero f ∣ Fa ∣₀) (-ₕ (coHomFun zero g ∣ Fb ∣₀)) (coHomFun zero g ∣ Fb ∣₀)) +-- ∙ cong (λ x → x +ₕ (coHomFun zero g ∣ Fb ∣₀)) p +-- ∙ lUnitₕ (coHomFun zero g ∣ Fb ∣₀)) +-- helper (suc n) Fa Fb p = Iso.fun (PathIdTrunc₀Iso) +-- (sym (rUnitₕ (coHomFun (suc n) f ∣ Fa ∣₀)) +-- ∙ (λ i → coHomFun (suc n) f ∣ Fa ∣₀ +ₕ (lCancelₕ (coHomFun (suc n) g ∣ Fb ∣₀) (~ i))) +-- ∙ sym (assocₕ (coHomFun (suc n) f ∣ Fa ∣₀) (-ₕ (coHomFun (suc n) g ∣ Fb ∣₀)) (coHomFun (suc n) g ∣ Fb ∣₀)) +-- ∙ cong (λ x → x +ₕ (coHomFun (suc n) g ∣ Fb ∣₀)) p +-- ∙ lUnitₕ (coHomFun (suc n) g ∣ Fb ∣₀)) + +-- helpFun : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) +-- → ((c : C) → Fa (f c) ≡ Fb (g c)) +-- → (Pushout f g) → coHomK n +-- helpFun n Fa Fb p (inl x) = Fa x +-- helpFun n Fa Fb p (inr x) = Fb x +-- helpFun n Fa Fb p (push a i) = p a i + +-- anotherHelper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) +-- → (q : Path (C → coHomK n) (λ c → Fa (f c)) (λ c → Fb (g c))) +-- → morph.fun (i n) ∣ helpFun n Fa Fb (λ x i₁ → q i₁ x) ∣₀ ≡ (∣ Fa ∣₀ , ∣ Fb ∣₀) +-- anotherHelper zero Fa Fb q = refl +-- anotherHelper (suc n) Fa Fb q = refl + + +-- Ker-d⊂Im-Δ : (n : ℕ) (a : coHom n C) +-- → isInKer (coHomGr n C) (coHomGr (suc n) (Pushout f g)) (d n) a +-- → isInIm (×coHomGr n A B) (coHomGr n C) (Δ n) a +-- Ker-d⊂Im-Δ n = +-- sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) +-- λ Fc p → pRec propTruncIsProp (λ p → ∣ (∣ (λ a → ΩKn+1→Kn (cong (λ f → f (inl a)) p)) ∣₀ , +-- ∣ (λ b → ΩKn+1→Kn (cong (λ f → f (inr b)) p)) ∣₀) , +-- Iso.inv (PathIdTrunc₀Iso) ∣ funExt (λ c → helper2 n Fc p c) ∣₋₁ ∣₋₁) +-- (Iso.fun (PathIdTrunc₀Iso) p) + +-- where +-- distrHelper : (n : ℕ) (p q : _) +-- → ΩKn+1→Kn {n = n} p +ₖ (-ₖ ΩKn+1→Kn q) ≡ ΩKn+1→Kn (p ∙ sym q) +-- distrHelper n p q i = +-- ΩKn+1→Kn (Iso.rightInv (Iso3-Kn-ΩKn+1 n) p i +-- ∙ Iso.rightInv (Iso3-Kn-ΩKn+1 n) (sym (Iso.rightInv (Iso3-Kn-ΩKn+1 n) q i)) i) + + +-- helper2 : (n : ℕ) (Fc : C → coHomK n) +-- (p : d-pre n Fc ≡ (λ _ → ∣ north ∣)) (c : C) +-- → ΩKn+1→Kn (λ i₁ → p i₁ (inl (f c))) +ₖ (-ₖ ΩKn+1→Kn (λ i₁ → p i₁ (inr (g c)))) ≡ Fc c +-- helper2 zero Fc p c = +-- distrHelper zero (cong (λ F → F (inl (f c))) p) (cong (λ F → F (inr (g c))) p) +-- ∙ cong ΩKn+1→Kn (sym ((PathP→compPathR (cong (λ f → cong f (push c)) p)) +-- ∙ (λ i → (λ i₁ → p i₁ (inl (f c))) +-- ∙ (lUnit (sym (λ i₁ → p i₁ (inr (g c)))) (~ i))))) +-- ∙ Iso.leftInv (Iso3-Kn-ΩKn+1 zero) (Fc c) +-- helper2 (suc n) Fc p c = +-- distrHelper (suc n) (cong (λ F → F (inl (f c))) p) (cong (λ F → F (inr (g c))) p) +-- ∙ cong ΩKn+1→Kn (sym ((PathP→compPathR (cong (λ f → cong f (push c)) p)) +-- ∙ (λ i → (λ i₁ → p i₁ (inl (f c))) +-- ∙ (lUnit (sym (λ i₁ → p i₁ (inr (g c)))) (~ i))))) +-- ∙ Iso.leftInv (Iso3-Kn-ΩKn+1 (suc n)) (Fc c) + +-- Im-Δ⊂Ker-d : (n : ℕ) (a : coHom n C) +-- → isInIm (×coHomGr n A B) (coHomGr n C) (Δ n) a +-- → isInKer (coHomGr n C) (coHomGr (suc n) (Pushout f g)) (d n) a +-- Im-Δ⊂Ker-d n = +-- sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- λ Fc → pRec (isOfHLevelPath' 1 setTruncIsSet _ _) +-- (sigmaProdElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- λ Fa Fb p → pRec (isOfHLevelPath' 1 setTruncIsSet _ _) +-- (λ q → ((λ i → morph.fun (d n) ∣ (q (~ i)) ∣₀) ∙ dΔ-Id n Fa Fb)) +-- (Iso.fun (PathIdTrunc₀Iso) p)) + +-- where + +-- d-preLeftId : (n : ℕ) (Fa : A → coHomK n)(d : (Pushout f g)) +-- → d-pre n (Fa ∘ f) d ≡ 0ₖ +-- d-preLeftId n Fa (inl x) = Kn→ΩKn+1 n (Fa x) +-- d-preLeftId n Fa (inr x) = refl +-- d-preLeftId zero Fa (push a i) j = Kn→ΩKn+1 zero (Fa (f a)) (j ∨ i) +-- d-preLeftId (suc n) Fa (push a i) j = Kn→ΩKn+1 (suc n) (Fa (f a)) (j ∨ i) + +-- d-preRightId : (n : ℕ) (Fb : B → coHomK n) (d : (Pushout f g)) +-- → d-pre n (Fb ∘ g) d ≡ 0ₖ +-- d-preRightId n Fb (inl x) = refl +-- d-preRightId n Fb (inr x) = sym (Kn→ΩKn+1 n (Fb x)) +-- d-preRightId zero Fb (push a i) j = Kn→ΩKn+1 zero (Fb (g a)) (~ j ∧ i) +-- d-preRightId (suc n) Fb (push a i) j = Kn→ΩKn+1 (suc n) (Fb (g a)) (~ j ∧ i) + +-- dΔ-Id : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) +-- → morph.fun (d n) (morph.fun (Δ n) (∣ Fa ∣₀ , ∣ Fb ∣₀)) ≡ 0ₕ +-- dΔ-Id zero Fa Fb = +-- morph.ismorph (d zero) ∣ (λ x → Fa (f x)) ∣₀ (-ₕ ∣ (λ x → Fb (g x)) ∣₀) +-- ∙ (λ i → morph.fun (d zero) ∣ (λ x → Fa (f x)) ∣₀ +ₕ morphMinus (coHomGr zero C) (coHomGr 1 (Pushout f g)) (d zero) ∣ (λ x → Fb (g x)) ∣₀ i) +-- ∙ (λ i → morph.fun (d zero) ∣ (λ x → Fa (f x)) ∣₀ +ₕ -ₕ (morph.fun (d zero) ∣ (λ x → Fb (g x)) ∣₀)) +-- ∙ (λ i → ∣ funExt (d-preLeftId zero Fa) i ∣₀ +ₕ -ₕ ∣ funExt (d-preRightId zero Fb) i ∣₀) +-- ∙ rCancelₕ 0ₕ +-- dΔ-Id (suc n) Fa Fb = +-- morph.ismorph (d (suc n)) ∣ (λ x → Fa (f x)) ∣₀ (-ₕ ∣ (λ x → Fb (g x)) ∣₀) +-- ∙ (λ i → morph.fun (d (suc n)) ∣ (λ x → Fa (f x)) ∣₀ +ₕ morphMinus (coHomGr (suc n) C) (coHomGr (2 + n) (Pushout f g)) (d (suc n)) ∣ (λ x → Fb (g x)) ∣₀ i) +-- ∙ (λ i → morph.fun (d (suc n)) ∣ (λ x → Fa (f x)) ∣₀ +ₕ -ₕ (morph.fun (d (suc n)) ∣ (λ x → Fb (g x)) ∣₀)) +-- ∙ (λ i → ∣ funExt (d-preLeftId (suc n) Fa) i ∣₀ +ₕ -ₕ ∣ funExt (d-preRightId (suc n) Fb) i ∣₀) +-- ∙ rCancelₕ 0ₕ + + +-- -- d-morph : ∀ {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) (n : ℕ) +-- -- → morph (coHomGr n C) (coHomGr (suc n) (Pushout f g)) +-- -- morph.fun (d-morph A B C f g n) = MV.d A B C f g n +-- -- morph.ismorph (d-morph A B C f g n) = MV.dIsHom A B C f g n + +-- -- i-morph : ∀ {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) (n : ℕ) +-- -- → morph (coHomGr n (Pushout f g)) (×coHomGr n A B) +-- -- morph.fun (i-morph A B C f g n) = MV.i A B C f g n +-- -- morph.ismorph (i-morph A B C f g n) = MV.iIsHom A B C f g n + +-- -- Δ-morph : ∀ {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) (n : ℕ) +-- -- → morph (×coHomGr n A B) (coHomGr n C) +-- -- morph.fun (Δ-morph A B C f g n) = MV.Δ A B C f g n +-- -- morph.ismorph (Δ-morph A B C f g n) = MV.ΔIsHom A B C f g n From da3234b523870e553807b21e506817f513e63981 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Wed, 17 Jun 2020 02:49:52 +0200 Subject: [PATCH 46/89] private --- Cubical/HITs/S1/Base.agda | 180 +++++++++++++++++++------------------- 1 file changed, 90 insertions(+), 90 deletions(-) diff --git a/Cubical/HITs/S1/Base.agda b/Cubical/HITs/S1/Base.agda index e07fa76c4..9b94a89fc 100644 --- a/Cubical/HITs/S1/Base.agda +++ b/Cubical/HITs/S1/Base.agda @@ -170,96 +170,96 @@ private ; (j = i1) → loop (i ∧ (~ t)) }) (inS (x j)) l -ΩS¹→basedΩS¹ : (i : I) → ΩS¹ → basedΩS¹ (loop i) -ΩS¹→basedΩS¹ i x j = ΩS¹→basedΩS¹-filler i1 i x j - -basedΩS¹→ΩS¹ : (i : I) → basedΩS¹ (loop i) → ΩS¹ -basedΩS¹→ΩS¹ i x j = basedΩS¹→ΩS¹-filler i1 i x j - - - -basedΩS¹→ΩS¹→basedΩS¹ : (i : I) → (x : basedΩS¹ (loop i)) - → ΩS¹→basedΩS¹ i (basedΩS¹→ΩS¹ i x) ≡ x -basedΩS¹→ΩS¹→basedΩS¹ i x j k = - hcomp (λ t → λ { (j = i1) → basedΩS¹→ΩS¹-filler (~ t) i x k - ; (j = i0) → ΩS¹→basedΩS¹ i (basedΩS¹→ΩS¹ i x) k - ; (k = i0) → loop (i ∧ (t ∨ (~ j))) - ; (k = i1) → loop (i ∧ (t ∨ (~ j))) }) - (ΩS¹→basedΩS¹-filler (~ j) i (basedΩS¹→ΩS¹ i x) k) - -ΩS¹→basedΩS¹→ΩS¹ : (i : I) → (x : ΩS¹) - → basedΩS¹→ΩS¹ i (ΩS¹→basedΩS¹ i x) ≡ x -ΩS¹→basedΩS¹→ΩS¹ i x j k = - hcomp (λ t → λ { (j = i1) → ΩS¹→basedΩS¹-filler (~ t) i x k - ; (j = i0) → basedΩS¹→ΩS¹ i (ΩS¹→basedΩS¹ i x) k - ; (k = i0) → loop (i ∧ ((~ t) ∧ j)) - ; (k = i1) → loop (i ∧ ((~ t) ∧ j)) }) - (basedΩS¹→ΩS¹-filler (~ j) i (ΩS¹→basedΩS¹ i x) k) - --- from the existence of our quasi-inverse, we deduce that the basechange is an equivalence --- for all loop i - -basedΩS¹→ΩS¹-isequiv : (i : I) → isEquiv (basedΩS¹→ΩS¹ i) -basedΩS¹→ΩS¹-isequiv i = isoToIsEquiv (iso (basedΩS¹→ΩS¹ i) (ΩS¹→basedΩS¹ i) - (ΩS¹→basedΩS¹→ΩS¹ i) (basedΩS¹→ΩS¹→basedΩS¹ i)) - - --- now extend the basechange so that both ends match --- (and therefore we get a basechange for any x : S¹) - -loop-conjugation : basedΩS¹→ΩS¹ i1 ≡ λ x → x -loop-conjugation i x = - let p = (doubleCompPath-elim loop x (sym loop)) - ∙ (λ i → (lUnit loop i ∙ x) ∙ sym loop) - in - ((sym (decodeEncode base (basedΩS¹→ΩS¹ i1 x))) - ∙ (λ t → intLoop (winding (p t))) - ∙ (λ t → intLoop (winding-hom (intLoop (pos (suc zero)) ∙ x) - (intLoop (negsuc zero)) t)) - ∙ (λ t → intLoop ((winding-hom (intLoop (pos (suc zero))) x t) - + (windingIntLoop (negsuc zero) t))) - ∙ (λ t → intLoop (((windingIntLoop (pos (suc zero)) t) + (winding x)) + (negsuc zero))) - ∙ (λ t → intLoop ((+-comm (pos (suc zero)) (winding x) t) + (negsuc zero))) - ∙ (λ t → intLoop (+-assoc (winding x) (pos (suc zero)) (negsuc zero) (~ t))) - ∙ (decodeEncode base x)) i - -refl-conjugation : basedΩS¹→ΩS¹ i0 ≡ λ x → x -refl-conjugation i x j = - hfill (λ t → λ { (j = i0) → base - ; (j = i1) → base }) - (inS (x j)) (~ i) - -basechange : (x : S¹) → basedΩS¹ x → ΩS¹ -basechange base y = y -basechange (loop i) y = - hcomp (λ t → λ { (i = i0) → refl-conjugation t y - ; (i = i1) → loop-conjugation t y }) - (basedΩS¹→ΩS¹ i y) - --- for any loop i, the old basechange is equal to the new one -basedΩS¹→ΩS¹≡basechange : (i : I) → basedΩS¹→ΩS¹ i ≡ basechange (loop i) -basedΩS¹→ΩS¹≡basechange i j y = - hfill (λ t → λ { (i = i0) → refl-conjugation t y - ; (i = i1) → loop-conjugation t y }) - (inS (basedΩS¹→ΩS¹ i y)) j - --- so for any loop i, the extended basechange is an equivalence -basechange-isequiv-aux : (i : I) → isEquiv (basechange (loop i)) -basechange-isequiv-aux i = - transport (λ j → isEquiv (basedΩS¹→ΩS¹≡basechange i j)) (basedΩS¹→ΩS¹-isequiv i) - - --- as being an equivalence is contractible, basechange is an equivalence for all x : S¹ -basechange-isequiv : (x : S¹) → isEquiv (basechange x) -basechange-isequiv base = basechange-isequiv-aux i0 -basechange-isequiv (loop i) = - hcomp (λ t → λ { (i = i0) → basechange-isequiv-aux i0 - ; (i = i1) → isPropIsEquiv (basechange base) (basechange-isequiv-aux i1) - (basechange-isequiv-aux i0) t }) - (basechange-isequiv-aux i) - -basedΩS¹≡ΩS¹ : (x : S¹) → basedΩS¹ x ≡ ΩS¹ -basedΩS¹≡ΩS¹ x = ua (basechange x , basechange-isequiv x) + ΩS¹→basedΩS¹ : (i : I) → ΩS¹ → basedΩS¹ (loop i) + ΩS¹→basedΩS¹ i x j = ΩS¹→basedΩS¹-filler i1 i x j + + basedΩS¹→ΩS¹ : (i : I) → basedΩS¹ (loop i) → ΩS¹ + basedΩS¹→ΩS¹ i x j = basedΩS¹→ΩS¹-filler i1 i x j + + + + basedΩS¹→ΩS¹→basedΩS¹ : (i : I) → (x : basedΩS¹ (loop i)) + → ΩS¹→basedΩS¹ i (basedΩS¹→ΩS¹ i x) ≡ x + basedΩS¹→ΩS¹→basedΩS¹ i x j k = + hcomp (λ t → λ { (j = i1) → basedΩS¹→ΩS¹-filler (~ t) i x k + ; (j = i0) → ΩS¹→basedΩS¹ i (basedΩS¹→ΩS¹ i x) k + ; (k = i0) → loop (i ∧ (t ∨ (~ j))) + ; (k = i1) → loop (i ∧ (t ∨ (~ j))) }) + (ΩS¹→basedΩS¹-filler (~ j) i (basedΩS¹→ΩS¹ i x) k) + + ΩS¹→basedΩS¹→ΩS¹ : (i : I) → (x : ΩS¹) + → basedΩS¹→ΩS¹ i (ΩS¹→basedΩS¹ i x) ≡ x + ΩS¹→basedΩS¹→ΩS¹ i x j k = + hcomp (λ t → λ { (j = i1) → ΩS¹→basedΩS¹-filler (~ t) i x k + ; (j = i0) → basedΩS¹→ΩS¹ i (ΩS¹→basedΩS¹ i x) k + ; (k = i0) → loop (i ∧ ((~ t) ∧ j)) + ; (k = i1) → loop (i ∧ ((~ t) ∧ j)) }) + (basedΩS¹→ΩS¹-filler (~ j) i (ΩS¹→basedΩS¹ i x) k) + + -- from the existence of our quasi-inverse, we deduce that the basechange is an equivalence + -- for all loop i + + basedΩS¹→ΩS¹-isequiv : (i : I) → isEquiv (basedΩS¹→ΩS¹ i) + basedΩS¹→ΩS¹-isequiv i = isoToIsEquiv (iso (basedΩS¹→ΩS¹ i) (ΩS¹→basedΩS¹ i) + (ΩS¹→basedΩS¹→ΩS¹ i) (basedΩS¹→ΩS¹→basedΩS¹ i)) + + + -- now extend the basechange so that both ends match + -- (and therefore we get a basechange for any x : S¹) + + loop-conjugation : basedΩS¹→ΩS¹ i1 ≡ λ x → x + loop-conjugation i x = + let p = (doubleCompPath-elim loop x (sym loop)) + ∙ (λ i → (lUnit loop i ∙ x) ∙ sym loop) + in + ((sym (decodeEncode base (basedΩS¹→ΩS¹ i1 x))) + ∙ (λ t → intLoop (winding (p t))) + ∙ (λ t → intLoop (winding-hom (intLoop (pos (suc zero)) ∙ x) + (intLoop (negsuc zero)) t)) + ∙ (λ t → intLoop ((winding-hom (intLoop (pos (suc zero))) x t) + + (windingIntLoop (negsuc zero) t))) + ∙ (λ t → intLoop (((windingIntLoop (pos (suc zero)) t) + (winding x)) + (negsuc zero))) + ∙ (λ t → intLoop ((+-comm (pos (suc zero)) (winding x) t) + (negsuc zero))) + ∙ (λ t → intLoop (+-assoc (winding x) (pos (suc zero)) (negsuc zero) (~ t))) + ∙ (decodeEncode base x)) i + + refl-conjugation : basedΩS¹→ΩS¹ i0 ≡ λ x → x + refl-conjugation i x j = + hfill (λ t → λ { (j = i0) → base + ; (j = i1) → base }) + (inS (x j)) (~ i) + + basechange : (x : S¹) → basedΩS¹ x → ΩS¹ + basechange base y = y + basechange (loop i) y = + hcomp (λ t → λ { (i = i0) → refl-conjugation t y + ; (i = i1) → loop-conjugation t y }) + (basedΩS¹→ΩS¹ i y) + + -- for any loop i, the old basechange is equal to the new one + basedΩS¹→ΩS¹≡basechange : (i : I) → basedΩS¹→ΩS¹ i ≡ basechange (loop i) + basedΩS¹→ΩS¹≡basechange i j y = + hfill (λ t → λ { (i = i0) → refl-conjugation t y + ; (i = i1) → loop-conjugation t y }) + (inS (basedΩS¹→ΩS¹ i y)) j + + -- so for any loop i, the extended basechange is an equivalence + basechange-isequiv-aux : (i : I) → isEquiv (basechange (loop i)) + basechange-isequiv-aux i = + transport (λ j → isEquiv (basedΩS¹→ΩS¹≡basechange i j)) (basedΩS¹→ΩS¹-isequiv i) + + + -- as being an equivalence is contractible, basechange is an equivalence for all x : S¹ + basechange-isequiv : (x : S¹) → isEquiv (basechange x) + basechange-isequiv base = basechange-isequiv-aux i0 + basechange-isequiv (loop i) = + hcomp (λ t → λ { (i = i0) → basechange-isequiv-aux i0 + ; (i = i1) → isPropIsEquiv (basechange base) (basechange-isequiv-aux i1) + (basechange-isequiv-aux i0) t }) + (basechange-isequiv-aux i) + + basedΩS¹≡ΩS¹ : (x : S¹) → basedΩS¹ x ≡ ΩS¹ + basedΩS¹≡ΩS¹ x = ua (basechange x , basechange-isequiv x) basedΩS¹≡Int : (x : S¹) → basedΩS¹ x ≡ Int basedΩS¹≡Int x = (basedΩS¹≡ΩS¹ x) ∙ ΩS¹≡Int From 5490f4c7e7d2a7bdf277f85caa108a5326b10043 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Wed, 17 Jun 2020 03:28:15 +0200 Subject: [PATCH 47/89] emptylines --- Cubical/Foundations/GroupoidLaws.agda | 3 +++ Cubical/Foundations/Isomorphism.agda | 3 --- Cubical/Foundations/Prelude.agda | 1 - Cubical/HITs/S1/Base.agda | 10 +-------- Cubical/HITs/SetQuotients/Properties.agda | 27 ----------------------- Cubical/HITs/Truncation/Properties.agda | 3 +-- Cubical/ZCohomology/Groups/Torus.agda | 6 ++--- Cubical/ZCohomology/Groups/Unit.agda | 2 +- Cubical/ZCohomology/Properties.agda | 10 ++++----- 9 files changed, 14 insertions(+), 51 deletions(-) diff --git a/Cubical/Foundations/GroupoidLaws.agda b/Cubical/Foundations/GroupoidLaws.agda index d793b5d03..e7e3da46b 100644 --- a/Cubical/Foundations/GroupoidLaws.agda +++ b/Cubical/Foundations/GroupoidLaws.agda @@ -240,6 +240,7 @@ hcomp-cong : ∀ {ℓ} {A : Type ℓ} {φ} → (u : I → Partial φ A) → (u0 → (hcomp u (outS u0) ≡ hcomp u' (outS u0')) [ φ ↦ (\ { (φ = i1) → ueq i1 1=1 }) ] hcomp-cong u u0 u' u0' ueq 0eq = inS (\ j → hcomp (\ i o → ueq i o j) (outS 0eq j)) + invSides-filler : {x y z : A} (p : x ≡ y) (q : x ≡ z) → PathP (λ j → q j ≡ p (~ j)) p (sym q) invSides-filler {A = A} {x = x} p q i j = hcomp (λ k → λ { (i = i0) → p (k ∧ j) @@ -288,6 +289,8 @@ symDistr : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) → symDistr p q i j = symDistr-filler p q j i i1 + + pentagonIdentity : (p : x ≡ y) → (q : y ≡ z) → (r : z ≡ w) → (s : w ≡ v) → (assoc p q (r ∙ s) ∙ assoc (p ∙ q) r s) diff --git a/Cubical/Foundations/Isomorphism.agda b/Cubical/Foundations/Isomorphism.agda index fb8e20a8a..753c0ea45 100644 --- a/Cubical/Foundations/Isomorphism.agda +++ b/Cubical/Foundations/Isomorphism.agda @@ -43,9 +43,6 @@ record Iso {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') : Type (ℓ-max ℓ ℓ') w isIso : (A → B) → Type _ isIso {A = A} {B = B} f = Σ[ g ∈ (B → A) ] Σ[ _ ∈ section f g ] retract f g -symIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → Iso A B → Iso B A -symIso (iso f g rinv linv) = iso g f linv rinv - isoFunInjective : (f : Iso A B) → (x y : A) → Iso.fun f x ≡ Iso.fun f y → x ≡ y isoFunInjective f x y h = sym (Iso.leftInv f x) ∙∙ cong (Iso.inv f) h ∙∙ Iso.leftInv f y diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index 3eea4ded7..570e31b06 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -66,7 +66,6 @@ cong₂ : ∀ {C : (a : A) → (b : B a) → Type ℓ} → cong₂ f p q i = f (p i) (q i) {-# INLINE cong₂ #-} - {- The most natural notion of homogenous path composition in a cubical setting is double composition: diff --git a/Cubical/HITs/S1/Base.agda b/Cubical/HITs/S1/Base.agda index 9b94a89fc..46310c9eb 100644 --- a/Cubical/HITs/S1/Base.agda +++ b/Cubical/HITs/S1/Base.agda @@ -52,7 +52,6 @@ intLoop (pos (suc n)) = intLoop (pos n) ∙ loop intLoop (negsuc zero) = sym loop intLoop (negsuc (suc n)) = intLoop (negsuc n) ∙ sym loop - decodeSquare : (n : Int) → PathP (λ i → base ≡ loop i) (intLoop (predInt n)) (intLoop n) decodeSquare (pos zero) i j = loop (i ∨ ~ j) decodeSquare (pos (suc n)) i j = hfill (λ k → λ { (j = i0) → base @@ -76,7 +75,6 @@ decode (loop i) y j = decodeEncode : (x : S¹) (p : base ≡ x) → decode x (encode x p) ≡ p decodeEncode x p = J (λ y q → decode y (encode y q) ≡ q) (λ _ → refl) p - isSetΩS¹ : isSet ΩS¹ isSetΩS¹ p q r s j i = hcomp (λ k → λ { (i = i0) → decodeEncode base p k @@ -85,9 +83,6 @@ isSetΩS¹ p q r s j i = ; (j = i1) → decodeEncode base (s i) k }) (decode base (isSetInt (winding p) (winding q) (cong winding r) (cong winding s) j i)) - - - -- This proof does not rely on rewriting hcomp with empty systems in -- Int as ghcomp has been implemented! windingIntLoop : (n : Int) → winding (intLoop n) ≡ n @@ -265,7 +260,7 @@ basedΩS¹≡Int : (x : S¹) → basedΩS¹ x ≡ Int basedΩS¹≡Int x = (basedΩS¹≡ΩS¹ x) ∙ ΩS¹≡Int ----- A shorter proof of the same thing ----- +---- Alternative proof of the same thing ----- toPropElim : ∀ {ℓ} {B : S¹ → Type ℓ} → ((s : S¹) → isProp (B s)) @@ -458,6 +453,3 @@ rotInv-4 base base i = base rotInv-4 base (loop k) i = rotInv-aux-4 i0 k (~ i) i1 rotInv-4 (loop j) base i = loop j rotInv-4 (loop j) (loop k) i = rotInv-aux-4 j k (~ i) i1 - - - diff --git a/Cubical/HITs/SetQuotients/Properties.agda b/Cubical/HITs/SetQuotients/Properties.agda index 24983db74..80fb2fe9f 100644 --- a/Cubical/HITs/SetQuotients/Properties.agda +++ b/Cubical/HITs/SetQuotients/Properties.agda @@ -37,10 +37,6 @@ private C : A / R → A / R → Type ℓ'' D : A / R → A / R → A / R → Type ℓ'' -setQuotientIsSet : isSet (A / R) -setQuotientIsSet a b p q = squash/ a b p q - - elimEq/ : (Bprop : (x : A / R ) → isProp (B x)) {x y : A / R} (eq : x ≡ y) @@ -94,28 +90,6 @@ elim Bset f feq (squash/ x y p q i j) = where g = elim Bset f feq --- elim2 : {B : A / R → A / R → Type ℓ} → --- (Bset : (x y : A / R) → isSet (B x y)) → --- (f : (a : A) (b : A) → (B [ a ] [ b ])) → --- (feq1 : (a a₁ b : A) (r : R a₁ b) → --- PathP (λ i → B [ a ] (eq/ a₁ b r i)) (f a a₁) (f a b)) → --- (feq2 : (a a₁ b : A) (r : R a₁ a) → --- PathP (λ i → B (eq/ a₁ a r i) [ b ]) (f a₁ b) (f a b)) → --- {!(a b c d : A) (r1 : R a b) (r2 : R a b) → !} → --- (x y : A / R) → B x y - - --- elim2 {A = A} {B = B} Bset f feq1 feq2 _ = --- elim (λ _ → isOfHLevelΠ 2 λ _ → Bset _ _) --- (λ a → elim (λ _ → Bset _ _) (λ b → f a b) (feq1 _)) --- λ p q r i → elim (Bset (eq/ p q r i)) (λ b → feq2 q p b r i) (λ a b r2 → {!isOfHLevel→isOfHLevelDep 2 {A = A} {λ x y → B [ y ] [ y ]}!}) --- {- --- Goal: PathP (λ i → (y : A / R) → B (eq/ p q r i) y) --- (elim (λ z → Bset [ p ] z) (λ b → f p b) (feq p)) --- (elim (λ z → Bset [ q ] z) (λ b → f q b) (feq q)) --- -} - - setQuotUniversal : {B : Type ℓ} (Bset : isSet B) → (A / R → B) ≃ (Σ[ f ∈ (A → B) ] ((a b : A) → R a b → f a ≡ f b)) setQuotUniversal Bset = isoToEquiv (iso intro out outRightInv outLeftInv) @@ -187,4 +161,3 @@ discreteSetQuotients {A = A} {R = R} Adis Rprop Req Rdec = J (λ b ab → ∀ k → PathP (λ i → (y : A / R) → Dec (ab i ≡ y)) (discreteSetQuotients' a) k) (λ k → funExt (λ x → isPropDec (squash/ _ _) _ _)) (eq/ a b ab) (discreteSetQuotients' b) - diff --git a/Cubical/HITs/Truncation/Properties.agda b/Cubical/HITs/Truncation/Properties.agda index 21e2011f5..d34019f94 100644 --- a/Cubical/HITs/Truncation/Properties.agda +++ b/Cubical/HITs/Truncation/Properties.agda @@ -131,7 +131,6 @@ elim : {n : HLevel} B x elim hB = Null.elim (λ x → isOfHLevel→isSnNull (hB x)) - elim2 : {n : HLevel} {B : hLevelTrunc n A → hLevelTrunc n A → Type ℓ'} (hB : ((x y : hLevelTrunc n A) → isOfHLevel n (B x y))) @@ -360,7 +359,7 @@ PathΩ n = PathIdTrunc n PathIdTrunc₀Iso : {a b : A} → Iso (∣ a ∣₂ ≡ ∣ b ∣₂) ∥ a ≡ b ∥₁ PathIdTrunc₀Iso = compIso (congIso setTruncTrunc2Iso) (compIso (ΩTrunc.IsoFinal _ ∣ _ ∣ ∣ _ ∣) - (symIso propTruncTrunc1Iso)) + (invIso propTruncTrunc1Iso)) ------------------------- diff --git a/Cubical/ZCohomology/Groups/Torus.agda b/Cubical/ZCohomology/Groups/Torus.agda index 17d592e78..e172ab72d 100644 --- a/Cubical/ZCohomology/Groups/Torus.agda +++ b/Cubical/ZCohomology/Groups/Torus.agda @@ -327,11 +327,11 @@ H²-T²≡ℤ = compIso helper' (compIso - (symIso (prodIso (groupIso→Iso H²-S¹≅0) + (invIso (prodIso (groupIso→Iso H²-S¹≅0) (groupIso→Iso (invGrIso (Hⁿ-Sⁿ≅ℤ 0))))) (compIso - (symIso setTruncOfProdIso) - (symIso + (invIso setTruncOfProdIso) + (invIso (setTruncIso (compIso curryIso diff --git a/Cubical/ZCohomology/Groups/Unit.agda b/Cubical/ZCohomology/Groups/Unit.agda index 3c7ff98c5..4ab14fda3 100644 --- a/Cubical/ZCohomology/Groups/Unit.agda +++ b/Cubical/ZCohomology/Groups/Unit.agda @@ -67,7 +67,7 @@ Hⁿ-contrType≅0 {A = A} n contr = where helper1 : Iso (coHom (suc n) A) (coHom (suc n) Unit) helper1 = compIso (setTruncIso (ContrToTypeIso contr)) - (setTruncIso (symIso (ContrToTypeIso isContrUnit))) + (setTruncIso (invIso (ContrToTypeIso isContrUnit))) helper : isContr (coHom (suc n) A) helper = (Iso.inv helper1 0ₕ) diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index 3c51d0db2..e46634727 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -26,7 +26,7 @@ open import Cubical.Homotopy.Loopspace open import Cubical.Homotopy.Connected open import Cubical.Homotopy.Freudenthal open import Cubical.HITs.SmashProduct.Base -open import Cubical.Data.Group.Base hiding (_≃_ ; Iso) +open import Cubical.Data.Group.Base hiding (_≃_ ; Iso ; invIso) open import Cubical.Foundations.Equiv.HalfAdjoint @@ -315,19 +315,19 @@ isCommΩK-based (suc n) x p q = Iso.leftInv (congIsoHelper x) a = assocₖ a x (-ₖ x) ∙∙ cong (a +ₖ_) (rCancelₖ x) ∙∙ rUnitₖ a typId : (x : coHomK (suc n)) → (x ≡ x) ≡ Path (coHomK (suc n)) 0ₖ 0ₖ - typId x = isoToPath (congIso (symIso (congIsoHelper x))) ∙ λ i → rCancelₖ x i ≡ rCancelₖ x i + typId x = isoToPath (congIso (invIso (congIsoHelper x))) ∙ λ i → rCancelₖ x i ≡ rCancelₖ x i transpTypId : (p q : (x ≡ x)) → transport (λ i → typId x i) (p ∙ q) ≡ transport (λ i → typId x i) p ∙ transport (λ i → typId x i) q transpTypId p q = - ((substComposite (λ x → x) (isoToPath ((congIso (symIso (congIsoHelper x))))) + ((substComposite (λ x → x) (isoToPath ((congIso (invIso (congIsoHelper x))))) (λ i → rCancelₖ x i ≡ rCancelₖ x i) (p ∙ q)) ∙∙ (λ i → transport (λ i → rCancelₖ x i ≡ rCancelₖ x i) (transportRefl (congFunct (_+ₖ (-ₖ x)) p q i) i)) ∙∙ overPathFunct (cong (_+ₖ (-ₖ x)) p) (cong (_+ₖ (-ₖ x)) q) (rCancelₖ x)) ∙∙ cong₂ (λ y z → transport (λ i → rCancelₖ x i ≡ rCancelₖ x i) y ∙ transport (λ i → rCancelₖ x i ≡ rCancelₖ x i) z) (sym (transportRefl (cong (_+ₖ (-ₖ x)) p))) (sym (transportRefl (cong (_+ₖ (-ₖ x)) q))) - ∙∙ cong₂ (_∙_) (sym (substComposite (λ x → x) (isoToPath ((congIso (symIso (congIsoHelper x))))) (λ i → rCancelₖ x i ≡ rCancelₖ x i) p)) - (sym (substComposite (λ x → x) (isoToPath ((congIso (symIso (congIsoHelper x))))) (λ i → rCancelₖ x i ≡ rCancelₖ x i) q)) + ∙∙ cong₂ (_∙_) (sym (substComposite (λ x → x) (isoToPath ((congIso (invIso (congIsoHelper x))))) (λ i → rCancelₖ x i ≡ rCancelₖ x i) p)) + (sym (substComposite (λ x → x) (isoToPath ((congIso (invIso (congIsoHelper x))))) (λ i → rCancelₖ x i ≡ rCancelₖ x i) q)) open import Cubical.Data.Int renaming (_+_ to _ℤ+_) From 0780eee2d0f9e7fa1a3137541bbaab2850ac5b69 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Wed, 17 Jun 2020 03:53:37 +0200 Subject: [PATCH 48/89] whitespace:-) --- Cubical/Data/Group/Properties.agda | 13 ++++--------- Cubical/Foundations/Prelude.agda | 10 +++++----- Cubical/HITs/SetTruncation/Properties.agda | 3 +-- Cubical/ZCohomology/Groups/Sn.agda | 14 +++++++------- 4 files changed, 17 insertions(+), 23 deletions(-) diff --git a/Cubical/Data/Group/Properties.agda b/Cubical/Data/Group/Properties.agda index 79982400f..5fdd04806 100644 --- a/Cubical/Data/Group/Properties.agda +++ b/Cubical/Data/Group/Properties.agda @@ -7,19 +7,14 @@ open import Cubical.Foundations.HLevels open import Cubical.Data.Prod open import Cubical.HITs.PropositionalTruncation hiding (map) open import Cubical.Data.Group.Base -open import Cubical.Data.Sigma.Base hiding (comp) - -import Cubical.Foundations.Isomorphism as I -import Cubical.Foundations.Equiv as E -import Cubical.Foundations.Equiv.HalfAdjoint as HAE - +open import Cubical.Data.Sigma.Base hiding (comp) {- Given the following diagram a ↦ (a , 0) ψ ϕ A --> A × A -------> B ---> C -If ψ is an isomorphism and ϕ is surjective with ker ϕ ≡ {ψ (a , a) ∣ a ∈ A}, then C ≅ B +If ψ is an isomorphism and ϕ is surjective with ker ϕ ≡ {ψ (a , a) ∣ a ∈ A}, then C ≅ B -} diagonalIso1 : ∀ {ℓ ℓ' ℓ''} (A : Group ℓ) (B : Group ℓ') (C : Group ℓ'') @@ -85,14 +80,14 @@ diagonalIso1 A' B' C' ψ' ϕ' issurj ker→diag diag→ker = Given the following diagram ϕ -A × A ⇉ B +A × A ⇉ B ^ | | a ↦ (a , 0) | A ≃ C -If ϕ is surjective with ker ϕ ≡ {(a , a) ∣ a ∈ A}, then C ≅ B +If ϕ is surjective with ker ϕ ≡ {(a , a) ∣ a ∈ A}, then C ≅ B -} diagonalIso2 : ∀ {ℓ ℓ' ℓ''} (A : Group ℓ) (B : Group ℓ') (C : Group ℓ'') diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index 570e31b06..1734922ee 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -28,7 +28,7 @@ open import Cubical.Core.Primitives public infixr 30 _∙_ infix 3 _∎ infixr 2 _≡⟨_⟩_ -infixr 2.5 _≡⟨_⟩≡⟨_⟩_ +infixr 2.5 _≡⟨_⟩≡⟨_⟩_ -- Basic theory about paths. These proofs should typically be -- inlined. This module also makes equational reasoning work with @@ -403,10 +403,10 @@ liftExt x i = lift (x i) q⁻¹ z — — — > x ^ ^ - q | | p⁻¹ - | | - x — — — > y - p + q | | p⁻¹ + | | + x — — — > y + p -} doubleCompPath-filler∙ : {a b c d : A} (p : a ≡ b) (q : b ≡ c) (r : c ≡ d) diff --git a/Cubical/HITs/SetTruncation/Properties.agda b/Cubical/HITs/SetTruncation/Properties.agda index 6590ea11a..1ac8034fd 100644 --- a/Cubical/HITs/SetTruncation/Properties.agda +++ b/Cubical/HITs/SetTruncation/Properties.agda @@ -149,7 +149,7 @@ prodElim2 isset f = prodElim (λ _ → isOfHLevelΠ 2 λ _ → isset _ _) λ c d → f a b c d -setTruncOfProdIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → Iso ∥ A × B ∥₂ (∥ A ∥₂ × ∥ B ∥₂) +setTruncOfProdIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → Iso ∥ A × B ∥₂ (∥ A ∥₂ × ∥ B ∥₂) Iso.fun setTruncOfProdIso = rec (isOfHLevelProd 2 setTruncIsSet setTruncIsSet) λ { (a , b) → ∣ a ∣₂ , ∣ b ∣₂ } Iso.inv setTruncOfProdIso = prodElim (λ _ → setTruncIsSet) λ a b → ∣ a , b ∣₂ Iso.rightInv setTruncOfProdIso = @@ -158,4 +158,3 @@ Iso.rightInv setTruncOfProdIso = Iso.leftInv setTruncOfProdIso = elim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ {(a , b) → refl} - diff --git a/Cubical/ZCohomology/Groups/Sn.agda b/Cubical/ZCohomology/Groups/Sn.agda index fce40492e..5a8d380ad 100644 --- a/Cubical/ZCohomology/Groups/Sn.agda +++ b/Cubical/ZCohomology/Groups/Sn.agda @@ -63,7 +63,7 @@ morph.fun (grIso.inv (coHomPushout≅coHomSn n m)) = (λ f → ∣ (λ {north → f (inl tt) ; south → f (inr tt) ; (merid a i) → f (push a i)}) ∣₂) -morph.ismorph (grIso.inv (coHomPushout≅coHomSn n m)) = +morph.ismorph (grIso.inv (coHomPushout≅coHomSn n m)) = sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ a b → cong ∣_∣₂ (funExt λ {north → refl ; south → refl @@ -139,7 +139,7 @@ H¹-S¹≅ℤ = (∣ (λ _ → x) ∣₂ , ∣ (λ _ → 0) ∣₂) , (cong ∣_∣₂ (funExt (surjHelper x))) ∙ sym id ∣₋₁ }) surj) ) (invGrIso (coHomPushout≅coHomSn 0 1)) - + where surjHelper : (x : Int) (x₁ : S₊ 0) → x +ₖ (-ₖ pos 0) ≡ S0→Int (x , x) x₁ surjHelper x north = cong (x +ₖ_) (-0ₖ) ∙ rUnitₖ x @@ -166,7 +166,7 @@ H¹-S¹≅ℤ = helper2 : (f g : Unit → Int) → Σ[ x ∈ Int ] morph.fun (grIso.inv H⁰-S⁰≅ℤ×ℤ) (x , x) ≡ morph.fun (MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) (∣ f ∣₂ , ∣ g ∣₂) - helper2 f g = (f _ +ₖ (-ₖ g _) ) , cong ∣_∣₂ (funExt (λ {north → refl ; south → refl})) + helper2 f g = (f _ +ₖ (-ₖ g _) ) , cong ∣_∣₂ (funExt (λ {north → refl ; south → refl})) ---------------------------- Hⁿ(Sⁿ) ≅ ℤ , n ≥ 1 ------------------- @@ -199,7 +199,7 @@ Hⁿ-Sⁿ≅ℤ (suc n) = ------------------------- H¹(S⁰) ≅ 0 ------------------------------- -H¹-S⁰≅0 : grIso (coHomGr 1 (S₊ 0)) trivialGroup +H¹-S⁰≅0 : grIso (coHomGr 1 (S₊ 0)) trivialGroup H¹-S⁰≅0 = Iso'→Iso (iso' @@ -216,7 +216,7 @@ H¹-S⁰≅0 = helper f = elim2 (λ _ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelPlus {n = 1} 2 propTruncIsProp) (suspToPropRec2 north (λ _ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → propTruncIsProp) - λ id id2 → ∣ funExt (λ {north → id ; south → id2}) ∣₋₁) + λ id id2 → ∣ funExt (λ {north → id ; south → id2}) ∣₋₁) ------------------------- H²(S¹) ≅ 0 ------------------------------- @@ -305,7 +305,7 @@ Iso.leftInv S¹→S¹≡S¹×Int f = funExt λ { base → S¹map-id (f base) (λ i → ∣ basechange2 (S¹map (f base)) (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁)))))) i ∣) (cong f loop) - helper2 i j = + helper2 i j = hcomp (λ k → λ { (i = i0) → cong ∣_∣ (basechange2 (S¹map (f base)) (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁))))))) j ; (i = i1) → S¹map-id (f (loop j)) k @@ -374,7 +374,7 @@ coHom1S1≃ℤ = (λ a → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹≡S¹×Int (base , a)) ∣₂) -- inv (λ a → (λ i → proj₂ (Iso.fun S¹→S¹≡S¹×Int (Iso.fun S¹→S¹≡S1→S1 (Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹≡S¹×Int (base , a)))))) -- rInv ∙ (λ i → proj₂ (Iso.fun S¹→S¹≡S¹×Int (Iso.rightInv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹≡S¹×Int (base , a)) i))) - ∙ λ i → proj₂ (Iso.rightInv S¹→S¹≡S¹×Int (base , a) i)) + ∙ λ i → proj₂ (Iso.rightInv S¹→S¹≡S¹×Int (base , a) i)) (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) -- lInv λ f → (λ i → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹≡S¹×Int (base , proj₂ (Iso.fun S¹→S¹≡S¹×Int (Iso.fun S¹→S¹≡S1→S1 f)))) ∣₂) ∙∙ ((λ i → sRec setTruncIsSet From 5fbcf36759ea7fb314f98ddb7693db872e02330f Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 6 Jul 2020 02:22:34 +0200 Subject: [PATCH 49/89] backup --- Cubical/Data/Group/Properties.agda | 2 +- Cubical/Foundations/Prelude.agda | 1 + Cubical/HITs/S1/Base.agda | 4 +- Cubical/Homotopy/Connected.agda | 11 +- Cubical/Homotopy/Freudenthal.agda | 6 +- Cubical/Homotopy/Loopspace.agda | 4 +- Cubical/ZCohomology/Groups/Sn.agda | 20 ++- Cubical/ZCohomology/Groups/Unit.agda | 76 +++++++++ Cubical/ZCohomology/Groups/Wedge.agda | 15 +- .../ZCohomology/Groups/WedgeOfSpheres.agda | 4 + Cubical/ZCohomology/KcompPrelims.agda | 1 + .../ZCohomology/MayerVietorisUnreduced.agda | 144 +++++++++--------- 12 files changed, 197 insertions(+), 91 deletions(-) diff --git a/Cubical/Data/Group/Properties.agda b/Cubical/Data/Group/Properties.agda index 5fdd04806..dba4f9bc0 100644 --- a/Cubical/Data/Group/Properties.agda +++ b/Cubical/Data/Group/Properties.agda @@ -80,7 +80,7 @@ diagonalIso1 A' B' C' ψ' ϕ' issurj ker→diag diag→ker = Given the following diagram ϕ -A × A ⇉ B +A × A ↠ B ^ | | a ↦ (a , 0) diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index 1734922ee..f494481da 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -181,6 +181,7 @@ compPathP-filler {A = A} {x = x} {B = B} p q j i = (inS (p i)) j + compPathP'-filler : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} {x y z : A} {x' : B x} {y' : B y} {z' : B z} {p : x ≡ y} {q : y ≡ z} (P : PathP (λ i → B (p i)) x' y') (Q : PathP (λ i → B (q i)) y' z') → PathP (λ j → PathP (λ i → B (compPath-filler p q j i)) x' (Q j)) P (compPathP' {x = x} {y = y} {x' = x'} {y' = y'} P Q) diff --git a/Cubical/HITs/S1/Base.agda b/Cubical/HITs/S1/Base.agda index 46310c9eb..d5c4fa1fb 100644 --- a/Cubical/HITs/S1/Base.agda +++ b/Cubical/HITs/S1/Base.agda @@ -282,17 +282,15 @@ basechange2 (loop i) p = ∙ cong (_∙ p) (lCancel loop) ∙ sym (lUnit _)) k }) ((λ j → loop (~ j ∧ i)) ∙ p ∙ λ j → loop (j ∧ i)) - basechange2⁻ : (x : S¹) → (x ≡ x) → ΩS¹ basechange2⁻ base p = p -basechange2⁻ (loop i) p = +basechange2⁻ (loop i) p = hcomp (λ k → λ { (i = i0) → lUnit (rUnit p (~ k)) (~ k) ; (i = i1) → (cong (loop ∙_) (comm-ΩS¹ p (sym loop)) ∙ assoc loop (sym loop) p ∙ cong (_∙ p) (rCancel loop) ∙ sym (lUnit _)) k }) ((λ j → loop (i ∧ j)) ∙ p ∙ λ j → loop (i ∧ (~ j))) - basechange2-sect : (x : S¹) → section (basechange2 x) (basechange2⁻ x) basechange2-sect = toPropElim (λ _ → isOfHLevelΠ 1 λ _ → isSetΩx _ _ _ ) diff --git a/Cubical/Homotopy/Connected.agda b/Cubical/Homotopy/Connected.agda index 061b6866a..bdba8c8b6 100644 --- a/Cubical/Homotopy/Connected.agda +++ b/Cubical/Homotopy/Connected.agda @@ -88,7 +88,7 @@ module elim {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} (f : A → B) (n : isConnectedPrecompose : ((P : B → TypeOfHLevel (ℓ-max ℓ ℓ') n) → hasSection (λ(s : (b : B) → P b .fst) → s ∘ f)) → isConnectedFun n f - isConnectedPrecompose P→sect b = c n P→sect b , λ y → sym (fun n P→sect b y) -- (c n P→sect b) , λ y → sym (fun n P→sect b y) + isConnectedPrecompose P→sect b = c n P→sect b , λ y → sym (fun n P→sect b y) where P : (n : HLevel) → ((P : B → TypeOfHLevel ℓ n) → hasSection (λ(s : (b : B) → P b .fst) → s ∘ f)) @@ -183,7 +183,7 @@ isConnectedRetract : ∀ {ℓ ℓ'} (n : HLevel) (f : A → B) (g : B → A) (h : (x : A) → g (f x) ≡ x) → isConnected n B → isConnected n A -isConnectedRetract n f g h = +isConnectedRetract n f g h = isContrRetract (Trunc.map f) (Trunc.map g) @@ -198,7 +198,6 @@ isConnectedPoint n connA a₀ a = isConnectedRetract n snd (_ ,_) (λ _ → refl) (isConnectedPath n connA a₀ a) - isConnectedPoint2 : ∀ {ℓ} (n : HLevel) {A : Type ℓ} (a : A) → isConnectedFun n (λ(_ : Unit) → a) → isConnected (suc n) A @@ -231,6 +230,7 @@ connectedTruncIso {A = A} {B = B} (suc n) f con = g back : B → hLevelTrunc (suc n) A back y = map fst ((con y) .fst) + backSection : (b : B) → Path (hLevelTrunc (suc n) B) (Trunc.rec (isOfHLevelTrunc (suc n)) (λ a → ∣ f a ∣) @@ -250,9 +250,6 @@ connectedTruncIso {A = A} {B = B} (suc n) f con = g → P (map fst p) helper P hlev pf = Trunc.elim hlev λ pair → pf (fst pair) (snd pair) - backRetract : (a : A) → map fst (con (f a) .fst) ≡ ∣ a ∣ - backRetract a = cong (map fst) (con (f a) .snd ∣ a , refl ∣) - g : Iso (hLevelTrunc (suc n) A) (hLevelTrunc (suc n) B) Iso.fun g = map f Iso.inv g = Trunc.rec (isOfHLevelTrunc _) back @@ -281,7 +278,7 @@ inrConnected : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ → isConnectedFun n f → isConnectedFun n {A = B} {B = Pushout f g} inr inrConnected {A = A} {B = B} {C = C} n f g iscon = - elim.isConnectedPrecompose inr n λ P → (λ l → k P l) , λ b → refl + elim.isConnectedPrecompose inr n λ P → (k P) , λ b → refl where module _ {ℓ : Level} (P : (Pushout f g) → TypeOfHLevel ℓ n) (h : (b : B) → typ (P (inr b))) diff --git a/Cubical/Homotopy/Freudenthal.agda b/Cubical/Homotopy/Freudenthal.agda index e82d8fbc8..9d3433193 100644 --- a/Cubical/Homotopy/Freudenthal.agda +++ b/Cubical/Homotopy/Freudenthal.agda @@ -55,7 +55,7 @@ module _ {ℓ} (n : HLevel) {A : Pointed ℓ} (connA : isConnected (suc (suc n)) .equiv-proof (λ _ → Trunc.elim (λ _ → isProp→isOfHLevelSuc (n + suc n) isPropIsContr) - (λ fib → + (λ fib → subst (λ k → isContr (fiber k ∣ fib ∣)) (cong (Trunc.rec (isOfHLevelTrunc 2n+2) ∘ uncurry) (funExt (WC.right p) ⁻¹)) @@ -63,7 +63,7 @@ module _ {ℓ} (n : HLevel) {A : Pointed ℓ} (connA : isConnected (suc (suc n)) (funExt (Trunc.mapId) ⁻¹) (idIsEquiv _) .equiv-proof ∣ fib ∣) - )) + )) .fst .fst a interpolate : (a : typ A) @@ -85,7 +85,7 @@ module _ {ℓ} (n : HLevel) {A : Pointed ℓ} (connA : isConnected (suc (suc n)) encode y = J Code ∣ pt A , rCancel' (merid (pt A)) ∣ encodeMerid : (a : typ A) → encode south (merid a) ≡ ∣ a , refl ∣ - encodeMerid a = + encodeMerid a = cong (transport (λ i → gluePath i)) (funExt⁻ (WC.left refl a) _ ∙ cong ∣_∣ (cong (a ,_) (lem _ _))) ∙ transport (PathP≡Path gluePath _ _) diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda index bb7450f54..fd45e71a6 100644 --- a/Cubical/Homotopy/Loopspace.agda +++ b/Cubical/Homotopy/Loopspace.agda @@ -26,8 +26,8 @@ open import Cubical.HITs.SetTruncation Ω→ (f , f∙) = (λ p → (sym f∙ ∙ cong f p) ∙ f∙) , cong (λ q → q ∙ f∙) (sym (rUnit (sym f∙))) ∙ lCancel f∙ generalEH : {ℓ : Level} {A : Type ℓ} {a b c : A} {p q : a ≡ b} {r s : b ≡ c} (α : p ≡ q) (β : r ≡ s) - → (cong (λ x → x ∙ r) α) ∙ (cong (λ x → q ∙ x) β) - ≡ (cong (λ x → p ∙ x) β) ∙ (cong (λ x → x ∙ s) α) + → (cong (_∙ r) α) ∙ (cong (q ∙_) β) + ≡ (cong (p ∙_) β) ∙ (cong (_∙ s) α) generalEH {p = p} {r = r} α β j i = hcomp (λ k → λ { (i = i0) → p ∙ r ; (i = i1) → α (k ∨ ~ j) ∙ β (k ∨ j) }) diff --git a/Cubical/ZCohomology/Groups/Sn.agda b/Cubical/ZCohomology/Groups/Sn.agda index 5a8d380ad..583490d44 100644 --- a/Cubical/ZCohomology/Groups/Sn.agda +++ b/Cubical/ZCohomology/Groups/Sn.agda @@ -31,7 +31,8 @@ open import Cubical.Data.Unit open import Cubical.Data.Group.Base renaming (Iso to grIso ; compIso to compGrIso ; invIso to invGrIso ; idIso to idGrIso) open import Cubical.Data.Group.Properties - +open import Cubical.Homotopy.Connected +open import Cubical.Foundations.Equiv H⁰-Sⁿ≅ℤ : (n : ℕ) → grIso (coHomGr 0 (S₊ (suc n))) intGroup H⁰-Sⁿ≅ℤ n = @@ -413,10 +414,25 @@ coHom1S1≃ℤ = base base (λ x → S¹map (trMap S1→S¹ (ΩKn+1→Kn x))) - (λ x → Kn→ΩKn+1 1 (f (S¹→S1 x))) ((λ x → Kn→ΩKn+1 1 (g (S¹→S1 x)))) + (λ x → Kn→ΩKn+1 1 (f (S¹→S1 x))) + ((λ x → Kn→ΩKn+1 1 (g (S¹→S1 x)))) (cong (Kn→ΩKn+1 1) reflf ∙ Kn→ΩKn+10ₖ 1) (cong (Kn→ΩKn+1 1) reflg ∙ Kn→ΩKn+10ₖ 1)) ∙ λ j → basechange2⁻ (S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 base)) j))) (λ i → S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 (loop i))) j))) ∙ basechange2⁻ (S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 base)) j))) (λ i → S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 (loop i))) j)))) + + + + + +-- -- grIso (coHomGr 1 (S₊ 1)) intGroup +-- test : coHom 1 (S₊ 1) → Int +-- test = morph.fun (grIso.fun coHom1S1≃ℤ) + +-- test2 : Int → Int +-- test2 = morph.fun (grIso.fun coHom1S1≃ℤ) ∘ morph.fun (grIso.inv coHom1S1≃ℤ) + +-- test3 : Int → Int → Int +-- test3 x y = test ((morph.fun (grIso.inv coHom1S1≃ℤ) x) +ₕ (morph.fun (grIso.inv coHom1S1≃ℤ) y)) diff --git a/Cubical/ZCohomology/Groups/Unit.agda b/Cubical/ZCohomology/Groups/Unit.agda index 4ab14fda3..bfa2239ea 100644 --- a/Cubical/ZCohomology/Groups/Unit.agda +++ b/Cubical/ZCohomology/Groups/Unit.agda @@ -73,3 +73,79 @@ Hⁿ-contrType≅0 {A = A} n contr = helper = (Iso.inv helper1 0ₕ) , λ y → cong (Iso.inv helper1) (isOfHLevelSuc 0 (isContrHⁿ-Unit n) 0ₕ (Iso.fun helper1 y) ) ∙ Iso.leftInv helper1 y + +open import Cubical.Data.Bool +open import Cubical.Data.Sigma +open import Cubical.Foundations.GroupoidLaws +test : ∀ {ℓ} {A : Type ℓ} (a : A) → Iso (Unit → A) (Σ[ f ∈ (Bool → A) ] f true ≡ a) +Iso.fun (test {A = A} a) x = fhel , refl + module _ where + fhel : Bool → A + fhel false = x _ + fhel true = a +Iso.inv (test a) (f , id) _ = f false +Iso.rightInv (test a) (f , id) = λ i → (funExt helper i) , λ j → id ((~ i) ∨ j) + where + helper : (x : Bool) → fhel a (λ _ → f false) x ≡ f x + helper false = refl + helper true = sym id +Iso.leftInv (test a) x = refl + +test2 : Iso (Susp Unit) Unit +Iso.fun test2 _ = _ +Iso.inv test2 _ = north +Iso.rightInv test2 a = refl +Iso.leftInv test2 north = refl +Iso.leftInv test2 south = merid tt +Iso.leftInv test2 (merid tt i) j = merid tt (i ∧ j) + + +-- open import Cubical.Foundations.Pointed +-- test3 : ∀ {ℓ} {A : Type ℓ} (a : A) → Iso (a ≡ a) ((Σ[ f ∈ (Unit → A) ] f tt ≡ a)) +-- Iso.fun (test3 a) p = (λ _ → a) , p +-- Iso.inv (test3 a) (f , id) = refl +-- Iso.rightInv (test3 a) = {!!} +-- Iso.leftInv (test3 a) x = {!!} + +-- open import Cubical.HITs.S1 +-- test4 : {A : Type₀} (a : A) → Iso ((S¹ , base) →∙ (A , a)) (a ≡ a) +-- Iso.fun (test4 a) (f , id) = sym id ∙∙ (cong f loop) ∙∙ id +-- Iso.inv (test4 a) p = (λ {base → a ; (loop i) → p i}) , refl +-- Iso.rightInv (test4 a) b = sym (doubleCompPath-filler refl b refl) +-- Iso.leftInv (test4 a) (f , id) i = (funExt helper i) , λ j → id (~ i ∨ j) +-- where +-- helper : (x : S¹) → fst (Iso.inv (test4 a) (Iso.fun (test4 a) (f , id))) x ≡ f x +-- helper base = sym id +-- helper (loop i) j = doubleCompPath-filler (sym id) (cong f loop) (id) (~ j) i + + +-- test5 : {A B : Type₀} (a : A) (b : B) → Iso ((Susp A , north) →∙ (B , b)) ((A , a) →∙ ((b ≡ b) , refl)) +-- Iso.fun (test5 a b) (f , p) = (λ a' → sym p ∙∙ cong f ((merid a') ∙ sym (merid a)) ∙∙ p) , {!!} +-- Iso.inv (test5 {A = A} {B = B} a b) (f , p) = helper , refl +-- where +-- helper : Susp A → B +-- helper north = b +-- helper south = b +-- helper (merid x i) = f x i +-- Iso.rightInv (test5 a b) (f , p) = {!!} +-- Iso.leftInv (test5 a b) = {!!} + + +-- open import Cubical.Foundations.Equiv +-- open import Cubical.Homotopy.Connected +-- H⁰-connected : ∀ {ℓ} {A : Type ℓ} (a : A) → isConnected 2 A → Iso (coHom 0 A) Int +-- Iso.fun (H⁰-connected a con) = sRec isSetInt λ f → f a +-- Iso.inv (H⁰-connected a con) b = ∣ (λ x → b) ∣₂ +-- Iso.rightInv (H⁰-connected a con) b = refl +-- Iso.leftInv (H⁰-connected a con) = +-- sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- λ f → cong ∣_∣₂ (funExt (equiv-proof (elim.isEquivPrecompose (λ (x : Unit) → a) 1 (λ x → _ , isSetInt _ _) (isConnectedPoint _ con _)) (λ _ → refl) .fst .fst )) + +-- open import Cubical.HITs.Truncation renaming (rec to trRec) +-- H⁰-connected' : ∀ {ℓ} {A : Type ℓ} (a : A) → isConnected 2 A → Iso (coHom 0 A) Int +-- Iso.fun (H⁰-connected' a con) = sRec isSetInt λ f → f a +-- Iso.inv (H⁰-connected' a con) b = ∣ (λ x → b) ∣₂ +-- Iso.rightInv (H⁰-connected' a con) b = refl +-- Iso.leftInv (H⁰-connected' a con) = +-- sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) +-- λ f → cong ∣_∣₂ (funExt λ x → trRec (isSetInt _ _) (cong f) (isConnectedPath 1 con a x .fst)) diff --git a/Cubical/ZCohomology/Groups/Wedge.agda b/Cubical/ZCohomology/Groups/Wedge.agda index 75ffde648..c65d2e785 100644 --- a/Cubical/ZCohomology/Groups/Wedge.agda +++ b/Cubical/ZCohomology/Groups/Wedge.agda @@ -40,7 +40,7 @@ private private H¹-⋁'' : Iso'' (coHomGr 1 (A ⋁ B)) (×coHomGr 1 (typ A) (typ B)) Iso''.ϕ (H¹-⋁'' {A = A} {B = B}) = (MV.i (typ A) (typ B) Unit (λ _ → pt A) (λ _ → pt B) 1) - Iso''.inj (H¹-⋁'' {A = A} {B = B}) = -- Requires a "useless" elimination to terminate + Iso''.inj (H¹-⋁'' {A = A} {B = B}) = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ f inker → helper ∣ f ∣₂ (MV.Ker-i⊂Im-d _ _ Unit (λ _ → pt A) (λ _ → pt B) 0 ∣ f ∣₂ inker) where @@ -73,3 +73,16 @@ Hⁿ-⋁ {A = A} {B = B} (suc n) = (i {A = A} {B = B} (suc (suc n))) (MV.Ker-i⊂Im-d (typ A) (typ B) Unit (λ _ → pt A) (λ _ → pt B) (suc n)) (MV.Ker-Δ⊂Im-i (typ A) (typ B) Unit (λ _ → pt A) (λ _ → pt B) (suc (suc n)))) + +open import Cubical.Foundations.Isomorphism +wedgetrunc : ((x : typ A) → ∥ pt A ≡ x ∥) → ((x : typ B) → ∥ pt B ≡ x ∥) → Iso (∥ (A ⋁ B) ∥₂) Unit +Iso.fun (wedgetrunc {A = A , ptA} {B = B , ptB} conA conB) _ = _ +Iso.inv (wedgetrunc {A = A , ptA} {B = B , ptB} conA conB) _ = ∣ inl ptA ∣₂ +Iso.rightInv (wedgetrunc {A = A , ptA} {B = B , ptB} conA conB) _ = refl +Iso.leftInv (wedgetrunc {A = A , ptA} {B = B , ptB} conA conB) = + sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + (PushoutToProp (λ _ → setTruncIsSet _ _) + (λ a → pRec (setTruncIsSet _ _) + (cong (λ x → ∣ inl x ∣₂)) (conA a)) + (λ b → pRec (setTruncIsSet _ _) + (λ p → cong ∣_∣₂ (push tt) ∙ cong (λ x → ∣ inr x ∣₂) p) (conB b))) diff --git a/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda b/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda index 2196ad6fb..9e48f22e7 100644 --- a/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda +++ b/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda @@ -76,6 +76,10 @@ H¹-S²⋁S¹⋁S¹ = compGrIso (Hⁿ-⋁ 0) (dirProdIso H¹-S²≅0 H¹-S¹⋁S¹) lUnitGroupIso) +test : Int × Int → Int × Int +test x = morph.fun (grIso.fun H¹-S²⋁S¹⋁S¹) (morph.fun (grIso.inv H¹-S²⋁S¹⋁S¹) x) + + ------------- H²(S²⋁S¹⋁S¹) --------- H²-S²⋁S¹⋁S¹ : grIso (coHomGr 2 S²⋁S¹⋁S¹) intGroup H²-S²⋁S¹⋁S¹ = compGrIso (Hⁿ-⋁ 1) diff --git a/Cubical/ZCohomology/KcompPrelims.agda b/Cubical/ZCohomology/KcompPrelims.agda index 5bb3e9bfe..30c94aecd 100644 --- a/Cubical/ZCohomology/KcompPrelims.agda +++ b/Cubical/ZCohomology/KcompPrelims.agda @@ -292,3 +292,4 @@ Iso.fun (Iso3-Kn-ΩKn+1 n) = Kn→ΩKn+1 n Iso.inv (Iso3-Kn-ΩKn+1 n) = absInv n Iso.rightInv (Iso3-Kn-ΩKn+1 n) = absSect n Iso.leftInv (Iso3-Kn-ΩKn+1 n) = absRetr n + diff --git a/Cubical/ZCohomology/MayerVietorisUnreduced.agda b/Cubical/ZCohomology/MayerVietorisUnreduced.agda index 2842d5ef4..b9548966a 100644 --- a/Cubical/ZCohomology/MayerVietorisUnreduced.agda +++ b/Cubical/ZCohomology/MayerVietorisUnreduced.agda @@ -129,11 +129,11 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : λ δ b → (λ i → sElim (λ _ → isOfHLevelProd 2 setTruncIsSet setTruncIsSet) (λ δ → ∣ (λ x → δ (inl x)) ∣₂ , ∣ (λ x → δ (inr x)) ∣₂ ) (b (~ i)))) - abstract - Ker-i⊂Im-d : (n : ℕ) (x : Group.type (coHomGr (suc n) (Pushout f g))) + + Ker-i⊂Im-d : (n : ℕ) (x : Group.type (coHomGr (suc n) (Pushout f g))) → isInKer (coHomGr (suc n) (Pushout f g)) (×coHomGr (suc n) A B) (i (suc n)) x → isInIm (coHomGr n C) (coHomGr (suc n) (Pushout f g)) (d n) x - Ker-i⊂Im-d n = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) + Ker-i⊂Im-d n = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) λ a p → pRec {A = (λ x → a (inl x)) ≡ λ _ → 0ₖ} (isOfHLevelΠ 1 (λ _ → propTruncIsProp)) (λ p1 → pRec propTruncIsProp λ p2 → ∣ ∣ (λ c → ΩKn+1→Kn (sym (cong (λ F → F (f c)) p1) ∙∙ cong a (push c) @@ -178,76 +178,76 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : ; (j = i1) → F (push a i)}) (pushFiller (suc n) F p1 p2 a j i) - abstract - Im-i⊂Ker-Δ : (n : ℕ) (x : Group.type (×coHomGr n A B)) - → isInIm (coHomGr n (Pushout f g)) (×coHomGr n A B) (i n) x - → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) x - Im-i⊂Ker-Δ n (Fa , Fb) = - sElim {B = λ Fa → (Fb : _) → isInIm (coHomGr n (Pushout f g)) (×coHomGr n A B) (i n) (Fa , Fb) - → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) (Fa , Fb)} - (λ _ → isOfHLevelΠ 2 λ _ → (isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _)) - (λ Fa → sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) - λ Fb → pRec (setTruncIsSet _ _) - (sigmaElim (λ x → isOfHLevelSuc 1 (setTruncIsSet _ _)) - λ Fd p → helper n Fa Fb Fd p)) - Fa - Fb - where - helper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) (Fd : (Pushout f g) → coHomK n) - → (morph.fun (i n) ∣ Fd ∣₂ ≡ (∣ Fa ∣₂ , ∣ Fb ∣₂)) - → (morph.fun (Δ n)) (∣ Fa ∣₂ , ∣ Fb ∣₂) ≡ 0ₕ - helper zero Fa Fb Fd p = cong (morph.fun (Δ zero)) (sym p) - ∙∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₂ +ₕ -ₕ ∣ (λ x → Fd (push x (~ i))) ∣₂) - ∙∙ rCancelₕ ∣ (λ x → Fd (inl (f x))) ∣₂ - helper (suc n) Fa Fb Fd p = cong (morph.fun (Δ (suc n))) (sym p) - ∙∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₂ +ₕ -ₕ ∣ (λ x → Fd (push x (~ i))) ∣₂) - ∙∙ rCancelₕ ∣ (λ x → Fd (inl (f x))) ∣₂ - - - Ker-Δ⊂Im-i : (n : ℕ) (a : Group.type (×coHomGr n A B)) - → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) a - → isInIm (coHomGr n (Pushout f g)) (×coHomGr n A B) (i n) a - Ker-Δ⊂Im-i n (Fa , Fb) = - sElim {B = λ Fa → (Fb : _) → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) (Fa , Fb) - → isInIm (coHomGr n (Pushout f g)) (×coHomGr n A B) (i n) (Fa , Fb)} - (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) - (λ Fa → sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) - λ Fb p → pRec propTruncIsProp - (λ q → ∣ ∣ helpFun n Fa Fb (funExt⁻ q) ∣₂ - , anotherHelper n Fa Fb q ∣₋₁) - (helper n Fa Fb p)) - Fa - Fb + abstract + Im-i⊂Ker-Δ : (n : ℕ) (x : Group.type (×coHomGr n A B)) + → isInIm (coHomGr n (Pushout f g)) (×coHomGr n A B) (i n) x + → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) x + Im-i⊂Ker-Δ n (Fa , Fb) = + sElim {B = λ Fa → (Fb : _) → isInIm (coHomGr n (Pushout f g)) (×coHomGr n A B) (i n) (Fa , Fb) + → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) (Fa , Fb)} + (λ _ → isOfHLevelΠ 2 λ _ → (isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _)) + (λ Fa → sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ Fb → pRec (setTruncIsSet _ _) + (sigmaElim (λ x → isOfHLevelSuc 1 (setTruncIsSet _ _)) + λ Fd p → helper n Fa Fb Fd p)) + Fa + Fb + where + helper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) (Fd : (Pushout f g) → coHomK n) + → (morph.fun (i n) ∣ Fd ∣₂ ≡ (∣ Fa ∣₂ , ∣ Fb ∣₂)) + → (morph.fun (Δ n)) (∣ Fa ∣₂ , ∣ Fb ∣₂) ≡ 0ₕ + helper zero Fa Fb Fd p = cong (morph.fun (Δ zero)) (sym p) + ∙∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₂ +ₕ -ₕ ∣ (λ x → Fd (push x (~ i))) ∣₂) + ∙∙ rCancelₕ ∣ (λ x → Fd (inl (f x))) ∣₂ + helper (suc n) Fa Fb Fd p = cong (morph.fun (Δ (suc n))) (sym p) + ∙∙ (λ i → ∣ (λ x → Fd (inl (f x))) ∣₂ +ₕ -ₕ ∣ (λ x → Fd (push x (~ i))) ∣₂) + ∙∙ rCancelₕ ∣ (λ x → Fd (inl (f x))) ∣₂ + + + Ker-Δ⊂Im-i : (n : ℕ) (a : Group.type (×coHomGr n A B)) + → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) a + → isInIm (coHomGr n (Pushout f g)) (×coHomGr n A B) (i n) a + Ker-Δ⊂Im-i n (Fa , Fb) = + sElim {B = λ Fa → (Fb : _) → isInKer (×coHomGr n A B) (coHomGr n C) (Δ n) (Fa , Fb) + → isInIm (coHomGr n (Pushout f g)) (×coHomGr n A B) (i n) (Fa , Fb)} + (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) + (λ Fa → sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) + λ Fb p → pRec propTruncIsProp + (λ q → ∣ ∣ helpFun n Fa Fb (funExt⁻ q) ∣₂ + , anotherHelper n Fa Fb q ∣₋₁) + (helper n Fa Fb p)) + Fa + Fb - where - helper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) - → (morph.fun (Δ n)) (∣ Fa ∣₂ , ∣ Fb ∣₂) ≡ 0ₕ - → ∥ (Path (_ → _) (λ c → Fa (f c)) (λ c → Fb (g c))) ∥₋₁ - helper zero Fa Fb p = Iso.fun (PathIdTrunc₀Iso) - ((sym (rUnitₕ (coHomFun zero f ∣ Fa ∣₂)) - ∙∙ (λ i → coHomFun zero f ∣ Fa ∣₂ +ₕ (lCancelₕ (coHomFun zero g ∣ Fb ∣₂) (~ i))) - ∙∙ sym (assocₕ (coHomFun zero f ∣ Fa ∣₂) (-ₕ (coHomFun zero g ∣ Fb ∣₂)) (coHomFun zero g ∣ Fb ∣₂))) - ∙∙ cong (λ x → x +ₕ (coHomFun zero g ∣ Fb ∣₂)) p - ∙∙ lUnitₕ (coHomFun zero g ∣ Fb ∣₂)) - helper (suc n) Fa Fb p = Iso.fun (PathIdTrunc₀Iso) - ((sym (rUnitₕ (coHomFun (suc n) f ∣ Fa ∣₂)) - ∙∙ (λ i → coHomFun (suc n) f ∣ Fa ∣₂ +ₕ (lCancelₕ (coHomFun (suc n) g ∣ Fb ∣₂) (~ i))) - ∙∙ sym (assocₕ (coHomFun (suc n) f ∣ Fa ∣₂) (-ₕ (coHomFun (suc n) g ∣ Fb ∣₂)) (coHomFun (suc n) g ∣ Fb ∣₂))) - ∙∙ cong (λ x → x +ₕ (coHomFun (suc n) g ∣ Fb ∣₂)) p - ∙∙ lUnitₕ (coHomFun (suc n) g ∣ Fb ∣₂)) - - helpFun : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) - → ((c : C) → Fa (f c) ≡ Fb (g c)) - → (Pushout f g) → coHomK n - helpFun n Fa Fb p (inl x) = Fa x - helpFun n Fa Fb p (inr x) = Fb x - helpFun n Fa Fb p (push a i) = p a i - - anotherHelper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) - → (q : Path (C → coHomK n) (λ c → Fa (f c)) (λ c → Fb (g c))) - → morph.fun (i n) ∣ helpFun n Fa Fb (λ x i₁ → q i₁ x) ∣₂ ≡ (∣ Fa ∣₂ , ∣ Fb ∣₂) - anotherHelper zero Fa Fb q = refl - anotherHelper (suc n) Fa Fb q = refl + where + helper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) + → (morph.fun (Δ n)) (∣ Fa ∣₂ , ∣ Fb ∣₂) ≡ 0ₕ + → ∥ (Path (_ → _) (λ c → Fa (f c)) (λ c → Fb (g c))) ∥₋₁ + helper zero Fa Fb p = Iso.fun (PathIdTrunc₀Iso) + ((sym (rUnitₕ (coHomFun zero f ∣ Fa ∣₂)) + ∙∙ (λ i → coHomFun zero f ∣ Fa ∣₂ +ₕ (lCancelₕ (coHomFun zero g ∣ Fb ∣₂) (~ i))) + ∙∙ sym (assocₕ (coHomFun zero f ∣ Fa ∣₂) (-ₕ (coHomFun zero g ∣ Fb ∣₂)) (coHomFun zero g ∣ Fb ∣₂))) + ∙∙ cong (λ x → x +ₕ (coHomFun zero g ∣ Fb ∣₂)) p + ∙∙ lUnitₕ (coHomFun zero g ∣ Fb ∣₂)) + helper (suc n) Fa Fb p = Iso.fun (PathIdTrunc₀Iso) + ((sym (rUnitₕ (coHomFun (suc n) f ∣ Fa ∣₂)) + ∙∙ (λ i → coHomFun (suc n) f ∣ Fa ∣₂ +ₕ (lCancelₕ (coHomFun (suc n) g ∣ Fb ∣₂) (~ i))) + ∙∙ sym (assocₕ (coHomFun (suc n) f ∣ Fa ∣₂) (-ₕ (coHomFun (suc n) g ∣ Fb ∣₂)) (coHomFun (suc n) g ∣ Fb ∣₂))) + ∙∙ cong (λ x → x +ₕ (coHomFun (suc n) g ∣ Fb ∣₂)) p + ∙∙ lUnitₕ (coHomFun (suc n) g ∣ Fb ∣₂)) + + helpFun : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) + → ((c : C) → Fa (f c) ≡ Fb (g c)) + → (Pushout f g) → coHomK n + helpFun n Fa Fb p (inl x) = Fa x + helpFun n Fa Fb p (inr x) = Fb x + helpFun n Fa Fb p (push a i) = p a i + + anotherHelper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) + → (q : Path (C → coHomK n) (λ c → Fa (f c)) (λ c → Fb (g c))) + → morph.fun (i n) ∣ helpFun n Fa Fb (λ x i₁ → q i₁ x) ∣₂ ≡ (∣ Fa ∣₂ , ∣ Fb ∣₂) + anotherHelper zero Fa Fb q = refl + anotherHelper (suc n) Fa Fb q = refl Ker-d⊂Im-Δ : (n : ℕ) (a : coHom n C) From 7d815fb92b1d6a7f62c8457a170047de427374d8 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Fri, 10 Jul 2020 02:08:34 +0200 Subject: [PATCH 50/89] last fixes --- Cubical/Data/Sigma/Properties.agda | 4 +- Cubical/Data/Unit/Properties.agda | 2 + Cubical/Foundations/Prelude.agda | 10 - Cubical/HITs/S1/Properties.agda | 5 - Cubical/Structures/Group/Algebra.agda | 119 +++---- Cubical/Structures/Group/Base.agda | 90 ++--- Cubical/Structures/Group/Morphism.agda | 22 +- .../Structures/Group/MorphismProperties.agda | 80 ++--- Cubical/Structures/Group/no-Eta/Algebra.agda | 188 ---------- Cubical/Structures/Group/no-Eta/Base.agda | 216 ------------ Cubical/Structures/Group/no-Eta/Morphism.agda | 64 ---- .../Group/no-Eta/MorphismProperties.agda | 80 ----- .../Structures/Group/no-Eta/Properties.agda | 66 ---- Cubical/Structures/Group/no-Eta/no-Eta.agda | 8 - Cubical/Structures/Monoid.agda | 7 +- Cubical/Structures/Semigroup.agda | 24 +- Cubical/ZCohomology/Groups/Connected.agda | 36 +- Cubical/ZCohomology/Groups/Prelims.agda | 321 +++++++++++++++++ Cubical/ZCohomology/Groups/Sn.agda | 283 ++++++--------- Cubical/ZCohomology/Groups/Torus.agda | 331 +++++------------- Cubical/ZCohomology/Groups/Unit.agda | 116 ++---- Cubical/ZCohomology/Groups/Wedge.agda | 12 +- .../ZCohomology/Groups/WedgeOfSpheres.agda | 24 +- .../ZCohomology/MayerVietorisUnreduced.agda | 14 +- Cubical/ZCohomology/Properties.agda | 33 +- 25 files changed, 748 insertions(+), 1407 deletions(-) delete mode 100644 Cubical/Structures/Group/no-Eta/Algebra.agda delete mode 100644 Cubical/Structures/Group/no-Eta/Base.agda delete mode 100644 Cubical/Structures/Group/no-Eta/Morphism.agda delete mode 100644 Cubical/Structures/Group/no-Eta/MorphismProperties.agda delete mode 100644 Cubical/Structures/Group/no-Eta/Properties.agda delete mode 100644 Cubical/Structures/Group/no-Eta/no-Eta.agda create mode 100644 Cubical/ZCohomology/Groups/Prelims.agda diff --git a/Cubical/Data/Sigma/Properties.agda b/Cubical/Data/Sigma/Properties.agda index be74cbd2d..42686cccc 100644 --- a/Cubical/Data/Sigma/Properties.agda +++ b/Cubical/Data/Sigma/Properties.agda @@ -280,6 +280,6 @@ Iso.leftInv toProdIso b = funExt λ _ → refl curryIso : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} → Iso (A × B → C) (A → B → C) Iso.fun curryIso f a b = f (a , b) -Iso.inv curryIso f (a , b) = f a b +Iso.inv curryIso f a = f (fst a) (snd a) Iso.rightInv curryIso a = refl -Iso.leftInv curryIso f = funExt λ {(a , b) → refl} +Iso.leftInv curryIso f = funExt λ _ → refl diff --git a/Cubical/Data/Unit/Properties.agda b/Cubical/Data/Unit/Properties.agda index 00283ed40..7f4c7f109 100644 --- a/Cubical/Data/Unit/Properties.agda +++ b/Cubical/Data/Unit/Properties.agda @@ -46,6 +46,8 @@ fibId A = (λ _ → refl) (λ a i → fst a , isOfHLevelSuc 1 isPropUnit _ _ (snd a) refl i)) +isContr→≃Unit : ∀ {ℓ} {A : Type ℓ} → isContr A → A ≃ Unit +isContr→≃Unit contr = isoToEquiv (iso (λ _ → tt) (λ _ → fst contr) (λ _ → refl) λ _ → snd contr _) isContr→≡Unit : {A : Type₀} → isContr A → A ≡ Unit isContr→≡Unit contr = isoToPath (iso (λ _ → tt) (λ _ → fst contr) (λ _ → refl) λ _ → snd contr _) diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index be6f313c7..43b99890a 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -430,16 +430,6 @@ open Lift public liftExt : ∀ {A : Type ℓ} {a b : Lift {ℓ} {ℓ'} A} → (lower a ≡ lower b) → a ≡ b liftExt x i = lift (x i) -{- filler of - q⁻¹ - z — — — > x - ^ ^ - q | | p⁻¹ - | | - x — — — > y - p --} - doubleCompPath-filler∙ : {a b c d : A} (p : a ≡ b) (q : b ≡ c) (r : c ≡ d) → PathP (λ i → p i ≡ r (~ i)) (p ∙ q ∙ r) q doubleCompPath-filler∙ {A = A} {b = b} p q r j i = diff --git a/Cubical/HITs/S1/Properties.agda b/Cubical/HITs/S1/Properties.agda index 156ef7f86..13a8987ca 100644 --- a/Cubical/HITs/S1/Properties.agda +++ b/Cubical/HITs/S1/Properties.agda @@ -10,10 +10,6 @@ open import Cubical.Foundations.Univalence open import Cubical.HITs.S1.Base open import Cubical.HITs.PropositionalTruncation as PropTrunc -open import Cubical.HITs.SetTruncation as s -open import Cubical.Data.Int -open import Cubical.Data.Nat -open import Cubical.Data.Prod isConnectedS¹ : (s : S¹) → ∥ base ≡ s ∥ isConnectedS¹ base = ∣ refl ∣ @@ -29,4 +25,3 @@ isGroupoidS¹ s t = (λ q → subst (λ t → isSet (base ≡ t)) q isSetΩS¹) (isConnectedS¹ t))) (isConnectedS¹ s) - diff --git a/Cubical/Structures/Group/Algebra.agda b/Cubical/Structures/Group/Algebra.agda index 673fca0ec..2d614df0e 100644 --- a/Cubical/Structures/Group/Algebra.agda +++ b/Cubical/Structures/Group/Algebra.agda @@ -2,21 +2,12 @@ module Cubical.Structures.Group.Algebra where open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Isomorphism as I -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.SIP open import Cubical.Foundations.Function using (_∘_) open import Cubical.Foundations.GroupoidLaws -open import Cubical.Functions.Embedding open import Cubical.Data.Sigma - -open import Cubical.Structures.Axioms -open import Cubical.Structures.Pointed -open import Cubical.Structures.Semigroup hiding (⟨_⟩) -open import Cubical.Structures.Monoid hiding (⟨_⟩) +open import Cubical.Data.Unit open import Cubical.Structures.Group.Base open import Cubical.Structures.Group.Properties @@ -35,11 +26,14 @@ private variable ℓ ℓ' ℓ'' : Level +------- elementary properties of morphisms -------- - +-- (- 0) = 0 -0≡0 : ∀ {ℓ} {G : Group {ℓ}} → (- G) (0g G) ≡ (0g G) -- - 0 ≡ 0 -0≡0 {G = G} = sym (IsGroup.lid (isGroup G) _) ∙ fst (IsGroup.inverse (isGroup G) _) + +-- ϕ(0) ≡ 0 morph0→0 : ∀ {ℓ ℓ'} (G : Group {ℓ}) (H : Group {ℓ'}) (f : GroupHom G H) → fun f (0g G) ≡ 0g H morph0→0 G H f = @@ -54,6 +48,7 @@ morph0→0 G H f = module H = Group H f' = fun f +-- ϕ(- x) = - ϕ(x) morphMinus : ∀ {ℓ ℓ'} (G : Group {ℓ}) (H : Group {ℓ'}) → (ϕ : GroupHom G H) → (g : ⟨ G ⟩) → fun ϕ ((- G) g) ≡ (- H) (fun ϕ g) morphMinus G H ϕ g = @@ -70,8 +65,7 @@ morphMinus G H ϕ g = helper : (f (G.- g) H.+ f g) ≡ 0g H helper = sym (isHom ϕ (G.- g) g) ∙∙ cong f (invl G g) ∙∙ morph0→0 G H ϕ --- ----------- Equivalent notions of isomorphisms -------------- - +-- ----------- Alternative notions of isomorphisms -------------- record GroupIso {ℓ ℓ'} (G : Group {ℓ}) (H : Group {ℓ'}) : Type (ℓ-max ℓ ℓ') where no-eta-equality constructor iso @@ -83,7 +77,36 @@ record GroupIso {ℓ ℓ'} (G : Group {ℓ}) (H : Group {ℓ'}) : Type (ℓ-max infixr 35 _◆_ +record BijectionIso {ℓ ℓ'} (A : Group {ℓ}) (B : Group {ℓ'}) : Type (ℓ-max ℓ ℓ') where + constructor bij-iso + field + map' : GroupHom A B + inj : isInjective A B map' + surj : isSurjective A B map' + +-- "Very" short exact sequences +-- i.e. an exact sequence A → B → C → D where A and D are trivial +record vSES {ℓ ℓ' ℓ'' ℓ'''} (A : Group {ℓ}) (B : Group {ℓ'}) (leftGr : Group {ℓ''}) (rightGr : Group {ℓ'''}) + : Type (ℓ-suc (ℓ-max ℓ (ℓ-max ℓ' (ℓ-max ℓ'' ℓ''')))) where + constructor ses + field + isTrivialLeft : isProp ⟨ leftGr ⟩ + isTrivialRight : isProp ⟨ rightGr ⟩ + + left : GroupHom leftGr A + right : GroupHom B rightGr + ϕ : GroupHom A B + + Ker-ϕ⊂Im-left : (x : ⟨ A ⟩) + → isInKer A B ϕ x + → isInIm leftGr A left x + Ker-right⊂Im-ϕ : (x : ⟨ B ⟩) + → isInKer B rightGr right x + → isInIm A B ϕ x + +open BijectionIso open GroupIso +open vSES compGroupIso : ∀ {ℓ''} {G : Group {ℓ}} {H : Group {ℓ'}} {A : Group {ℓ''}} → GroupIso G H → GroupIso H A → GroupIso G A map (compGroupIso iso1 iso2) = compGroupHom (map iso1) (map iso2) @@ -118,8 +141,8 @@ _◆_ = compGroupIso dirProdGroupIso : ∀ {ℓ'' ℓ'''} {G : Group {ℓ}} {H : Group {ℓ'}} {A : Group {ℓ''}} {B : Group {ℓ'''}} → GroupIso G H → GroupIso A B → GroupIso (dirProd G A) (dirProd H B) -map (dirProdGroupIso iso1 iso2) = grouphom (λ prod → fun (map iso1) (fst prod) , fun (map iso2) (snd prod)) - λ a b → ΣPathP (isHom (map iso1) (fst a) (fst b) , isHom (map iso2) (snd a) (snd b)) +fun (map (dirProdGroupIso iso1 iso2)) prod = fun (map iso1) (fst prod) , fun (map iso2) (snd prod) +isHom (map (dirProdGroupIso iso1 iso2)) a b = ΣPathP (isHom (map iso1) (fst a) (fst b) , isHom (map iso2) (snd a) (snd b)) inv (dirProdGroupIso iso1 iso2) prod = (inv iso1) (fst prod) , (inv iso2) (snd prod) rightInv (dirProdGroupIso iso1 iso2) a = ΣPathP (rightInv iso1 (fst a) , (rightInv iso2 (snd a))) leftInv (dirProdGroupIso iso1 iso2) a = ΣPathP (leftInv iso1 (fst a) , (leftInv iso2 (snd a))) @@ -128,38 +151,9 @@ GrIsoToGrEquiv : {G : Group {ℓ}} {H : Group {ℓ'}} → GroupIso G H → Group GroupEquiv.eq (GrIsoToGrEquiv i) = isoToEquiv (iso (fun (map i)) (inv i) (rightInv i) (leftInv i)) GroupEquiv.isHom (GrIsoToGrEquiv i) = isHom (map i) -record Iso' {ℓ ℓ'} (A : Group {ℓ}) (B : Group {ℓ'}) : Type (ℓ-max ℓ ℓ') where - constructor iso' - field - map' : GroupHom A B - inj : isInjective A B map' - surj : isSurjective A B map' - -open Iso' - --- "Very" short exact sequences -record vSES {ℓ ℓ' ℓ'' ℓ'''} (A : Group {ℓ}) (B : Group {ℓ'}) (leftGr : Group {ℓ''}) (rightGr : Group {ℓ'''}) - : Type (ℓ-suc (ℓ-max ℓ (ℓ-max ℓ' (ℓ-max ℓ'' ℓ''')))) where - constructor ses - field - isTrivialLeft : isProp ⟨ leftGr ⟩ - isTrivialRight : isProp ⟨ rightGr ⟩ - - left : GroupHom leftGr A - right : GroupHom B rightGr - ϕ : GroupHom A B - - Ker-ϕ⊂Im-left : (x : ⟨ A ⟩) - → isInKer A B ϕ x - → isInIm leftGr A left x - Ker-right⊂Im-ϕ : (x : ⟨ B ⟩) - → isInKer B rightGr right x - → isInIm A B ϕ x - ---- Proofs that different notions of ismomorphisms agree --- - -Iso'ToGroupIso : {A : Group {ℓ}} {B : Group {ℓ'}} → Iso' A B → GroupIso A B -Iso'ToGroupIso {A = A} {B = B} i = grIso +--- Proofs that BijectionIso and vSES both induce isomorphisms --- +BijectionIsoToGroupIso : {A : Group {ℓ}} {B : Group {ℓ'}} → BijectionIso A B → GroupIso A B +BijectionIsoToGroupIso {A = A} {B = B} i = grIso where module A = Group A module B = Group B @@ -189,36 +183,29 @@ Iso'ToGroupIso {A = A} {B = B} i = grIso rightInv grIso b = (rec (helper b) (λ a → a) (surj i b)) .snd leftInv grIso b j = rec (helper (f b)) (λ a → a) (propTruncIsProp (surj i (f b)) ∣ b , refl ∣ j) .fst -Iso'ToGroupEquiv : {A : Group {ℓ}} {B : Group {ℓ'}} → Iso' A B → GroupEquiv A B -Iso'ToGroupEquiv i = GrIsoToGrEquiv (Iso'ToGroupIso i) +BijectionIsoToGroupEquiv : {A : Group {ℓ}} {B : Group {ℓ'}} → BijectionIso A B → GroupEquiv A B +BijectionIsoToGroupEquiv i = GrIsoToGrEquiv (BijectionIsoToGroupIso i) vSES→GroupIso : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group {ℓ}} {B : Group {ℓ'}} (leftGr : Group {ℓ''}) (rightGr : Group {ℓ'''}) → vSES A B leftGr rightGr → GroupIso A B -vSES→GroupIso {A = A} lGr rGr vses = Iso'ToGroupIso theIso +vSES→GroupIso {A = A} lGr rGr vses = BijectionIsoToGroupIso theIso where - theIso : Iso' _ _ + theIso : BijectionIso _ _ map' theIso = vSES.ϕ vses inj theIso a inker = rec (isSetCarrier A _ _) (λ (a , p) → sym p - ∙∙ cong (fun (vSES.left vses)) (vSES.isTrivialLeft vses a _) - ∙∙ morph0→0 lGr A (vSES.left vses)) - (vSES.Ker-ϕ⊂Im-left vses a inker) - surj theIso a = vSES.Ker-right⊂Im-ϕ vses a (vSES.isTrivialRight vses _ _) + ∙∙ cong (fun (left vses)) (isTrivialLeft vses a _) + ∙∙ morph0→0 lGr A (left vses)) + (Ker-ϕ⊂Im-left vses a inker) + surj theIso a = Ker-right⊂Im-ϕ vses a (isTrivialRight vses _ _) vSES→GroupEquiv : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group {ℓ}} {B : Group {ℓ'}} (leftGr : Group {ℓ''}) (rightGr : Group {ℓ'''}) → vSES A B leftGr rightGr → GroupEquiv A B vSES→GroupEquiv {A = A} lGr rGr vses = GrIsoToGrEquiv (vSES→GroupIso lGr rGr vses) -dirProdEquiv : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group {ℓ}} {B : Group {ℓ'}} {C : Group {ℓ''}} {D : Group {ℓ'''}} - → GroupEquiv A C → GroupEquiv B D - → GroupEquiv (dirProd A B) (dirProd C D) -GroupEquiv.eq (dirProdEquiv eq1 eq2) = ≃-× (GroupEquiv.eq eq1) (GroupEquiv.eq eq2) -GroupEquiv.isHom (dirProdEquiv eq1 eq2) = isHom (×hom (GroupEquiv.hom eq1) (GroupEquiv.hom eq2)) - - ---- +-- The trivial group is a unit. open import Cubical.Data.Unit lUnitGroupIso : ∀ {ℓ} {G : Group {ℓ}} → GroupEquiv (dirProd trivialGroup G) G lUnitGroupIso = @@ -236,3 +223,7 @@ rUnitGroupIso = (λ g → g , tt) (λ _ → refl) λ _ → refl) + +contrGroup≅trivialGroup : {G : Group {ℓ}} → isContr ⟨ G ⟩ → GroupEquiv G trivialGroup +GroupEquiv.eq (contrGroup≅trivialGroup contr) = isContr→≃Unit contr +GroupEquiv.isHom (contrGroup≅trivialGroup contr) _ _ = refl diff --git a/Cubical/Structures/Group/Base.agda b/Cubical/Structures/Group/Base.agda index 315179b40..da39eb4f9 100644 --- a/Cubical/Structures/Group/Base.agda +++ b/Cubical/Structures/Group/Base.agda @@ -33,11 +33,11 @@ record IsGroup {G : Type ℓ} invr : (x : G) → x + (- x) ≡ 0g invr x = inverse x .fst -η-isGroup : {G : Type ℓ} {0g 0g' : G} {_+_ _+'_ : G → G → G} { -_ -'_ : G → G} - → 0g ≡ 0g' - → _+_ ≡ _+'_ - → -_ ≡ -'_ - → IsGroup 0g _+_ -_ ≡ IsGroup 0g' _+'_ -'_ +η-isGroup : {G : Type ℓ} {0g 0g' : G} {_+_ _+'_ : G → G → G} { -_ -'_ : G → G} + → 0g ≡ 0g' + → _+_ ≡ _+'_ + → -_ ≡ -'_ + → IsGroup 0g _+_ -_ ≡ IsGroup 0g' _+'_ -'_ η-isGroup id1 id2 id3 i = IsGroup (id1 i) (id2 i) (id3 i) record Group : Type (ℓ-suc ℓ) where @@ -72,8 +72,8 @@ makeIsGroup : {G : Type ℓ} {0g : G} {_+_ : G → G → G} { -_ : G → G} (rinv : (x : G) → x + (- x) ≡ 0g) (linv : (x : G) → (- x) + x ≡ 0g) → IsGroup 0g _+_ -_ -makeIsGroup is-setG assoc rid lid rinv linv = - isgroup (makeIsMonoid is-setG assoc rid lid) λ x → rinv x , linv x +IsGroup.isMonoid (makeIsGroup is-setG assoc rid lid rinv linv) = makeIsMonoid is-setG assoc rid lid +IsGroup.inverse (makeIsGroup is-setG assoc rid lid rinv linv) = λ x → rinv x , linv x makeGroup : {G : Type ℓ} (0g : G) (_+_ : G → G → G) (-_ : G → G) (is-setG : isSet G) @@ -83,8 +83,11 @@ makeGroup : {G : Type ℓ} (0g : G) (_+_ : G → G → G) (-_ : G → G) (rinv : (x : G) → x + (- x) ≡ 0g) (linv : (x : G) → (- x) + x ≡ 0g) → Group -makeGroup 0g _+_ -_ is-setG assoc rid lid rinv linv = - group _ 0g _+_ -_ (makeIsGroup is-setG assoc rid lid rinv linv) +Group.Carrier (makeGroup 0g _+_ -_ is-setG assoc rid lid rinv linv) = _ +Group.0g (makeGroup 0g _+_ -_ is-setG assoc rid lid rinv linv) = 0g +Group._+_ (makeGroup 0g _+_ -_ is-setG assoc rid lid rinv linv) = _+_ +Group.- makeGroup 0g _+_ -_ is-setG assoc rid lid rinv linv = -_ +Group.isGroup (makeGroup 0g _+_ -_ is-setG assoc rid lid rinv linv) = makeIsGroup is-setG assoc rid lid rinv linv makeGroup-right : ∀ {ℓ} {A : Type ℓ} → (id : A) @@ -199,68 +202,43 @@ isSetCarrier : ∀ {ℓ} → (G : Group {ℓ}) → isSet ⟨ G ⟩ isSetCarrier G = IsSemigroup.is-set (IsMonoid.isSemigroup (isMonoid (isGroup G))) open import Cubical.Foundations.HLevels - +open IsMonoid +open IsSemigroup dirProd : ∀ {ℓ ℓ'} → Group {ℓ} → Group {ℓ'} → Group Carrier (dirProd A B) = Carrier A × Carrier B 0g (dirProd A B) = (0g A) , (0g B) _+_ (dirProd A B) a b = (_+_ A (fst a) (fst b)) , _+_ B (snd a) (snd b) -_ (dirProd A B) a = (-_ A (fst a)) , (-_ B (snd a)) -IsSemigroup.is-set (IsMonoid.isSemigroup (isMonoid (isGroup (dirProd A B)))) = +is-set (isSemigroup (isMonoid (isGroup (dirProd A B)))) = isOfHLevelΣ 2 (isSetCarrier A) λ _ → isSetCarrier B -IsSemigroup.assoc (IsMonoid.isSemigroup (isMonoid (isGroup (dirProd A B)))) x y z i = +assoc (isSemigroup (isMonoid (isGroup (dirProd A B)))) x y z i = assoc (isGroup A) (fst x) (fst y) (fst z) i , assoc (isGroup B) (snd x) (snd y) (snd z) i -IsMonoid.identity (isMonoid (isGroup (dirProd A B))) x = +identity (isMonoid (isGroup (dirProd A B))) x = (λ i → IsGroup.rid (isGroup A) (fst x) i , IsGroup.rid (isGroup B) (snd x) i) , λ i → IsGroup.lid (isGroup A) (fst x) i , IsGroup.lid (isGroup B) (snd x) i inverse (isGroup (dirProd A B)) x = (λ i → (fst (inverse (isGroup A) (fst x)) i) , (fst (inverse (isGroup B) (snd x)) i)) , λ i → (snd (inverse (isGroup A) (fst x)) i) , (snd (inverse (isGroup B) (snd x)) i) -{- - makeGroup {G = Carrier A × Carrier B} - ((0g A) , (0g B)) - (λ a b → (_+_ A (fst a) (fst b)) , _+_ B (snd a) (snd b)) - (λ a → (-_ A (fst a)) , (-_ B (snd a))) - (isOfHLevelΣ 2 (isSetCarrier A) λ _ → isSetCarrier B) - (λ x y z i → assoc (isGroup A) (fst x) (fst y) (fst z) i , assoc (isGroup B) (snd x) (snd y) (snd z) i) - (λ x i → IsGroup.rid (isGroup A) (fst x) i , IsGroup.rid (isGroup B) (snd x) i) - (λ x i → IsGroup.lid (isGroup A) (fst x) i , IsGroup.lid (isGroup B) (snd x) i) - (λ x i → (fst (inverse (isGroup A) (fst x)) i) , (fst (inverse (isGroup B) (snd x)) i)) - (λ x i → (snd (inverse (isGroup A) (fst x)) i) , (snd (inverse (isGroup B) (snd x)) i)) --} --- makeGroup ((0g A) , (0g B)) --- (λ a b → (_+_ A (proj₁ a) (proj₂ b)) , _+_ B (proj₁ a) (proj₂ b)) --- (λ a → (-_ A (proj₂ a)) , (-_ B (proj₂ a))) --- (isOfHLevelProd 2 (isSetCarrier A) (isSetCarrier B)) --- (λ x y z i → assoc (isGroup A) (proj₁ x) (proj₁ y) (proj₁ z) i , assoc (isGroup B) (proj₂ x) (proj₂ y) (proj₂ z) i) --- (λ x i → IsGroup.rid (isGroup A) (proj₁ x) i , IsGroup.rid (isGroup B) (proj₂ x) i) --- (λ x i → IsGroup.lid (isGroup A) (proj₁ x) i , IsGroup.lid (isGroup B) (proj₂ x) i) --- (λ x i → (fst (inverse (isGroup A) (proj₁ x)) i) , (fst (inverse (isGroup B) (proj₂ x)) i)) --- (λ x i → (snd (inverse (isGroup A) (proj₁ x)) i) , (snd (inverse (isGroup B) (proj₂ x)) i)) open import Cubical.Data.Unit trivialGroup : Group₀ -trivialGroup = - makeGroup {G = Unit} - _ - (λ _ _ → _) - (λ _ → _) - isSetUnit - (λ _ _ _ → refl) - (λ _ → refl) - (λ _ → refl) - (λ _ → refl) - λ _ → refl +Carrier trivialGroup = Unit +0g trivialGroup = _ +_+_ trivialGroup _ _ = _ +-_ trivialGroup _ = _ +is-set (isSemigroup (isMonoid (isGroup trivialGroup))) = isSetUnit +assoc (isSemigroup (isMonoid (isGroup trivialGroup))) _ _ _ = refl +identity (isMonoid (isGroup trivialGroup)) _ = refl , refl +inverse (isGroup trivialGroup) _ = refl , refl + open import Cubical.Data.Int renaming (_+_ to _+Int_ ; _-_ to _-Int_) intGroup : Group₀ -intGroup = - makeGroup {G = Int} - 0 - _+Int_ - (pos 0 -Int_) - isSetInt - +-assoc - (λ _ → refl) - (λ x → +-comm (pos 0) x) - (λ x → +-comm x (pos 0 -Int x) ∙ minusPlus x 0) - λ x → minusPlus x 0 +Carrier intGroup = Int +0g intGroup = 0 +_+_ intGroup = _+Int_ +- intGroup = 0 -Int_ +is-set (isSemigroup (isMonoid (isGroup intGroup))) = isSetInt +assoc (isSemigroup (isMonoid (isGroup intGroup))) = +-assoc +identity (isMonoid (isGroup intGroup)) x = refl , (+-comm (pos 0) x) +inverse (isGroup intGroup) x = +-comm x (pos 0 -Int x) ∙ minusPlus x 0 , (minusPlus x 0) diff --git a/Cubical/Structures/Group/Morphism.agda b/Cubical/Structures/Group/Morphism.agda index 02b128ad5..b3ac1811a 100644 --- a/Cubical/Structures/Group/Morphism.agda +++ b/Cubical/Structures/Group/Morphism.agda @@ -38,22 +38,23 @@ record GroupEquiv (G : Group {ℓ}) (H : Group {ℓ'}) : Type (ℓ-max ℓ ℓ') hom : GroupHom G H hom = grouphom (equivFun eq) isHom +open GroupHom +open GroupEquiv + η-hom : {G : Group {ℓ}} {H : Group {ℓ'}} → (a : GroupHom G H) - → grouphom (GroupHom.fun a) (GroupHom.isHom a) ≡ a -GroupHom.fun (η-hom a i) = GroupHom.fun a -GroupHom.isHom (η-hom a i) = GroupHom.isHom a + → grouphom (fun a) (isHom a) ≡ a +fun (η-hom a i) = fun a +isHom (η-hom a i) = isHom a η-equiv : {G : Group {ℓ}} {H : Group {ℓ'}} → (a : GroupEquiv G H) - → groupequiv (GroupEquiv.eq a) (GroupEquiv.isHom a) ≡ a -GroupEquiv.eq (η-equiv a i) = GroupEquiv.eq a -GroupEquiv.isHom (η-equiv a i) = GroupEquiv.isHom a - -open GroupHom + → groupequiv (eq a) (isHom a) ≡ a +eq (η-equiv a i) = eq a +isHom (η-equiv a i) = isHom a ×hom : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group {ℓ}} {B : Group {ℓ'}} {C : Group {ℓ''}} {D : Group {ℓ'''}} → GroupHom A C → GroupHom B D → GroupHom (dirProd A B) (dirProd C D) -GroupHom.fun (×hom mf1 mf2) = map-× (fun mf1) (fun mf2) -GroupHom.isHom (×hom mf1 mf2) a b = ≡-× (isHom mf1 _ _) (isHom mf2 _ _) +fun (×hom mf1 mf2) = map-× (fun mf1) (fun mf2) +isHom (×hom mf1 mf2) a b = ≡-× (isHom mf1 _ _) (isHom mf2 _ _) open Group @@ -70,4 +71,3 @@ isSurjective G H ϕ = (x : ⟨ H ⟩) → isInIm G H ϕ x isInjective : ∀ {ℓ ℓ'} (G : Group {ℓ}) (H : Group {ℓ'}) → GroupHom G H → Type (ℓ-max ℓ ℓ') isInjective G H ϕ = (x : ⟨ G ⟩) → isInKer G H ϕ x → x ≡ 0g G - diff --git a/Cubical/Structures/Group/MorphismProperties.agda b/Cubical/Structures/Group/MorphismProperties.agda index 75f1a8403..69e283727 100644 --- a/Cubical/Structures/Group/MorphismProperties.agda +++ b/Cubical/Structures/Group/MorphismProperties.agda @@ -30,21 +30,23 @@ private infixr 31 _□_ +open Group +open GroupHom +open IsMonoid +open IsSemigroup +open IsGroup isPropIsGroupHom : (G : Group {ℓ}) (H : Group {ℓ'}) {f : ⟨ G ⟩ → ⟨ H ⟩} → isProp (isGroupHom G H f) isPropIsGroupHom G H {f} = isPropΠ2 λ a b → Group.is-set H _ _ isSetGroupHom : {G : Group {ℓ}} {H : Group {ℓ'}} → isSet (GroupHom G H) isSetGroupHom {G = G} {H = H} = isOfHLevelRespectEquiv 2 equiv (isSetΣ (isSetΠ λ _ → is-set H) λ _ → isProp→isSet (isPropIsGroupHom G H)) where - open Group - open GroupHom equiv : (Σ[ g ∈ (Carrier G → Carrier H) ] (isGroupHom G H g)) ≃ (GroupHom G H) equiv = isoToEquiv (iso (λ (g , m) → grouphom g m) (λ hom → fun hom , isHom hom) (λ a → η-hom _) λ _ → refl) -- Morphism composition -open GroupHom isGroupHomComp : {F : Group {ℓ}} {G : Group {ℓ'}} {H : Group {ℓ''}} → - (f : GroupHom F G) → (g : GroupHom G H) → isGroupHom F H (GroupHom.fun g ∘ GroupHom.fun f) + (f : GroupHom F G) → (g : GroupHom G H) → isGroupHom F H (fun g ∘ fun f) isGroupHomComp f g x y = cong (fun g) (isHom f _ _) ∙ isHom g _ _ @@ -54,8 +56,8 @@ isHom (compGroupHom f g) = isGroupHomComp f g open GroupEquiv compGroupEquiv : {F : Group {ℓ}} {G : Group {ℓ'}} {H : Group {ℓ''}} → GroupEquiv F G → GroupEquiv G H → GroupEquiv F H -GroupEquiv.eq (compGroupEquiv f g) = compEquiv (eq f) (eq g) -GroupEquiv.isHom (compGroupEquiv f g) = isHom (compGroupHom (hom f) (hom g)) +eq (compGroupEquiv f g) = compEquiv (eq f) (eq g) +isHom (compGroupEquiv f g) = isHom (compGroupHom (hom f) (hom g)) _□_ : _ _□_ = compGroupEquiv @@ -74,8 +76,8 @@ isGroupHomInv {G = G} {H = H} f h h' = isInj-f _ _ ( f' (g h ⋆¹ g h') ∎) where f' = fst (eq f) - _⋆¹_ = Group._+_ G - _⋆²_ = Group._+_ H + _⋆¹_ = _+_ G + _⋆²_ = _+_ H g = invEq (eq f) isInj-f : (x y : ⟨ G ⟩) → f' x ≡ f' y → x ≡ y @@ -85,8 +87,13 @@ invGroupEquiv : {G : Group {ℓ}} {H : Group {ℓ'}} → GroupEquiv G H → Grou eq (invGroupEquiv f) = invEquiv (eq f) isHom (invGroupEquiv f) = isGroupHomInv f +dirProdEquiv : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group {ℓ}} {B : Group {ℓ'}} {C : Group {ℓ''}} {D : Group {ℓ'''}} + → GroupEquiv A C → GroupEquiv B D + → GroupEquiv (dirProd A B) (dirProd C D) +eq (dirProdEquiv eq1 eq2) = ≃-× (eq eq1) (eq eq2) +isHom (dirProdEquiv eq1 eq2) = isHom (×hom (GroupEquiv.hom eq1) (GroupEquiv.hom eq2)) -groupHomEq : {G : Group {ℓ}} {H : Group {ℓ'}} {f g : GroupHom G H} → (GroupHom.fun f ≡ GroupHom.fun g) → f ≡ g +groupHomEq : {G : Group {ℓ}} {H : Group {ℓ'}} {f g : GroupHom G H} → (fun f ≡ fun g) → f ≡ g fun (groupHomEq p i) = p i isHom (groupHomEq {G = G} {H = H} {f = f} {g = g} p i) = p-hom i where @@ -94,7 +101,7 @@ isHom (groupHomEq {G = G} {H = H} {f = f} {g = g} p i) = p-hom i p-hom = toPathP (isPropIsGroupHom G H _ _) -groupEquivEq : {G : Group {ℓ}} {H : Group {ℓ'}} {f g : GroupEquiv G H} → (GroupEquiv.eq f ≡ GroupEquiv.eq g) → f ≡ g +groupEquivEq : {G : Group {ℓ}} {H : Group {ℓ'}} {f g : GroupEquiv G H} → (eq f ≡ eq g) → f ≡ g eq (groupEquivEq {G = G} {H = H} {f} {g} p i) = p i isHom (groupEquivEq {G = G} {H = H} {f} {g} p i) = p-hom i where @@ -134,7 +141,6 @@ module GroupΣTheory {ℓ} where GroupEquivStr = AxiomsEquivStr RawGroupEquivStr GroupAxioms open MonoidTheory - isPropGroupAxioms : (G : Type ℓ) → (s : RawGroupStructure G) → isProp (GroupAxioms G s) @@ -144,29 +150,26 @@ module GroupΣTheory {ℓ} where isProp (Σ[ e ∈ G ] ((x : G) → (x + e ≡ x) × (e + x ≡ x)) × ((x : G) → Σ[ x' ∈ G ] (x + x' ≡ e) × (x' + x ≡ e))) γ h (e , P , _) (e' , Q , _) = - Σ≡Prop (λ x → isPropΣ (isPropΠ λ _ → isProp× (isSetG _ _) (isSetG _ _)) (β x)) + Σ≡Prop (λ x → isPropΣ (isPropΠ λ _ → isProp× ((is-set h) _ _) ((is-set h) _ _)) (β x)) (sym (fst (Q e)) ∙ snd (P e')) where - isSetG : isSet G - isSetG = IsSemigroup.is-set h - β : (e : G) → ((x : G) → (x + e ≡ x) × (e + x ≡ x)) → isProp ((x : G) → Σ[ x' ∈ G ] (x + x' ≡ e) × (x' + x ≡ e)) β e He = isPropΠ λ { x (x' , _ , P) (x'' , Q , _) → - Σ≡Prop (λ _ → isProp× (isSetG _ _) (isSetG _ _)) + Σ≡Prop (λ _ → isProp× ((is-set h) _ _) ((is-set h) _ _)) (inv-lemma ℳ x x' x'' P Q) } where ℳ : Monoid - ℳ = makeMonoid e _+_ isSetG (IsSemigroup.assoc h) (λ x → He x .fst) (λ x → He x .snd) + ℳ = makeMonoid e _+_ (is-set h) (IsSemigroup.assoc h) (λ x → He x .fst) (λ x → He x .snd) Group→GroupΣ : Group → GroupΣ Group→GroupΣ G = _ , _ , - ((IsMonoid.isSemigroup (IsGroup.isMonoid (Group.isGroup G))) + ((isSemigroup (isMonoid (isGroup G))) , _ - , (IsMonoid.identity (IsGroup.isMonoid (Group.isGroup G))) - , λ x → Group.-_ G x - , IsGroup.inverse (Group.isGroup G) x) + , (identity (isMonoid (isGroup G))) + , λ x → (- G) x + , inverse (isGroup G) x) GroupΣ→Group : GroupΣ → Group GroupΣ→Group (G , _ , SG , _ , H0g , w ) = @@ -176,16 +179,16 @@ module GroupΣTheory {ℓ} where GroupIsoGroupΣ = iso Group→GroupΣ GroupΣ→Group (λ _ → refl) helper where helper : retract (λ z → Group→GroupΣ z) GroupΣ→Group - Group.Carrier (helper a i) = ⟨ a ⟩ - Group.0g (helper a i) = Group.0g a - Group._+_ (helper a i) = Group._+_ a - Group.- helper a i = Group.- a - IsMonoid.isSemigroup (IsGroup.isMonoid (Group.isGroup (helper a i))) = - IsMonoid.isSemigroup (IsGroup.isMonoid (Group.isGroup a)) - IsMonoid.identity (IsGroup.isMonoid (Group.isGroup (helper a i))) = - IsMonoid.identity (IsGroup.isMonoid (Group.isGroup a)) - IsGroup.inverse (Group.isGroup (helper a i)) = IsGroup.inverse (Group.isGroup a) - + Carrier (helper a i) = ⟨ a ⟩ + 0g (helper a i) = Group.0g a + _+_ (helper a i) = Group._+_ a + - helper a i = Group.- a + isSemigroup (isMonoid (isGroup (helper a i))) = + isSemigroup (isMonoid (isGroup a)) + IsMonoid.identity (isMonoid (isGroup (helper a i))) = + identity (isMonoid (Group.isGroup a)) + inverse (isGroup (helper a i)) = inverse (Group.isGroup a) + groupUnivalentStr : UnivalentStr GroupStructure GroupEquivStr groupUnivalentStr = axiomsUnivalentStr _ isPropGroupAxioms rawGroupUnivalentStr @@ -202,7 +205,7 @@ module GroupΣTheory {ℓ} where rightInv GroupIsoΣPath _ = refl leftInv GroupIsoΣPath _ = η-equiv _ - + GroupPath : (G H : Group) → (GroupEquiv G H) ≃ (G ≡ H) GroupPath G H = GroupEquiv G H ≃⟨ isoToEquiv GroupIsoΣPath ⟩ @@ -213,16 +216,16 @@ module GroupΣTheory {ℓ} where RawGroupΣ : Type (ℓ-suc ℓ) RawGroupΣ = TypeWithStr ℓ RawGroupStructure - open Group + open Group Group→RawGroupΣ : Group → RawGroupΣ Group→RawGroupΣ G = ⟨ G ⟩ , (_+_ G) - InducedGroup : (G : Group) (H : RawGroupΣ) (e : G .Group.Carrier ≃ H .fst) + InducedGroup : (G : Group) (H : RawGroupΣ) (e : G .Carrier ≃ H .fst) → RawGroupEquivStr (Group→RawGroupΣ G) H e → Group InducedGroup G H e r = GroupΣ→Group (transferAxioms rawGroupUnivalentStr (Group→GroupΣ G) H (e , r)) - InducedGroupPath : (G : Group {ℓ}) (H : RawGroupΣ) (e : G .Group.Carrier ≃ H .fst) + InducedGroupPath : (G : Group {ℓ}) (H : RawGroupΣ) (e : G .Carrier ≃ H .fst) (E : RawGroupEquivStr (Group→RawGroupΣ G) H e) → G ≡ InducedGroup G H e E InducedGroupPath G H e E = @@ -233,12 +236,12 @@ module GroupΣTheory {ℓ} where GroupPath : (G H : Group {ℓ}) → (GroupEquiv G H) ≃ (G ≡ H) GroupPath = GroupΣTheory.GroupPath -InducedGroup : (G : Group {ℓ}) (H : GroupΣTheory.RawGroupΣ) (e : G .Group.Carrier ≃ H .fst) +InducedGroup : (G : Group {ℓ}) (H : GroupΣTheory.RawGroupΣ) (e : G .Carrier ≃ H .fst) → GroupΣTheory.RawGroupEquivStr (GroupΣTheory.Group→RawGroupΣ G) H e → Group InducedGroup = GroupΣTheory.InducedGroup -InducedGroupPath : (G : Group {ℓ}) (H : GroupΣTheory.RawGroupΣ) (e : G .Group.Carrier ≃ H .fst) +InducedGroupPath : (G : Group {ℓ}) (H : GroupΣTheory.RawGroupΣ) (e : G .Carrier ≃ H .fst) (E : GroupΣTheory.RawGroupEquivStr (GroupΣTheory.Group→RawGroupΣ G) H e) → G ≡ InducedGroup G H e E InducedGroupPath = GroupΣTheory.InducedGroupPath @@ -247,12 +250,11 @@ InducedGroupPath = GroupΣTheory.InducedGroupPath uaGroup : {G H : Group {ℓ}} → GroupEquiv G H → G ≡ H uaGroup {G = G} {H = H} = equivFun (GroupPath G H) -carac-uaGroup : {G H : Group {ℓ}} (f : GroupEquiv G H) → cong Group.Carrier (uaGroup f) ≡ ua (GroupEquiv.eq f) +carac-uaGroup : {G H : Group {ℓ}} (f : GroupEquiv G H) → cong Carrier (uaGroup f) ≡ ua (GroupEquiv.eq f) carac-uaGroup f = ua (eq f) ∙ refl ≡⟨ sym (rUnit _) ⟩ ua (eq f) ∎ -- Group-ua functoriality -open Group Group≡ : (G H : Group {ℓ}) → ( Σ[ p ∈ Carrier G ≡ Carrier H ] diff --git a/Cubical/Structures/Group/no-Eta/Algebra.agda b/Cubical/Structures/Group/no-Eta/Algebra.agda deleted file mode 100644 index 8b3d51c65..000000000 --- a/Cubical/Structures/Group/no-Eta/Algebra.agda +++ /dev/null @@ -1,188 +0,0 @@ -{-# OPTIONS --cubical --no-import-sorts --safe #-} -module Cubical.Structures.Group.no-Eta.Algebra where - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Isomorphism as I -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Equiv.HalfAdjoint -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.SIP -open import Cubical.Foundations.Function using (_∘_) -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Functions.Embedding -open import Cubical.Data.Sigma - -open import Cubical.Structures.Axioms -open import Cubical.Structures.Pointed -open import Cubical.Structures.Semigroup hiding (⟨_⟩) -open import Cubical.Structures.Monoid hiding (⟨_⟩) - -open import Cubical.Structures.Group.no-Eta.Base -open import Cubical.Structures.Group.no-Eta.Properties -open import Cubical.Structures.Group.no-Eta.Morphism -open import Cubical.Structures.Group.no-Eta.MorphismProperties - -open import Cubical.HITs.PropositionalTruncation hiding (map) - --- open import Cubical.Data.Group.Base - -open Iso -open Group -open GroupHom - -private - variable - ℓ ℓ' ℓ'' : Level - - - --0≡0 : ∀ {ℓ} {G : Group {ℓ}} → (- G) (0g G) ≡ (0g G) -- - 0 ≡ 0 --0≡0 {G = G} = sym (IsGroup.lid (isGroup G) _) ∙ fst (IsGroup.inverse (isGroup G) _) - -morph0→0 : ∀ {ℓ ℓ'} (G : Group {ℓ}) (H : Group {ℓ'}) (f : GroupHom G H) - → fun f (0g G) ≡ 0g H -morph0→0 G H f = - (fun f) (0g G) ≡⟨ sym (IsGroup.rid (isGroup H) _) ⟩ - (f' (0g G) H.+ 0g H) ≡⟨ (λ i → f' (0g G) H.+ invr H (f' (0g G)) (~ i)) ⟩ - (f' (0g G) H.+ (f' (0g G) H.+ (H.- f' (0g G)))) ≡⟨ (Group.assoc H (f' (0g G)) (f' (0g G)) (H.- (f' (0g G)))) ⟩ - ((f' (0g G) H.+ f' (0g G)) H.+ (H.- f' (0g G))) ≡⟨ sym (cong (λ x → x H.+ (H.- f' (0g G))) (sym (cong f' (IsGroup.lid (isGroup G) _)) ∙ isHom f (0g G) (0g G))) ⟩ - (f' (0g G)) H.+ (H.- (f' (0g G))) ≡⟨ invr H (f' (0g G)) ⟩ - 0g H ∎ - where - module G = Group G - module H = Group H - f' = fun f - -morphMinus : ∀ {ℓ ℓ'} (G : Group {ℓ}) (H : Group {ℓ'}) → (ϕ : GroupHom G H) - → (g : ⟨ G ⟩) → fun ϕ ((- G) g) ≡ (- H) (fun ϕ g) -morphMinus G H ϕ g = - f (G.- g) ≡⟨ sym (IsGroup.rid (isGroup H) (f (G.- g))) ⟩ - (f (G.- g) H.+ 0g H) ≡⟨ cong (f (G.- g) H.+_) (sym (invr H (f g))) ⟩ - (f (G.- g) H.+ (f g H.+ (H.- f g))) ≡⟨ Group.assoc H (f (G.- g)) (f g) (H.- f g) ⟩ - ((f (G.- g) H.+ f g) H.+ (H.- f g)) ≡⟨ cong (H._+ (H.- f g)) helper ⟩ - (0g H H.+ (H.- f g)) ≡⟨ IsGroup.lid (isGroup H) (H.- (f g))⟩ - H.- (f g) ∎ - where - module G = Group G - module H = Group H - f = fun ϕ - helper : (f (G.- g) H.+ f g) ≡ 0g H - helper = sym (isHom ϕ (G.- g) g) ∙∙ cong f (invl G g) ∙∙ morph0→0 G H ϕ - --- ----------- Equivalent notions of isomorphisms -------------- - -record GroupIso {ℓ ℓ'} (G : Group {ℓ}) (H : Group {ℓ'}) : Type (ℓ-max ℓ ℓ') where - constructor iso - field - map : GroupHom G H - inv : ⟨ H ⟩ → ⟨ G ⟩ - rightInv : I.section (fun map) inv - leftInv : I.retract (fun map) inv - -open GroupIso -GrIsoToGrEquiv : {G : Group {ℓ}} {H : Group {ℓ'}} → GroupIso G H → GroupEquiv G H -GroupEquiv.eq (GrIsoToGrEquiv i) = isoToEquiv (iso (fun (map i)) (inv i) (rightInv i) (leftInv i)) -GroupEquiv.isHom (GrIsoToGrEquiv i) = isHom (map i) - -record Iso' {ℓ ℓ'} (A : Group {ℓ}) (B : Group {ℓ'}) : Type (ℓ-max ℓ ℓ') where - constructor iso' - field - map' : GroupHom A B - inj : isInjective A B map' - surj : isSurjective A B map' - -open Iso' - --- "Very" short exact sequences -record vSES {ℓ ℓ' ℓ'' ℓ'''} (A : Group {ℓ}) (B : Group {ℓ'}) (leftGr : Group {ℓ''}) (rightGr : Group {ℓ'''}) - : Type (ℓ-suc (ℓ-max ℓ (ℓ-max ℓ' (ℓ-max ℓ'' ℓ''')))) where - no-eta-equality - constructor ses - field - isTrivialLeft : isProp ⟨ leftGr ⟩ - isTrivialRight : isProp ⟨ rightGr ⟩ - - left : GroupHom leftGr A - right : GroupHom B rightGr - ϕ : GroupHom A B - - Ker-ϕ⊂Im-left : (x : ⟨ A ⟩) - → isInKer A B ϕ x - → isInIm leftGr A left x - Ker-right⊂Im-ϕ : (x : ⟨ B ⟩) - → isInKer B rightGr right x - → isInIm A B ϕ x - ---- Proofs that different notions of ismomorphisms agree --- - -Iso'ToGroupEquiv : {A : Group {ℓ}} {B : Group {ℓ'}} → Iso' A B → GroupEquiv A B -Iso'ToGroupEquiv {A = A} {B = B} i = GrIsoToGrEquiv grIso - where - module A = Group A - module B = Group B - f = fun (map' i) - - helper : (b : _) → isProp (Σ[ a ∈ ⟨ A ⟩ ] f a ≡ b) - helper _ a b = - Σ≡Prop (λ _ → isSetCarrier B _ _) - (fst a ≡⟨ sym (IsGroup.rid (isGroup A) (fst a)) ⟩ - ((fst a) A.+ 0g A) ≡⟨ cong ((fst a) A.+_) (sym (invl A (fst b))) ⟩ - ((fst a) A.+ ((A.- fst b) A.+ fst b)) ≡⟨ Group.assoc A _ _ _ ⟩ - (((fst a) A.+ (A.- fst b)) A.+ fst b) ≡⟨ cong (A._+ fst b) idHelper ⟩ - (0g A A.+ fst b) ≡⟨ IsGroup.lid (isGroup A) (fst b) ⟩ - fst b ∎) - where - idHelper : fst a A.+ (A.- fst b) ≡ 0g A - idHelper = - inj i _ - (isHom (map' i) (fst a) (A.- (fst b)) - ∙ (cong (f (fst a) B.+_) (morphMinus A B (map' i) (fst b)) - ∙∙ cong (B._+ (B.- f (fst b))) (snd a ∙ sym (snd b)) - ∙∙ invr B (f (fst b)))) - - grIso : GroupIso A B - map grIso = map' i - inv grIso b = (rec (helper b) (λ a → a) (surj i b)) .fst - rightInv grIso b = (rec (helper b) (λ a → a) (surj i b)) .snd - leftInv grIso b j = rec (helper (f b)) (λ a → a) (propTruncIsProp (surj i (f b)) ∣ b , refl ∣ j) .fst - -vSES→GroupEquiv : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group {ℓ}} {B : Group {ℓ'}} (leftGr : Group {ℓ''}) (rightGr : Group {ℓ'''}) - → vSES A B leftGr rightGr - → GroupEquiv A B -vSES→GroupEquiv {A = A} lGr rGr vses = - Iso'ToGroupEquiv - (iso' (vSES.ϕ vses) - (λ a inker → rec (isSetCarrier A _ _) - (λ {(a , p) → sym p - ∙∙ cong (fun (vSES.left vses)) (vSES.isTrivialLeft vses a _) - ∙∙ morph0→0 lGr A (vSES.left vses)}) - (vSES.Ker-ϕ⊂Im-left vses a inker)) - λ a → vSES.Ker-right⊂Im-ϕ vses a (vSES.isTrivialRight vses _ _)) - - -dirProdEquiv : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group {ℓ}} {B : Group {ℓ'}} {C : Group {ℓ''}} {D : Group {ℓ'''}} - → GroupEquiv A C → GroupEquiv B D - → GroupEquiv (dirProd A B) (dirProd C D) -GroupEquiv.eq (dirProdEquiv eq1 eq2) = ≃-× (GroupEquiv.eq eq1) (GroupEquiv.eq eq2) -GroupEquiv.isHom (dirProdEquiv eq1 eq2) = isHom (×hom (GroupEquiv.hom eq1) (GroupEquiv.hom eq2)) - - ---- -open import Cubical.Data.Unit -lUnitGroupIso : ∀ {ℓ} {G : Group {ℓ}} → GroupEquiv (dirProd trivialGroup G) G -lUnitGroupIso = - GrIsoToGrEquiv - (iso (grouphom snd (λ a b → refl)) - (λ g → tt , g) - (λ _ → refl) - λ _ → refl) - -rUnitGroupIso : ∀ {ℓ} {G : Group {ℓ}} → GroupEquiv (dirProd G trivialGroup) G -rUnitGroupIso = - GrIsoToGrEquiv - (iso - (grouphom fst λ _ _ → refl) - (λ g → g , tt) - (λ _ → refl) - λ _ → refl) diff --git a/Cubical/Structures/Group/no-Eta/Base.agda b/Cubical/Structures/Group/no-Eta/Base.agda deleted file mode 100644 index 4ea0ad1b2..000000000 --- a/Cubical/Structures/Group/no-Eta/Base.agda +++ /dev/null @@ -1,216 +0,0 @@ -{-# OPTIONS --cubical --no-import-sorts --safe #-} -module Cubical.Structures.Group.no-Eta.Base where - -open import Cubical.Foundations.Prelude -open import Cubical.Data.Sigma -open import Cubical.Structures.Monoid hiding (⟨_⟩) -open import Cubical.Structures.Semigroup hiding (⟨_⟩) - -private - variable - ℓ : Level - -record IsGroup {G : Type ℓ} - (0g : G) (_+_ : G → G → G) (-_ : G → G) : Type ℓ where - no-eta-equality - constructor isgroup - - field - isMonoid : IsMonoid 0g _+_ - inverse : (x : G) → (x + (- x) ≡ 0g) × ((- x) + x ≡ 0g) - - open IsMonoid isMonoid public - - infixl 6 _-_ - - _-_ : G → G → G - x - y = x + (- y) - - invl : (x : G) → (- x) + x ≡ 0g - invl x = inverse x .snd - - invr : (x : G) → x + (- x) ≡ 0g - invr x = inverse x .fst - -record Group : Type (ℓ-suc ℓ) where - no-eta-equality - constructor group - - field - Carrier : Type ℓ - 0g : Carrier - _+_ : Carrier → Carrier → Carrier - -_ : Carrier → Carrier - isGroup : IsGroup 0g _+_ -_ - - infix 8 -_ - infixr 7 _+_ - - open IsGroup isGroup public - -Group₀ : Type₁ -Group₀ = Group {ℓ-zero} - --- Extractor for the carrier type -⟨_⟩ : Group → Type ℓ -⟨_⟩ = Group.Carrier - -makeIsGroup : {G : Type ℓ} {0g : G} {_+_ : G → G → G} { -_ : G → G} - (is-setG : isSet G) - (assoc : (x y z : G) → x + (y + z) ≡ (x + y) + z) - (rid : (x : G) → x + 0g ≡ x) - (lid : (x : G) → 0g + x ≡ x) - (rinv : (x : G) → x + (- x) ≡ 0g) - (linv : (x : G) → (- x) + x ≡ 0g) - → IsGroup 0g _+_ -_ -makeIsGroup is-setG assoc rid lid rinv linv = - isgroup (makeIsMonoid is-setG assoc rid lid) λ x → rinv x , linv x - -makeGroup : {G : Type ℓ} (0g : G) (_+_ : G → G → G) (-_ : G → G) - (is-setG : isSet G) - (assoc : (x y z : G) → x + (y + z) ≡ (x + y) + z) - (rid : (x : G) → x + 0g ≡ x) - (lid : (x : G) → 0g + x ≡ x) - (rinv : (x : G) → x + (- x) ≡ 0g) - (linv : (x : G) → (- x) + x ≡ 0g) - → Group -makeGroup 0g _+_ -_ is-setG assoc rid lid rinv linv = - group _ 0g _+_ -_ (makeIsGroup is-setG assoc rid lid rinv linv) - -makeGroup-right : ∀ {ℓ} {A : Type ℓ} - → (id : A) - → (comp : A → A → A) - → (inv : A → A) - → (set : isSet A) - → (assoc : ∀ a b c → comp a (comp b c) ≡ comp (comp a b) c) - → (rUnit : ∀ a → comp a id ≡ a) - → (rCancel : ∀ a → comp a (inv a) ≡ id) - → Group -makeGroup-right {A = A} id comp inv set assoc rUnit rCancel = - makeGroup id comp inv set assoc rUnit lUnit rCancel lCancel - where - _⨀_ = comp - abstract - lCancel : ∀ a → comp (inv a) a ≡ id - lCancel a = - inv a ⨀ a - ≡⟨ sym (rUnit (comp (inv a) a)) ⟩ - (inv a ⨀ a) ⨀ id - ≡⟨ cong (comp (comp (inv a) a)) (sym (rCancel (inv a))) ⟩ - (inv a ⨀ a) ⨀ (inv a ⨀ (inv (inv a))) - ≡⟨ assoc _ _ _ ⟩ - ((inv a ⨀ a) ⨀ (inv a)) ⨀ (inv (inv a)) - ≡⟨ cong (λ □ → □ ⨀ _) (sym (assoc _ _ _)) ⟩ - (inv a ⨀ (a ⨀ inv a)) ⨀ (inv (inv a)) - ≡⟨ cong (λ □ → (inv a ⨀ □) ⨀ (inv (inv a))) (rCancel a) ⟩ - (inv a ⨀ id) ⨀ (inv (inv a)) - ≡⟨ cong (λ □ → □ ⨀ (inv (inv a))) (rUnit (inv a)) ⟩ - inv a ⨀ (inv (inv a)) - ≡⟨ rCancel (inv a) ⟩ - id - ∎ - - lUnit : ∀ a → comp id a ≡ a - lUnit a = - id ⨀ a - ≡⟨ cong (λ b → comp b a) (sym (rCancel a)) ⟩ - (a ⨀ inv a) ⨀ a - ≡⟨ sym (assoc _ _ _) ⟩ - a ⨀ (inv a ⨀ a) - ≡⟨ cong (comp a) (lCancel a) ⟩ - a ⨀ id - ≡⟨ rUnit a ⟩ - a - ∎ - -makeGroup-left : ∀ {ℓ} {A : Type ℓ} - → (id : A) - → (comp : A → A → A) - → (inv : A → A) - → (set : isSet A) - → (assoc : ∀ a b c → comp a (comp b c) ≡ comp (comp a b) c) - → (lUnit : ∀ a → comp id a ≡ a) - → (lCancel : ∀ a → comp (inv a) a ≡ id) - → Group -makeGroup-left {A = A} id comp inv set assoc lUnit lCancel = - makeGroup id comp inv set assoc rUnit lUnit rCancel lCancel - where - abstract - rCancel : ∀ a → comp a (inv a) ≡ id - rCancel a = - comp a (inv a) - ≡⟨ sym (lUnit (comp a (inv a))) ⟩ - comp id (comp a (inv a)) - ≡⟨ cong (λ b → comp b (comp a (inv a))) (sym (lCancel (inv a))) ⟩ - comp (comp (inv (inv a)) (inv a)) (comp a (inv a)) - ≡⟨ sym (assoc (inv (inv a)) (inv a) (comp a (inv a))) ⟩ - comp (inv (inv a)) (comp (inv a) (comp a (inv a))) - ≡⟨ cong (comp (inv (inv a))) (assoc (inv a) a (inv a)) ⟩ - comp (inv (inv a)) (comp (comp (inv a) a) (inv a)) - ≡⟨ cong (λ b → comp (inv (inv a)) (comp b (inv a))) (lCancel a) ⟩ - comp (inv (inv a)) (comp id (inv a)) - ≡⟨ cong (comp (inv (inv a))) (lUnit (inv a)) ⟩ - comp (inv (inv a)) (inv a) - ≡⟨ lCancel (inv a) ⟩ - id - ∎ - - rUnit : ∀ a → comp a id ≡ a - rUnit a = - comp a id - ≡⟨ cong (comp a) (sym (lCancel a)) ⟩ - comp a (comp (inv a) a) - ≡⟨ assoc a (inv a) a ⟩ - comp (comp a (inv a)) a - ≡⟨ cong (λ b → comp b a) (rCancel a) ⟩ - comp id a - ≡⟨ lUnit a ⟩ - a - ∎ - - -open Group -open IsGroup - -isSetCarrier : ∀ {ℓ} → (G : Group {ℓ}) → isSet ⟨ G ⟩ -isSetCarrier G = IsSemigroup.is-set (IsMonoid.isSemigroup (isMonoid (isGroup G))) - -open import Cubical.Foundations.HLevels -dirProd : ∀ {ℓ ℓ'} → Group {ℓ} → Group {ℓ'} → Group -dirProd A B = - makeGroup ((0g A) , (0g B)) - (λ a b → (_+_ A (fst a) (fst b)) , _+_ B (snd a) (snd b)) - (λ a → (-_ A (fst a)) , (-_ B (snd a))) - (isOfHLevelΣ 2 (isSetCarrier A) λ _ → isSetCarrier B) - (λ x y z i → assoc (isGroup A) (fst x) (fst y) (fst z) i , assoc (isGroup B) (snd x) (snd y) (snd z) i) - (λ x i → IsGroup.rid (isGroup A) (fst x) i , IsGroup.rid (isGroup B) (snd x) i) - (λ x i → IsGroup.lid (isGroup A) (fst x) i , IsGroup.lid (isGroup B) (snd x) i) - (λ x i → (fst (inverse (isGroup A) (fst x)) i) , (fst (inverse (isGroup B) (snd x)) i)) - (λ x i → (snd (inverse (isGroup A) (fst x)) i) , (snd (inverse (isGroup B) (snd x)) i)) - -open import Cubical.Data.Unit -trivialGroup : Group₀ -trivialGroup = - makeGroup _ - (λ _ _ → _) - (λ _ → _) - isSetUnit - (λ _ _ _ → refl) - (λ _ → refl) - (λ _ → refl) - (λ _ → refl) - λ _ → refl - -open import Cubical.Data.Int renaming (_+_ to _+Int_ ; _-_ to _-Int_) -intGroup : Group₀ -intGroup = - makeGroup {G = Int} - 0 - _+Int_ - (pos 0 -Int_) - isSetInt - +-assoc - (λ _ → refl) - (λ x → +-comm (pos 0) x) - (λ x → +-comm x (pos 0 -Int x) ∙ minusPlus x 0) - λ x → minusPlus x 0 diff --git a/Cubical/Structures/Group/no-Eta/Morphism.agda b/Cubical/Structures/Group/no-Eta/Morphism.agda deleted file mode 100644 index a82d6b6ef..000000000 --- a/Cubical/Structures/Group/no-Eta/Morphism.agda +++ /dev/null @@ -1,64 +0,0 @@ -{-# OPTIONS --cubical --no-import-sorts --safe #-} - -module Cubical.Structures.Group.no-Eta.Morphism where - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Equiv -open import Cubical.Structures.Group.no-Eta.Base -open import Cubical.HITs.PropositionalTruncation hiding (map) -open import Cubical.Data.Sigma - -private - variable - ℓ ℓ' : Level - --- The following definition of GroupHom and GroupEquiv are level-wise heterogeneous. --- This allows for example to deduce that G ≡ F from a chain of isomorphisms --- G ≃ H ≃ F, even if H does not lie in the same level as G and F. - -isGroupHom : (G : Group {ℓ}) (H : Group {ℓ'}) (f : ⟨ G ⟩ → ⟨ H ⟩) → Type _ -isGroupHom G H f = (x y : ⟨ G ⟩) → f (x G.+ y) ≡ (f x H.+ f y) where - module G = Group G - module H = Group H - -record GroupHom (G : Group {ℓ}) (H : Group {ℓ'}) : Type (ℓ-max ℓ ℓ') where - no-eta-equality - constructor grouphom - - field - fun : ⟨ G ⟩ → ⟨ H ⟩ - isHom : isGroupHom G H fun - -record GroupEquiv (G : Group {ℓ}) (H : Group {ℓ'}) : Type (ℓ-max ℓ ℓ') where - no-eta-equality - constructor groupequiv - - field - eq : ⟨ G ⟩ ≃ ⟨ H ⟩ - isHom : isGroupHom G H (equivFun eq) - - hom : GroupHom G H - hom = grouphom (equivFun eq) isHom - -open GroupHom - -×hom : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group {ℓ}} {B : Group {ℓ'}} {C : Group {ℓ''}} {D : Group {ℓ'''}} - → GroupHom A C → GroupHom B D → GroupHom (dirProd A B) (dirProd C D) -GroupHom.fun (×hom mf1 mf2) = map-× (fun mf1) (fun mf2) -GroupHom.isHom (×hom mf1 mf2) a b = ≡-× (isHom mf1 _ _) (isHom mf2 _ _) - -open Group - -isInIm : ∀ {ℓ ℓ'} (G : Group {ℓ}) (H : Group {ℓ'}) → (GroupHom G H) - → ⟨ H ⟩ → Type (ℓ-max ℓ ℓ') -isInIm G H ϕ h = ∃[ g ∈ ⟨ G ⟩ ] (fun ϕ) g ≡ h - -isInKer : ∀ {ℓ ℓ'} (G : Group {ℓ}) (H : Group {ℓ'}) → (GroupHom G H) - → ⟨ G ⟩ → Type ℓ' -isInKer G H ϕ g = (fun ϕ) g ≡ 0g H - -isSurjective : ∀ {ℓ ℓ'} (G : Group {ℓ}) (H : Group {ℓ'}) → GroupHom G H → Type (ℓ-max ℓ ℓ') -isSurjective G H ϕ = (x : ⟨ H ⟩) → isInIm G H ϕ x - -isInjective : ∀ {ℓ ℓ'} (G : Group {ℓ}) (H : Group {ℓ'}) → GroupHom G H → Type (ℓ-max ℓ ℓ') -isInjective G H ϕ = (x : ⟨ G ⟩) → isInKer G H ϕ x → x ≡ 0g G diff --git a/Cubical/Structures/Group/no-Eta/MorphismProperties.agda b/Cubical/Structures/Group/no-Eta/MorphismProperties.agda deleted file mode 100644 index 07d1836d4..000000000 --- a/Cubical/Structures/Group/no-Eta/MorphismProperties.agda +++ /dev/null @@ -1,80 +0,0 @@ -{-# OPTIONS --cubical --no-import-sorts --safe #-} -module Cubical.Structures.Group.no-Eta.MorphismProperties where - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Equiv.HalfAdjoint -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.SIP -open import Cubical.Foundations.Function using (_∘_) -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Functions.Embedding -open import Cubical.Data.Sigma - -open import Cubical.Structures.Axioms -open import Cubical.Structures.Pointed -open import Cubical.Structures.Semigroup hiding (⟨_⟩) -open import Cubical.Structures.Monoid hiding (⟨_⟩) - -open import Cubical.Structures.Group.no-Eta.Base -open import Cubical.Structures.Group.no-Eta.Properties -open import Cubical.Structures.Group.no-Eta.Morphism - -open Iso - -private - variable - ℓ ℓ' ℓ'' : Level - -infixr 31 _□_ - -isPropIsGroupHom : (G : Group {ℓ}) (H : Group {ℓ'}) {f : ⟨ G ⟩ → ⟨ H ⟩} → isProp (isGroupHom G H f) -isPropIsGroupHom G H {f} = isPropΠ2 λ a b → Group.is-set H _ _ - - --- Morphism composition - -open GroupHom -open GroupEquiv -isGroupHomComp : {F : Group {ℓ}} {G : Group {ℓ'}} {H : Group {ℓ''}} → - (f : GroupHom F G) → (g : GroupHom G H) → isGroupHom F H (GroupHom.fun g ∘ GroupHom.fun f) -isGroupHomComp F G x y = - cong (fun G) (isHom F _ _) ∙ isHom G _ _ - -compGroupHom : {F : Group {ℓ}} {G : Group {ℓ'}} {H : Group {ℓ''}} → GroupHom F G → GroupHom G H → GroupHom F H -fun (compGroupHom {F = F} {G = G} {H = H} f g) = fun g ∘ fun f -isHom (compGroupHom {F = F} {G = G} {H = H} f g) = isGroupHomComp f g - -compGroupEquiv : {F : Group {ℓ}} {G : Group {ℓ'}} {H : Group {ℓ''}} → GroupEquiv F G → GroupEquiv G H → GroupEquiv F H -GroupEquiv.eq (compGroupEquiv {F = F} {G = G} {H = H} f g) = compEquiv (eq f) (eq g) -GroupEquiv.isHom (compGroupEquiv {F = F} {G = G} {H = H} f g) = isHom (compGroupHom (hom f) (hom g)) - -_□_ : _ -_□_ = compGroupEquiv - - -idGroupEquiv : (G : Group {ℓ}) → GroupEquiv G G -eq (idGroupEquiv G) = idEquiv ⟨ G ⟩ -isHom (idGroupEquiv G) = λ _ _ → refl - --- Isomorphism inversion -isGroupHomInv : {G : Group {ℓ}} {H : Group {ℓ'}} (f : GroupEquiv G H) → isGroupHom H G (invEq (GroupEquiv.eq f)) -isGroupHomInv {G = G} {H = H} f h h' = isInj-f _ _ ( - f' (g (h ⋆² h')) ≡⟨ retEq (eq f) _ ⟩ - (h ⋆² h') ≡⟨ sym (cong₂ _⋆²_ (retEq (eq f) h) (retEq (eq f) h')) ⟩ - (f' (g h) ⋆² f' (g h')) ≡⟨ sym (isHom (hom f) _ _) ⟩ - f' (g h ⋆¹ g h') ∎) - where - f' = fst (eq f) - _⋆¹_ = Group._+_ G - _⋆²_ = Group._+_ H - g = invEq (eq f) - - isInj-f : (x y : ⟨ G ⟩) → f' x ≡ f' y → x ≡ y - isInj-f x y = invEq (_ , isEquiv→isEmbedding (snd (eq f)) x y) - -invGroupEquiv : {G : Group {ℓ}} {H : Group {ℓ'}} → GroupEquiv G H → GroupEquiv H G -eq (invGroupEquiv f) = invEquiv (eq f) -isHom (invGroupEquiv f) = isGroupHomInv f diff --git a/Cubical/Structures/Group/no-Eta/Properties.agda b/Cubical/Structures/Group/no-Eta/Properties.agda deleted file mode 100644 index 0fc078010..000000000 --- a/Cubical/Structures/Group/no-Eta/Properties.agda +++ /dev/null @@ -1,66 +0,0 @@ -{-# OPTIONS --cubical --no-import-sorts --safe #-} - -module Cubical.Structures.Group.no-Eta.Properties where - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.HLevels -open import Cubical.Data.Sigma -open import Cubical.Structures.Semigroup -open import Cubical.Structures.Monoid -open import Cubical.Structures.Group.no-Eta.Base - -private - variable - ℓ ℓ' ℓ'' : Level - -module GroupLemmas (G : Group {ℓ}) where - open Group G public - abstract - simplL : (a : Carrier) {b c : Carrier} → a + b ≡ a + c → b ≡ c - simplL a {b} {c} p = - b - ≡⟨ sym (lid b) ∙ cong (_+ b) (sym (invl a)) ∙ sym (assoc _ _ _) ⟩ - - a + (a + b) - ≡⟨ cong (- a +_) p ⟩ - - a + (a + c) - ≡⟨ assoc _ _ _ ∙ cong (_+ c) (invl a) ∙ lid c ⟩ - c ∎ - - simplR : {a b : Carrier} (c : Carrier) → a + c ≡ b + c → a ≡ b - simplR {a} {b} c p = - a - ≡⟨ sym (rid a) ∙ cong (a +_) (sym (invr c)) ∙ assoc _ _ _ ⟩ - (a + c) - c - ≡⟨ cong (_- c) p ⟩ - (b + c) - c - ≡⟨ sym (assoc _ _ _) ∙ cong (b +_) (invr c) ∙ rid b ⟩ - b ∎ - - invInvo : (a : Carrier) → - (- a) ≡ a - invInvo a = simplL (- a) (invr (- a) ∙ sym (invl a)) - - invId : - 0g ≡ 0g - invId = simplL 0g (invr 0g ∙ sym (lid 0g)) - - idUniqueL : {e : Carrier} (x : Carrier) → e + x ≡ x → e ≡ 0g - idUniqueL {e} x p = simplR x (p ∙ sym (lid _)) - - idUniqueR : (x : Carrier) {e : Carrier} → x + e ≡ x → e ≡ 0g - idUniqueR x {e} p = simplL x (p ∙ sym (rid _)) - - invUniqueL : {g h : Carrier} → g + h ≡ 0g → g ≡ - h - invUniqueL {g} {h} p = simplR h (p ∙ sym (invl h)) - - invUniqueR : {g h : Carrier} → g + h ≡ 0g → h ≡ - g - invUniqueR {g} {h} p = simplL g (p ∙ sym (invr g)) - - invDistr : (a b : Carrier) → - (a + b) ≡ - b - a - invDistr a b = sym (invUniqueR γ) where - γ : (a + b) + (- b - a) ≡ 0g - γ = (a + b) + (- b - a) - ≡⟨ sym (assoc _ _ _) ⟩ - a + b + (- b) + (- a) - ≡⟨ cong (a +_) (assoc _ _ _ ∙ cong (_+ (- a)) (invr b)) ⟩ - a + (0g - a) - ≡⟨ cong (a +_) (lid (- a)) ∙ invr a ⟩ - 0g ∎ diff --git a/Cubical/Structures/Group/no-Eta/no-Eta.agda b/Cubical/Structures/Group/no-Eta/no-Eta.agda deleted file mode 100644 index afa1f3f6a..000000000 --- a/Cubical/Structures/Group/no-Eta/no-Eta.agda +++ /dev/null @@ -1,8 +0,0 @@ -{-# OPTIONS --cubical --no-import-sorts --safe #-} -module Cubical.Structures.Group.no-Eta.no-Eta where - -open import Cubical.Structures.Group.no-Eta.Base public -open import Cubical.Structures.Group.no-Eta.Properties public -open import Cubical.Structures.Group.no-Eta.Morphism public -open import Cubical.Structures.Group.no-Eta.MorphismProperties public -open import Cubical.Structures.Group.no-Eta.Algebra public diff --git a/Cubical/Structures/Monoid.agda b/Cubical/Structures/Monoid.agda index 74175013b..6f9a25431 100644 --- a/Cubical/Structures/Monoid.agda +++ b/Cubical/Structures/Monoid.agda @@ -59,6 +59,7 @@ record Monoid : Type (ℓ-suc ℓ) where -- open Semigroup semigrp public + -- Extractor for the carrier type ⟨_⟩ : Monoid → Type ℓ ⟨_⟩ = Monoid.Carrier @@ -76,8 +77,8 @@ makeIsMonoid : {M : Type ℓ} {ε : M} {_·_ : M → M → M} (rid : (x : M) → x · ε ≡ x) (lid : (x : M) → ε · x ≡ x) → IsMonoid ε _·_ -makeIsMonoid is-setM assoc rid lid = - ismonoid (issemigroup is-setM assoc) λ x → rid x , lid x +IsMonoid.isSemigroup (makeIsMonoid is-setM assoc rid lid) = issemigroup is-setM assoc +IsMonoid.identity (makeIsMonoid is-setM assoc rid lid) = λ x → rid x , lid x makeMonoid : {M : Type ℓ} (ε : M) (_·_ : M → M → M) (is-setM : isSet M) @@ -85,7 +86,7 @@ makeMonoid : {M : Type ℓ} (ε : M) (_·_ : M → M → M) (rid : (x : M) → x · ε ≡ x) (lid : (x : M) → ε · x ≡ x) → Monoid -makeMonoid ε _·_ is-setM assoc rid lid = +makeMonoid ε _·_ is-setM assoc rid lid = ? monoid _ ε _·_ (makeIsMonoid is-setM assoc rid lid) record MonoidEquiv (M N : Monoid {ℓ}) : Type ℓ where diff --git a/Cubical/Structures/Semigroup.agda b/Cubical/Structures/Semigroup.agda index 0eea157c1..2f4281d63 100644 --- a/Cubical/Structures/Semigroup.agda +++ b/Cubical/Structures/Semigroup.agda @@ -17,6 +17,7 @@ open import Cubical.Structures.Auto open Iso + private variable ℓ : Level @@ -67,16 +68,19 @@ record SemigroupEquiv (M N : Semigroup {ℓ}) : Type ℓ where e : ⟨ M ⟩ ≃ ⟨ N ⟩ isHom : (x y : ⟨ M ⟩) → equivFun e (x M.· y) ≡ equivFun e x N.· equivFun e y +open Semigroup +open IsSemigroup +open SemigroupEquiv η-isSemiGroup : {A : Type ℓ} {_·_ : A → A → A} (b : IsSemigroup _·_) - → issemigroup (IsSemigroup.is-set b) (IsSemigroup.assoc b) ≡ b -IsSemigroup.is-set (η-isSemiGroup b i) = IsSemigroup.is-set b -IsSemigroup.assoc (η-isSemiGroup b i) = IsSemigroup.assoc b + → issemigroup (is-set b) (assoc b) ≡ b +is-set (η-isSemiGroup b i) = is-set b +assoc (η-isSemiGroup b i) = assoc b η-SemigroupEquiv : {M N : Semigroup {ℓ}} (p : SemigroupEquiv M N) - → semigroupequiv (SemigroupEquiv.e p) (SemigroupEquiv.isHom p) ≡ p -SemigroupEquiv.e (η-SemigroupEquiv p i) = SemigroupEquiv.e p -SemigroupEquiv.isHom (η-SemigroupEquiv p i) = SemigroupEquiv.isHom p + → semigroupequiv (e p) (isHom p) ≡ p +e (η-SemigroupEquiv p i) = e p +isHom (η-SemigroupEquiv p i) = isHom p -- Develop some theory about Semigroups using various general results -- that are stated using Σ-types. For this we define Semigroup as a @@ -109,8 +113,6 @@ module SemigroupΣTheory {ℓ} where SemigroupEquivStr : StrEquiv SemigroupStructure ℓ SemigroupEquivStr = AxiomsEquivStr RawSemigroupEquivStr SemigroupAxioms - open IsSemigroup - SemigroupAxiomsIsoIsSemigroup : {A : Type ℓ} (_·_ : RawSemigroupStructure A) → Iso (SemigroupAxioms A _·_) (IsSemigroup _·_) fun (SemigroupAxiomsIsoIsSemigroup s) (x , y) = issemigroup x y @@ -135,9 +137,9 @@ module SemigroupΣTheory {ℓ} where iso Semigroup→SemigroupΣ SemigroupΣ→Semigroup (λ _ → refl) helper where helper : (a : Semigroup) → SemigroupΣ→Semigroup (Semigroup→SemigroupΣ a) ≡ a - Semigroup.Carrier (helper a i) = ⟨ a ⟩ - Semigroup._·_ (helper a i) = Semigroup._·_ a - Semigroup.isSemigroup (helper a i) = η-isSemiGroup (Semigroup.isSemigroup a) i + Carrier (helper a i) = ⟨ a ⟩ + _·_ (helper a i) = _·_ a + isSemigroup (helper a i) = η-isSemiGroup (isSemigroup a) i semigroupUnivalentStr : UnivalentStr SemigroupStructure SemigroupEquivStr semigroupUnivalentStr = axiomsUnivalentStr _ isPropSemigroupAxioms rawSemigroupUnivalentStr diff --git a/Cubical/ZCohomology/Groups/Connected.agda b/Cubical/ZCohomology/Groups/Connected.agda index 267d4c890..d953248ce 100644 --- a/Cubical/ZCohomology/Groups/Connected.agda +++ b/Cubical/ZCohomology/Groups/Connected.agda @@ -19,7 +19,7 @@ open import Cubical.HITs.Sn open import Cubical.HITs.S1 open import Cubical.HITs.Susp open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) -open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₋₁ ; ∣_∣ to ∣_∣₋₁) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₁ ; ∣_∣ to ∣_∣₁) open import Cubical.HITs.Nullification open import Cubical.Data.Sigma hiding (_×_) @@ -42,25 +42,17 @@ private sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ f → cong ∣_∣₂ (funExt λ x → trRec (isSetInt _ _) (cong f) (isConnectedPath 1 con a x .fst)) - --- H⁰-connected : ∀ {ℓ} {A : Type ℓ} (a : A) → isConnected 2 A → GroupEquiv (coHomGr 0 A) intGroup --- H⁰-connected {A = A} a con = invGroupEquiv invEq' --- where --- invEq' : GroupEquiv intGroup (coHomGr 0 A) --- GroupEquiv.eq invEq' = invEquiv (isoToEquiv (H⁰-connected-type a con)) --- GroupEquiv.isHom invEq' c d = cong ∣_∣₂ (funExt λ _ → sym (addLemma c d)) - - -H⁰-connected' : ∀ {ℓ} {A : Type ℓ} (a : A) → ((x : A) → ∥ a ≡ x ∥₋₁) → GroupIso (coHomGr 0 A) intGroup -H⁰-connected' a con = invGroupIso Iso⁻ - where - Iso⁻ : GroupIso _ _ - GroupHom.fun (GroupIso.map Iso⁻) b = ∣ (λ _ → b) ∣₂ - GroupHom.isHom (GroupIso.map Iso⁻) c d = cong ∣_∣₂ (funExt λ _ → sym (addLemma c d)) - GroupIso.inv Iso⁻ = sRec isSetInt (λ f → f a) - GroupIso.rightInv Iso⁻ = (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) - λ f → cong ∣_∣₂ (funExt λ x → pRec (isSetInt _ _) (cong f) (con x))) - GroupIso.leftInv Iso⁻ _ = refl - -H⁰-connected : ∀ {ℓ} {A : Type ℓ} (a : A) → ((x : A) → ∥ a ≡ x ∥₋₁) → GroupEquiv (coHomGr 0 A) intGroup +private + H⁰-connected' : ∀ {ℓ} {A : Type ℓ} (a : A) → ((x : A) → ∥ a ≡ x ∥₁) → GroupIso (coHomGr 0 A) intGroup + H⁰-connected' a con = invGroupIso Iso⁻ + where + Iso⁻ : GroupIso _ _ + GroupHom.fun (GroupIso.map Iso⁻) b = ∣ (λ _ → b) ∣₂ + GroupHom.isHom (GroupIso.map Iso⁻) c d = cong ∣_∣₂ (funExt λ _ → sym (addLemma c d)) + GroupIso.inv Iso⁻ = sRec isSetInt (λ f → f a) + GroupIso.rightInv Iso⁻ = (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ∣_∣₂ (funExt λ x → pRec (isSetInt _ _) (cong f) (con x))) + GroupIso.leftInv Iso⁻ _ = refl + +H⁰-connected : ∀ {ℓ} {A : Type ℓ} (a : A) → ((x : A) → ∥ a ≡ x ∥₁) → GroupEquiv (coHomGr 0 A) intGroup H⁰-connected a con = GrIsoToGrEquiv (H⁰-connected' a con) diff --git a/Cubical/ZCohomology/Groups/Prelims.agda b/Cubical/ZCohomology/Groups/Prelims.agda new file mode 100644 index 000000000..dd953f17a --- /dev/null +++ b/Cubical/ZCohomology/Groups/Prelims.agda @@ -0,0 +1,321 @@ +{-# OPTIONS --cubical --no-import-sorts --safe #-} +module Cubical.ZCohomology.Groups.Prelims where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.KcompPrelims + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 +open import Cubical.HITs.Nullification + +open import Cubical.Data.Sigma +open import Cubical.Data.Int renaming (_+_ to _+ℤ_; +-comm to +ℤ-comm ; +-assoc to +ℤ-assoc) +open import Cubical.Data.Nat +open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec) + +S¹map : hLevelTrunc 3 S¹ → S¹ +S¹map = trRec isGroupoidS¹ (idfun _) + +S¹map-id : (x : hLevelTrunc 3 S¹) → Path (hLevelTrunc 3 S¹) ∣ S¹map x ∣ x +S¹map-id = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → refl + +S1map : hLevelTrunc 3 (S₊ 1) → (S₊ 1) +S1map = trRec isGroupoidS1 (idfun _) + +S1→S1≡S¹→S¹ : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → hLevelTrunc 3 S¹) +Iso.fun S1→S1≡S¹→S¹ f x = trMap S1→S¹ (f (S¹→S1 x)) +Iso.inv S1→S1≡S¹→S¹ f x = trMap S¹→S1 (f (S1→S¹ x)) +Iso.rightInv S1→S1≡S¹→S¹ F = funExt λ x i → helper2 (F (S1→S¹-sect x i)) i + where + helper2 : (x : hLevelTrunc 3 S¹) → trMap S1→S¹ (trMap S¹→S1 x) ≡ x + helper2 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → cong ∣_∣ (S1→S¹-sect a) +Iso.leftInv S1→S1≡S¹→S¹ F = funExt λ x i → helper2 (F (S1→S¹-retr x i)) i + where + helper2 : (x : hLevelTrunc 3 (S₊ 1)) → trMap S¹→S1 (trMap S1→S¹ x) ≡ x + helper2 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → cong ∣_∣ (S1→S¹-retr a) + +{- Proof that (S¹ → ∥ S¹ ∥₁) ≃ S¹ × ℤ. Needed for H¹(S¹)) -} +-- We prove that (S¹ → ∥ S¹ ∥) ≃ S¹ × ℤ. Note that the truncation doesn't really matter, since S¹ is a groupoid. +-- Given a map f : S¹ → S¹, the idea is to send this to (f(base) , winding (f(loop))). For this to be +-- well-typed, we need to translate f(loop) into an element in Ω(S¹,base). +S¹→S¹≡S¹×Int : Iso (S¹ → hLevelTrunc 3 S¹) (S¹ × Int) +Iso.fun S¹→S¹≡S¹×Int f = S¹map (f base) + , winding (basechange2⁻ (S¹map (f base)) λ i → S¹map (f (loop i))) +Iso.inv S¹→S¹≡S¹×Int (s , int) base = ∣ s ∣ +Iso.inv S¹→S¹≡S¹×Int (s , int) (loop i) = ∣ basechange2 s (intLoop int) i ∣ +Iso.rightInv S¹→S¹≡S¹×Int (s , int) = ΣPathP (refl , ((λ i → winding (basechange2-retr s (λ i → intLoop int i) i)) + ∙ windingIntLoop int)) +Iso.leftInv S¹→S¹≡S¹×Int f = funExt λ { base → S¹map-id (f base) + ; (loop i) j → helper j i} + where + helper : PathP (λ i → S¹map-id (f base) i ≡ S¹map-id (f base) i) + (λ i → ∣ basechange2 (S¹map (f base)) + (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁)))))) i ∣) + (cong f loop) + helper i j = + hcomp (λ k → λ { (i = i0) → cong ∣_∣ (basechange2 (S¹map (f base)) + (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁))))))) j + ; (i = i1) → S¹map-id (f (loop j)) k + ; (j = i0) → S¹map-id (f base) (i ∧ k) + ; (j = i1) → S¹map-id (f base) (i ∧ k)}) + (helper2 i j) + where + helper2 : Path (Path (hLevelTrunc 3 _) _ _) + (cong ∣_∣ (basechange2 (S¹map (f base)) + (intLoop + (winding + (basechange2⁻ (S¹map (f base)) + (λ i₁ → S¹map (f (loop i₁)))))))) + λ i → ∣ S¹map (f (loop i)) ∣ + helper2 i = + cong ∣_∣ + ((cong (basechange2 (S¹map (f base))) + (decodeEncode base (basechange2⁻ (S¹map (f base)) + (λ i₁ → S¹map (f (loop i₁))))) + ∙ basechange2-sect (S¹map (f base)) + (λ i → S¹map (f (loop i)))) i) + + +{- Proof that (S¹ → K₁) ≃ K₁ × ℤ. Needed for H¹(T²) -} +S1→K₁≡S1×Int : Iso ((S₊ 1) → coHomK 1) (coHomK 1 × Int) +S1→K₁≡S1×Int = helper2 ⋄ S¹→S¹≡S¹×Int ⋄ helper + where + helper : Iso (S¹ × Int) (hLevelTrunc 3 (S₊ 1) × Int) + Iso.fun helper (s , int) = ∣ S¹→S1 s ∣ , int + Iso.inv helper (s , int) = (S1→S¹ (S1map s)) , int + Iso.rightInv helper (s , int) = + trElim {B = λ s → (∣ S¹→S1 (S1→S¹ (S1map s)) ∣ , int) ≡ (s , int)} + (λ _ → isOfHLevelPath 3 (isOfHLevelΣ 3 (isOfHLevelTrunc 3) (λ _ → isOfHLevelSuc 2 isSetInt)) _ _) + (λ a → ΣPathP ((cong ∣_∣ (S1→S¹-retr a)) , refl)) + s + Iso.leftInv helper (s , int) = ΣPathP ((S1→S¹-sect s) , refl) + + helper2 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → hLevelTrunc 3 S¹) + Iso.fun helper2 f x = trMap S1→S¹ (f (S¹→S1 x)) + Iso.inv helper2 f x = trMap S¹→S1 (f (S1→S¹ x)) + Iso.rightInv helper2 f = funExt λ x i → helper3 (f (S1→S¹-sect x i)) i + where + helper3 : (x : hLevelTrunc 3 S¹) → trMap S1→S¹ (trMap S¹→S1 x) ≡ x + helper3 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → cong ∣_∣ (S1→S¹-sect a) + Iso.leftInv helper2 f = funExt λ x i → helper3 (f (S1→S¹-retr x i)) i + where + helper3 : (x : hLevelTrunc 3 (S₊ 1)) → trMap S¹→S1 (trMap S1→S¹ x) ≡ x + helper3 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → cong ∣_∣ (S1→S¹-retr a) + + +{- Proof that S¹→K2 is isomorphic to K2×K1 (as types). Needed for H²(T²) -} +S1→K2≡K2×K1 : Iso (S₊ 1 → coHomK 2) (coHomK 2 × coHomK 1) +Iso.fun S1→K2≡K2×K1 f = f north , ΩKn+1→Kn (sym (rCancelₖ (f north)) + ∙ (λ i → f (merid south i) +ₖ (-ₖ f (merid north i))) + ∙ rCancelₖ (f south)) +Iso.inv S1→K2≡K2×K1 (a , b) north = a +ₖ 0ₖ +Iso.inv S1→K2≡K2×K1 (a , b) south = a +ₖ 0ₖ +Iso.inv S1→K2≡K2×K1 (a , b) (merid south i) = a +ₖ (Kn→ΩKn+1 1 b i) +Iso.inv S1→K2≡K2×K1 (a , b) (merid north i) = a +ₖ 0ₖ +Iso.rightInv S1→K2≡K2×K1 (a , b) = + ΣPathP ((rUnitₖ a) , + ((cong ΩKn+1→Kn (congHelper2 (Kn→ΩKn+1 1 b) (λ x → (a +ₖ x) +ₖ (-ₖ (a +ₖ 0ₖ))) (funExt (λ x → sym (cancelHelper a x))) (rCancelₖ (a +ₖ 0ₖ)))) + ∙ Iso.leftInv (Iso3-Kn-ΩKn+1 1) b)) + module _ where + cancelHelper : (a b : hLevelTrunc 4 (S₊ 2)) → (a +ₖ b) +ₖ (-ₖ (a +ₖ 0ₖ)) ≡ b + cancelHelper a b = + (a +ₖ b) +ₖ (-ₖ (a +ₖ 0ₖ)) ≡⟨ (λ i → (a +ₖ b) +ₖ (-ₖ (rUnitₖ a i))) ⟩≡⟨ cong (λ x → x +ₖ (-ₖ a)) (commₖ a b) ⟩ + (b +ₖ a) +ₖ (-ₖ a) ≡⟨ assocₖ b a (-ₖ a) ⟩≡⟨ cong (λ x → b +ₖ x) (rCancelₖ a) ⟩ + (b +ₖ 0ₖ ≡⟨ rUnitₖ b ⟩ + b ∎) + + moveabout : {x : hLevelTrunc 4 (S₊ 2)} (p q : x ≡ 0ₖ) (mid : 0ₖ ≡ 0ₖ) + → sym q ∙ (p ∙ mid ∙ sym p) ∙ q ≡ mid + moveabout p q mid = (assoc (sym q) _ _ + ∙∙ cong (_∙ q) (assoc (sym q) p (mid ∙ sym p)) + ∙∙ sym (assoc (sym q ∙ p) (mid ∙ sym p) q)) + ∙∙ (cong ((sym q ∙ p) ∙_) (sym (assoc mid (sym p) q)) + ∙∙ cong (λ x → (sym q ∙ p) ∙ (mid ∙ x)) (sym (symDistr (sym q) p)) + ∙∙ cong ((sym q ∙ p)∙_) (isCommΩK 2 mid _)) + ∙∙ (assoc _ _ _ + ∙∙ cong (_∙ mid) (rCancel (sym q ∙ p)) + ∙∙ sym (lUnit mid)) + + congHelper : ∀ {ℓ} {A : Type ℓ} {a1 : A} (p : a1 ≡ a1) (f : A → A) (id : (λ x → x) ≡ f) + → cong f p ≡ sym (funExt⁻ id a1) ∙ p ∙ funExt⁻ id a1 + congHelper {a1 = a1} p f id = + (λ i → lUnit (rUnit (cong f p) i) i) + ∙ (λ i → (λ j → id ((~ i) ∨ (~ j)) a1) ∙ cong (id (~ i)) p ∙ λ j → id (~ i ∨ j) a1) + + + congHelper2 : (p : 0ₖ ≡ 0ₖ) (f : coHomK 2 → coHomK 2) (id : (λ x → x) ≡ f) + → (q : (f 0ₖ) ≡ 0ₖ) + → (sym q) ∙ cong f p ∙ q ≡ p + congHelper2 p f id q = + cong (λ x → sym q ∙ x ∙ q) (congHelper p f id) ∙ + moveabout (sym (funExt⁻ id ∣ north ∣)) q p + +Iso.leftInv S1→K2≡K2×K1 a = + funExt λ {north → rUnitₖ (a north) + ; south → rUnitₖ (a north) ∙ cong a (merid north) + ; (merid south i) → m-south i + ; (merid north i) → m-north i} + where + + m-north : PathP (λ i → a north +ₖ ∣ north ∣ ≡ a (merid north i)) + (rUnitₖ (a north)) (rUnitₖ (a north) ∙ cong a (merid north)) + m-north i j = + hcomp (λ k → λ { (i = i0) → rUnitₖ (a north) j + ; (j = i1) → a (merid north (i ∧ k)) + ; (j = i0) → a north +ₖ ∣ north ∣}) + (rUnitₖ (a north) j) + + m-south : PathP (λ i → a north +ₖ Kn→ΩKn+1 1 (ΩKn+1→Kn (sym (rCancelₖ (a north)) + ∙ (λ i → a (merid south i) +ₖ (-ₖ a (merid north i))) + ∙ rCancelₖ (a south))) i ≡ a (merid south i)) (rUnitₖ (a north)) (rUnitₖ (a north) ∙ cong a (merid north)) + m-south j i = + hcomp (λ k → λ { (i = i0) → (helper ∙∙ together ∙∙ congFunct (_+ₖ 0ₖ) (cong a (merid south)) (cong a (sym (merid north)))) (~ k) j + ; (i = i1) → a (merid south j) + ; (j = i0) → rUnitₖ (a north) i + ; (j = i1) → ((λ j → rUnitₖ (a (merid north ((~ k) ∧ j))) (j ∧ k)) ∙ λ j → rUnitₖ (a (merid north ((~ k) ∨ j))) (k ∨ j)) i }) + (hcomp (λ k → λ { (i = i1) → a (merid south j) + ; (i = i0) → compPath-filler ((cong (λ x → a x +ₖ 0ₖ) (merid south))) (cong (λ x → a x +ₖ 0ₖ) (sym (merid north))) k j + ; (j = i0) → rUnitₖ (a north) i + ; (j = i1) → pathFiller2 (cong (_+ₖ 0ₖ) (cong a (merid north))) (rUnitₖ (a south)) k i}) + (rUnitₖ (a (merid south j)) i)) + + + where + pathFiller : (rUnitₖ (a north) ∙ cong a (merid north)) ≡ cong (_+ₖ 0ₖ) (cong a (merid north)) ∙ rUnitₖ (a south) + pathFiller = (λ i → (λ j → rUnitₖ (a (merid north (i ∧ j))) (j ∧ (~ i))) + ∙ λ j → rUnitₖ (a (merid north (i ∨ j))) (~ i ∨ j)) + + pathFiller2 : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) → PathP (λ i → p (~ i) ≡ z) q (p ∙ q) + pathFiller2 p q i = + hcomp (λ k → λ {(i = i0) → lUnit q (~ k) + ; (i = i1) → p ∙ q}) + ((λ j → p (~ i ∨ j)) ∙ q) + + helper : Path (_ ≡ _) (λ i → a north +ₖ Kn→ΩKn+1 1 (ΩKn+1→Kn (sym (rCancelₖ (a north)) + ∙ (λ i → a (merid south i) +ₖ (-ₖ a (merid north i))) + ∙ rCancelₖ (a south))) i) + --- + (cong (a north +ₖ_) (sym (rCancelₖ (a north))) + ∙ ((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) + ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north))) + ∙ cong (a north +ₖ_) (rCancelₖ (a north))) + helper = + ((λ j i → a north +ₖ Iso.rightInv (Iso3-Kn-ΩKn+1 1) ((sym (rCancelₖ (a north)) + ∙ (λ i → a (merid south i) +ₖ (-ₖ a (merid north i))) + ∙ rCancelₖ (a south))) j i) + ∙∙ congFunct (a north +ₖ_) (sym (rCancelₖ (a north))) ((λ i → a (merid south i) +ₖ (-ₖ a (merid north i))) ∙ rCancelₖ (a south)) + ∙∙ (λ j → cong (a north +ₖ_) (sym (rCancelₖ (a north))) + ∙ congFunct (a north +ₖ_) ((λ i → a (merid south i) +ₖ (-ₖ a (merid north i)))) (rCancelₖ (a south)) j)) + ∙∙ (λ j → cong (a north +ₖ_) (sym (rCancelₖ (a north))) + ∙ cong₂Funct (λ x y → a north +ₖ a x +ₖ (-ₖ (a y))) (merid south) (merid north) j ∙ cong (a north +ₖ_) ((rCancelₖ (a south)))) + ∙∙ (λ i → cong (a north +ₖ_) (sym (rCancelₖ (a north))) + ∙ ((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) + ∙ λ j → a north +ₖ a (merid north (~ i ∨ (~ j))) +ₖ (-ₖ (a (merid north (j ∧ (~ i)))))) + ∙ cong (a north +ₖ_) ((rCancelₖ (a (merid north (~ i)))))) + abstract + pathHelper : (a b : hLevelTrunc 4 (S₊ 2)) → a +ₖ b +ₖ (-ₖ a) ≡ b +ₖ 0ₖ + pathHelper a b = + a +ₖ b +ₖ (-ₖ a) ≡⟨ commₖ a (b +ₖ (-ₖ a)) ⟩≡⟨ assocₖ b (-ₖ a) a ⟩ + (b +ₖ (-ₖ a) +ₖ a ≡⟨ cong (b +ₖ_) (lCancelₖ a) ⟩ + b +ₖ 0ₖ ∎) + + helperFun : ∀ {ℓ} {A : Type ℓ} {a b : A} (p p' : a ≡ b) (q q' : b ≡ b) (r : a ≡ a) + → ((p q : a ≡ a) → p ∙ q ≡ q ∙ p) + → q ≡ q' + → PathP (λ i → p' (~ i) ≡ p' (~ i)) q' r + → p ∙ q ∙ sym p ≡ r + helperFun p p' q q' r comm qid dep = + p ∙ q ∙ sym p ≡⟨ cong (λ x → p ∙ x ∙ sym p) qid + ⟩≡⟨ cong (λ x → p ∙ x ∙ sym p) (λ i → lUnit (rUnit q' i) i) ⟩ + p ∙ (refl ∙ q' ∙ refl) ∙ sym p ≡⟨ cong (λ x → p ∙ x ∙ sym p) (λ i → (λ j → p' (~ i ∨ ~ j)) ∙ dep i ∙ λ j → p' (~ i ∨ j)) + ⟩≡⟨ assoc p (sym p' ∙ r ∙ p') (sym p) ⟩ + (p ∙ sym p' ∙ r ∙ p') ∙ sym p ≡⟨ cong (_∙ sym p) (assoc p (sym p') (r ∙ p')) + ⟩≡⟨ sym (assoc (p ∙ sym p') (r ∙ p') (sym p)) ⟩ + (p ∙ sym p') ∙ (r ∙ p') ∙ sym p ≡⟨ cong ((p ∙ sym p') ∙_) (sym (assoc r p' (sym p))) + ⟩≡⟨ cong (λ x → (p ∙ sym p') ∙ r ∙ x) (sym (symDistr p (sym p'))) ⟩ + (p ∙ sym p') ∙ r ∙ sym (p ∙ sym p') ≡⟨ cong ((p ∙ sym p') ∙_) (comm r (sym (p ∙ sym p'))) + ⟩≡⟨ assoc (p ∙ sym p') (sym (p ∙ sym p')) r ⟩ + ((p ∙ sym p') ∙ sym (p ∙ sym p')) ∙ r ≡⟨ cong (_∙ r) (rCancel (p ∙ sym p')) ⟩≡⟨ sym (lUnit r) ⟩ + r ∎ + + together : Path (_ ≡ _) (cong (a north +ₖ_) (sym (rCancelₖ (a north))) + ∙ ((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north))) + ∙ cong (a north +ₖ_) (rCancelₖ (a north))) + (cong(_+ₖ 0ₖ) (cong a (merid south) ∙ cong a (sym (merid north)))) + together = + helperFun (cong (a north +ₖ_) (sym (rCancelₖ (a north)))) + (λ i → pathHelper (a north) (a north) (~ i)) + ((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north))) + (((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north)))) + (cong(_+ₖ 0ₖ) (cong a (merid south) ∙ cong a (sym (merid north)))) + (isCommΩK-based 2 _) + refl + λ i → hcomp (λ k → λ {(i = i0) → (λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) + ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north)) + ; (i = i1) → congFunct (_+ₖ 0ₖ) (cong a (merid south)) (cong a (sym (merid north))) (~ k)}) + ((λ j → pathHelper (a north) (a (merid south j)) i) ∙ (λ j → pathHelper (a north) (a (merid north (~ j))) i)) + + + + +-- The translation mention above uses the basechange function. + +---------- lemmas on the baschange of ΩS¹ ---------- + +--The following lemma is used to prove the basechange2⁻ preserves +-- path composition (in a more general sense than what is proved in basechange2⁻-morph) + +basechange-lemma : ∀ {ℓ} {A : Type ℓ} {a : A} (x y : S¹) (F : a ≡ a → S¹) (f : S¹ → a ≡ a) (g : S¹ → a ≡ a) + → (f base ≡ refl) + → (g base ≡ refl) + → basechange2⁻ (F (f base ∙ g base)) (cong₂ {A = S¹} {B = λ x → S¹} (λ x y → F (f x ∙ g y)) loop loop) + ≡ basechange2⁻ (F (f base)) (cong (λ x → F (f x)) loop) ∙ basechange2⁻ (F (g base)) (cong (λ x → F (g x)) loop) +basechange-lemma x y F f g frefl grefl = + ((λ i → basechange2⁻ (F (f base ∙ g base)) (cong₂Funct (λ x y → F (f x ∙ g y)) loop loop i)) + ∙∙ (λ i → basechange2⁻ (F (f base ∙ g base)) (cong (λ x₁ → F (f x₁ ∙ g base)) loop ∙ cong (λ y₁ → F (f base ∙ g y₁)) loop)) + ∙∙ basechange2⁻-morph (F (f base ∙ g base)) _ _) + ∙∙ (λ j → basechange2⁻ (F (f base ∙ grefl j)) + (λ i → F (f (loop i) ∙ grefl j)) + ∙ basechange2⁻ (F (frefl j ∙ g base)) + (λ i → F (frefl j ∙ g (loop i)))) + ∙∙ ((λ j → basechange2⁻ (F (rUnit (f base) (~ j))) + (λ i → F (rUnit (f (loop i)) (~ j))) + ∙ basechange2⁻ (F (lUnit (g base) (~ j))) + (λ i → F (lUnit (g (loop i)) (~ j))))) + + +basechange-lemma2 : (f g : S¹ → hLevelTrunc 3 (S₊ 1)) (F : hLevelTrunc 3 (S₊ 1) → S¹) + → ((basechange2⁻ (F (f base +ₖ g base)) λ i → F ((f (loop i)) +ₖ g (loop i))) + ≡ basechange2⁻ (F (f base)) (cong (F ∘ f) loop) + ∙ basechange2⁻ (F (g base)) (cong (F ∘ g) loop)) +basechange-lemma2 f g F = coInd (f base) (g base) refl refl + where + coInd : (x y : hLevelTrunc 3 (S₊ 1)) + → f base ≡ x + → g base ≡ y + → ((basechange2⁻ (F (f base +ₖ g base)) λ i → F ((f (loop i)) +ₖ g (loop i))) + ≡ basechange2⁻ (F (f base)) (cong (F ∘ f) loop) + ∙ basechange2⁻ (F (g base)) (cong (F ∘ g) loop)) + coInd = + elim2 (λ _ _ → isGroupoidΠ2 λ _ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 (isGroupoidS¹ base base)) _ _ ) + (suspToPropRec2 north (λ _ _ → isPropΠ2 λ _ _ → isGroupoidS¹ _ _ _ _) + λ fb gb → basechange-lemma base base (F ∘ ΩKn+1→Kn) (Kn→ΩKn+1 1 ∘ f) (Kn→ΩKn+1 1 ∘ g) + (cong (Kn→ΩKn+1 1) fb ∙ Kn→ΩKn+10ₖ 1) + (cong (Kn→ΩKn+1 1) gb ∙ Kn→ΩKn+10ₖ 1) + ∙ cong₂ (_∙_) (λ j i → basechange2⁻ (F (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f base) j)) + (cong (λ x → F (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f x) j)) loop) i) + λ j i → basechange2⁻ (F (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g base) j)) + (cong (λ x → F (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g x) j)) loop) i) diff --git a/Cubical/ZCohomology/Groups/Sn.agda b/Cubical/ZCohomology/Groups/Sn.agda index 4263185a4..4b4a21996 100644 --- a/Cubical/ZCohomology/Groups/Sn.agda +++ b/Cubical/ZCohomology/Groups/Sn.agda @@ -7,8 +7,10 @@ open import Cubical.ZCohomology.MayerVietorisUnreduced open import Cubical.ZCohomology.Groups.Unit open import Cubical.ZCohomology.Groups.Connected open import Cubical.ZCohomology.KcompPrelims +open import Cubical.ZCohomology.Groups.Prelims open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv open import Cubical.Foundations.Function open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed @@ -20,7 +22,7 @@ open import Cubical.HITs.Sn open import Cubical.HITs.S1 open import Cubical.HITs.Susp open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) -open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₋₁ ; ∣_∣ to ∣_∣₋₁) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₁ ; ∣_∣ to ∣_∣₁) open import Cubical.HITs.Nullification open import Cubical.Data.Sigma @@ -28,17 +30,19 @@ open import Cubical.Data.Int renaming (_+_ to _+ℤ_; +-comm to +ℤ-comm ; +-as open import Cubical.Data.Nat open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec) open import Cubical.Data.Unit + open import Cubical.Structures.Group -open import Cubical.Homotopy.Connected -open import Cubical.Foundations.Equiv +open GroupEquiv +open vSES -Sn-connected : (n : ℕ) (x : S₊ (suc n)) → ∥ north ≡ x ∥₋₁ -Sn-connected n = suspToPropRec north (λ _ → propTruncIsProp) ∣ refl ∣₋₁ + +Sn-connected : (n : ℕ) (x : S₊ (suc n)) → ∥ north ≡ x ∥₁ +Sn-connected n = suspToPropRec north (λ _ → propTruncIsProp) ∣ refl ∣₁ H⁰-Sⁿ≅ℤ : (n : ℕ) → GroupEquiv (coHomGr 0 (S₊ (suc n))) intGroup -H⁰-Sⁿ≅ℤ n = GrIsoToGrEquiv (H⁰-connected' north (Sn-connected n)) +H⁰-Sⁿ≅ℤ n = H⁰-connected north (Sn-connected n) -- ---------------------------------------------------------------------- @@ -55,20 +59,18 @@ S0→Int a north = fst a S0→Int a south = snd a H⁰-S⁰≅ℤ×ℤ : GroupEquiv (coHomGr 0 (S₊ 0)) (dirProd intGroup intGroup) -GroupEquiv.eq H⁰-S⁰≅ℤ×ℤ = +eq H⁰-S⁰≅ℤ×ℤ = isoToEquiv (iso (sRec (isOfHLevelΣ 2 isSetInt λ _ → isSetInt) λ f → (f north) , (f south)) (λ a → ∣ S0→Int a ∣₂) (λ _ → refl) (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ f → cong ∣_∣₂ (funExt (λ {north → refl ; south → refl})))) -GroupEquiv.isHom H⁰-S⁰≅ℤ×ℤ = +isHom H⁰-S⁰≅ℤ×ℤ = sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelΣ 2 isSetInt (λ _ → isSetInt)) _ _) λ a b i → addLemma (a north) (b north) i , addLemma (a south) (b south) i -- --------------------------H¹(S¹) ----------------------------------- - - {- In order to apply Mayer-Vietoris, we need the following lemma. Given the following diagram @@ -77,7 +79,6 @@ Given the following diagram If ψ is an isomorphism and ϕ is surjective with ker ϕ ≡ {ψ (a , a) ∣ a ∈ A}, then C ≅ B -} -open GroupEquiv diagonalIso : ∀ {ℓ ℓ' ℓ''} {A : Group {ℓ}} (B : Group {ℓ'}) {C : Group {ℓ''}} (ψ : GroupEquiv (dirProd A A) B) (ϕ : GroupHom B C) → isSurjective _ _ ϕ @@ -87,8 +88,8 @@ diagonalIso : ∀ {ℓ ℓ' ℓ''} {A : Group {ℓ}} (B : Group {ℓ'}) {C : Gro → isInKer B C ϕ x) → GroupEquiv A C diagonalIso {A = A} B {C = C} ψ ϕ issurj ker→diag diag→ker = - Iso'ToGroupEquiv - (iso' (compGroupHom fstProj (compGroupHom (grouphom (fst (eq ψ)) (isHom ψ)) ϕ)) + BijectionIsoToGroupEquiv + (bij-iso (compGroupHom fstProj (compGroupHom (grouphom (fst (eq ψ)) (isHom ψ)) ϕ)) (λ a inker → pRec (isSetCarrier A _ _) (λ {(a' , id) → cong fst (sym (secEq (eq ψ) (a , 0g A)) ∙∙ cong (invEq (eq ψ)) id ∙∙ secEq (eq ψ) (a' , a')) @@ -99,7 +100,7 @@ diagonalIso {A = A} B {C = C} ψ ϕ issurj ker→diag diag→ker = , (sym (Group.rid C _) ∙∙ cong ((fun ϕ) (equivFun (eq ψ) (fst (ψ⁻ b) A.+ (A.- snd (ψ⁻ b)) , 0g A)) C.+_) (sym (diag→ker (equivFun (eq ψ) ((snd (ψ⁻ b)) , (snd (ψ⁻ b)))) - ∣ (snd (ψ⁻ b)) , refl ∣₋₁)) + ∣ (snd (ψ⁻ b)) , refl ∣₁)) ∙∙ sym ((isHom ϕ) _ _)) ∙∙ cong (fun ϕ) (sym ((isHom ψ) _ _) ∙∙ cong (equivFun (eq ψ)) (ΣPathP (sym (Group.assoc A _ _ _) @@ -107,7 +108,7 @@ diagonalIso {A = A} B {C = C} ψ ϕ issurj ker→diag diag→ker = ∙∙ Group.rid A _ , (Group.lid A _))) ∙∙ retEq (eq ψ) b) - ∙∙ id ∣₋₁ }) + ∙∙ id ∣₁ }) (issurj c)) where open Group @@ -122,80 +123,82 @@ diagonalIso {A = A} B {C = C} ψ ϕ issurj ker→diag diag→ker = ψ⁻ = fst (invEquiv (eq ψ)) fstProj : GroupHom A (dirProd A A) - GroupHom.fun fstProj a = a , 0g A - GroupHom.isHom fstProj g0 g1 i = (g0 A.+ g1) , Group.lid A (0g A) (~ i) + fun fstProj a = a , 0g A + isHom fstProj g0 g1 i = (g0 A.+ g1) , Group.lid A (0g A) (~ i) H¹-S¹≅ℤ : GroupEquiv intGroup (coHomGr 1 (S₊ 1)) H¹-S¹≅ℤ = diagonalIso (coHomGr 0 (S₊ 0)) (invGroupEquiv H⁰-S⁰≅ℤ×ℤ) - (MV.d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0) - (λ x → MV.Ker-i⊂Im-d _ _(S₊ 0) (λ _ → tt) (λ _ → tt) 0 x + (I.d 0) + (λ x → I.Ker-i⊂Im-d 0 x (ΣPathP (isOfHLevelSuc 0 (isContrHⁿ-Unit 0) _ _ , isOfHLevelSuc 0 (isContrHⁿ-Unit 0) _ _))) ((sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) (λ x inker → pRec propTruncIsProp (λ {((f , g) , id') → helper x f g id' inker}) - ((MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ x ∣₂ inker))))) + ((I.Ker-d⊂Im-Δ 0 ∣ x ∣₂ inker))))) ((sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ F surj → pRec (setTruncIsSet _ _) - (λ { (x , id) → MV.Im-Δ⊂Ker-d _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ F ∣₂ + (λ { (x , id) → I.Im-Δ⊂Ker-d 0 ∣ F ∣₂ ∣ (∣ (λ _ → x) ∣₂ , ∣ (λ _ → 0) ∣₂) , - (cong ∣_∣₂ (funExt (surjHelper x))) ∙ sym id ∣₋₁ }) + (cong ∣_∣₂ (funExt (surjHelper x))) ∙ sym id ∣₁ }) surj) ) □ invGroupEquiv (coHomPushout≅coHomSn 0 1) where + module I = MV Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) surjHelper : (x : Int) (x₁ : S₊ 0) → x +ₖ (-ₖ pos 0) ≡ S0→Int (x , x) x₁ surjHelper x north = cong (x +ₖ_) (-0ₖ) ∙ rUnitₖ x surjHelper x south = cong (x +ₖ_) (-0ₖ) ∙ rUnitₖ x helper : (F : S₊ 0 → Int) (f g : ∥ (Unit → Int) ∥₂) - (id : GroupHom.fun (MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) (f , g) ≡ ∣ F ∣₂) + (id : GroupHom.fun (I.Δ 0) (f , g) ≡ ∣ F ∣₂) → isInKer (coHomGr 0 (S₊ 0)) (coHomGr 1 (Pushout (λ _ → tt) (λ _ → tt))) - (MV.d Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) + (I.d 0) ∣ F ∣₂ → ∃[ x ∈ Int ] ∣ F ∣₂ ≡ equivFun (invEquiv (eq H⁰-S⁰≅ℤ×ℤ)) (x , x) helper F = sElim2 (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) λ f g id inker → pRec propTruncIsProp - (λ {((a , b) , id2) - → sElim2 {B = λ f g → GroupHom.fun (MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) (f , g) ≡ ∣ F ∣₂ → _ } + (λ ((a , b) , id2) + → sElim2 {B = λ f g → GroupHom.fun (I.Δ 0) (f , g) ≡ ∣ F ∣₂ → _ } (λ _ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) - (λ f g id → ∣ (helper2 f g .fst) , (sym id ∙ sym (helper2 f g .snd)) ∣₋₁) - a b id2}) + (λ f g id → ∣ (helper2 f g .fst) , (sym id ∙ sym (helper2 f g .snd)) ∣₁) + a b id2) (MV.Ker-d⊂Im-Δ _ _ (S₊ 0) (λ _ → tt) (λ _ → tt) 0 ∣ F ∣₂ inker) where helper2 : (f g : Unit → Int) → Σ[ x ∈ Int ] (equivFun (invEquiv (eq H⁰-S⁰≅ℤ×ℤ))) (x , x) - ≡ GroupHom.fun (MV.Δ Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) 0) (∣ f ∣₂ , ∣ g ∣₂) + ≡ GroupHom.fun (I.Δ 0) (∣ f ∣₂ , ∣ g ∣₂) helper2 f g = (f _ +ₖ (-ₖ g _) ) , cong ∣_∣₂ (funExt (λ {north → refl ; south → refl})) ------------------------- H¹(S⁰) ≅ 0 ------------------------------- - -H¹-S⁰≅0 : GroupEquiv (coHomGr 1 (S₊ 0)) trivialGroup -eq H¹-S⁰≅0 = - isoToEquiv - (iso (λ _ → tt) - (λ _ → 0ₕ) - (λ _ → refl) - (sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) - (λ f → pRec (setTruncIsSet _ _) - (λ id → cong ∣_∣₂ (sym id)) - (helper f (f north) (f south) refl refl)))) - where - helper : (f : S₊ 0 → coHomK 1) → (x y : coHomK 1) → (f north ≡ x) → (f south ≡ y) → ∥ f ≡ (λ _ → 0ₖ) ∥₋₁ - helper f = - elim2 (λ _ _ → isOfHLevelΠ 3 λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelPlus {n = 1} 2 propTruncIsProp) - (suspToPropRec2 north (λ _ _ → isOfHLevelΠ 1 λ _ → isOfHLevelΠ 1 λ _ → propTruncIsProp) - λ id id2 → ∣ funExt (λ {north → id ; south → id2}) ∣₋₁) -isHom H¹-S⁰≅0 _ _ = refl +private + Hⁿ-S0≃Kₙ×Kₙ : (n : ℕ) → Iso (S₊ 0 → coHomK (suc n)) (coHomK (suc n) × coHomK (suc n)) + Iso.fun (Hⁿ-S0≃Kₙ×Kₙ n) f = (f north) , (f south) + Iso.inv (Hⁿ-S0≃Kₙ×Kₙ n) (a , b) north = a + Iso.inv (Hⁿ-S0≃Kₙ×Kₙ n) (a , b) south = b + Iso.rightInv (Hⁿ-S0≃Kₙ×Kₙ n) a = refl + Iso.leftInv (Hⁿ-S0≃Kₙ×Kₙ n) b = funExt λ {north → refl ; south → refl} + + isContrHⁿ-S0 : (n : ℕ) → isContr (coHom (suc n) (S₊ 0)) + isContrHⁿ-S0 n = + transport (λ i → isContr ∥ isoToPath (Hⁿ-S0≃Kₙ×Kₙ n) (~ i) ∥₂) + (transport (λ i → isContr (isoToPath (setTruncOfProdIso {A = coHomK (suc n)} {B = coHomK (suc n)} ) (~ i))) + ((∣ 0ₖ ∣₂ , ∣ 0ₖ ∣₂) + , prodElim (λ _ → isOfHLevelSuc 1 (isOfHLevelΣ 2 setTruncIsSet (λ _ → setTruncIsSet) _ _)) + (elim2 (λ _ _ → isProp→isOfHLevelSuc (2 + n) (isOfHLevelΣ 2 setTruncIsSet (λ _ → setTruncIsSet) _ _)) + (suspToPropRec2 north (λ _ _ → isOfHLevelΣ 2 setTruncIsSet (λ _ → setTruncIsSet) _ _) refl)))) + +H¹-S⁰≅0 : (n : ℕ) → GroupEquiv (coHomGr (suc n) (S₊ 0)) trivialGroup +H¹-S⁰≅0 n = contrGroup≅trivialGroup (isContrHⁿ-S0 n) ------------------------- H²(S¹) ≅ 0 ------------------------------- @@ -203,32 +206,29 @@ H²-S¹≅0 : GroupEquiv (coHomGr 2 (S₊ 1)) trivialGroup H²-S¹≅0 = coHomPushout≅coHomSn 0 2 □ (invGroupEquiv (vSES→GroupEquiv _ _ vSES-helper)) - □ H¹-S⁰≅0 + □ (H¹-S⁰≅0 0) where module I = MV Unit Unit (S₊ 0) (λ _ → tt) (λ _ → tt) vSES-helper : vSES (coHomGr 1 (S₊ 0)) (coHomGr 2 (Pushout (λ _ → tt) (λ _ → tt))) _ _ - vSES.isTrivialLeft vSES-helper = isOfHLevelSuc 0 (isOfHLevelΣ 0 (isContrHⁿ-Unit 0) (λ _ → isContrHⁿ-Unit 0)) - vSES.isTrivialRight vSES-helper = isOfHLevelSuc 0 (isOfHLevelΣ 0 (isContrHⁿ-Unit 1) (λ _ → isContrHⁿ-Unit 1)) - vSES.left vSES-helper = I.Δ 1 - vSES.right vSES-helper = I.i 2 + isTrivialLeft vSES-helper = isOfHLevelSuc 0 (isOfHLevelΣ 0 (isContrHⁿ-Unit 0) (λ _ → isContrHⁿ-Unit 0)) + isTrivialRight vSES-helper = isOfHLevelSuc 0 (isOfHLevelΣ 0 (isContrHⁿ-Unit 1) (λ _ → isContrHⁿ-Unit 1)) + left vSES-helper = I.Δ 1 + right vSES-helper = I.i 2 vSES.ϕ vSES-helper = I.d 1 - vSES.Ker-ϕ⊂Im-left vSES-helper = I.Ker-d⊂Im-Δ 1 - vSES.Ker-right⊂Im-ϕ vSES-helper = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) -- doesn't terminate without elimination + Ker-ϕ⊂Im-left vSES-helper = I.Ker-d⊂Im-Δ 1 + Ker-right⊂Im-ϕ vSES-helper = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) -- doesn't terminate without elimination λ a → I.Ker-i⊂Im-d 1 ∣ a ∣₂ - -------------------------------------------------------------------- - --------------- H¹(S²) -------------------------------------------- H¹-S²≅0 : GroupEquiv (coHomGr 1 (S₊ 2)) trivialGroup H¹-S²≅0 = coHomPushout≅coHomSn 1 1 - □ Iso'ToGroupEquiv - (iso' (I.i 1) - helper - λ x → ∣ 0ₕ , isOfHLevelSuc 0 (isOfHLevelΣ 0 (isContrHⁿ-Unit 0) (λ _ → isContrHⁿ-Unit 0)) _ x ∣₋₁) + □ BijectionIsoToGroupEquiv + (bij-iso (I.i 1) + helper + λ x → ∣ 0ₕ , isOfHLevelSuc 0 (isOfHLevelΣ 0 (isContrHⁿ-Unit 0) (λ _ → isContrHⁿ-Unit 0)) _ x ∣₁) □ dirProdEquiv (Hⁿ-Unit≅0 0) (Hⁿ-Unit≅0 0) □ lUnitGroupIso where @@ -239,7 +239,7 @@ H¹-S²≅0 = λ f → ∣ (∣ (λ _ → f north) ∣₂ , 0ₕ) , (cong ∣_∣₂ (funExt (suspToPropRec north (λ _ → isSetInt _ _) - (cong (f north +ₖ_) -0ₖ ∙ rUnitₖ (f north))))) ∣₋₁ + (cong (f north +ₖ_) -0ₖ ∙ rUnitₖ (f north))))) ∣₁ helper : isInjective _ _ (I.i 1) helper = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 (setTruncIsSet _ _)) -- useless elimination speeds things up significantly @@ -251,127 +251,56 @@ H¹-S²≅0 = --------- Direct proof of H¹(S¹) ≅ ℤ without Mayer-Vietoris ------- -S¹map : hLevelTrunc 3 S¹ → S¹ -S¹map = trRec isGroupoidS¹ (idfun _) - -S¹map-id : (x : hLevelTrunc 3 S¹) → Path (hLevelTrunc 3 S¹) ∣ S¹map x ∣ x -S¹map-id = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) - λ a → refl - -S1map : hLevelTrunc 3 (S₊ 1) → (S₊ 1) -S1map = trRec isGroupoidS1 (idfun _) - - -S¹→S¹≡S¹×Int : Iso (S¹ → hLevelTrunc 3 S¹) (S¹ × Int) -Iso.fun S¹→S¹≡S¹×Int f = S¹map (f base) - , winding (basechange2⁻ (S¹map (f base)) λ i → S¹map (f (loop i))) -Iso.inv S¹→S¹≡S¹×Int (s , int) base = ∣ s ∣ -Iso.inv S¹→S¹≡S¹×Int (s , int) (loop i) = ∣ basechange2 s (intLoop int) i ∣ -Iso.rightInv S¹→S¹≡S¹×Int (s , int) = ΣPathP (refl , ((λ i → winding (basechange2-retr s (λ i → intLoop int i) i)) - ∙ windingIntLoop int)) -Iso.leftInv S¹→S¹≡S¹×Int f = funExt λ { base → S¹map-id (f base) - ; (loop i) j → helper2 j i} - where - helper2 : PathP (λ i → S¹map-id (f base) i ≡ S¹map-id (f base) i) - (λ i → ∣ basechange2 (S¹map (f base)) - (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁)))))) i ∣) - (cong f loop) - helper2 i j = - hcomp (λ k → λ { (i = i0) → cong ∣_∣ (basechange2 (S¹map (f base)) - (intLoop (winding (basechange2⁻ (S¹map (f base)) (λ i₁ → S¹map (f (loop i₁))))))) j - ; (i = i1) → S¹map-id (f (loop j)) k - ; (j = i0) → S¹map-id (f base) (i ∧ k) - ; (j = i1) → S¹map-id (f base) (i ∧ k)}) - (helper4 i j) - where - helper4 : Path (Path (hLevelTrunc 3 _) _ _) - (cong ∣_∣ (basechange2 (S¹map (f base)) - (intLoop - (winding - (basechange2⁻ (S¹map (f base)) - (λ i₁ → S¹map (f (loop i₁)))))))) - λ i → ∣ S¹map (f (loop i)) ∣ - helper4 i = - cong ∣_∣ - ((cong (basechange2 (S¹map (f base))) - (decodeEncode base (basechange2⁻ (S¹map (f base)) - (λ i₁ → S¹map (f (loop i₁))))) - ∙ basechange2-sect (S¹map (f base)) - (λ i → S¹map (f (loop i)))) i) - -S¹→S¹≡S1→S1 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → hLevelTrunc 3 S¹) -Iso.fun S¹→S¹≡S1→S1 f x = trMap S1→S¹ (f (S¹→S1 x)) -Iso.inv S¹→S¹≡S1→S1 f x = trMap S¹→S1 (f (S1→S¹ x)) -Iso.rightInv S¹→S¹≡S1→S1 F = funExt λ x i → helper2 (F (S1→S¹-sect x i)) i - where - helper2 : (x : hLevelTrunc 3 S¹) → trMap S1→S¹ (trMap S¹→S1 x) ≡ x - helper2 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) - λ a → cong ∣_∣ (S1→S¹-sect a) -Iso.leftInv S¹→S¹≡S1→S1 F = funExt λ x i → helper2 (F (S1→S¹-retr x i)) i - where - helper2 : (x : hLevelTrunc 3 (S₊ 1)) → trMap S¹→S1 (trMap S1→S¹ x) ≡ x - helper2 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) - λ a → cong ∣_∣ (S1→S¹-retr a) - - - - -basechange-lemma : ∀ {ℓ} {A : Type ℓ} {a : A} (x y : S¹) (F : a ≡ a → S¹) (f : S¹ → a ≡ a) (g : S¹ → a ≡ a) - → (f base ≡ refl) - → (g base ≡ refl) - → basechange2⁻ (F (f base ∙ g base)) (cong₂ {A = S¹} {B = λ x → S¹} (λ x y → F (f x ∙ g y)) loop loop) - ≡ basechange2⁻ (F (f base)) (cong (λ x → F (f x)) loop) ∙ basechange2⁻ (F (g base)) (cong (λ x → F (g x)) loop) -basechange-lemma x y F f g frefl grefl = - ((λ i → basechange2⁻ (F (f base ∙ g base)) (cong₂Funct (λ x y → F (f x ∙ g y)) loop loop i)) - ∙∙ (λ i → basechange2⁻ (F (f base ∙ g base)) (cong (λ x₁ → F (f x₁ ∙ g base)) loop ∙ cong (λ y₁ → F (f base ∙ g y₁)) loop)) - ∙∙ basechange2⁻-morph (F (f base ∙ g base)) _ _) - ∙∙ (λ j → basechange2⁻ (F (f base ∙ grefl j)) - (λ i → F (f (loop i) ∙ grefl j)) - ∙ basechange2⁻ (F (frefl j ∙ g base)) - (λ i → F (frefl j ∙ g (loop i)))) - ∙∙ ((λ j → basechange2⁻ (F (rUnit (f base) (~ j))) - (λ i → F (rUnit (f (loop i)) (~ j))) - ∙ basechange2⁻ (F (lUnit (g base) (~ j))) - (λ i → F (lUnit (g (loop i)) (~ j))))) +-- The strategy is to use the proof that ΩS¹ ≃ ℤ. Since we only have this for S¹ with the base/loop definition +-- we begin with some functions translating between H¹(S₊ 1) and ∥ S¹ → S¹ ∥₀. The latter type is easy to characterise, +-- by (S¹ → S¹) ≃ S¹ × ℤ (see Cubical.ZCohomology.Groups.Prelims). Truncating this leaves only ℤ, since S¹ is connected. +-- The translation mention above uses the basechange function. We use basechange-lemma (Cubical.ZCohomology.Groups.Prelims) to prove the basechange2⁻ preserves +-- path composition (in a more general sense than what is proved in basechange2⁻-morph) +-- We can now give the group equivalence. The first bit is just a big composition of our previously defined translations and is pretty uninteresting. +-- The harder step is proving that the equivalence is a morphism. This relies heavily on the fact that addition the cohomology groups essentially is defined using an +-- application of cong₂, which allows us to use basechange-lemma. coHom1S1≃ℤ : GroupEquiv (coHomGr 1 (S₊ 1)) intGroup eq coHom1S1≃ℤ = isoToEquiv - (iso (sRec isSetInt (λ f → snd (Iso.fun S¹→S¹≡S¹×Int (Iso.fun S¹→S¹≡S1→S1 f)))) -- fun - (λ a → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹≡S¹×Int (base , a)) ∣₂) -- inv - (λ a → (λ i → snd (Iso.fun S¹→S¹≡S¹×Int (Iso.fun S¹→S¹≡S1→S1 (Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹≡S¹×Int (base , a)))))) -- rInv - ∙ (λ i → snd (Iso.fun S¹→S¹≡S¹×Int (Iso.rightInv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹≡S¹×Int (base , a)) i))) + (iso (sRec isSetInt (λ f → snd (Iso.fun S¹→S¹≡S¹×Int (Iso.fun S1→S1≡S¹→S¹ f)))) -- fun + (λ a → ∣ Iso.inv S1→S1≡S¹→S¹ (Iso.inv S¹→S¹≡S¹×Int (base , a)) ∣₂) -- inv + (λ a → (λ i → snd (Iso.fun S¹→S¹≡S¹×Int (Iso.fun S1→S1≡S¹→S¹ (Iso.inv S1→S1≡S¹→S¹ (Iso.inv S¹→S¹≡S¹×Int (base , a)))))) -- rInv + ∙ (λ i → snd (Iso.fun S¹→S¹≡S¹×Int (Iso.rightInv S1→S1≡S¹→S¹ (Iso.inv S¹→S¹≡S¹×Int (base , a)) i))) ∙ λ i → snd (Iso.rightInv S¹→S¹≡S¹×Int (base , a) i)) ((sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) -- lInv - λ f → (λ i → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.inv S¹→S¹≡S¹×Int (base , snd (Iso.fun S¹→S¹≡S¹×Int (Iso.fun S¹→S¹≡S1→S1 f)))) ∣₂) + λ f → (λ i → ∣ Iso.inv S1→S1≡S¹→S¹ (Iso.inv S¹→S¹≡S¹×Int (base , snd (Iso.fun S¹→S¹≡S¹×Int (Iso.fun S1→S1≡S¹→S¹ f)))) ∣₂) ∙∙ ((λ i → sRec setTruncIsSet - (λ x → ∣ Iso.inv S¹→S¹≡S1→S1 x ∣₂) + (λ x → ∣ Iso.inv S1→S1≡S¹→S¹ x ∣₂) (sRec setTruncIsSet - (λ x → ∣ Iso.inv S¹→S¹≡S¹×Int (x , (snd (Iso.fun S¹→S¹≡S¹×Int (Iso.fun S¹→S¹≡S1→S1 f)))) ∣₂) - (Iso.inv PathIdTrunc₀Iso (isConnectedS¹ (fst (Iso.fun S¹→S¹≡S¹×Int (Iso.fun S¹→S¹≡S1→S1 f)))) i))) - ∙ (λ i → ∣ Iso.inv S¹→S¹≡S1→S1 (Iso.leftInv S¹→S¹≡S¹×Int (Iso.fun S¹→S¹≡S1→S1 f) i) ∣₂)) - ∙∙ (λ i → ∣ Iso.leftInv S¹→S¹≡S1→S1 f i ∣₂)))) + (λ x → ∣ Iso.inv S¹→S¹≡S¹×Int (x , (snd (Iso.fun S¹→S¹≡S¹×Int (Iso.fun S1→S1≡S¹→S¹ f)))) ∣₂) + (Iso.inv PathIdTrunc₀Iso (isConnectedS¹ (fst (Iso.fun S¹→S¹≡S¹×Int (Iso.fun S1→S1≡S¹→S¹ f)))) i))) + ∙ (λ i → ∣ Iso.inv S1→S1≡S¹→S¹ (Iso.leftInv S¹→S¹≡S¹×Int (Iso.fun S1→S1≡S¹→S¹ f) i) ∣₂)) + ∙∙ (λ i → ∣ Iso.leftInv S1→S1≡S¹→S¹ f i ∣₂)))) isHom coHom1S1≃ℤ = - (sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) -- isMorph - λ f g → (λ i → winding (basechange2⁻ (S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)) ∙ Kn→ΩKn+1 1 (g (S¹→S1 base)))))) - (λ i → S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 (loop i))) ∙ Kn→ΩKn+1 1 (g (S¹→S1 (loop i))))))))) + sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) -- isMorph + λ f g → (λ i → winding (guy (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)) ∙ Kn→ΩKn+1 1 (g (S¹→S1 base)))) + (λ i → pre-guy (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 (loop i))) ∙ Kn→ΩKn+1 1 (g (S¹→S1 (loop i)))))))) ∙∙ cong winding (helper (f (S¹→S1 base)) (g (S¹→S1 base)) f g refl refl) - ∙∙ winding-hom ((basechange2⁻ (S¹map (trMap S1→S¹ (f north))) - (λ i → S¹map (trMap S1→S¹ (f (S¹→S1 (loop i))))))) - ((basechange2⁻ (S¹map (trMap S1→S¹ (g north))) - (λ i → S¹map (trMap S1→S¹ (g (S¹→S1 (loop i)))))))) + ∙∙ winding-hom (guy (f north) + (λ i → pre-guy (f (S¹→S1 (loop i))))) + (guy (g north) + (λ i → pre-guy (g (S¹→S1 (loop i))))) where + pre-guy = S¹map ∘ trMap S1→S¹ + guy = basechange2⁻ ∘ pre-guy + helper : (x y : coHomK 1) (f g : S₊ 1 → coHomK 1) → (f (S¹→S1 base)) ≡ x → (g (S¹→S1 base)) ≡ y - → (basechange2⁻ (S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)) ∙ Kn→ΩKn+1 1 (g (S¹→S1 base))))))) - (λ i → S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 (loop i))) ∙ Kn→ΩKn+1 1 (g (S¹→S1 (loop i))))))) - ≡ (basechange2⁻ (S¹map (trMap S1→S¹ ((f (S¹→S1 base)))))) - (λ i → S¹map (trMap S1→S¹ (f (S¹→S1 (loop i))))) - ∙ (basechange2⁻ (S¹map (trMap S1→S¹ ((g (S¹→S1 base))))) - (λ i → S¹map (trMap S1→S¹ ((g (S¹→S1 (loop i))))))) + → (guy (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)) ∙ Kn→ΩKn+1 1 (g (S¹→S1 base)))) + (λ i → S¹map (trMap S1→S¹ (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 (loop i))) ∙ Kn→ΩKn+1 1 (g (S¹→S1 (loop i)))))))) + ≡ (guy (f (S¹→S1 base)) + (λ i → pre-guy (f (S¹→S1 (loop i))))) + ∙ (guy (g (S¹→S1 base)) + (λ i → pre-guy ((g (S¹→S1 (loop i)))))) helper = elim2 (λ _ _ → isGroupoidΠ4 λ _ _ _ _ → isOfHLevelPath 3 (isOfHLevelSuc 3 (isGroupoidS¹) base base) _ _) (suspToPropRec2 {A = S₊ 0} north (λ _ _ → isPropΠ4 λ _ _ _ _ → isGroupoidS¹ _ _ _ _) @@ -384,10 +313,10 @@ isHom coHom1S1≃ℤ = ((λ x → Kn→ΩKn+1 1 (g (S¹→S1 x)))) (cong (Kn→ΩKn+1 1) reflf ∙ Kn→ΩKn+10ₖ 1) (cong (Kn→ΩKn+1 1) reflg ∙ Kn→ΩKn+10ₖ 1)) - ∙ λ j → basechange2⁻ (S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 base)) j))) - (λ i → S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 (loop i))) j))) - ∙ basechange2⁻ (S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 base)) j))) - (λ i → S¹map (trMap S1→S¹ (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 (loop i))) j)))) + ∙ λ j → guy (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 base)) j) + (λ i → pre-guy (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 (loop i))) j)) + ∙ guy (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 base)) j) + (λ i → pre-guy (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 (loop i))) j))) @@ -405,12 +334,12 @@ Hⁿ-Sⁿ≅ℤ (suc n) = (Pushout {A = S₊ (suc n)} (λ _ → tt) (λ _ → tt))) _ _ - vSES.isTrivialLeft theIso p q = ΣPathP (isOfHLevelSuc 0 (isContrHⁿ-Unit n) (fst p) (fst q) + isTrivialLeft theIso p q = ΣPathP (isOfHLevelSuc 0 (isContrHⁿ-Unit n) (fst p) (fst q) , isOfHLevelSuc 0 (isContrHⁿ-Unit n) (snd p) (snd q)) - vSES.isTrivialRight theIso p q = ΣPathP (isOfHLevelSuc 0 (isContrHⁿ-Unit (suc n)) (fst p) (fst q) + isTrivialRight theIso p q = ΣPathP (isOfHLevelSuc 0 (isContrHⁿ-Unit (suc n)) (fst p) (fst q) , isOfHLevelSuc 0 (isContrHⁿ-Unit (suc n)) (snd p) (snd q)) - vSES.left theIso = I.Δ (suc n) - vSES.right theIso = I.i (2 + n) + left theIso = I.Δ (suc n) + right theIso = I.i (2 + n) vSES.ϕ theIso = I.d (suc n) - vSES.Ker-ϕ⊂Im-left theIso = I.Ker-d⊂Im-Δ (suc n) - vSES.Ker-right⊂Im-ϕ theIso = I.Ker-i⊂Im-d (suc n) + Ker-ϕ⊂Im-left theIso = I.Ker-d⊂Im-Δ (suc n) + Ker-right⊂Im-ϕ theIso = I.Ker-i⊂Im-d (suc n) diff --git a/Cubical/ZCohomology/Groups/Torus.agda b/Cubical/ZCohomology/Groups/Torus.agda index 28ddcea44..5d651a2f0 100644 --- a/Cubical/ZCohomology/Groups/Torus.agda +++ b/Cubical/ZCohomology/Groups/Torus.agda @@ -7,10 +7,12 @@ open import Cubical.ZCohomology.Groups.Connected open import Cubical.ZCohomology.MayerVietorisUnreduced open import Cubical.ZCohomology.Groups.Unit open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.Groups.Prelims open import Cubical.ZCohomology.KcompPrelims open import Cubical.Foundations.HLevels open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed open import Cubical.Foundations.Isomorphism @@ -27,238 +29,24 @@ open import Cubical.HITs.S1 open import Cubical.HITs.Sn open import Cubical.HITs.Susp open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) -open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim2 to pElim2 ; ∣_∣ to ∣_∣₋₁) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim2 to pElim2 ; ∣_∣ to ∣_∣₁) open import Cubical.HITs.Nullification -open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec) - - -{- Proof that S¹→K2 is isomorphic to K2×K1 (as types). Needed for H²(T²) -} -S1→K2≡K2×K1 : Iso (S₊ 1 → coHomK 2) (coHomK 2 × coHomK 1) -Iso.fun S1→K2≡K2×K1 f = f north , ΩKn+1→Kn (sym (rCancelₖ (f north)) - ∙ (λ i → f (merid south i) +ₖ (-ₖ f (merid north i))) - ∙ rCancelₖ (f south)) -Iso.inv S1→K2≡K2×K1 (a , b) north = a +ₖ 0ₖ -Iso.inv S1→K2≡K2×K1 (a , b) south = a +ₖ 0ₖ -Iso.inv S1→K2≡K2×K1 (a , b) (merid south i) = a +ₖ (Kn→ΩKn+1 1 b i) -Iso.inv S1→K2≡K2×K1 (a , b) (merid north i) = a +ₖ 0ₖ -Iso.rightInv S1→K2≡K2×K1 (a , b) = - ΣPathP ((rUnitₖ a) , - ((cong ΩKn+1→Kn (congHelper2 (Kn→ΩKn+1 1 b) (λ x → (a +ₖ x) +ₖ (-ₖ (a +ₖ 0ₖ))) (funExt (λ x → sym (cancelHelper a x))) (rCancelₖ (a +ₖ 0ₖ)))) - ∙ Iso.leftInv (Iso3-Kn-ΩKn+1 1) b)) - module _ where - cancelHelper : (a b : hLevelTrunc 4 (S₊ 2)) → (a +ₖ b) +ₖ (-ₖ (a +ₖ 0ₖ)) ≡ b - cancelHelper a b = - (a +ₖ b) +ₖ (-ₖ (a +ₖ 0ₖ)) ≡⟨ (λ i → (a +ₖ b) +ₖ (-ₖ (rUnitₖ a i))) ⟩≡⟨ cong (λ x → x +ₖ (-ₖ a)) (commₖ a b) ⟩ - (b +ₖ a) +ₖ (-ₖ a) ≡⟨ assocₖ b a (-ₖ a) ⟩≡⟨ cong (λ x → b +ₖ x) (rCancelₖ a) ⟩ - (b +ₖ 0ₖ ≡⟨ rUnitₖ b ⟩ - b ∎) - - moveabout : {x : hLevelTrunc 4 (S₊ 2)} (p q : x ≡ 0ₖ) (mid : 0ₖ ≡ 0ₖ) - → sym q ∙ (p ∙ mid ∙ sym p) ∙ q ≡ mid - moveabout p q mid = (assoc (sym q) _ _ - ∙∙ cong (_∙ q) (assoc (sym q) p (mid ∙ sym p)) - ∙∙ sym (assoc (sym q ∙ p) (mid ∙ sym p) q)) - ∙∙ (cong ((sym q ∙ p) ∙_) (sym (assoc mid (sym p) q)) - ∙∙ cong (λ x → (sym q ∙ p) ∙ (mid ∙ x)) (sym (symDistr (sym q) p)) - ∙∙ cong ((sym q ∙ p)∙_) (isCommΩK 2 mid _)) - ∙∙ (assoc _ _ _ - ∙∙ cong (_∙ mid) (rCancel (sym q ∙ p)) - ∙∙ sym (lUnit mid)) - - congHelper : ∀ {ℓ} {A : Type ℓ} {a1 : A} (p : a1 ≡ a1) (f : A → A) (id : (λ x → x) ≡ f) - → cong f p ≡ sym (funExt⁻ id a1) ∙ p ∙ funExt⁻ id a1 - congHelper {a1 = a1} p f id = - (λ i → lUnit (rUnit (cong f p) i) i) - ∙ (λ i → (λ j → id ((~ i) ∨ (~ j)) a1) ∙ cong (id (~ i)) p ∙ λ j → id (~ i ∨ j) a1) - - - congHelper2 : (p : 0ₖ ≡ 0ₖ) (f : coHomK 2 → coHomK 2) (id : (λ x → x) ≡ f) - → (q : (f 0ₖ) ≡ 0ₖ) - → (sym q) ∙ cong f p ∙ q ≡ p - congHelper2 p f id q = - cong (λ x → sym q ∙ x ∙ q) (congHelper p f id) ∙ - moveabout (sym (funExt⁻ id ∣ north ∣)) q p - -Iso.leftInv S1→K2≡K2×K1 a = - funExt λ {north → rUnitₖ (a north) - ; south → rUnitₖ (a north) ∙ cong a (merid north) - ; (merid south i) → m-south i - ; (merid north i) → m-north i} - where - - m-north : PathP (λ i → a north +ₖ ∣ north ∣ ≡ a (merid north i)) - (rUnitₖ (a north)) (rUnitₖ (a north) ∙ cong a (merid north)) - m-north i j = - hcomp (λ k → λ { (i = i0) → rUnitₖ (a north) j - ; (j = i1) → a (merid north (i ∧ k)) - ; (j = i0) → a north +ₖ ∣ north ∣}) - (rUnitₖ (a north) j) - - m-south : PathP (λ i → a north +ₖ Kn→ΩKn+1 1 (ΩKn+1→Kn (sym (rCancelₖ (a north)) - ∙ (λ i → a (merid south i) +ₖ (-ₖ a (merid north i))) - ∙ rCancelₖ (a south))) i ≡ a (merid south i)) (rUnitₖ (a north)) (rUnitₖ (a north) ∙ cong a (merid north)) - m-south j i = - hcomp (λ k → λ { (i = i0) → (helper ∙∙ together ∙∙ congFunct (_+ₖ 0ₖ) (cong a (merid south)) (cong a (sym (merid north)))) (~ k) j - ; (i = i1) → a (merid south j) - ; (j = i0) → rUnitₖ (a north) i - ; (j = i1) → ((λ j → rUnitₖ (a (merid north ((~ k) ∧ j))) (j ∧ k)) ∙ λ j → rUnitₖ (a (merid north ((~ k) ∨ j))) (k ∨ j)) i }) - (hcomp (λ k → λ { (i = i1) → a (merid south j) - ; (i = i0) → compPath-filler ((cong (λ x → a x +ₖ 0ₖ) (merid south))) (cong (λ x → a x +ₖ 0ₖ) (sym (merid north))) k j - ; (j = i0) → rUnitₖ (a north) i - ; (j = i1) → pathFiller2 (cong (_+ₖ 0ₖ) (cong a (merid north))) (rUnitₖ (a south)) k i}) - (rUnitₖ (a (merid south j)) i)) - - - where - pathFiller : (rUnitₖ (a north) ∙ cong a (merid north)) ≡ cong (_+ₖ 0ₖ) (cong a (merid north)) ∙ rUnitₖ (a south) - pathFiller = (λ i → (λ j → rUnitₖ (a (merid north (i ∧ j))) (j ∧ (~ i))) - ∙ λ j → rUnitₖ (a (merid north (i ∨ j))) (~ i ∨ j)) - - pathFiller2 : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) → PathP (λ i → p (~ i) ≡ z) q (p ∙ q) - pathFiller2 p q i = - hcomp (λ k → λ {(i = i0) → lUnit q (~ k) - ; (i = i1) → p ∙ q}) - ((λ j → p (~ i ∨ j)) ∙ q) - - helper : Path (_ ≡ _) (λ i → a north +ₖ Kn→ΩKn+1 1 (ΩKn+1→Kn (sym (rCancelₖ (a north)) - ∙ (λ i → a (merid south i) +ₖ (-ₖ a (merid north i))) - ∙ rCancelₖ (a south))) i) - --- - (cong (a north +ₖ_) (sym (rCancelₖ (a north))) - ∙ ((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) - ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north))) - ∙ cong (a north +ₖ_) (rCancelₖ (a north))) - helper = - ((λ j i → a north +ₖ Iso.rightInv (Iso3-Kn-ΩKn+1 1) ((sym (rCancelₖ (a north)) - ∙ (λ i → a (merid south i) +ₖ (-ₖ a (merid north i))) - ∙ rCancelₖ (a south))) j i) - ∙∙ congFunct (a north +ₖ_) (sym (rCancelₖ (a north))) ((λ i → a (merid south i) +ₖ (-ₖ a (merid north i))) ∙ rCancelₖ (a south)) - ∙∙ (λ j → cong (a north +ₖ_) (sym (rCancelₖ (a north))) - ∙ congFunct (a north +ₖ_) ((λ i → a (merid south i) +ₖ (-ₖ a (merid north i)))) (rCancelₖ (a south)) j)) - ∙∙ (λ j → cong (a north +ₖ_) (sym (rCancelₖ (a north))) - ∙ cong₂Funct (λ x y → a north +ₖ a x +ₖ (-ₖ (a y))) (merid south) (merid north) j ∙ cong (a north +ₖ_) ((rCancelₖ (a south)))) - ∙∙ (λ i → cong (a north +ₖ_) (sym (rCancelₖ (a north))) - ∙ ((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) - ∙ λ j → a north +ₖ a (merid north (~ i ∨ (~ j))) +ₖ (-ₖ (a (merid north (j ∧ (~ i)))))) - ∙ cong (a north +ₖ_) ((rCancelₖ (a (merid north (~ i)))))) - abstract - pathHelper : (a b : hLevelTrunc 4 (S₊ 2)) → a +ₖ b +ₖ (-ₖ a) ≡ b +ₖ 0ₖ - pathHelper a b = - a +ₖ b +ₖ (-ₖ a) ≡⟨ commₖ a (b +ₖ (-ₖ a)) ⟩≡⟨ assocₖ b (-ₖ a) a ⟩ - (b +ₖ (-ₖ a) +ₖ a ≡⟨ cong (b +ₖ_) (lCancelₖ a) ⟩ - b +ₖ 0ₖ ∎) - - helperFun : ∀ {ℓ} {A : Type ℓ} {a b : A} (p p' : a ≡ b) (q q' : b ≡ b) (r : a ≡ a) - → ((p q : a ≡ a) → p ∙ q ≡ q ∙ p) - → q ≡ q' - → PathP (λ i → p' (~ i) ≡ p' (~ i)) q' r - → p ∙ q ∙ sym p ≡ r - helperFun p p' q q' r comm qid dep = - p ∙ q ∙ sym p ≡⟨ cong (λ x → p ∙ x ∙ sym p) qid - ⟩≡⟨ cong (λ x → p ∙ x ∙ sym p) (λ i → lUnit (rUnit q' i) i) ⟩ - p ∙ (refl ∙ q' ∙ refl) ∙ sym p ≡⟨ cong (λ x → p ∙ x ∙ sym p) (λ i → (λ j → p' (~ i ∨ ~ j)) ∙ dep i ∙ λ j → p' (~ i ∨ j)) - ⟩≡⟨ assoc p (sym p' ∙ r ∙ p') (sym p) ⟩ - (p ∙ sym p' ∙ r ∙ p') ∙ sym p ≡⟨ cong (_∙ sym p) (assoc p (sym p') (r ∙ p')) - ⟩≡⟨ sym (assoc (p ∙ sym p') (r ∙ p') (sym p)) ⟩ - (p ∙ sym p') ∙ (r ∙ p') ∙ sym p ≡⟨ cong ((p ∙ sym p') ∙_) (sym (assoc r p' (sym p))) - ⟩≡⟨ cong (λ x → (p ∙ sym p') ∙ r ∙ x) (sym (symDistr p (sym p'))) ⟩ - (p ∙ sym p') ∙ r ∙ sym (p ∙ sym p') ≡⟨ cong ((p ∙ sym p') ∙_) (comm r (sym (p ∙ sym p'))) - ⟩≡⟨ assoc (p ∙ sym p') (sym (p ∙ sym p')) r ⟩ - ((p ∙ sym p') ∙ sym (p ∙ sym p')) ∙ r ≡⟨ cong (_∙ r) (rCancel (p ∙ sym p')) ⟩≡⟨ sym (lUnit r) ⟩ - r ∎ - - together : Path (_ ≡ _) (cong (a north +ₖ_) (sym (rCancelₖ (a north))) - ∙ ((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north))) - ∙ cong (a north +ₖ_) (rCancelₖ (a north))) - (cong(_+ₖ 0ₖ) (cong a (merid south) ∙ cong a (sym (merid north)))) - together = - helperFun (cong (a north +ₖ_) (sym (rCancelₖ (a north)))) - (λ i → pathHelper (a north) (a north) (~ i)) - ((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north))) - (((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north)))) - (cong(_+ₖ 0ₖ) (cong a (merid south) ∙ cong a (sym (merid north)))) - (isCommΩK-based 2 _) - refl - λ i → hcomp (λ k → λ {(i = i0) → (λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) - ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north)) - ; (i = i1) → congFunct (_+ₖ 0ₖ) (cong a (merid south)) (cong a (sym (merid north))) (~ k)}) - ((λ j → pathHelper (a north) (a (merid south j)) i) ∙ (λ j → pathHelper (a north) (a (merid north (~ j))) i)) +open import Cubical.HITs.Truncation renaming (elim to trElim ; elim2 to trElim2 ; map to trMap ; rec to trRec) --------- H⁰(T²) ------------ H⁰-T²≅0 : GroupEquiv (coHomGr 0 (S₊ 1 × S₊ 1)) intGroup -H⁰-T²≅0 = H⁰-connected (north , north) - (∣ north , north ∣ - , trElim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) - λ {(a , b) → pRec (isOfHLevelTrunc 2 _ _) - (λ pa → pRec (isOfHLevelTrunc 2 _ _) - (λ pb → cong ∣_∣ (ΣPathP (pa , pb))) - (Sn-connected 0 b)) - (Sn-connected 0 a)}) - --- Note : the above could of course be abbreviated by transporting over isContr ∥ S¹ × S¹ ∥₂ ≡ isContr (∥ S ∥₂ × ∥ S ∥₂). --- This proof should however compute a bit better. +H⁰-T²≅0 = + H⁰-connected (north , north) + λ (a , b) → pRec propTruncIsProp + (λ id1 → pRec propTruncIsProp + (λ id2 → ∣ ΣPathP (id1 , id2) ∣₁) + (Sn-connected 0 b) ) + (Sn-connected 0 a) ------------------ H¹(T²) ------------------------------- --- We first need the following technical lemma -basechange-lemma2 : (f g : S¹ → hLevelTrunc 3 (S₊ 1)) (F : hLevelTrunc 3 (S₊ 1) → S¹) - → ((basechange2⁻ (F (f base +ₖ g base)) λ i → F ((f (loop i)) +ₖ g (loop i))) - ≡ basechange2⁻ (F (f base)) (cong (F ∘ f) loop) - ∙ basechange2⁻ (F (g base)) (cong (F ∘ g) loop)) -basechange-lemma2 f g F = coInd f g F (f base) (g base) refl refl - where - coInd : (f g : S¹ → hLevelTrunc 3 (S₊ 1)) (F : hLevelTrunc 3 (S₊ 1) → S¹) (x y : hLevelTrunc 3 (S₊ 1)) - → f base ≡ x - → g base ≡ y - → ((basechange2⁻ (F (f base +ₖ g base)) λ i → F ((f (loop i)) +ₖ g (loop i))) - ≡ basechange2⁻ (F (f base)) (cong (F ∘ f) loop) - ∙ basechange2⁻ (F (g base)) (cong (F ∘ g) loop)) - coInd f g F = - elim2 (λ _ _ → isGroupoidΠ2 λ _ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 (isGroupoidS¹ base base)) _ _ ) - (suspToPropRec2 - north - (λ _ _ → isPropΠ2 λ _ _ → isGroupoidS¹ _ _ _ _) - λ fid gid → - ((λ i → basechange2⁻ (F (f base +ₖ g base)) (cong₂Funct (λ x y → F (f x +ₖ g y)) loop loop i)) - ∙ (basechange2⁻-morph (F (f base +ₖ g base)) - (cong (λ x → F (f x +ₖ g base)) loop) - (cong (λ x → F (f base +ₖ g x)) loop))) - ∙∙ (λ i → basechange2⁻ (F (f base +ₖ gid i)) (cong (λ x → F (f x +ₖ gid i)) loop) - ∙ basechange2⁻ (F (fid i +ₖ g base)) (cong (λ x → F (fid i +ₖ g x)) loop)) - ∙∙ (λ i → basechange2⁻ (F (rUnitₖ (f base) i)) (cong (λ x → F (rUnitₖ (f x) i)) loop) - ∙ basechange2⁻ (F (lUnitₖ (g base) i)) (cong (λ x → F (lUnitₖ (g x) i)) loop))) - - -S1→K₁≡S1×Int : Iso ((S₊ 1) → hLevelTrunc 3 (S₊ 1)) ((hLevelTrunc 3 (S₊ 1)) × Int) -S1→K₁≡S1×Int = helper2 ⋄ S¹→S¹≡S¹×Int ⋄ helper - where - helper : Iso (S¹ × Int) (hLevelTrunc 3 (S₊ 1) × Int) - Iso.fun helper (s , int) = ∣ S¹→S1 s ∣ , int - Iso.inv helper (s , int) = (S1→S¹ (S1map s)) , int - Iso.rightInv helper (s , int) = - trElim {B = λ s → (∣ S¹→S1 (S1→S¹ (S1map s)) ∣ , int) ≡ (s , int)} - (λ _ → isOfHLevelPath 3 (isOfHLevelΣ 3 (isOfHLevelTrunc 3) (λ _ → isOfHLevelSuc 2 isSetInt)) _ _) - (λ a → ΣPathP ((cong ∣_∣ (S1→S¹-retr a)) , refl)) - s - Iso.leftInv helper (s , int) = ΣPathP ((S1→S¹-sect s) , refl) - - helper2 : Iso (S₊ 1 → hLevelTrunc 3 (S₊ 1)) (S¹ → hLevelTrunc 3 S¹) - Iso.fun helper2 f x = trMap S1→S¹ (f (S¹→S1 x)) - Iso.inv helper2 f x = trMap S¹→S1 (f (S1→S¹ x)) - Iso.rightInv helper2 f = funExt λ x i → helper3 (f (S1→S¹-sect x i)) i - where - helper3 : (x : hLevelTrunc 3 S¹) → trMap S1→S¹ (trMap S¹→S1 x) ≡ x - helper3 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) - λ a → cong ∣_∣ (S1→S¹-sect a) - Iso.leftInv helper2 f = funExt λ x i → helper3 (f (S1→S¹-retr x i)) i - where - helper3 : (x : hLevelTrunc 3 (S₊ 1)) → trMap S¹→S1 (trMap S1→S¹ x) ≡ x - helper3 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) - λ a → cong ∣_∣ (S1→S¹-retr a) - - H¹-T²≅ℤ×ℤ : GroupEquiv (coHomGr 1 ((S₊ 1) × (S₊ 1))) (dirProd intGroup intGroup) H¹-T²≅ℤ×ℤ = groupequiv (isoToEquiv (setTruncIso (curryIso ⋄ codomainIso S1→K₁≡S1×Int ⋄ toProdIso) @@ -268,7 +56,7 @@ H¹-T²≅ℤ×ℤ = (funExt (λ x → helper (f (x , S¹→S1 base) +ₖ g (x , S¹→S1 base)) ∙ sym (cong₂ (λ x y → x +ₖ y) (helper (f (x , S¹→S1 base))) - (helper (g (x , S¹→S1 base))))))) , + (helper (g (x , S¹→S1 base))))))) , (cong ∣_∣₂ (funExt (suspToPropRec @@ -295,7 +83,7 @@ H¹-T²≅ℤ×ℤ = (basechange2⁻ (S¹map (trMap S1→S¹ (g (north , S¹→S1 base)))) (λ i → S¹map (trMap S1→S¹ (g (north , S¹→S1 (loop i)))))))))))))) - □ dirProdEquiv (invGroupEquiv (Hⁿ-Sⁿ≅ℤ 0)) (invGroupEquiv H⁰-S¹≅ℤ) + □ dirProdEquiv (invGroupEquiv (Hⁿ-Sⁿ≅ℤ 0)) (H⁰-Sⁿ≅ℤ 0) where helper : (x : hLevelTrunc 3 (S₊ 1)) → ∣ S¹→S1 (S¹map (trMap S1→S¹ x)) ∣ ≡ x @@ -307,22 +95,81 @@ H¹-T²≅ℤ×ℤ = ----------------------- H²(T²) ------------------------------ open import Cubical.Foundations.Equiv -H²-T²≡ℤ : Int ≃ (coHom 2 (S₊ 1 × S₊ 1)) -H²-T²≡ℤ = - compEquiv (isoToEquiv helper') - (compEquiv (invEquiv (≃-× (GroupEquiv.eq H²-S¹≅0) - (invEquiv (GroupEquiv.eq (Hⁿ-Sⁿ≅ℤ 0))))) - (isoToEquiv - (invIso setTruncOfProdIso - ⋄ invIso (setTruncIso (curryIso - ⋄ codomainIso S1→K2≡K2×K1 - ⋄ toProdIso))))) +H²-T²≅ℤ : GroupEquiv (coHomGr 2 (S₊ 1 × S₊ 1)) intGroup +H²-T²≅ℤ = invGroupEquiv ℤ≅H²-T² where - helper' : Iso Int (Unit × Int) - Iso.inv helper' = snd - Iso.fun helper' x = tt , x - Iso.leftInv helper' _ = refl - Iso.rightInv helper' _ = refl + ℤ≅H²-T² : GroupEquiv intGroup (coHomGr 2 (S₊ 1 × S₊ 1)) + GroupEquiv.eq ℤ≅H²-T² = + compEquiv (isoToEquiv helper') + (compEquiv (invEquiv (≃-× (GroupEquiv.eq H²-S¹≅0) + (invEquiv (GroupEquiv.eq (Hⁿ-Sⁿ≅ℤ 0))))) + (isoToEquiv + (invIso setTruncOfProdIso + ⋄ invIso (setTruncIso (curryIso + ⋄ codomainIso S1→K2≡K2×K1 + ⋄ toProdIso))))) + where + helper' : Iso Int (Unit × Int) + Iso.inv helper' = snd + Iso.fun helper' x = tt , x + Iso.leftInv helper' _ = refl + Iso.rightInv helper' _ = refl + + GroupEquiv.isHom ℤ≅H²-T² a b = + (λ i → f (GroupEquiv.isHom (invGroupEquiv (dirProdEquiv H²-S¹≅0 (invGroupEquiv (Hⁿ-Sⁿ≅ℤ 0)))) (_ , a) (_ , b) i)) + ∙ ((λ i → f (guyId i , (g a +ₕ g b))) + ∙∙ helper (g a) (g b) + ∙∙ (λ i → f (guyId2 (~ i) , g a) +ₕ f (guyId2 (~ i) , g b))) + where + f = Iso.fun (((invIso setTruncOfProdIso ⋄ invIso (setTruncIso (curryIso ⋄ codomainIso S1→K2≡K2×K1 ⋄ toProdIso))))) + g = invEq (GroupEquiv.eq (invGroupEquiv (Hⁿ-Sⁿ≅ℤ 0))) --- Showing that the map from ℤ to H²(T²) is a morphism should be easy, but it gets somewhat messy. --- Posponing it for now. + isPropH²-S¹ : isProp (coHom 2 (S₊ 1)) + isPropH²-S¹ = transport (λ i → isProp (ua (GroupEquiv.eq H²-S¹≅0) (~ i))) isPropUnit + + guyId : ∣ _ ∣₂ ≡ 0ₕ + guyId = isPropH²-S¹ _ _ + + guyId2 : ∣ _ ∣₂ ≡ 0ₕ + guyId2 = isPropH²-S¹ _ _ + + helper : (x y : ∥ ((S₊ 1) → (hLevelTrunc 3 (S₊ 1) )) ∥₂) → + f (0ₕ , x +ₕ y) ≡ f (0ₕ , x) +ₕ f (0ₕ , y) + helper = + sElim2 (λ _ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f g i → ∣ (λ x → helper2 (f (fst x)) (g (fst x)) (snd x) i) ∣₂ + where + helper2 : (x y : coHomK 1) (s : S₊ 1) + → Iso.inv S1→K2≡K2×K1 (0ₖ , x +ₖ y) s ≡ (Iso.inv S1→K2≡K2×K1 (0ₖ , x)) s +ₖ (Iso.inv S1→K2≡K2×K1 (0ₖ , y)) s + helper2 = + trElim2 (λ _ _ → isOfHLevelΠ 3 λ _ → isOfHLevelTrunc 4 _ _) + λ a b → + λ {north → cong₂ (_+ₖ_) (sym (lUnitₖ 0ₖ)) (sym (lUnitₖ 0ₖ)) + ; south → cong₂ (_+ₖ_) (sym (lUnitₖ 0ₖ)) (sym (lUnitₖ 0ₖ)) + ; (merid south i) j → + hcomp (λ k → λ {(i = i0) → cong₂ (_+ₖ_) (sym (lUnitₖ 0ₖ)) (sym (lUnitₖ 0ₖ)) (j ∧ k) + ; (i = i1) → cong₂ (_+ₖ_) (sym (lUnitₖ 0ₖ)) (sym (lUnitₖ 0ₖ)) (j ∧ k) + ; (j = i0) → 0ₖ +ₖ Kn→ΩKn+1 1 (∣ a ∣ +ₖ ∣ b ∣) i + ; (j = i1) → cong₂ (_+ₖ_) (sym (lUnitₖ (Kn→ΩKn+1 1 ∣ a ∣ i))) (sym (lUnitₖ (Kn→ΩKn+1 1 ∣ b ∣ i))) k}) + (helper3 ∣ a ∣ ∣ b ∣ j i) + ; (merid north i) → cong₂ (_+ₖ_) (sym (lUnitₖ 0ₖ)) (sym (lUnitₖ 0ₖ)) } + where + helper3 : (a b : coHomK 1) → cong (0ₖ +ₖ_) (Kn→ΩKn+1 1 (a +ₖ b)) ≡ cong₂ _+ₖ_ (Kn→ΩKn+1 1 a) (Kn→ΩKn+1 1 b) + helper3 a b = (cong (cong (0ₖ +ₖ_)) helper4 + ∙ congFunct (0ₖ +ₖ_) (Kn→ΩKn+1 1 a) (Kn→ΩKn+1 1 b)) + ∙∙ cong (_∙ cong (0ₖ +ₖ_) (Kn→ΩKn+1 1 b)) (helper5 (Kn→ΩKn+1 1 a)) + ∙∙ sym (cong₂Funct (_+ₖ_) (Kn→ΩKn+1 1 a) (Kn→ΩKn+1 1 b)) + where + helper4 : Kn→ΩKn+1 1 (a +ₖ b) ≡ Kn→ΩKn+1 1 a ∙ Kn→ΩKn+1 1 b -- Termination issues unless this is put as a separate lemma... + helper4 = +ₖ→∙ 1 a b + + helper6 : ∀{ℓ} {A : Type ℓ} {a b : A} (p : a ≡ b) → p ≡ (refl ∙ refl) ∙ p ∙ refl ∙ refl + helper6 p = (λ i → rUnit p i) ∙∙ (λ i → lUnit (p ∙ rUnit refl i) i) ∙∙ λ i → rUnit refl i ∙ p ∙ refl ∙ refl + helper5 : (a : Path (coHomK 2) 0ₖ 0ₖ) → cong (0ₖ +ₖ_) a ≡ cong (_+ₖ 0ₖ) a + helper5 a = (((helper6 (cong (0ₖ +ₖ_) a)) + ∙∙ (λ i → ((λ j → lUnitₖ 0ₖ (i ∧ j)) ∙ refl) ∙ cong (λ x → lUnitₖ x i) a ∙ refl ∙ λ j → lUnitₖ 0ₖ (i ∧ (~ j))) + ∙∙ λ i → (lUnitₖ 0ₖ ∙ λ j → rUnitₖ 0ₖ ((~ i) ∨ (~ j))) ∙ cong (λ x → rUnitₖ x (~ i)) a ∙ (λ j → rUnitₖ 0ₖ (~ i ∨ j)) ∙ sym (lUnitₖ 0ₖ)) + ∙∙ (λ i → (lUnitₖ 0ₖ ∙ sym (rUnitₖ 0ₖ)) ∙ isCommΩK-based 2 (0ₖ +ₖ 0ₖ) (cong (_+ₖ 0ₖ) a) (rUnitₖ 0ₖ ∙ sym (lUnitₖ 0ₖ)) i) + ∙∙ λ i → assoc (lUnitₖ 0ₖ ∙ sym (rUnitₖ 0ₖ)) (symDistr ((lUnitₖ 0ₖ)) (sym (rUnitₖ 0ₖ)) (~ i)) (cong (_+ₖ 0ₖ) a) i) + ∙∙ cong (_∙ (cong (_+ₖ 0ₖ) a)) (rCancel (lUnitₖ 0ₖ ∙ sym (rUnitₖ 0ₖ))) + ∙∙ sym (lUnit (cong (_+ₖ 0ₖ) a)) diff --git a/Cubical/ZCohomology/Groups/Unit.agda b/Cubical/ZCohomology/Groups/Unit.agda index 878ca7ee7..2ab01ec7a 100644 --- a/Cubical/ZCohomology/Groups/Unit.agda +++ b/Cubical/ZCohomology/Groups/Unit.agda @@ -19,12 +19,15 @@ open import Cubical.Data.Unit open import Cubical.Structures.Group -- H⁰(Unit) -H⁰-Unit≅ℤ' : GroupIso (coHomGr 0 Unit) intGroup -GroupHom.fun (GroupIso.map H⁰-Unit≅ℤ') = sRec isSetInt (λ f → f tt) -GroupHom.isHom (GroupIso.map H⁰-Unit≅ℤ') = sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) λ a b → addLemma (a tt) (b tt) -GroupIso.inv H⁰-Unit≅ℤ' a = ∣ (λ _ → a) ∣₂ -GroupIso.rightInv H⁰-Unit≅ℤ' _ = refl -GroupIso.leftInv H⁰-Unit≅ℤ' = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ a → refl +open GroupHom +open GroupIso +private + H⁰-Unit≅ℤ' : GroupIso (coHomGr 0 Unit) intGroup + fun (GroupIso.map H⁰-Unit≅ℤ') = sRec isSetInt (λ f → f tt) + isHom (GroupIso.map H⁰-Unit≅ℤ') = sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) λ a b → addLemma (a tt) (b tt) + inv H⁰-Unit≅ℤ' a = ∣ (λ _ → a) ∣₂ + rightInv H⁰-Unit≅ℤ' _ = refl + leftInv H⁰-Unit≅ℤ' = sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) λ a → refl H⁰-Unit≅ℤ : GroupEquiv (coHomGr 0 Unit) intGroup H⁰-Unit≅ℤ = GrIsoToGrEquiv H⁰-Unit≅ℤ' @@ -63,95 +66,20 @@ private Hⁿ-contrTypeIso n contr = compIso (setTruncIso (ContrToTypeIso contr)) (setTruncIso (invIso (ContrToTypeIso isContrUnit))) -Hⁿ-contrType≅0' : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → isContr A - → GroupIso (coHomGr (suc n) A) trivialGroup -GroupHom.fun (GroupIso.map (Hⁿ-contrType≅0' _ _)) _ = _ -GroupHom.isHom (GroupIso.map (Hⁿ-contrType≅0' _ _)) _ _ = refl -GroupIso.inv (Hⁿ-contrType≅0' _ _) _ = 0ₕ -GroupIso.rightInv (Hⁿ-contrType≅0' _ _) _ = refl -GroupIso.leftInv (Hⁿ-contrType≅0' {A = A} n contr) _ = isOfHLevelSuc 0 helper _ _ - where - helper : isContr (coHom (suc n) A) - helper = (Iso.inv (Hⁿ-contrTypeIso n contr) 0ₕ) - , λ y → cong (Iso.inv (Hⁿ-contrTypeIso n contr)) (isOfHLevelSuc 0 (isContrHⁿ-Unit n) 0ₕ (Iso.fun (Hⁿ-contrTypeIso n contr) y)) - ∙ Iso.leftInv (Hⁿ-contrTypeIso n contr) y + Hⁿ-contrType≅0' : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → isContr A + → GroupIso (coHomGr (suc n) A) trivialGroup + fun (GroupIso.map (Hⁿ-contrType≅0' _ _)) _ = _ + isHom (GroupIso.map (Hⁿ-contrType≅0' _ _)) _ _ = refl + inv (Hⁿ-contrType≅0' _ _) _ = 0ₕ + rightInv (Hⁿ-contrType≅0' _ _) _ = refl + leftInv (Hⁿ-contrType≅0' {A = A} n contr) _ = isOfHLevelSuc 0 helper _ _ + where + helper : isContr (coHom (suc n) A) + helper = (Iso.inv (Hⁿ-contrTypeIso n contr) 0ₕ) + , λ y → cong (Iso.inv (Hⁿ-contrTypeIso n contr)) + (isOfHLevelSuc 0 (isContrHⁿ-Unit n) 0ₕ (Iso.fun (Hⁿ-contrTypeIso n contr) y)) + ∙ Iso.leftInv (Hⁿ-contrTypeIso n contr) y Hⁿ-contrType≅0 : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → isContr A → GroupEquiv (coHomGr (suc n) A) trivialGroup Hⁿ-contrType≅0 n contr = GrIsoToGrEquiv (Hⁿ-contrType≅0' n contr) - --- open import Cubical.Data.Bool --- open import Cubical.Data.Sigma --- open import Cubical.Foundations.GroupoidLaws --- test : ∀ {ℓ} {A : Type ℓ} (a : A) → Iso (Unit → A) (Σ[ f ∈ (Bool → A) ] f true ≡ a) --- Iso.fun (test {A = A} a) x = fhel , refl --- module _ where --- fhel : Bool → A --- fhel false = x _ --- fhel true = a --- Iso.inv (test a) (f , id) _ = f false --- Iso.rightInv (test a) (f , id) = λ i → (funExt helper i) , λ j → id ((~ i) ∨ j) --- where --- helper : (x : Bool) → fhel a (λ _ → f false) x ≡ f x --- helper false = refl --- helper true = sym id --- Iso.leftInv (test a) x = refl - --- test2 : Iso (Susp Unit) Unit --- Iso.fun test2 _ = _ --- Iso.inv test2 _ = north --- Iso.rightInv test2 a = refl --- Iso.leftInv test2 north = refl --- Iso.leftInv test2 south = merid tt --- Iso.leftInv test2 (merid tt i) j = merid tt (i ∧ j) - - --- -- open import Cubical.Foundations.Pointed --- -- test3 : ∀ {ℓ} {A : Type ℓ} (a : A) → Iso (a ≡ a) ((Σ[ f ∈ (Unit → A) ] f tt ≡ a)) --- -- Iso.fun (test3 a) p = (λ _ → a) , p --- -- Iso.inv (test3 a) (f , id) = refl --- -- Iso.rightInv (test3 a) = {!!} --- -- Iso.leftInv (test3 a) x = {!!} - --- -- open import Cubical.HITs.S1 --- -- test4 : {A : Type₀} (a : A) → Iso ((S¹ , base) →∙ (A , a)) (a ≡ a) --- -- Iso.fun (test4 a) (f , id) = sym id ∙∙ (cong f loop) ∙∙ id --- -- Iso.inv (test4 a) p = (λ {base → a ; (loop i) → p i}) , refl --- -- Iso.rightInv (test4 a) b = sym (doubleCompPath-filler refl b refl) --- -- Iso.leftInv (test4 a) (f , id) i = (funExt helper i) , λ j → id (~ i ∨ j) --- -- where --- -- helper : (x : S¹) → fst (Iso.inv (test4 a) (Iso.fun (test4 a) (f , id))) x ≡ f x --- -- helper base = sym id --- -- helper (loop i) j = doubleCompPath-filler (sym id) (cong f loop) (id) (~ j) i - - --- -- test5 : {A B : Type₀} (a : A) (b : B) → Iso ((Susp A , north) →∙ (B , b)) ((A , a) →∙ ((b ≡ b) , refl)) --- -- Iso.fun (test5 a b) (f , p) = (λ a' → sym p ∙∙ cong f ((merid a') ∙ sym (merid a)) ∙∙ p) , {!!} --- -- Iso.inv (test5 {A = A} {B = B} a b) (f , p) = helper , refl --- -- where --- -- helper : Susp A → B --- -- helper north = b --- -- helper south = b --- -- helper (merid x i) = f x i --- -- Iso.rightInv (test5 a b) (f , p) = {!!} --- -- Iso.leftInv (test5 a b) = {!!} - - --- -- open import Cubical.Foundations.Equiv --- -- open import Cubical.Homotopy.Connected --- -- H⁰-connected : ∀ {ℓ} {A : Type ℓ} (a : A) → isConnected 2 A → Iso (coHom 0 A) Int --- -- Iso.fun (H⁰-connected a con) = sRec isSetInt λ f → f a --- -- Iso.inv (H⁰-connected a con) b = ∣ (λ x → b) ∣₂ --- -- Iso.rightInv (H⁰-connected a con) b = refl --- -- Iso.leftInv (H⁰-connected a con) = --- -- sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) --- -- λ f → cong ∣_∣₂ (funExt (equiv-proof (elim.isEquivPrecompose (λ (x : Unit) → a) 1 (λ x → _ , isSetInt _ _) (isConnectedPoint _ con _)) (λ _ → refl) .fst .fst )) - --- -- open import Cubical.HITs.Truncation renaming (rec to trRec) --- -- H⁰-connected' : ∀ {ℓ} {A : Type ℓ} (a : A) → isConnected 2 A → Iso (coHom 0 A) Int --- -- Iso.fun (H⁰-connected' a con) = sRec isSetInt λ f → f a --- -- Iso.inv (H⁰-connected' a con) b = ∣ (λ x → b) ∣₂ --- -- Iso.rightInv (H⁰-connected' a con) b = refl --- -- Iso.leftInv (H⁰-connected' a con) = --- -- sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) --- -- λ f → cong ∣_∣₂ (funExt λ x → trRec (isSetInt _ _) (cong f) (isConnectedPath 1 con a x .fst)) diff --git a/Cubical/ZCohomology/Groups/Wedge.agda b/Cubical/ZCohomology/Groups/Wedge.agda index f6cf9391e..3100db2ad 100644 --- a/Cubical/ZCohomology/Groups/Wedge.agda +++ b/Cubical/ZCohomology/Groups/Wedge.agda @@ -9,7 +9,7 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed open import Cubical.HITs.Wedge open import Cubical.HITs.SetTruncation renaming (elim to sElim ; elim2 to sElim2) -open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; ∣_∣ to ∣_∣₋₁) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; ∣_∣ to ∣_∣₁) open import Cubical.Data.Nat open import Cubical.Data.Prod open import Cubical.Data.Unit @@ -28,8 +28,8 @@ module _ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') where Hⁿ-⋁ : (n : ℕ) → GroupEquiv (coHomGr (suc n) (A ⋁ B)) (×coHomGr (suc n) (typ A) (typ B)) Hⁿ-⋁ zero = - Iso'ToGroupEquiv - (iso' + BijectionIsoToGroupEquiv + (bij-iso (grouphom (GroupHom.fun (I.i 1)) (sElim2 (λ _ _ → isOfHLevelPath 2 (isOfHLevelΣ 2 setTruncIsSet λ _ → setTruncIsSet) _ _) @@ -43,7 +43,7 @@ module _ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') where → isInIm _ _ (I.Δ 0) x surj-helper = sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) - λ f → ∣ (∣ (λ _ → f tt) ∣₂ , 0ₕ) , cong ∣_∣₂ (funExt (λ _ → cong ((f tt) +ₖ_) -0ₖ ∙ rUnitₖ (f tt))) ∣₋₁ + λ f → ∣ (∣ (λ _ → f tt) ∣₂ , 0ₕ) , cong ∣_∣₂ (funExt (λ _ → cong ((f tt) +ₖ_) -0ₖ ∙ rUnitₖ (f tt))) ∣₁ helper : (x : coHom 1 (A ⋁ B)) → isInIm _ _ (I.d 0) x → x ≡ 0ₕ @@ -67,5 +67,5 @@ module _ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') where wedgeConnected : ((x : typ A) → ∥ pt A ≡ x ∥) → ((x : typ B) → ∥ pt B ≡ x ∥) → (x : A ⋁ B) → ∥ (inl (pt A)) ≡ x ∥ wedgeConnected conA conB = PushoutToProp (λ _ → propTruncIsProp) - (λ a → pRec propTruncIsProp (λ p → ∣ cong inl p ∣₋₁) (conA a)) - λ b → pRec propTruncIsProp (λ p → ∣ push tt ∙ cong inr p ∣₋₁) (conB b) + (λ a → pRec propTruncIsProp (λ p → ∣ cong inl p ∣₁) (conA a)) + λ b → pRec propTruncIsProp (λ p → ∣ push tt ∙ cong inr p ∣₁) (conB b) diff --git a/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda b/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda index 96e20abd6..c9416010d 100644 --- a/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda +++ b/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda @@ -9,23 +9,12 @@ open import Cubical.ZCohomology.Groups.Wedge open import Cubical.ZCohomology.Groups.Connected open import Cubical.HITs.Sn -open import Cubical.Foundations.HLevels open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.GroupoidLaws open import Cubical.HITs.Susp open import Cubical.HITs.Wedge open import Cubical.HITs.Pushout -open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) -open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; ∣_∣ to ∣_∣₋₁) -open import Cubical.HITs.Nullification -open import Cubical.Data.Int -open import Cubical.Data.Sigma open import Cubical.HITs.Truncation renaming (elim to trElim) --- open import Cubical.Data.Group.Base renaming (Iso to grIso ; compIso to compGrIso ; invIso to invGrIso ; idIso to idGrIso) open import Cubical.Structures.Group -open import Cubical.Data.Unit S¹⋁S¹ : Type₀ S¹⋁S¹ = S₊∙ 1 ⋁ S₊∙ 1 @@ -43,13 +32,12 @@ H¹-S¹⋁S¹ = (Hⁿ-⋁ _ _ 0) □ dirProdEquiv coHom1S1≃ℤ coHom1S1≃ℤ ------------- H⁰(S²⋁S¹⋁S¹) --------- H⁰-S²⋁S¹⋁S¹ : GroupEquiv (coHomGr 0 S²⋁S¹⋁S¹) intGroup -H⁰-S²⋁S¹⋁S¹ = - H⁰-connected (inl north) - (wedgeConnected _ _ - (Sn-connected _) - (wedgeConnected _ _ - (Sn-connected _) - (Sn-connected _))) +H⁰-S²⋁S¹⋁S¹ = H⁰-connected (inl north) + (wedgeConnected _ _ + (Sn-connected _) + (wedgeConnected _ _ + (Sn-connected _) + (Sn-connected _))) ------------- H¹(S²⋁S¹⋁S¹) --------- H¹-S²⋁S¹⋁S¹ : GroupEquiv (coHomGr 1 S²⋁S¹⋁S¹) (dirProd intGroup intGroup) diff --git a/Cubical/ZCohomology/MayerVietorisUnreduced.agda b/Cubical/ZCohomology/MayerVietorisUnreduced.agda index 3939ab238..db27289d1 100644 --- a/Cubical/ZCohomology/MayerVietorisUnreduced.agda +++ b/Cubical/ZCohomology/MayerVietorisUnreduced.agda @@ -20,17 +20,11 @@ open import Cubical.HITs.Nullification open import Cubical.Data.Nat open import Cubical.Data.Prod hiding (_×_) open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) --- open import Cubical.Data.Group.Base renaming (Iso to grIso) - open import Cubical.Structures.Group -open import Cubical.Structures.Group.Algebra - -coHomFun : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) (f : A → B) → coHom n B → coHom n A -coHomFun n f = sRec setTruncIsSet λ β → ∣ β ∘ f ∣₂ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : C → A) (g : C → B) where - ---- Proof from Brunerie 2016. - ---- We first define the three morphisms involved: i, Δ and d. + -- Proof from Brunerie 2016. + -- We first define the three morphisms involved: i, Δ and d. private i* : (n : ℕ) → coHom n (Pushout f g) → coHom n A × coHom n B @@ -133,7 +127,7 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : λ δ b → (λ i → sElim (λ _ → isOfHLevelΣ 2 setTruncIsSet (λ _ → setTruncIsSet)) (λ δ → ∣ (λ x → δ (inl x)) ∣₂ , ∣ (λ x → δ (inr x)) ∣₂ ) (b (~ i)))) - + Ker-i⊂Im-d : (n : ℕ) (x : ⟨ coHomGr (suc n) (Pushout f g) ⟩) → isInKer (coHomGr (suc n) (Pushout f g)) (×coHomGr (suc n) A B) (i (suc n)) x → isInIm (coHomGr n C) (coHomGr (suc n) (Pushout f g)) (d n) x @@ -182,7 +176,7 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : ; (j = i1) → F (push a i)}) (pushFiller (suc n) F p1 p2 a j i) - open GroupHom + open GroupHom abstract Im-i⊂Ker-Δ : (n : ℕ) (x : ⟨ ×coHomGr n A B ⟩) diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index ef8daabab..b5c1137f5 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -5,6 +5,7 @@ open import Cubical.ZCohomology.Base open import Cubical.HITs.S1 open import Cubical.HITs.Sn open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function open import Cubical.Foundations.Equiv open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed @@ -20,7 +21,6 @@ open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; e open import Cubical.HITs.Nullification open import Cubical.Data.Int hiding (_+_) open import Cubical.Data.Nat -open import Cubical.Data.Prod open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) open import Cubical.Homotopy.Loopspace open import Cubical.Homotopy.Connected @@ -110,8 +110,9 @@ Kn→ΩKn+10ₖ (suc n) = (λ i → cong ∣_∣ (rCancel (merid north) i)) -- c ∙∙ (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ n (~ i))) ∙∙ Iso.leftInv (Iso3-Kn-ΩKn+1 n) 0ₖ -+ₖ→∙ : (n : ℕ) (a b : coHomK n) → Kn→ΩKn+1 n (a +ₖ b) ≡ Kn→ΩKn+1 n a ∙ Kn→ΩKn+1 n b -+ₖ→∙ n a b = Iso.rightInv (Iso3-Kn-ΩKn+1 n) (Kn→ΩKn+1 n a ∙ Kn→ΩKn+1 n b) +abstract + +ₖ→∙ : (n : ℕ) (a b : coHomK n) → Kn→ΩKn+1 n (a +ₖ b) ≡ Kn→ΩKn+1 n a ∙ Kn→ΩKn+1 n b + +ₖ→∙ n a b = Iso.rightInv (Iso3-Kn-ΩKn+1 n) (Kn→ΩKn+1 n a ∙ Kn→ΩKn+1 n b) lUnitₖ : {n : ℕ} (x : coHomK n) → 0ₖ +ₖ x ≡ x lUnitₖ {n = zero} x = cong ΩKn+1→Kn (sym (lUnit (Kn→ΩKn+1 zero x))) ∙ @@ -270,24 +271,26 @@ rUnitlUnit0 {n = suc n} = 0ₕ∙ zero = ∣ (λ _ → 0ₖ) , refl ∣₂ 0ₕ∙ (suc n) = ∣ (λ _ → 0ₖ) , refl ∣₂ - +open import Cubical.Structures.Semigroup +open import Cubical.Structures.Monoid +open IsSemigroup +open IsMonoid +open Group coHomGr : ∀ {ℓ} (n : ℕ) (A : Type ℓ) → Group -Group.Carrier (coHomGr n A) = coHom n A -Group.0g (coHomGr n A) = 0ₕ +Carrier (coHomGr n A) = coHom n A +0g (coHomGr n A) = 0ₕ Group._+_ (coHomGr n A) = _+ₕ_ Group.- coHomGr n A = -ₕ -Group.isGroup (coHomGr n A) = - makeIsGroup - § - (λ x y z → sym (assocₕ x y z)) - rUnitₕ - lUnitₕ - rCancelₕ - lCancelₕ +is-set (isSemigroup (IsGroup.isMonoid (Group.isGroup (coHomGr n A)))) = § +IsSemigroup.assoc (isSemigroup (IsGroup.isMonoid (Group.isGroup (coHomGr n A)))) x y z = sym (assocₕ x y z) +identity (IsGroup.isMonoid (Group.isGroup (coHomGr n A))) x = (rUnitₕ x) , (lUnitₕ x) +IsGroup.inverse (Group.isGroup (coHomGr n A)) x = (rCancelₕ x) , (lCancelₕ x) ×coHomGr : (n : ℕ) (A : Type ℓ) (B : Type ℓ') → Group -×coHomGr n A B = dirProd (coHomGr n A) (coHomGr n B) -- dirProd (coHomGr n A) (coHomGr n B) +×coHomGr n A B = dirProd (coHomGr n A) (coHomGr n B) +coHomFun : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) (f : A → B) → coHom n B → coHom n A +coHomFun n f = sRec § λ β → ∣ β ∘ f ∣₂ --- ΩKₙ is commutative isCommΩK : (n : ℕ) → (p q : typ (Ω (coHomK n , coHom-pt n))) → p ∙ q ≡ q ∙ p From af86c970e5875165c830dda67bb5210118970379 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Fri, 10 Jul 2020 02:18:17 +0200 Subject: [PATCH 51/89] merge --- Cubical/Structures/Monoid.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Cubical/Structures/Monoid.agda b/Cubical/Structures/Monoid.agda index 6f9a25431..d4e5ecbeb 100644 --- a/Cubical/Structures/Monoid.agda +++ b/Cubical/Structures/Monoid.agda @@ -86,7 +86,7 @@ makeMonoid : {M : Type ℓ} (ε : M) (_·_ : M → M → M) (rid : (x : M) → x · ε ≡ x) (lid : (x : M) → ε · x ≡ x) → Monoid -makeMonoid ε _·_ is-setM assoc rid lid = ? +makeMonoid ε _·_ is-setM assoc rid lid = monoid _ ε _·_ (makeIsMonoid is-setM assoc rid lid) record MonoidEquiv (M N : Monoid {ℓ}) : Type ℓ where From 7d43c5f93c26c61b1067ce2ab00a605b1931a10a Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Fri, 10 Jul 2020 03:23:43 +0200 Subject: [PATCH 52/89] whiteSpaceFiller3 --- Cubical/Data/Group/Base.agda | 487 ---------------------- Cubical/Data/Group/Properties.agda | 144 ------- Cubical/Homotopy/Connected.agda | 2 +- Cubical/Homotopy/Freudenthal.agda | 2 +- Cubical/Structures/AbGroup.agda | 27 +- Cubical/Structures/CommRing.agda | 46 +- Cubical/Structures/Ring.agda | 2 +- Cubical/ZCohomology/Groups/Connected.agda | 14 +- Cubical/ZCohomology/Groups/Wedge.agda | 5 +- Cubical/ZCohomology/Properties.agda | 1 - 10 files changed, 50 insertions(+), 680 deletions(-) delete mode 100644 Cubical/Data/Group/Base.agda delete mode 100644 Cubical/Data/Group/Properties.agda diff --git a/Cubical/Data/Group/Base.agda b/Cubical/Data/Group/Base.agda deleted file mode 100644 index 23b3af70b..000000000 --- a/Cubical/Data/Group/Base.agda +++ /dev/null @@ -1,487 +0,0 @@ -{-# OPTIONS --cubical --no-import-sorts --safe #-} -module Cubical.Data.Group.Base where - -open import Cubical.Foundations.Prelude renaming (comp to comp') -open import Cubical.Foundations.HLevels -open import Cubical.Data.Prod -open import Cubical.HITs.PropositionalTruncation hiding (map) -open import Cubical.Data.Sigma hiding (_×_ ; comp) - -import Cubical.Foundations.Isomorphism as I -import Cubical.Foundations.Equiv as E -import Cubical.Foundations.Equiv.HalfAdjoint as HAE - -record isGroup {ℓ} (A : Type ℓ) : Type ℓ where - no-eta-equality - constructor group-struct - field - id : A - inv : A → A - comp : A → A → A - lUnit : ∀ a → comp id a ≡ a - rUnit : ∀ a → comp a id ≡ a - assoc : ∀ a b c → comp (comp a b) c ≡ comp a (comp b c) - lCancel : ∀ a → comp (inv a) a ≡ id - rCancel : ∀ a → comp a (inv a) ≡ id - -record Group ℓ : Type (ℓ-suc ℓ) where - no-eta-equality - constructor group - field - type : Type ℓ - setStruc : isSet type - groupStruc : isGroup type - -isMorph : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (f : (Group.type G → Group.type H)) → Type (ℓ-max ℓ ℓ') -isMorph G H f = (g0 g1 : Group.type G) → f (isGroup.comp (Group.groupStruc G) g0 g1) ≡ isGroup.comp (Group.groupStruc H) (f g0) (f g1) - -record morph {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') : Type (ℓ-max ℓ ℓ') where - no-eta-equality - constructor mph - field - fun : Group.type G → Group.type H - ismorph : isMorph G H fun - -isInIm : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (morph G H) - → Group.type H → Type (ℓ-max ℓ ℓ') -isInIm G H ϕ h = ∃[ g ∈ Group.type G ] (morph.fun ϕ g) ≡ h - -isInKer : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → morph G H - → Group.type G → Type ℓ' -isInKer G H ϕ g = (morph.fun ϕ g) ≡ isGroup.id (Group.groupStruc H) - -isSurjective : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → morph G H → Type (ℓ-max ℓ ℓ') -isSurjective G H ϕ = (x : Group.type H) → isInIm G H ϕ x - -isInjective : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → morph G H → Type (ℓ-max ℓ ℓ') -isInjective G H ϕ = (x : Group.type G) → isInKer G H ϕ x → x ≡ isGroup.id (Group.groupStruc G) - --0≡0 : ∀ {ℓ} {G : Group ℓ} → isGroup.inv (Group.groupStruc G) (isGroup.id (Group.groupStruc G)) ≡ isGroup.id (Group.groupStruc G) --0≡0 {G = G} = sym (isGroup.lUnit (Group.groupStruc G) _) ∙ isGroup.rCancel (Group.groupStruc G) _ - -morph0→0 : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (f : morph G H) - → morph.fun f (isGroup.id (Group.groupStruc G)) ≡ isGroup.id (Group.groupStruc H) -morph0→0 G' H' f' = 0→0 - where - open Group - open isGroup - G = groupStruc G' - H = groupStruc H' - f = morph.fun f' - ismorph = morph.ismorph f' - - 0→0 : f (id G) ≡ id H - 0→0 = - f (id G) ≡⟨ sym (rUnit H (f (id G))) ⟩ - comp H (f (id G)) (id H) ≡⟨ (λ i → comp H (f (id G)) (rCancel H (f (id G)) (~ i))) ⟩ - comp H (f (id G)) (comp H (f (id G)) (inv H (f (id G)))) ≡⟨ sym (assoc H (f (id G)) (f (id G)) (inv H (f (id G)))) ⟩ - comp H (comp H (f (id G)) (f (id G))) (inv H (f (id G))) ≡⟨ sym (cong (λ x → comp H x (inv H (f (id G)))) (sym (cong f (lUnit G (id G))) ∙ ismorph (id G) (id G))) ⟩ - comp H (f (id G)) (inv H (f (id G))) ≡⟨ rCancel H (f (id G)) ⟩ - id H ∎ - -{- a morphism ϕ satisfies ϕ(- a) = - ϕ(a) -} -morphMinus : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') → (ϕ : morph G H) - → (g : Group.type G) → morph.fun ϕ (isGroup.inv (Group.groupStruc G) g) ≡ isGroup.inv (Group.groupStruc H) (morph.fun ϕ g) -morphMinus G' H' ϕ g = mainLemma - where - open Group - open isGroup - H = groupStruc H' - G = groupStruc G' - f = morph.fun ϕ - - helper : comp H (f (inv G g)) (f g) ≡ id H - helper = sym (morph.ismorph ϕ (inv G g) g) ∙∙ cong f (lCancel G g) ∙∙ morph0→0 G' H' ϕ - - mainLemma : f (inv G g) ≡ inv H (f g) - mainLemma = f (inv G g) ≡⟨ sym (rUnit H (f (inv G g))) ⟩ - comp H (f (inv G g)) (id H) ≡⟨ cong (comp H (f (inv G g))) (sym (rCancel H (f g))) ⟩ - comp H (f (inv G g)) (comp H (f g) (inv H (f g))) ≡⟨ sym (assoc H (f (inv G g)) (f g) (inv H (f g))) ⟩ - comp H (comp H (f (inv G g)) (f g)) (inv H (f g)) ≡⟨ cong (λ x → comp H x (inv H (f g))) helper ⟩ - comp H (id H) (inv H (f g)) ≡⟨ lUnit H (inv H (f g)) ⟩ - inv H (f g) ∎ - -rightist-group-struct : ∀ {ℓ} {A : Type ℓ} - → (id : A) (inv : A → A) (comp : A → A → A) - → (rUnit : ∀ a → comp a id ≡ a) - → (assoc : ∀ a b c → comp (comp a b) c ≡ comp a (comp b c)) - → (rCancel : ∀ a → comp a (inv a) ≡ id) - → isGroup A -rightist-group-struct id inv comp rUnit assoc rCancel = - group-struct id inv comp lUnit rUnit assoc lCancel rCancel - where - abstract - lCancel : ∀ a → comp (inv a) a ≡ id - lCancel a = - comp (inv a) a - ≡⟨ sym (rUnit (comp (inv a) a)) ⟩ - comp (comp (inv a) a) id - ≡⟨ cong (comp (comp (inv a) a)) (sym (rCancel (inv a))) ⟩ - comp (comp (inv a) a) (comp (inv a) (inv (inv a))) - ≡⟨ sym (assoc (comp (inv a) a) (inv a) (inv (inv a))) ⟩ - comp (comp (comp (inv a) a) (inv a)) (inv (inv a)) - ≡⟨ cong (λ b → comp b (inv (inv a))) (assoc (inv a) a (inv a)) ⟩ - comp (comp (inv a) (comp a (inv a))) (inv (inv a)) - ≡⟨ cong (λ b → comp (comp (inv a) b) (inv (inv a))) (rCancel a) ⟩ - comp (comp (inv a) id) (inv (inv a)) - ≡⟨ cong (λ b → comp b (inv (inv a))) (rUnit (inv a)) ⟩ - comp (inv a) (inv (inv a)) - ≡⟨ rCancel (inv a) ⟩ - id - ∎ - - lUnit : ∀ a → comp id a ≡ a - lUnit a = - comp id a - ≡⟨ cong (λ b → comp b a) (sym (rCancel a)) ⟩ - comp (comp a (inv a)) a - ≡⟨ assoc a (inv a) a ⟩ - comp a (comp (inv a) a) - ≡⟨ cong (comp a) (lCancel a) ⟩ - comp a id - ≡⟨ rUnit a ⟩ - a - ∎ - -rightist-group : ∀ {ℓ} {A : Type ℓ} (Aset : isSet A) - → (id : A) (inv : A → A) (comp : A → A → A) - → (rUnit : ∀ a → comp a id ≡ a) - → (assoc : ∀ a b c → comp (comp a b) c ≡ comp a (comp b c)) - → (rCancel : ∀ a → comp a (inv a) ≡ id) - → Group ℓ -rightist-group Aset id inv comp rUnit assoc rCancel = - group _ Aset (rightist-group-struct id inv comp rUnit assoc rCancel) - -leftist-group-struct : ∀ {ℓ} {A : Type ℓ} - → (id : A) (inv : A → A) (comp : A → A → A) - → (lUnit : ∀ a → comp id a ≡ a) - → (assoc : ∀ a b c → comp (comp a b) c ≡ comp a (comp b c)) - → (lCancel : ∀ a → comp (inv a) a ≡ id) - → isGroup A -leftist-group-struct id inv comp lUnit assoc lCancel = - group-struct id inv comp lUnit rUnit assoc lCancel rCancel - where - abstract - rCancel : ∀ a → comp a (inv a) ≡ id - rCancel a = - comp a (inv a) - ≡⟨ sym (lUnit (comp a (inv a))) ⟩ - comp id (comp a (inv a)) - ≡⟨ cong (λ b → comp b (comp a (inv a))) (sym (lCancel (inv a))) ⟩ - comp (comp (inv (inv a)) (inv a)) (comp a (inv a)) - ≡⟨ assoc (inv (inv a)) (inv a) (comp a (inv a)) ⟩ - comp (inv (inv a)) (comp (inv a) (comp a (inv a))) - ≡⟨ cong (comp (inv (inv a))) (sym (assoc (inv a) a (inv a))) ⟩ - comp (inv (inv a)) (comp (comp (inv a) a) (inv a)) - ≡⟨ cong (λ b → comp (inv (inv a)) (comp b (inv a))) (lCancel a) ⟩ - comp (inv (inv a)) (comp id (inv a)) - ≡⟨ cong (comp (inv (inv a))) (lUnit (inv a)) ⟩ - comp (inv (inv a)) (inv a) - ≡⟨ lCancel (inv a) ⟩ - id - ∎ - - rUnit : ∀ a → comp a id ≡ a - rUnit a = - comp a id - ≡⟨ cong (comp a) (sym (lCancel a)) ⟩ - comp a (comp (inv a) a) - ≡⟨ sym (assoc a (inv a) a) ⟩ - comp (comp a (inv a)) a - ≡⟨ cong (λ b → comp b a) (rCancel a) ⟩ - comp id a - ≡⟨ lUnit a ⟩ - a - ∎ - -leftist-group : ∀ {ℓ} {A : Type ℓ} (Aset : isSet A) - → (id : A) (inv : A → A) (comp : A → A → A) - → (lUnit : ∀ a → comp id a ≡ a) - → (assoc : ∀ a b c → comp (comp a b) c ≡ comp a (comp b c)) - → (lCancel : ∀ a → comp (inv a) a ≡ id) - → Group ℓ -leftist-group Aset id inv comp lUnit assoc lCancel = - group _ Aset (leftist-group-struct id inv comp lUnit assoc lCancel) - ------------ Equivalent notions of isomorphisms -------------- - -record Iso {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') : Type (ℓ-max ℓ ℓ') where - constructor iso - field - fun : morph G H - inv : morph H G - rightInv : I.section (morph.fun fun) (morph.fun inv) - leftInv : I.retract (morph.fun fun) (morph.fun inv) - -record Iso' {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') : Type (ℓ-max ℓ ℓ') where - constructor iso' - field - isoSet : I.Iso (Group.type G) (Group.type H) - isoSetMorph : isMorph G H (I.Iso.fun isoSet) - -record Iso'' {ℓ ℓ'} (A : Group ℓ) (B : Group ℓ') : Type (ℓ-max ℓ ℓ') where - constructor iso'' - field - ϕ : morph A B - inj : isInjective A B ϕ - surj : isSurjective A B ϕ - -record Iso''' {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') : Type (ℓ-max ℓ ℓ') where -- Should perhaps replace Iso'. Appears to compute somewhat better. - constructor iso''' - field - fun : morph G H - inv : Group.type H → Group.type G - rightInv : I.section (morph.fun fun) inv - leftInv : I.retract (morph.fun fun) inv - -record SES {ℓ ℓ' ℓ'' ℓ'''} (A : Group ℓ) (B : Group ℓ') (leftGr : Group ℓ'') (rightGr : Group ℓ''') - : Type (ℓ-suc (ℓ-max ℓ (ℓ-max ℓ' (ℓ-max ℓ'' ℓ''')))) where - constructor ses - field - isTrivialLeft : isProp (Group.type leftGr) - isTrivialRight : isProp (Group.type rightGr) - - left : morph leftGr A - right : morph B rightGr - ϕ : morph A B - - Ker-ϕ⊂Im-left : (x : Group.type A) -- - → isInKer A B ϕ x - → isInIm leftGr A left x - Ker-right⊂Im-ϕ : (x : Group.type B) -- - → isInKer B rightGr right x - → isInIm A B ϕ x - -_≃_ : ∀ {ℓ ℓ'} (A : Group ℓ) (B : Group ℓ') → Type (ℓ-max ℓ ℓ') -A ≃ B = Σ (morph A B) \ f → (E.isEquiv (morph.fun f)) - ------------ -compMorph : ∀ {ℓ ℓ' ℓ''} {F : Group ℓ} {G : Group ℓ'} {H : Group ℓ''} (I : morph F G) (J : morph G H) → morph F H -morph.fun (compMorph I J) x = morph.fun J (morph.fun I x) -morph.ismorph (compMorph {F = F} {G = G} {H = H} I J) g0 g1 = - cong (morph.fun J) (morph.ismorph I g0 g1) - ∙ morph.ismorph J (morph.fun I g0) (morph.fun I g1) - -compIso : ∀ {ℓ} {F G H : Group ℓ} (I : Iso F G) (J : Iso G H) → Iso F H -Iso.fun (compIso iso1 iso2) = compMorph (Iso.fun iso1) (Iso.fun iso2) -Iso.inv (compIso iso1 iso2) = compMorph (Iso.inv iso2) (Iso.inv iso1) -Iso.rightInv (compIso iso1 iso2) b = - cong (morph.fun (Iso.fun iso2)) (Iso.rightInv iso1 _) ∙ Iso.rightInv iso2 _ -Iso.leftInv (compIso iso1 iso2) b = - cong (morph.fun (Iso.inv iso1)) (Iso.leftInv iso2 _) ∙ Iso.leftInv iso1 _ - -idIso : ∀ {ℓ} (G : Group ℓ) → Iso G G -Iso.fun (idIso G) = mph (λ x → x) (λ _ _ → refl) -Iso.inv (idIso G) = mph (λ x → x) (λ _ _ → refl) -Iso.rightInv (idIso G) _ = refl -Iso.leftInv (idIso G) _ = refl - -invIso : ∀ {ℓ ℓ'} {G : Group ℓ} {H : Group ℓ'} → Iso G H → Iso H G -Iso.fun (invIso is) = Iso.inv is -Iso.inv (invIso is) = Iso.fun is -Iso.rightInv (invIso is) = Iso.leftInv is -Iso.leftInv (invIso is) = Iso.rightInv is - -groupIso→Iso : ∀ {ℓ ℓ'} {G : Group ℓ} {H : Group ℓ'} → Iso G H → I.Iso (Group.type G) (Group.type H) -I.Iso.fun (groupIso→Iso i) = morph.fun (Iso.fun i) -I.Iso.inv (groupIso→Iso i) = morph.fun (Iso.inv i) -I.Iso.rightInv (groupIso→Iso i) = Iso.rightInv i -I.Iso.leftInv (groupIso→Iso i) = Iso.leftInv i - - ---- Proofs that different notions of ismomorphisms agree --- -Equiv→Iso' : ∀ {ℓ ℓ'} {G : Group ℓ} {H : Group ℓ'} → G ≃ H → Iso' G H -Iso'.isoSet (Equiv→Iso' {G = G} {H = H} e) = E.equivToIso (morph.fun (e .fst) , e .snd) -Iso'.isoSetMorph (Equiv→Iso' {G = G} {H = H} e) = morph.ismorph (e .fst) - -open isGroup -Iso'''→Iso : ∀ {ℓ ℓ'} {G : Group ℓ} {H : Group ℓ'} → Iso''' G H → Iso G H -Iso.fun (Iso'''→Iso is) = Iso'''.fun is -morph.fun (Iso.inv (Iso'''→Iso is)) = Iso'''.inv is -morph.ismorph (Iso.inv (Iso'''→Iso {G = G} {H = H} is)) a b = - cong₂ (λ x y → (Iso'''.inv is) (comp (Group.groupStruc H) x y)) - (sym (Iso'''.rightInv is _)) - (sym (Iso'''.rightInv is _)) - ∙ cong (Iso'''.inv is) (sym (morph.ismorph (Iso'''.fun is) _ _)) - ∙ Iso'''.leftInv is _ -Iso.rightInv (Iso'''→Iso is) = Iso'''.rightInv is -Iso.leftInv (Iso'''→Iso is) = Iso'''.leftInv is - -Iso'→Iso : ∀ {ℓ ℓ'} {G : Group ℓ} {H : Group ℓ'} → Iso' G H → Iso G H -morph.fun (Iso.fun (Iso'→Iso iso1)) = I.Iso.fun (Iso'.isoSet iso1) -morph.ismorph (Iso.fun (Iso'→Iso iso1)) = Iso'.isoSetMorph iso1 -morph.fun (Iso.inv (Iso'→Iso iso1)) = I.Iso.inv (Iso'.isoSet iso1) -morph.ismorph (Iso.inv (Iso'→Iso {G = G} {H = H} iso1)) a b = - cong₂ (λ x y → ψ (comp H' x y)) - (sym (I.Iso.rightInv (Iso'.isoSet iso1) _)) - (sym (I.Iso.rightInv (Iso'.isoSet iso1) _)) - ∙ cong ψ (sym (Iso'.isoSetMorph iso1 _ _)) - ∙ I.Iso.leftInv (Iso'.isoSet iso1) _ - where - H' = Group.groupStruc H - ψ = I.Iso.inv (Iso'.isoSet iso1) - -Iso.rightInv (Iso'→Iso iso1) = I.Iso.rightInv (Iso'.isoSet iso1) -Iso.leftInv (Iso'→Iso iso1) = I.Iso.leftInv (Iso'.isoSet iso1) - -open import Cubical.Data.Sigma hiding (comp ; _×_) - -Iso''→Iso : ∀ {ℓ ℓ'} {A : Group ℓ} {B : Group ℓ'} → Iso'' A B → Iso A B -Iso''→Iso {A = A} {B = B} is = Iso'''→Iso theIso - where - helper : (b : _) → isProp (Σ (Group.type A) (λ a → (morph.fun (Iso''.ϕ is)) a ≡ b)) - helper _ a b = - Σ≡Prop (λ _ → isOfHLevelPath' 1 (Group.setStruc B) _ _) - fstId - where - open Group - open morph - A' = groupStruc A - B' = groupStruc B - - fstIdHelper : comp (A') (fst a) (inv (A') (fst b)) - ≡ id (A') - fstIdHelper = - Iso''.inj is _ - (morph.ismorph (Iso''.ϕ is) (fst a) (inv A' (fst b)) - ∙ cong (λ x → comp B' (morph.fun (Iso''.ϕ is) (fst a)) x) (morphMinus A B (Iso''.ϕ is) (fst b)) - ∙ cong (λ x → comp B' x (inv B' (morph.fun (Iso''.ϕ is) (fst b)))) ((snd a) ∙ sym (snd b)) - ∙ rCancel B' (morph.fun (Iso''.ϕ is) (fst b))) - fstId : fst a ≡ fst b - fstId = - (fst a) ≡⟨ sym (rUnit A' (fst a)) ⟩ - comp A' (fst a) (id A') ≡⟨ cong (λ x → comp A' (fst a) x) (sym (lCancel A' (fst b))) ⟩ - comp A' (fst a) (comp A' (inv A' (fst b)) (fst b)) ≡⟨ sym (assoc A' (fst a) (inv A' (fst b)) (fst b)) ⟩ - comp A' (comp A' (fst a) (inv A' (fst b))) (fst b) ≡⟨ cong (λ x → comp A' x (fst b)) fstIdHelper ⟩ - comp A' (id A') (fst b) ≡⟨ lUnit A' (fst b) ⟩ - (fst b) ∎ - - theIso : Iso''' A B - Iso'''.fun theIso = Iso''.ϕ is - Iso'''.inv theIso b = rec (helper b) (λ a → a) (Iso''.surj is b) .fst - Iso'''.rightInv theIso b = rec (helper b) (λ a → a) (Iso''.surj is b) .snd - Iso'''.leftInv theIso b i = rec (helper (morph.fun (Iso''.ϕ is) b)) (λ a → a) - (propTruncIsProp (Iso''.surj is (morph.fun (Iso''.ϕ is) b)) ∣ b , refl ∣ i) .fst - - -SES→Iso : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group ℓ} {B : Group ℓ'} (leftGr : Group ℓ'') (rightGr : Group ℓ''') - → SES A B leftGr rightGr - → Iso A B -SES→Iso {A = A} lGr rGr sess = - Iso''→Iso - (iso'' (SES.ϕ sess) - (λ a inker → rec (Group.setStruc A _ _) - (λ {(a , p) → sym p ∙ cong (morph.fun (SES.left sess)) (SES.isTrivialLeft sess a _) - ∙ morph0→0 lGr A (SES.left sess)}) - (SES.Ker-ϕ⊂Im-left sess a inker)) - λ a → SES.Ker-right⊂Im-ϕ sess a (SES.isTrivialRight sess _ _)) - ---- Some elementary groups --- - -open import Cubical.Data.Unit -open Group -trivialGroup : Group ℓ-zero -type trivialGroup = Unit -setStruc trivialGroup = isOfHLevelSuc 1 isPropUnit -id (groupStruc trivialGroup) = tt -inv (groupStruc trivialGroup) _ = tt -comp (groupStruc trivialGroup) _ _ = tt -lUnit (groupStruc trivialGroup) _ = refl -rUnit (groupStruc trivialGroup) _ = refl -assoc (groupStruc trivialGroup) _ _ _ = refl -lCancel (groupStruc trivialGroup) _ = refl -rCancel (groupStruc trivialGroup) _ = refl - -open import Cubical.Data.Int - -intGroup : Group ℓ-zero -type intGroup = Int -setStruc intGroup = isSetInt -id (groupStruc intGroup) = pos 0 -inv (groupStruc intGroup) = pos 0 -_ -comp (groupStruc intGroup) = _+_ -lUnit (groupStruc intGroup) a = +-comm (pos 0) a -rUnit (groupStruc intGroup) _ = refl -assoc (groupStruc intGroup) a b c = sym (+-assoc a b c) -lCancel (groupStruc intGroup) a = minusPlus a 0 -rCancel (groupStruc intGroup) a = +-comm a (pos 0 - a) ∙ minusPlus a 0 - -dirProd : ∀ {ℓ ℓ'} (A : Group ℓ) (B : Group ℓ') → Group (ℓ-max ℓ ℓ') -type (dirProd A B) = type A × type B -setStruc (dirProd A B) = isOfHLevelProd 2 (setStruc A) (setStruc B) -id (groupStruc (dirProd A B)) = id (groupStruc A) , id (groupStruc B) -inv (groupStruc (dirProd A B)) = ×Inv - where - ×Inv : _ - ×Inv (a , b) = (inv (groupStruc A) a) , (inv (groupStruc B) b) -comp (groupStruc (dirProd A B)) = ×comp - where - ×comp : _ - ×comp (a , b) (c , d) = (comp (groupStruc A) a c) , comp (groupStruc B) b d -lUnit (groupStruc (dirProd A B)) = ×lUnit - where - ×lUnit : _ - ×lUnit (a , b) i = lUnit (groupStruc A) a i , lUnit (groupStruc B) b i -rUnit (groupStruc (dirProd A B)) = ×rUnit - where - ×rUnit : _ - ×rUnit (a , b) i = rUnit (groupStruc A) a i , rUnit (groupStruc B) b i -assoc (groupStruc (dirProd A B)) = ×assoc - where - ×assoc : _ - ×assoc (a , b) (c , d) (e , f) i = (assoc (groupStruc A) a c e i) , (assoc (groupStruc B) b d f i) -lCancel (groupStruc (dirProd A B)) = ×lCancel - where - ×lCancel : _ - ×lCancel (a , b) i = (lCancel (groupStruc A) a i) , lCancel (groupStruc B) b i -rCancel (groupStruc (dirProd A B)) = ×rCancel - where - ×rCancel : _ - ×rCancel (a , b) i = (rCancel (groupStruc A) a i) , rCancel (groupStruc B) b i - - -×morph : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group ℓ} {B : Group ℓ'} {C : Group ℓ''} {D : Group ℓ'''} - → morph A B → morph C D → morph (dirProd A C) (dirProd B D) -morph.fun (×morph mf1 mf2) = fun - where - fun : _ - fun (a , b) = morph.fun mf1 a , morph.fun mf2 b -morph.ismorph (×morph mf1 mf2) = ismf - where - ismf : _ - ismf (a , b) (c , d) i = morph.ismorph mf1 a c i , morph.ismorph mf2 b d i - -dirProdIso : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group ℓ} {B : Group ℓ'} {C : Group ℓ''} {D : Group ℓ'''} - → Iso A C → Iso B D - → Iso (dirProd A B) (dirProd C D) -Iso.fun (dirProdIso isoAC isoBD) = ×morph (Iso.fun isoAC) (Iso.fun isoBD) -Iso.inv (dirProdIso isoAC isoBD) = ×morph (Iso.inv isoAC) (Iso.inv isoBD) -Iso.rightInv (dirProdIso isoAC isoBD) = rInv - where - rInv : _ - rInv (a , b) = ×≡ (Iso.rightInv isoAC a) (Iso.rightInv isoBD b) -Iso.leftInv (dirProdIso isoAC isoBD) = lInv - where - lInv : _ - lInv (a , b) = ×≡ (Iso.leftInv isoAC a) (Iso.leftInv isoBD b) - - - ---- -lUnitGroupIso : ∀ {ℓ} {G : Group ℓ} → Iso (dirProd trivialGroup G) G -lUnitGroupIso = - Iso'''→Iso - (iso''' - (mph proj₂ λ { (x , y) (z , w) → refl}) - (λ g → tt , g) - (λ _ → refl) - λ _ → ×≡ refl refl) - -rUnitGroupIso : ∀ {ℓ} {G : Group ℓ} → Iso (dirProd G trivialGroup) G -rUnitGroupIso = - Iso'''→Iso - (iso''' - (mph proj₁ λ { (x , y) (z , w) → refl}) - (λ g → g , tt) - (λ _ → refl) - λ _ → ×≡ refl refl) diff --git a/Cubical/Data/Group/Properties.agda b/Cubical/Data/Group/Properties.agda deleted file mode 100644 index dba4f9bc0..000000000 --- a/Cubical/Data/Group/Properties.agda +++ /dev/null @@ -1,144 +0,0 @@ -{-# OPTIONS --cubical --no-import-sorts --safe #-} - -module Cubical.Data.Group.Properties where - -open import Cubical.Foundations.Prelude hiding (comp) -open import Cubical.Foundations.HLevels -open import Cubical.Data.Prod -open import Cubical.HITs.PropositionalTruncation hiding (map) -open import Cubical.Data.Group.Base -open import Cubical.Data.Sigma.Base hiding (comp) - - -{- -Given the following diagram - a ↦ (a , 0) ψ ϕ - A --> A × A -------> B ---> C -If ψ is an isomorphism and ϕ is surjective with ker ϕ ≡ {ψ (a , a) ∣ a ∈ A}, then C ≅ B --} - -diagonalIso1 : ∀ {ℓ ℓ' ℓ''} (A : Group ℓ) (B : Group ℓ') (C : Group ℓ'') - (ψ : Iso (dirProd A A) B) (ϕ : morph B C) - → isSurjective _ _ ϕ - → ((x : Group.type B) → isInKer B C ϕ x - → ∃[ y ∈ Group.type A ] x ≡ morph.fun (Iso.fun ψ) (y , y)) - → ((x : Group.type B) → (∃[ y ∈ Group.type A ] x ≡ morph.fun (Iso.fun ψ) (y , y)) - → isInKer B C ϕ x) - → Iso A C -diagonalIso1 A' B' C' ψ' ϕ' issurj ker→diag diag→ker = - Iso''→Iso - (iso'' (compMorph fstProj (compMorph (Iso.fun ψ') ϕ')) - (λ a inker - → rec (Group.setStruc A' _ _) - (λ {(a' , id') → cong proj₁ (sym (Iso.leftInv ψ' (a , id A)) - ∙ (cong (morph.fun (Iso.inv ψ')) id') - ∙ Iso.leftInv ψ' (a' , a')) - ∙ cong proj₂ (sym ((sym (Iso.leftInv ψ' (a , id A)) - ∙ (cong (morph.fun (Iso.inv ψ')) id') - ∙ Iso.leftInv ψ' (a' , a'))))}) - (ker→diag (ψ (a , id A)) inker)) - λ c → rec propTruncIsProp - (λ { (b , id') → ∣ comp A (proj₁ (ψ⁻ b)) (inv A (proj₂ (ψ⁻ b))) - , sym (rUnit C _) - ∙ cong (comp C _) (sym (diag→ker (ψ (proj₂ (ψ⁻ b) , proj₂ (ψ⁻ b))) - ∣ proj₂ (ψ⁻ b) , refl ∣)) - ∙ sym (ϕmf _ _) - ∙ cong ϕ (sym (Iso.rightInv ψ' _)) - ∙ cong (λ x → ϕ (ψ x)) (morph.ismorph (Iso.inv ψ') _ _) - ∙ cong (λ x → ϕ (ψ x)) (cong₂ (comp (groupStruc (dirProd A' A'))) - (Iso.leftInv ψ' _) - (Iso.leftInv ψ' _)) - ∙ cong (λ x → ϕ (ψ x)) (×≡ {a = _ , _} {b = _ , _} - (assoc A _ _ _) - (lUnit A _)) - ∙ cong (λ x → ϕ (ψ x)) (×≡ {a = _ , _} {b = _ , _} - (cong (comp A (proj₁ (ψ⁻ b))) (lCancel A _)) - refl) - ∙ cong (λ x → ϕ (ψ x)) (×≡ {a = _ , _} {b = _ , _} - (rUnit A _) - refl) - ∙ cong (λ x → ϕ (ψ x)) (sym (×-η (ψ⁻ b))) - ∙ cong ϕ (Iso.rightInv ψ' b) - ∙ id' ∣ }) - (issurj c)) - where - open Group - open isGroup - A = groupStruc A' - B = groupStruc B' - C = groupStruc C' - ϕ = morph.fun ϕ' - ϕmf = morph.ismorph ϕ' - ψ = morph.fun (Iso.fun ψ') - ψ⁻ = morph.fun (Iso.inv ψ') - - fstProj : morph A' (dirProd A' A') - morph.fun fstProj = λ a → a , (id A) - morph.ismorph fstProj = λ g0 g1 i → comp A g0 g1 , lUnit A (id A) (~ i) - -{- -Given the following diagram - - ϕ -A × A ↠ B - ^ - | - | a ↦ (a , 0) - | - A ≃ C - -If ϕ is surjective with ker ϕ ≡ {(a , a) ∣ a ∈ A}, then C ≅ B --} - -diagonalIso2 : ∀ {ℓ ℓ' ℓ''} (A : Group ℓ) (B : Group ℓ') (C : Group ℓ'') - → (ϕ : morph (dirProd A A) B) - → ((x : Group.type (dirProd A A)) → isInKer (dirProd A A) B ϕ x → ∃[ y ∈ Group.type A ] x ≡ (y , y)) - → ((x : Group.type (dirProd A A)) → ∃[ y ∈ Group.type A ] x ≡ (y , y) → isInKer (dirProd A A) B ϕ x) - → isSurjective (dirProd A A) B ϕ - → Iso C A - → Iso C B -diagonalIso2 A' B' C' ϕ' diagKer1 diagKer2 issurj I = - Iso''→Iso - (iso'' (compMorph (compMorph (Iso.fun I) fstProj) ϕ') - inj - surj) - where - open Group - open isGroup - A = groupStruc A' - B = groupStruc B' - C = groupStruc C' - ϕ = morph.fun ϕ' - ϕmf = morph.ismorph ϕ' - C→A = morph.fun (Iso.fun I) - C→A-mf = morph.ismorph (Iso.fun I) - - fstProj : morph A' (dirProd A' A') - morph.fun fstProj = λ a → a , (id A) - morph.ismorph fstProj = λ g0 g1 i → comp A g0 g1 , lUnit A (id A) (~ i) - - inj : (a : Group.type C') → isInKer C' B' (compMorph (compMorph (Iso.fun I) fstProj) ϕ') a - → a ≡ id C - inj a inker = rec (Group.setStruc C' _ _) - (λ {(b , id') → sym (Iso.leftInv I a) - ∙ cong (morph.fun (Iso.inv I)) (cong proj₁ id') - ∙ cong (morph.fun (Iso.inv I)) (cong proj₂ (sym id')) - ∙ morph0→0 A' C' (Iso.inv I)}) - (diagKer1 (C→A a , isGroup.id A) inker) - - surj : (b : Group.type B') → ∃[ c ∈ Group.type C' ] ϕ (morph.fun fstProj (C→A c)) ≡ b - surj b = rec propTruncIsProp - (λ {(a , id') → ∣ morph.fun (Iso.inv I) a , cong ϕ (×≡ {a = _ , _} {b = _ , _} (Iso.rightInv I a) refl) ∙ id' ∣}) - (idhelper b) - where - idhelper : (b : Group.type B') → ∃[ a ∈ Group.type A' ] ϕ (a , (id A)) ≡ b - idhelper b = rec propTruncIsProp - (λ {((a1 , a2) , id') → ∣ (comp A a1 (inv A a2)) - , sym (rUnit B (ϕ (comp A a1 (inv A a2) , id A))) - ∙ cong (comp B _) (sym (diagKer2 (a2 , a2) ∣ a2 , refl ∣)) - ∙ sym (ϕmf _ _) - ∙ cong ϕ (×≡ {a = _ , _} {b = _ , _} - (assoc A a1 (inv A a2) a2 ∙ cong (comp A a1) (lCancel A a2) ∙ rUnit A a1) - (lUnit A a2)) - ∙ id' ∣}) - (issurj b) diff --git a/Cubical/Homotopy/Connected.agda b/Cubical/Homotopy/Connected.agda index 24beb371d..67aa9a2f1 100644 --- a/Cubical/Homotopy/Connected.agda +++ b/Cubical/Homotopy/Connected.agda @@ -183,7 +183,7 @@ isConnectedRetract : ∀ {ℓ ℓ'} (n : HLevel) (f : A → B) (g : B → A) (h : (x : A) → g (f x) ≡ x) → isConnected n B → isConnected n A -isConnectedRetract n f g h = +isConnectedRetract n f g h = isContrRetract (Trunc.map f) (Trunc.map g) diff --git a/Cubical/Homotopy/Freudenthal.agda b/Cubical/Homotopy/Freudenthal.agda index 9d3433193..106267de5 100644 --- a/Cubical/Homotopy/Freudenthal.agda +++ b/Cubical/Homotopy/Freudenthal.agda @@ -85,7 +85,7 @@ module _ {ℓ} (n : HLevel) {A : Pointed ℓ} (connA : isConnected (suc (suc n)) encode y = J Code ∣ pt A , rCancel' (merid (pt A)) ∣ encodeMerid : (a : typ A) → encode south (merid a) ≡ ∣ a , refl ∣ - encodeMerid a = + encodeMerid a = cong (transport (λ i → gluePath i)) (funExt⁻ (WC.left refl a) _ ∙ cong ∣_∣ (cong (a ,_) (lem _ _))) ∙ transport (PathP≡Path gluePath _ _) diff --git a/Cubical/Structures/AbGroup.agda b/Cubical/Structures/AbGroup.agda index 846c7ea37..3622629bd 100644 --- a/Cubical/Structures/AbGroup.agda +++ b/Cubical/Structures/AbGroup.agda @@ -52,6 +52,7 @@ record AbGroup : Type (ℓ-suc ℓ) where open IsAbGroup isAbGroup public + -- Extractor for the carrier type ⟨_⟩ : AbGroup → Type ℓ ⟨_⟩ = AbGroup.Carrier @@ -116,21 +117,27 @@ module AbGroupΣTheory {ℓ} where AbGroupΣ→AbGroup (_ , _ , G , C) = abgroup _ _ _ _ (isabgroup (GroupΣ→Group (_ , _ , G) .Group.isGroup) C) + open AbGroup + open IsAbGroup + open Monoid + open IsMonoid + open IsGroup + AbGroupIsoAbGroupΣ : Iso AbGroup AbGroupΣ - AbGroupIsoAbGroupΣ = iso AbGroup→AbGroupΣ AbGroupΣ→AbGroup helper helper2 -- iso AbGroup→AbGroupΣ AbGroupΣ→AbGroup (λ _ → ?) + AbGroupIsoAbGroupΣ = iso AbGroup→AbGroupΣ AbGroupΣ→AbGroup helper helper2 where helper : _ fst (helper b i) = fst b snd (helper b i) = snd b helper2 : _ - AbGroup.Carrier (helper2 a i) = ⟨ a ⟩ - AbGroup.0g (helper2 a i) = AbGroup.0g a - AbGroup._+_ (helper2 a i) = AbGroup._+_ a - AbGroup.- helper2 a i = AbGroup.- a - IsGroup.isMonoid (IsAbGroup.isGroup (AbGroup.isAbGroup (helper2 a i))) = η-isMonoid (IsGroup.isMonoid (IsAbGroup.isGroup (AbGroup.isAbGroup a))) i - IsGroup.inverse (IsAbGroup.isGroup (AbGroup.isAbGroup (helper2 a i))) = IsGroup.inverse (IsAbGroup.isGroup (AbGroup.isAbGroup a)) - IsAbGroup.comm (AbGroup.isAbGroup (helper2 a i)) = IsAbGroup.comm (AbGroup.isAbGroup a) + Carrier (helper2 a i) = ⟨ a ⟩ + 0g (helper2 a i) = 0g a + _+_ (helper2 a i) = _+_ a + - helper2 a i = - a + isMonoid (isGroup (isAbGroup (helper2 a i))) = η-isMonoid (isMonoid (isGroup (isAbGroup a))) i + inverse (isGroup (isAbGroup (helper2 a i))) = inverse (isGroup (isAbGroup a)) + comm (isAbGroup (helper2 a i)) = comm (isAbGroup a) abGroupUnivalentStr : UnivalentStr AbGroupStructure AbGroupEquivStr abGroupUnivalentStr = axiomsUnivalentStr _ isPropAbGroupAxioms rawGroupUnivalentStr @@ -151,12 +158,12 @@ module AbGroupΣTheory {ℓ} where AbGroup→RawGroupΣ : AbGroup {ℓ} → RawGroupΣ AbGroup→RawGroupΣ (abgroup G _ _+_ _ _) = G , _+_ - InducedAbGroup : (G : AbGroup) (H : RawGroupΣ) (e : G .AbGroup.Carrier ≃ H .fst) + InducedAbGroup : (G : AbGroup) (H : RawGroupΣ) (e : G .Carrier ≃ H .fst) → RawGroupEquivStr (AbGroup→RawGroupΣ G) H e → AbGroup InducedAbGroup G H e r = AbGroupΣ→AbGroup (transferAxioms rawGroupUnivalentStr (AbGroup→AbGroupΣ G) H (e , r)) - InducedAbGroupPath : (G : AbGroup {ℓ}) (H : RawGroupΣ) (e : G .AbGroup.Carrier ≃ H .fst) + InducedAbGroupPath : (G : AbGroup {ℓ}) (H : RawGroupΣ) (e : G .Carrier ≃ H .fst) (E : RawGroupEquivStr (AbGroup→RawGroupΣ G) H e) → G ≡ InducedAbGroup G H e E InducedAbGroupPath G H e E = diff --git a/Cubical/Structures/CommRing.agda b/Cubical/Structures/CommRing.agda index dc810d47d..17f69bfd9 100644 --- a/Cubical/Structures/CommRing.agda +++ b/Cubical/Structures/CommRing.agda @@ -131,27 +131,33 @@ module CommRingΣTheory {ℓ} where where open import Cubical.Structures.Group.Base hiding (⟨_⟩) + open CommRing + open IsCommRing + open IsGroup + open Group + open IsMonoid + open IsAbGroup + open IsRing + helper : _ - CommRing.Carrier (helper a i) = ⟨ a ⟩ - CommRing.0r (helper a i) = CommRing.0r a - CommRing.1r (helper a i) = CommRing.1r a - CommRing._+_ (helper a i) = CommRing._+_ a - CommRing._·_ (helper a i) = CommRing._·_ a - CommRing.- helper a i = CommRing.- a - IsGroup.isMonoid (IsAbGroup.isGroup (IsRing.+-isAbGroup (IsCommRing.isRing (CommRing.isCommRing (helper a i))))) = - η-isMonoid (IsGroup.isMonoid (IsAbGroup.isGroup (IsRing.+-isAbGroup (IsCommRing.isRing (CommRing.isCommRing a))))) i - IsGroup.inverse (IsAbGroup.isGroup (IsRing.+-isAbGroup (IsCommRing.isRing (CommRing.isCommRing (helper a i))))) = - IsGroup.inverse (IsAbGroup.isGroup (IsRing.+-isAbGroup (IsCommRing.isRing (CommRing.isCommRing a)))) - IsAbGroup.comm (IsRing.+-isAbGroup (IsCommRing.isRing (CommRing.isCommRing (helper a i)))) = - IsAbGroup.comm (IsRing.+-isAbGroup (IsCommRing.isRing (CommRing.isCommRing a))) - IsSemigroup.is-set (IsMonoid.isSemigroup (IsRing.·-isMonoid (IsCommRing.isRing (CommRing.isCommRing (helper a i))))) = - IsSemigroup.is-set (IsMonoid.isSemigroup (IsRing.·-isMonoid (IsCommRing.isRing (CommRing.isCommRing a)))) - IsSemigroup.assoc (IsMonoid.isSemigroup (IsRing.·-isMonoid (IsCommRing.isRing (CommRing.isCommRing (helper a i))))) = - IsSemigroup.assoc (IsMonoid.isSemigroup (IsRing.·-isMonoid (IsCommRing.isRing (CommRing.isCommRing a)))) - IsMonoid.identity (IsRing.·-isMonoid (IsCommRing.isRing (CommRing.isCommRing (helper a i)))) = - IsMonoid.identity (IsRing.·-isMonoid (IsCommRing.isRing (CommRing.isCommRing a))) - IsRing.dist (IsCommRing.isRing (CommRing.isCommRing (helper a i))) = IsRing.dist (IsCommRing.isRing (CommRing.isCommRing a)) - IsCommRing.·-comm (CommRing.isCommRing (helper a i)) = IsCommRing.·-comm (CommRing.isCommRing a) + Carrier (helper a i) = ⟨ a ⟩ + 0r (helper a i) = 0r a + 1r (helper a i) = 1r a + _+_ (helper a i) = CommRing._+_ a + _·_ (helper a i) = _·_ a + - helper a i = - a + isMonoid (isGroup (+-isAbGroup (isRing (isCommRing (helper a i))))) = + η-isMonoid (isMonoid (isGroup (+-isAbGroup (isRing (isCommRing a))))) i + inverse (isGroup (+-isAbGroup (isRing (isCommRing (helper a i))))) = + inverse (isGroup (+-isAbGroup (isRing (isCommRing a)))) + comm (+-isAbGroup (isRing (isCommRing (helper a i)))) = + comm (+-isAbGroup (isRing (isCommRing a))) + isSemigroup (·-isMonoid (isRing (isCommRing (helper a i)))) = + isSemigroup (·-isMonoid (isRing (isCommRing a))) + identity (·-isMonoid (isRing (isCommRing (helper a i)))) = + identity (·-isMonoid (isRing (isCommRing a))) + dist (isRing (isCommRing (helper a i))) = dist (isRing (isCommRing a)) + ·-comm (isCommRing (helper a i)) = ·-comm (isCommRing a) commRingUnivalentStr : UnivalentStr CommRingStructure CommRingEquivStr commRingUnivalentStr = axiomsUnivalentStr _ isPropCommRingAxioms rawRingUnivalentStr diff --git a/Cubical/Structures/Ring.agda b/Cubical/Structures/Ring.agda index 3d39d01b2..a59eaa9d6 100644 --- a/Cubical/Structures/Ring.agda +++ b/Cubical/Structures/Ring.agda @@ -210,7 +210,7 @@ module RingΣTheory {ℓ} where Cubical.Structures.Group.Base.IsGroup.isMonoid (IsAbGroup.isGroup (IsRing.+-isAbGroup (Ring.isRing (helper a i)))) = η-isMonoid (Cubical.Structures.Group.Base.IsGroup.isMonoid (IsAbGroup.isGroup (IsRing.+-isAbGroup (Ring.isRing a)))) i Cubical.Structures.Group.Base.IsGroup.inverse (IsAbGroup.isGroup (IsRing.+-isAbGroup (Ring.isRing (helper a i)))) = - Cubical.Structures.Group.Base.IsGroup.inverse (IsAbGroup.isGroup (IsRing.+-isAbGroup (Ring.isRing a))) + Cubical.Structures.Group.Base.IsGroup.inverse (IsAbGroup.isGroup (IsRing.+-isAbGroup (Ring.isRing a))) IsAbGroup.comm (IsRing.+-isAbGroup (Ring.isRing (helper a i))) = IsAbGroup.comm (IsRing.+-isAbGroup (Ring.isRing a)) IsSemigroup.is-set (IsMonoid.isSemigroup (IsRing.·-isMonoid (Ring.isRing (helper a i)))) = IsSemigroup.is-set (IsMonoid.isSemigroup (IsRing.·-isMonoid (Ring.isRing a))) diff --git a/Cubical/ZCohomology/Groups/Connected.agda b/Cubical/ZCohomology/Groups/Connected.agda index d953248ce..e0de46c97 100644 --- a/Cubical/ZCohomology/Groups/Connected.agda +++ b/Cubical/ZCohomology/Groups/Connected.agda @@ -3,31 +3,21 @@ module Cubical.ZCohomology.Groups.Connected where open import Cubical.ZCohomology.Base open import Cubical.ZCohomology.Properties -open import Cubical.ZCohomology.MayerVietorisUnreduced open import Cubical.ZCohomology.Groups.Unit -open import Cubical.ZCohomology.KcompPrelims open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Function open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.GroupoidLaws -open import Cubical.HITs.Pushout -open import Cubical.HITs.Sn -open import Cubical.HITs.S1 -open import Cubical.HITs.Susp open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) -open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₁ ; ∣_∣ to ∣_∣₁) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; ∥_∥ to ∥_∥₁ ; ∣_∣ to ∣_∣₁) open import Cubical.HITs.Nullification open import Cubical.Data.Sigma hiding (_×_) open import Cubical.Data.Int renaming (_+_ to _+ℤ_; +-comm to +ℤ-comm ; +-assoc to +ℤ-assoc) open import Cubical.Data.Nat -open import Cubical.Data.Prod -open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec) -open import Cubical.Data.Unit +open import Cubical.HITs.Truncation renaming (rec to trRec) open import Cubical.Structures.Group open import Cubical.Homotopy.Connected diff --git a/Cubical/ZCohomology/Groups/Wedge.agda b/Cubical/ZCohomology/Groups/Wedge.agda index 3100db2ad..e3ee9f169 100644 --- a/Cubical/ZCohomology/Groups/Wedge.agda +++ b/Cubical/ZCohomology/Groups/Wedge.agda @@ -13,7 +13,6 @@ open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; ∣_∣ open import Cubical.Data.Nat open import Cubical.Data.Prod open import Cubical.Data.Unit --- open import Cubical.Data.Group.Base renaming (Iso to grIso ; compIso to compGrIso ; invIso to invGrIso ; idIso to idGrIso) open import Cubical.Structures.Group open import Cubical.ZCohomology.Groups.Unit @@ -24,7 +23,7 @@ open import Cubical.HITs.Pushout --- This module contains a proof that Hⁿ(A ⋁ B) ≅ Hⁿ(A) × Hⁿ(B), n ≥ 1 module _ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') where - module I = MV (typ A) (typ B) Unit (λ _ → pt A) (λ _ → pt B) + module I = MV (typ A) (typ B) Unit (λ _ → pt A) (λ _ → pt B) Hⁿ-⋁ : (n : ℕ) → GroupEquiv (coHomGr (suc n) (A ⋁ B)) (×coHomGr (suc n) (typ A) (typ B)) Hⁿ-⋁ zero = @@ -47,7 +46,7 @@ module _ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') where helper : (x : coHom 1 (A ⋁ B)) → isInIm _ _ (I.d 0) x → x ≡ 0ₕ - helper x inim = + helper x inim = pRec (setTruncIsSet _ _) (λ p → sym (snd p) ∙ MV.Im-Δ⊂Ker-d _ _ Unit (λ _ → pt A) (λ _ → pt B) 0 (fst p) (surj-helper (fst p))) diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index b5c1137f5..57bfc284e 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -27,7 +27,6 @@ open import Cubical.Homotopy.Connected open import Cubical.Homotopy.Freudenthal open import Cubical.HITs.SmashProduct.Base open import Cubical.Structures.Group -open import Cubical.Data.Group.Base hiding (_≃_ ; Iso ; invIso ; Group ; dirProd) open import Cubical.Foundations.Equiv.HalfAdjoint From 1eeab4982c3d4fcdd9954cd322f6260d09db9590 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Fri, 10 Jul 2020 04:16:39 +0200 Subject: [PATCH 53/89] whiteSpaceFiller4 --- Cubical/HITs/S1/Base.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Cubical/HITs/S1/Base.agda b/Cubical/HITs/S1/Base.agda index d5c4fa1fb..65612e377 100644 --- a/Cubical/HITs/S1/Base.agda +++ b/Cubical/HITs/S1/Base.agda @@ -284,7 +284,7 @@ basechange2 (loop i) p = ((λ j → loop (~ j ∧ i)) ∙ p ∙ λ j → loop (j ∧ i)) basechange2⁻ : (x : S¹) → (x ≡ x) → ΩS¹ basechange2⁻ base p = p -basechange2⁻ (loop i) p = +basechange2⁻ (loop i) p = hcomp (λ k → λ { (i = i0) → lUnit (rUnit p (~ k)) (~ k) ; (i = i1) → (cong (loop ∙_) (comm-ΩS¹ p (sym loop)) ∙ assoc loop (sym loop) p From 3ab33410e9ce254200aa386ffa4db3c9025f0e28 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Thu, 16 Jul 2020 17:50:14 +0200 Subject: [PATCH 54/89] fixes --- Cubical/Data/Empty/Base.agda | 3 + Cubical/Data/Sigma/Properties.agda | 21 +++---- Cubical/Data/Unit/Properties.agda | 18 +++--- Cubical/Foundations/Isomorphism.agda | 5 -- Cubical/Foundations/Prelude.agda | 10 ++-- Cubical/HITs/S1/Base.agda | 15 +++-- Cubical/HITs/SetTruncation/Properties.agda | 19 +++---- Cubical/Homotopy/Loopspace.agda | 6 +- Cubical/Structures/AbGroup.agda | 10 +--- Cubical/Structures/CommRing.agda | 17 +++--- Cubical/Structures/Group/Algebra.agda | 25 ++++----- Cubical/Structures/Group/Base.agda | 27 ++++++--- .../Structures/Group/MorphismProperties.agda | 43 +++++--------- Cubical/Structures/Monoid.agda | 1 - Cubical/Structures/Semigroup.agda | 2 +- Cubical/ZCohomology/Groups/Prelims.agda | 56 +++++++++---------- Cubical/ZCohomology/Groups/Sn.agda | 6 +- Cubical/ZCohomology/Groups/Unit.agda | 6 +- .../ZCohomology/MayerVietorisUnreduced.agda | 10 ++-- Cubical/ZCohomology/Properties.agda | 12 ++-- 20 files changed, 149 insertions(+), 163 deletions(-) diff --git a/Cubical/Data/Empty/Base.agda b/Cubical/Data/Empty/Base.agda index 43cef0613..e55bc3860 100644 --- a/Cubical/Data/Empty/Base.agda +++ b/Cubical/Data/Empty/Base.agda @@ -10,3 +10,6 @@ rec () elim : ∀ {ℓ} {A : ⊥ → Type ℓ} → (x : ⊥) → A x elim () + +nope : ∀ {ℓ} (A : Type ℓ) → Type ℓ +nope A = A → ⊥ diff --git a/Cubical/Data/Sigma/Properties.agda b/Cubical/Data/Sigma/Properties.agda index 42686cccc..6d0d9ced0 100644 --- a/Cubical/Data/Sigma/Properties.agda +++ b/Cubical/Data/Sigma/Properties.agda @@ -248,16 +248,17 @@ PathΣ→ΣPathTransport a b = invEq (ΣPathTransport≃PathΣ a b) map-× (fst eq1) (fst eq2) , record { equiv-proof - = λ {(c , d) → ((equiv-proof (eq1 .snd) c .fst .fst - , equiv-proof (eq2 .snd) d .fst .fst) - , ≡-× (equiv-proof (eq1 .snd) c .fst .snd) - (equiv-proof (eq2 .snd) d .fst .snd)) - , λ {((a , b) , p) → ΣPathP (≡-× (cong fst (equiv-proof (eq1 .snd) c .snd (a , cong fst p))) - (cong fst (equiv-proof (eq2 .snd) d .snd (b , cong snd p))) - , λ i → ≡-× (snd ((equiv-proof (eq1 .snd) c .snd (a , cong fst p)) i)) - (snd ((equiv-proof (eq2 .snd) d .snd (b , cong snd p)) i)))}}} - - + = λ {(c , d) → ((eq1⁻ c .fst .fst + , eq2⁻ d .fst .fst) + , ≡-× (eq1⁻ c .fst .snd) + (eq2⁻ d .fst .snd)) + , λ {((a , b) , p) → ΣPathP (≡-× (cong fst (eq1⁻ c .snd (a , cong fst p))) + (cong fst (eq2⁻ d .snd (b , cong snd p))) + , λ i → ≡-× (snd ((eq1⁻ c .snd (a , cong fst p)) i)) + (snd ((eq2⁻ d .snd (b , cong snd p)) i)))}}} + where + eq1⁻ = equiv-proof (eq1 .snd) + eq2⁻ = equiv-proof (eq2 .snd) {- Some simple ismorphisms -} diff --git a/Cubical/Data/Unit/Properties.agda b/Cubical/Data/Unit/Properties.agda index 7f4c7f109..97a01f082 100644 --- a/Cubical/Data/Unit/Properties.agda +++ b/Cubical/Data/Unit/Properties.agda @@ -13,6 +13,7 @@ open import Cubical.Data.Prod.Base open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence isContrUnit : isContr Unit isContrUnit = tt , λ {tt → refl} @@ -26,14 +27,14 @@ isSetUnit = isProp→isSet isPropUnit isOfHLevelUnit : (n : HLevel) → isOfHLevel n Unit isOfHLevelUnit n = isContr→isOfHLevel n isContrUnit -UnitToTypeId : ∀ {ℓ} (A : Type ℓ) → (Unit → A) ≡ A -UnitToTypeId A = isoToPath (iso (λ f → f tt) (λ a _ → a) (λ _ → refl) λ _ → refl) +UnitToTypePath : ∀ {ℓ} (A : Type ℓ) → (Unit → A) ≡ A +UnitToTypePath A = isoToPath (iso (λ f → f tt) (λ a _ → a) (λ _ → refl) λ _ → refl) -ContrToTypeIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → isContr A → Iso (A → B) B -Iso.fun (ContrToTypeIso iscontr) f = f (fst iscontr) -Iso.inv (ContrToTypeIso iscontr) b _ = b -Iso.rightInv (ContrToTypeIso iscontr) _ = refl -Iso.leftInv (ContrToTypeIso iscontr) f = funExt λ x → cong f (snd iscontr x) +isContr→Iso2 : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → isContr A → Iso (A → B) B +Iso.fun (isContr→Iso2 iscontr) f = f (fst iscontr) +Iso.inv (isContr→Iso2 iscontr) b _ = b +Iso.rightInv (isContr→Iso2 iscontr) _ = refl +Iso.leftInv (isContr→Iso2 iscontr) f = funExt λ x → cong f (snd iscontr x) diagonal-unit : Unit ≡ Unit × Unit diagonal-unit = isoToPath (iso (λ x → tt , tt) (λ x → tt) (λ {(tt , tt) i → tt , tt}) λ {tt i → tt}) @@ -46,8 +47,9 @@ fibId A = (λ _ → refl) (λ a i → fst a , isOfHLevelSuc 1 isPropUnit _ _ (snd a) refl i)) + isContr→≃Unit : ∀ {ℓ} {A : Type ℓ} → isContr A → A ≃ Unit isContr→≃Unit contr = isoToEquiv (iso (λ _ → tt) (λ _ → fst contr) (λ _ → refl) λ _ → snd contr _) isContr→≡Unit : {A : Type₀} → isContr A → A ≡ Unit -isContr→≡Unit contr = isoToPath (iso (λ _ → tt) (λ _ → fst contr) (λ _ → refl) λ _ → snd contr _) +isContr→≡Unit contr = ua (isContr→≃Unit contr) diff --git a/Cubical/Foundations/Isomorphism.agda b/Cubical/Foundations/Isomorphism.agda index 236f4a1e2..753c0ea45 100644 --- a/Cubical/Foundations/Isomorphism.agda +++ b/Cubical/Foundations/Isomorphism.agda @@ -125,11 +125,6 @@ rightInv (compIso (iso _ _ rightInv _) (iso g ginv rightInv' _)) b = leftInv (compIso (iso f finv _ leftInv) (iso _ _ _ leftInv')) a = cong finv (leftInv' (f a)) ∙ leftInv a -infixr 33 _⋄_ - -_⋄_ : _ -_⋄_ = compIso - composesToId→Iso : (G : Iso A B) (g : B → A) → G .fun ∘ g ≡ idfun B → Iso B A fun (composesToId→Iso _ g _) = g inv (composesToId→Iso (iso f _ _ _) _ _) = f diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index 43b99890a..4e9946e35 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -168,7 +168,7 @@ compPathP' : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} {x y z : A} {x' compPathP' {B = B} {x' = x'} {p = p} {q = q} P Q i = comp (λ j → B (compPath-filler p q j i)) (λ j → λ { (i = i0) → x' ; - (i = i1) → Q j }) + (i = i1) → Q j }) (P i) @@ -188,10 +188,10 @@ compPathP'-filler : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} {x y z : → PathP (λ j → PathP (λ i → B (compPath-filler p q j i)) x' (Q j)) P (compPathP' {x = x} {y = y} {x' = x'} {y' = y'} P Q) compPathP'-filler {B = B} {x' = x'} {p = p} {q = q} P Q j i = fill (λ j → B (compPath-filler p q j i)) - (λ j → λ { (i = i0) → x' ; - (i = i1) → Q j }) - (inS (P i)) - j + (λ j → λ { (i = i0) → x' ; + (i = i1) → Q j }) + (inS (P i)) + j _≡⟨_⟩_ : (x : A) → x ≡ y → y ≡ z → x ≡ z _ ≡⟨ x≡y ⟩ y≡z = x≡y ∙ y≡z diff --git a/Cubical/HITs/S1/Base.agda b/Cubical/HITs/S1/Base.agda index 65612e377..259ff77f2 100644 --- a/Cubical/HITs/S1/Base.agda +++ b/Cubical/HITs/S1/Base.agda @@ -253,11 +253,11 @@ private (basechange-isequiv-aux i0) t }) (basechange-isequiv-aux i) - basedΩS¹≡ΩS¹ : (x : S¹) → basedΩS¹ x ≡ ΩS¹ - basedΩS¹≡ΩS¹ x = ua (basechange x , basechange-isequiv x) + basedΩS¹≡ΩS¹' : (x : S¹) → basedΩS¹ x ≡ ΩS¹ + basedΩS¹≡ΩS¹' x = ua (basechange x , basechange-isequiv x) -basedΩS¹≡Int : (x : S¹) → basedΩS¹ x ≡ Int -basedΩS¹≡Int x = (basedΩS¹≡ΩS¹ x) ∙ ΩS¹≡Int + basedΩS¹≡Int' : (x : S¹) → basedΩS¹ x ≡ Int + basedΩS¹≡Int' x = (basedΩS¹≡ΩS¹' x) ∙ ΩS¹≡Int ---- Alternative proof of the same thing ----- @@ -301,12 +301,15 @@ basechange2-retr = toPropElim (λ s → isOfHLevelΠ 1 λ x → isSetΩx _ _ _) λ _ → refl -basedΩS¹≡ΩS¹' : (x : S¹) → basedΩS¹ x ≡ ΩS¹ -basedΩS¹≡ΩS¹' x = isoToPath (iso (basechange2⁻ x) +basedΩS¹≡ΩS¹ : (x : S¹) → basedΩS¹ x ≡ ΩS¹ +basedΩS¹≡ΩS¹ x = isoToPath (iso (basechange2⁻ x) (basechange2 x) (basechange2-retr x) (basechange2-sect x)) +basedΩS¹≡Int : (x : S¹) → basedΩS¹ x ≡ Int +basedΩS¹≡Int x = (basedΩS¹≡ΩS¹ x) ∙ ΩS¹≡Int + -- baschange2⁻ is a morphism basechange2⁻-morph : (x : S¹) (p q : x ≡ x) → diff --git a/Cubical/HITs/SetTruncation/Properties.agda b/Cubical/HITs/SetTruncation/Properties.agda index bf874f086..746652321 100644 --- a/Cubical/HITs/SetTruncation/Properties.agda +++ b/Cubical/HITs/SetTruncation/Properties.agda @@ -21,9 +21,9 @@ open import Cubical.Data.Sigma private variable ℓ : Level - A : Type ℓ + A B C D : Type ℓ -rec : {B : Type ℓ} → isSet B → (A → B) → ∥ A ∥₂ → B +rec : isSet B → (A → B) → ∥ A ∥₂ → B rec Bset f ∣ x ∣₂ = f x rec Bset f (squash₂ x y p q i j) = Bset _ _ (cong (rec Bset f) p) (cong (rec Bset f) q) i j @@ -38,7 +38,7 @@ elim Bset g (squash₂ x y p q i j) = isOfHLevel→isOfHLevelDep 2 Bset _ _ (cong (elim Bset g) p) (cong (elim Bset g) q) (squash₂ x y p q) i j -setTruncUniversal : {B : Type ℓ} → isSet B → (∥ A ∥₂ → B) ≃ (A → B) +setTruncUniversal : isSet B → (∥ A ∥₂ → B) ≃ (A → B) setTruncUniversal {B = B} Bset = isoToEquiv (iso (λ h x → h ∣ x ∣₂) (rec Bset) (λ _ → refl) rinv) where @@ -76,13 +76,13 @@ setTruncIdempotent≃ {A = A} hA = isoToEquiv f setTruncIdempotent : isSet A → ∥ A ∥₂ ≡ A setTruncIdempotent hA = ua (setTruncIdempotent≃ hA) -isContr→isContrSetTrunc : ∀ {ℓ} {A : Type ℓ} → isContr A → isContr (∥ A ∥₂) +isContr→isContrSetTrunc : isContr A → isContr (∥ A ∥₂) isContr→isContrSetTrunc contr = ∣ fst contr ∣₂ , elim (λ _ → isOfHLevelPath 2 (setTruncIsSet) _ _) λ a → cong ∣_∣₂ (snd contr a) -setTruncIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → Iso A B → Iso ∥ A ∥₂ ∥ B ∥₂ +setTruncIso : Iso A B → Iso ∥ A ∥₂ ∥ B ∥₂ Iso.fun (setTruncIso is) = rec setTruncIsSet (λ x → ∣ Iso.fun is x ∣₂) Iso.inv (setTruncIso is) = rec setTruncIsSet (λ x → ∣ Iso.inv is x ∣₂) Iso.rightInv (setTruncIso is) = @@ -116,7 +116,7 @@ sigmaElim {B = B} {C = C} set g (x , y) = elim {B = λ x → (y : B x) → C (x (λ _ → isOfHLevelΠ 2 λ _ → set (_ , _)) g x y -sigmaProdElim : ∀ {ℓ ℓ' ℓ''} {B : Type ℓ} {C : ∥ A ∥₂ × ∥ B ∥₂ → Type ℓ'} {D : Σ (∥ A ∥₂ × ∥ B ∥₂) C → Type ℓ''} +sigmaProdElim : ∀ {ℓ ℓ'} {C : ∥ A ∥₂ × ∥ B ∥₂ → Type ℓ} {D : Σ (∥ A ∥₂ × ∥ B ∥₂) C → Type ℓ'} (Bset : (x : Σ (∥ A ∥₂ × ∥ B ∥₂) C) → isSet (D x)) (g : (a : A) (b : B) (c : C (∣ a ∣₂ , ∣ b ∣₂)) → D ((∣ a ∣₂ , ∣ b ∣₂) , c)) (x : Σ (∥ A ∥₂ × ∥ B ∥₂) C) → D x @@ -128,7 +128,7 @@ sigmaProdElim {B = B} {C = C} {D = D} set g ((x , y) , c) = x y c -prodElim : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : ∥ A ∥₂ × ∥ B ∥₂ → Type ℓ''} +prodElim : ∀ {ℓ} {C : ∥ A ∥₂ × ∥ B ∥₂ → Type ℓ} → ((x : ∥ A ∥₂ × ∥ B ∥₂) → isOfHLevel 2 (C x)) → ((a : A) (b : B) → C (∣ a ∣₂ , ∣ b ∣₂)) → (x : ∥ A ∥₂ × ∥ B ∥₂) → C x @@ -138,8 +138,7 @@ prodElim {A = A} {B = B} {C = C} hlevel ind (a , b) = schonf a b schonf = elim (λ x → isOfHLevelΠ 2 λ y → hlevel (_ , _)) λ a → elim (λ x → hlevel (_ , _)) λ b → ind a b -prodElim2 : ∀ {ℓ ℓ' ℓ'' ℓ''' ℓ''''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} - {E : (∥ A ∥₂ × ∥ B ∥₂) → (∥ C ∥₂ × ∥ D ∥₂) → Type ℓ''''} +prodElim2 : ∀ {ℓ} {E : (∥ A ∥₂ × ∥ B ∥₂) → (∥ C ∥₂ × ∥ D ∥₂) → Type ℓ} → ((x : ∥ A ∥₂ × ∥ B ∥₂) (y : ∥ C ∥₂ × ∥ D ∥₂) → isOfHLevel 2 (E x y)) → ((a : A) (b : B) (c : C) (d : D) → E (∣ a ∣₂ , ∣ b ∣₂) (∣ c ∣₂ , ∣ d ∣₂)) → ((x : ∥ A ∥₂ × ∥ B ∥₂) (y : ∥ C ∥₂ × ∥ D ∥₂) → (E x y)) @@ -148,7 +147,7 @@ prodElim2 isset f = prodElim (λ _ → isOfHLevelΠ 2 λ _ → isset _ _) λ c d → f a b c d -setTruncOfProdIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → Iso ∥ A × B ∥₂ (∥ A ∥₂ × ∥ B ∥₂) +setTruncOfProdIso : Iso ∥ A × B ∥₂ (∥ A ∥₂ × ∥ B ∥₂) Iso.fun setTruncOfProdIso = rec (isOfHLevelΣ 2 setTruncIsSet (λ _ → setTruncIsSet)) λ { (a , b) → ∣ a ∣₂ , ∣ b ∣₂ } Iso.inv setTruncOfProdIso = prodElim (λ _ → setTruncIsSet) λ a b → ∣ a , b ∣₂ Iso.rightInv setTruncOfProdIso = diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda index fd45e71a6..c6e99c623 100644 --- a/Cubical/Homotopy/Loopspace.agda +++ b/Cubical/Homotopy/Loopspace.agda @@ -26,13 +26,11 @@ open import Cubical.HITs.SetTruncation Ω→ (f , f∙) = (λ p → (sym f∙ ∙ cong f p) ∙ f∙) , cong (λ q → q ∙ f∙) (sym (rUnit (sym f∙))) ∙ lCancel f∙ generalEH : {ℓ : Level} {A : Type ℓ} {a b c : A} {p q : a ≡ b} {r s : b ≡ c} (α : p ≡ q) (β : r ≡ s) - → (cong (_∙ r) α) ∙ (cong (q ∙_) β) - ≡ (cong (p ∙_) β) ∙ (cong (_∙ s) α) -generalEH {p = p} {r = r} α β j i = + → (cong (_∙ r) α) ∙ (cong (q ∙_) β) ≡ (cong (p ∙_) β) ∙ (cong (_∙ s) α) +generalEH {p = p} {r = r} α β j i = hcomp (λ k → λ { (i = i0) → p ∙ r ; (i = i1) → α (k ∨ ~ j) ∙ β (k ∨ j) }) (α (~ j ∧ i) ∙ β (j ∧ i)) - Eckmann-Hilton : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (α β : typ ((Ω^ (2 + n)) A)) → α ∙ β ≡ β ∙ α Eckmann-Hilton {A = A} n α β i = diff --git a/Cubical/Structures/AbGroup.agda b/Cubical/Structures/AbGroup.agda index 3622629bd..57989a28e 100644 --- a/Cubical/Structures/AbGroup.agda +++ b/Cubical/Structures/AbGroup.agda @@ -118,10 +118,6 @@ module AbGroupΣTheory {ℓ} where abgroup _ _ _ _ (isabgroup (GroupΣ→Group (_ , _ , G) .Group.isGroup) C) open AbGroup - open IsAbGroup - open Monoid - open IsMonoid - open IsGroup AbGroupIsoAbGroupΣ : Iso AbGroup AbGroupΣ AbGroupIsoAbGroupΣ = iso AbGroup→AbGroupΣ AbGroupΣ→AbGroup helper helper2 @@ -135,9 +131,9 @@ module AbGroupΣTheory {ℓ} where 0g (helper2 a i) = 0g a _+_ (helper2 a i) = _+_ a - helper2 a i = - a - isMonoid (isGroup (isAbGroup (helper2 a i))) = η-isMonoid (isMonoid (isGroup (isAbGroup a))) i - inverse (isGroup (isAbGroup (helper2 a i))) = inverse (isGroup (isAbGroup a)) - comm (isAbGroup (helper2 a i)) = comm (isAbGroup a) + IsGroup.isMonoid (IsAbGroup.isGroup (isAbGroup (helper2 a i))) = η-isMonoid (isMonoid a) i + IsGroup.inverse (IsAbGroup.isGroup (isAbGroup (helper2 a i))) = inverse a + IsAbGroup.comm (isAbGroup (helper2 a i)) = comm a abGroupUnivalentStr : UnivalentStr AbGroupStructure AbGroupEquivStr abGroupUnivalentStr = axiomsUnivalentStr _ isPropAbGroupAxioms rawGroupUnivalentStr diff --git a/Cubical/Structures/CommRing.agda b/Cubical/Structures/CommRing.agda index 17f69bfd9..b85d961c1 100644 --- a/Cubical/Structures/CommRing.agda +++ b/Cubical/Structures/CommRing.agda @@ -132,12 +132,11 @@ module CommRingΣTheory {ℓ} where where open import Cubical.Structures.Group.Base hiding (⟨_⟩) open CommRing - open IsCommRing open IsGroup - open Group open IsMonoid open IsAbGroup open IsRing + open IsCommRing helper : _ Carrier (helper a i) = ⟨ a ⟩ @@ -147,17 +146,17 @@ module CommRingΣTheory {ℓ} where _·_ (helper a i) = _·_ a - helper a i = - a isMonoid (isGroup (+-isAbGroup (isRing (isCommRing (helper a i))))) = - η-isMonoid (isMonoid (isGroup (+-isAbGroup (isRing (isCommRing a))))) i + η-isMonoid (isMonoid (isGroup (+-isAbGroup a))) i inverse (isGroup (+-isAbGroup (isRing (isCommRing (helper a i))))) = - inverse (isGroup (+-isAbGroup (isRing (isCommRing a)))) + inverse (isGroup (+-isAbGroup a)) comm (+-isAbGroup (isRing (isCommRing (helper a i)))) = - comm (+-isAbGroup (isRing (isCommRing a))) + comm (+-isAbGroup a) isSemigroup (·-isMonoid (isRing (isCommRing (helper a i)))) = - isSemigroup (·-isMonoid (isRing (isCommRing a))) + isSemigroup (·-isMonoid a) identity (·-isMonoid (isRing (isCommRing (helper a i)))) = - identity (·-isMonoid (isRing (isCommRing a))) - dist (isRing (isCommRing (helper a i))) = dist (isRing (isCommRing a)) - ·-comm (isCommRing (helper a i)) = ·-comm (isCommRing a) + identity (·-isMonoid a) + dist (isRing (isCommRing (helper a i))) = dist a + ·-comm (isCommRing (helper a i)) = ·-comm a commRingUnivalentStr : UnivalentStr CommRingStructure CommRingEquivStr commRingUnivalentStr = axiomsUnivalentStr _ isPropCommRingAxioms rawRingUnivalentStr diff --git a/Cubical/Structures/Group/Algebra.agda b/Cubical/Structures/Group/Algebra.agda index 2d614df0e..1a4a18cc8 100644 --- a/Cubical/Structures/Group/Algebra.agda +++ b/Cubical/Structures/Group/Algebra.agda @@ -6,6 +6,7 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.HLevels open import Cubical.Foundations.Function using (_∘_) open import Cubical.Foundations.GroupoidLaws + open import Cubical.Data.Sigma open import Cubical.Data.Unit @@ -24,7 +25,7 @@ open GroupHom private variable - ℓ ℓ' ℓ'' : Level + ℓ ℓ₁ ℓ₂ ℓ₃ : Level ------- elementary properties of morphisms -------- @@ -75,8 +76,6 @@ record GroupIso {ℓ ℓ'} (G : Group {ℓ}) (H : Group {ℓ'}) : Type (ℓ-max rightInv : section (GroupHom.fun map) inv leftInv : retract (GroupHom.fun map) inv -infixr 35 _◆_ - record BijectionIso {ℓ ℓ'} (A : Group {ℓ}) (B : Group {ℓ'}) : Type (ℓ-max ℓ ℓ') where constructor bij-iso field @@ -108,13 +107,13 @@ open BijectionIso open GroupIso open vSES -compGroupIso : ∀ {ℓ''} {G : Group {ℓ}} {H : Group {ℓ'}} {A : Group {ℓ''}} → GroupIso G H → GroupIso H A → GroupIso G A +compGroupIso : {G : Group {ℓ}} {H : Group {ℓ₁}} {A : Group {ℓ₂}} → GroupIso G H → GroupIso H A → GroupIso G A map (compGroupIso iso1 iso2) = compGroupHom (map iso1) (map iso2) inv (compGroupIso iso1 iso2) = inv iso1 ∘ inv iso2 rightInv (compGroupIso iso1 iso2) a = cong (fun (map iso2)) (rightInv iso1 _) ∙ rightInv iso2 a leftInv (compGroupIso iso1 iso2) a = cong (inv iso1) (leftInv iso2 _) ∙ leftInv iso1 a -isGroupHomInv' : {G : Group {ℓ}} {H : Group {ℓ'}} (f : GroupIso G H) → isGroupHom H G (inv f) +isGroupHomInv' : {G : Group {ℓ}} {H : Group {ℓ₁}} (f : GroupIso G H) → isGroupHom H G (inv f) isGroupHomInv' {G = G} {H = H} f h h' = isInj-f _ _ ( f' (g (h ⋆² h')) ≡⟨ (rightInv f) _ ⟩ (h ⋆² h') ≡⟨ sym (cong₂ _⋆²_ (rightInv f h) (rightInv f h')) ⟩ @@ -129,17 +128,14 @@ isGroupHomInv' {G = G} {H = H} f h h' = isInj-f _ _ ( isInj-f : (x y : ⟨ G ⟩) → f' x ≡ f' y → x ≡ y isInj-f x y p = sym (leftInv f _) ∙∙ cong g p ∙∙ leftInv f _ -invGroupIso : {G : Group {ℓ}} {H : Group {ℓ'}} → GroupIso G H → GroupIso H G +invGroupIso : {G : Group {ℓ}} {H : Group {ℓ₁}} → GroupIso G H → GroupIso H G fun (map (invGroupIso iso1)) = inv iso1 isHom (map (invGroupIso iso1)) = isGroupHomInv' iso1 inv (invGroupIso iso1) = fun (map iso1) rightInv (invGroupIso iso1) = leftInv iso1 leftInv (invGroupIso iso1) = rightInv iso1 -_◆_ : _ -_◆_ = compGroupIso - -dirProdGroupIso : ∀ {ℓ'' ℓ'''} {G : Group {ℓ}} {H : Group {ℓ'}} {A : Group {ℓ''}} {B : Group {ℓ'''}} +dirProdGroupIso : {G : Group {ℓ}} {H : Group {ℓ₁}} {A : Group {ℓ₂}} {B : Group {ℓ₃}} → GroupIso G H → GroupIso A B → GroupIso (dirProd G A) (dirProd H B) fun (map (dirProdGroupIso iso1 iso2)) prod = fun (map iso1) (fst prod) , fun (map iso2) (snd prod) isHom (map (dirProdGroupIso iso1 iso2)) a b = ΣPathP (isHom (map iso1) (fst a) (fst b) , isHom (map iso2) (snd a) (snd b)) @@ -147,12 +143,12 @@ inv (dirProdGroupIso iso1 iso2) prod = (inv iso1) (fst prod) , (inv iso2) (snd p rightInv (dirProdGroupIso iso1 iso2) a = ΣPathP (rightInv iso1 (fst a) , (rightInv iso2 (snd a))) leftInv (dirProdGroupIso iso1 iso2) a = ΣPathP (leftInv iso1 (fst a) , (leftInv iso2 (snd a))) -GrIsoToGrEquiv : {G : Group {ℓ}} {H : Group {ℓ'}} → GroupIso G H → GroupEquiv G H +GrIsoToGrEquiv : {G : Group {ℓ}} {H : Group {ℓ₂}} → GroupIso G H → GroupEquiv G H GroupEquiv.eq (GrIsoToGrEquiv i) = isoToEquiv (iso (fun (map i)) (inv i) (rightInv i) (leftInv i)) GroupEquiv.isHom (GrIsoToGrEquiv i) = isHom (map i) --- Proofs that BijectionIso and vSES both induce isomorphisms --- -BijectionIsoToGroupIso : {A : Group {ℓ}} {B : Group {ℓ'}} → BijectionIso A B → GroupIso A B +BijectionIsoToGroupIso : {A : Group {ℓ}} {B : Group {ℓ₂}} → BijectionIso A B → GroupIso A B BijectionIsoToGroupIso {A = A} {B = B} i = grIso where module A = Group A @@ -183,7 +179,7 @@ BijectionIsoToGroupIso {A = A} {B = B} i = grIso rightInv grIso b = (rec (helper b) (λ a → a) (surj i b)) .snd leftInv grIso b j = rec (helper (f b)) (λ a → a) (propTruncIsProp (surj i (f b)) ∣ b , refl ∣ j) .fst -BijectionIsoToGroupEquiv : {A : Group {ℓ}} {B : Group {ℓ'}} → BijectionIso A B → GroupEquiv A B +BijectionIsoToGroupEquiv : {A : Group {ℓ}} {B : Group {ℓ₂}} → BijectionIso A B → GroupEquiv A B BijectionIsoToGroupEquiv i = GrIsoToGrEquiv (BijectionIsoToGroupIso i) vSES→GroupIso : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group {ℓ}} {B : Group {ℓ'}} (leftGr : Group {ℓ''}) (rightGr : Group {ℓ'''}) @@ -200,13 +196,12 @@ vSES→GroupIso {A = A} lGr rGr vses = BijectionIsoToGroupIso theIso (Ker-ϕ⊂Im-left vses a inker) surj theIso a = Ker-right⊂Im-ϕ vses a (isTrivialRight vses _ _) -vSES→GroupEquiv : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Group {ℓ}} {B : Group {ℓ'}} (leftGr : Group {ℓ''}) (rightGr : Group {ℓ'''}) +vSES→GroupEquiv : {A : Group {ℓ}} {B : Group {ℓ₁}} (leftGr : Group {ℓ₂}) (rightGr : Group {ℓ₃}) → vSES A B leftGr rightGr → GroupEquiv A B vSES→GroupEquiv {A = A} lGr rGr vses = GrIsoToGrEquiv (vSES→GroupIso lGr rGr vses) -- The trivial group is a unit. -open import Cubical.Data.Unit lUnitGroupIso : ∀ {ℓ} {G : Group {ℓ}} → GroupEquiv (dirProd trivialGroup G) G lUnitGroupIso = GrIsoToGrEquiv diff --git a/Cubical/Structures/Group/Base.agda b/Cubical/Structures/Group/Base.agda index da39eb4f9..28980a554 100644 --- a/Cubical/Structures/Group/Base.agda +++ b/Cubical/Structures/Group/Base.agda @@ -4,8 +4,12 @@ module Cubical.Structures.Group.Base where open import Cubical.Foundations.Prelude open import Cubical.Data.Sigma open import Cubical.Data.Prod renaming (_×_ to _×'_) +open import Cubical.Data.Int renaming (_+_ to _+Int_ ; _-_ to _-Int_) +open import Cubical.Data.Unit + open import Cubical.Structures.Monoid hiding (⟨_⟩) open import Cubical.Structures.Semigroup hiding (⟨_⟩) +open import Cubical.Foundations.HLevels private variable @@ -41,9 +45,9 @@ record IsGroup {G : Type ℓ} η-isGroup id1 id2 id3 i = IsGroup (id1 i) (id2 i) (id3 i) record Group : Type (ℓ-suc ℓ) where - no-eta-equality - constructor group + constructor group + no-eta-equality field Carrier : Type ℓ 0g : Carrier @@ -180,11 +184,9 @@ makeGroup-left {A = A} id comp inv set assoc lUnit lCancel = a ∎ - open Group open IsGroup -open Group η-Group : {G H : Group {ℓ}} → (p : ⟨ G ⟩ ≡ ⟨ H ⟩) → (q : PathP (λ i → p i) (0g G) (0g H)) @@ -201,7 +203,6 @@ isGroup (η-Group _ _ _ _ t i) = t i isSetCarrier : ∀ {ℓ} → (G : Group {ℓ}) → isSet ⟨ G ⟩ isSetCarrier G = IsSemigroup.is-set (IsMonoid.isSemigroup (isMonoid (isGroup G))) -open import Cubical.Foundations.HLevels open IsMonoid open IsSemigroup dirProd : ∀ {ℓ ℓ'} → Group {ℓ} → Group {ℓ'} → Group @@ -220,7 +221,6 @@ inverse (isGroup (dirProd A B)) x = (λ i → (fst (inverse (isGroup A) (fst x)) i) , (fst (inverse (isGroup B) (snd x)) i)) , λ i → (snd (inverse (isGroup A) (fst x)) i) , (snd (inverse (isGroup B) (snd x)) i) -open import Cubical.Data.Unit trivialGroup : Group₀ Carrier trivialGroup = Unit 0g trivialGroup = _ @@ -231,8 +231,6 @@ assoc (isSemigroup (isMonoid (isGroup trivialGroup))) _ _ _ = refl identity (isMonoid (isGroup trivialGroup)) _ = refl , refl inverse (isGroup trivialGroup) _ = refl , refl - -open import Cubical.Data.Int renaming (_+_ to _+Int_ ; _-_ to _-Int_) intGroup : Group₀ Carrier intGroup = Int 0g intGroup = 0 @@ -242,3 +240,16 @@ is-set (isSemigroup (isMonoid (isGroup intGroup))) = isSetInt assoc (isSemigroup (isMonoid (isGroup intGroup))) = +-assoc identity (isMonoid (isGroup intGroup)) x = refl , (+-comm (pos 0) x) inverse (isGroup intGroup) x = +-comm x (pos 0 -Int x) ∙ minusPlus x 0 , (minusPlus x 0) + + +dirProd' : Group₀ → Group₀ → Group₀ +dirProd' G H = + makeGroup (0g G , 0g H) + (λ x y → (_+_ G) (fst x) (fst y) , (_+_ H) (snd x) (snd y)) + (λ x → ((- G) (fst x)) , (- H) (snd x)) + (isOfHLevelΣ 2 (isSetCarrier G) (λ _ → isSetCarrier H)) + (λ x y z → ΣPathP ((assoc G _ _ _) , assoc H _ _ _)) + (λ x → ΣPathP ((Group.rid G (fst x)) , (Group.rid H (snd x)))) + (λ x → ΣPathP ((Group.lid G (fst x)) , (Group.lid H (snd x)))) + (λ x → ΣPathP ((Group.invr G _) , (Group.invr H _))) + λ x → ΣPathP ((Group.invl G _) , (Group.invl H _)) diff --git a/Cubical/Structures/Group/MorphismProperties.agda b/Cubical/Structures/Group/MorphismProperties.agda index 69e283727..ee2b8f27a 100644 --- a/Cubical/Structures/Group/MorphismProperties.agda +++ b/Cubical/Structures/Group/MorphismProperties.agda @@ -6,6 +6,7 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport open import Cubical.Foundations.Univalence open import Cubical.Foundations.SIP open import Cubical.Foundations.Function using (_∘_) @@ -22,19 +23,13 @@ open import Cubical.Structures.Group.Base open import Cubical.Structures.Group.Properties open import Cubical.Structures.Group.Morphism -open Iso - private variable ℓ ℓ' ℓ'' : Level -infixr 31 _□_ - +open Iso open Group open GroupHom -open IsMonoid -open IsSemigroup -open IsGroup isPropIsGroupHom : (G : Group {ℓ}) (H : Group {ℓ'}) {f : ⟨ G ⟩ → ⟨ H ⟩} → isProp (isGroupHom G H f) isPropIsGroupHom G H {f} = isPropΠ2 λ a b → Group.is-set H _ _ @@ -59,10 +54,6 @@ compGroupEquiv : {F : Group {ℓ}} {G : Group {ℓ'}} {H : Group {ℓ''}} → Gr eq (compGroupEquiv f g) = compEquiv (eq f) (eq g) isHom (compGroupEquiv f g) = isHom (compGroupHom (hom f) (hom g)) -_□_ : _ -_□_ = compGroupEquiv - - idGroupEquiv : (G : Group {ℓ}) → GroupEquiv G G eq (idGroupEquiv G) = idEquiv ⟨ G ⟩ isHom (idGroupEquiv G) = λ _ _ → refl @@ -150,26 +141,26 @@ module GroupΣTheory {ℓ} where isProp (Σ[ e ∈ G ] ((x : G) → (x + e ≡ x) × (e + x ≡ x)) × ((x : G) → Σ[ x' ∈ G ] (x + x' ≡ e) × (x' + x ≡ e))) γ h (e , P , _) (e' , Q , _) = - Σ≡Prop (λ x → isPropΣ (isPropΠ λ _ → isProp× ((is-set h) _ _) ((is-set h) _ _)) (β x)) + Σ≡Prop (λ x → isPropΣ (isPropΠ λ _ → isProp× ((IsSemigroup.is-set h) _ _) ((IsSemigroup.is-set h) _ _)) (β x)) (sym (fst (Q e)) ∙ snd (P e')) where β : (e : G) → ((x : G) → (x + e ≡ x) × (e + x ≡ x)) → isProp ((x : G) → Σ[ x' ∈ G ] (x + x' ≡ e) × (x' + x ≡ e)) β e He = isPropΠ λ { x (x' , _ , P) (x'' , Q , _) → - Σ≡Prop (λ _ → isProp× ((is-set h) _ _) ((is-set h) _ _)) + Σ≡Prop (λ _ → isProp× ((IsSemigroup.is-set h) _ _) ((IsSemigroup.is-set h) _ _)) (inv-lemma ℳ x x' x'' P Q) } where ℳ : Monoid - ℳ = makeMonoid e _+_ (is-set h) (IsSemigroup.assoc h) (λ x → He x .fst) (λ x → He x .snd) + ℳ = makeMonoid e _+_ (IsSemigroup.is-set h) (IsSemigroup.assoc h) (λ x → He x .fst) (λ x → He x .snd) Group→GroupΣ : Group → GroupΣ Group→GroupΣ G = _ , _ , - ((isSemigroup (isMonoid (isGroup G))) + ((isSemigroup G) , _ - , (identity (isMonoid (isGroup G))) + , (identity G) , λ x → (- G) x - , inverse (isGroup G) x) + , inverse G x) GroupΣ→Group : GroupΣ → Group GroupΣ→Group (G , _ , SG , _ , H0g , w ) = @@ -180,15 +171,12 @@ module GroupΣTheory {ℓ} where where helper : retract (λ z → Group→GroupΣ z) GroupΣ→Group Carrier (helper a i) = ⟨ a ⟩ - 0g (helper a i) = Group.0g a - _+_ (helper a i) = Group._+_ a - - helper a i = Group.- a - isSemigroup (isMonoid (isGroup (helper a i))) = - isSemigroup (isMonoid (isGroup a)) - IsMonoid.identity (isMonoid (isGroup (helper a i))) = - identity (isMonoid (Group.isGroup a)) - inverse (isGroup (helper a i)) = inverse (Group.isGroup a) - + 0g (helper a i) = 0g a + _+_ (helper a i) = (_+_) a + - helper a i = - a + IsMonoid.isSemigroup (IsGroup.isMonoid (isGroup (helper a i))) = isSemigroup a + IsMonoid.identity (IsGroup.isMonoid (isGroup (helper a i))) = identity a + IsGroup.inverse (isGroup (helper a i)) = inverse a groupUnivalentStr : UnivalentStr GroupStructure GroupEquivStr groupUnivalentStr = axiomsUnivalentStr _ isPropGroupAxioms rawGroupUnivalentStr @@ -279,8 +267,7 @@ Group≡ G H = isoToEquiv theIso isGroup (rightInv theIso a i i₁) = isGroup (a i₁) leftInv theIso _ = refl -open import Cubical.Foundations.Transport -caracGroup≡ : {G H : Group {ℓ}} (p q : G ≡ H) → cong Group.Carrier p ≡ cong Group.Carrier q → p ≡ q +caracGroup≡ : {G H : Group {ℓ}} (p q : G ≡ H) → cong ⟨_⟩ p ≡ cong ⟨_⟩ q → p ≡ q caracGroup≡ {G = G} {H = H} p q P = sym (transportTransport⁻ (ua (Group≡ G H)) p) ∙∙ cong (transport (ua (Group≡ G H))) helper ∙∙ transportTransport⁻ (ua (Group≡ G H)) q diff --git a/Cubical/Structures/Monoid.agda b/Cubical/Structures/Monoid.agda index d4e5ecbeb..455c60d04 100644 --- a/Cubical/Structures/Monoid.agda +++ b/Cubical/Structures/Monoid.agda @@ -41,7 +41,6 @@ record IsMonoid {A : Type ℓ} (ε : A) (_·_ : A → A → A) : Type ℓ where record Monoid : Type (ℓ-suc ℓ) where - no-eta-equality constructor monoid field diff --git a/Cubical/Structures/Semigroup.agda b/Cubical/Structures/Semigroup.agda index 2f4281d63..8b4845dc2 100644 --- a/Cubical/Structures/Semigroup.agda +++ b/Cubical/Structures/Semigroup.agda @@ -31,7 +31,7 @@ private -- would only contain isSet A if we had it. record IsSemigroup {A : Type ℓ} (_·_ : A → A → A) : Type ℓ where - no-eta-equality + -- no-eta-equality constructor issemigroup field is-set : isSet A diff --git a/Cubical/ZCohomology/Groups/Prelims.agda b/Cubical/ZCohomology/Groups/Prelims.agda index dd953f17a..c2a3c7a93 100644 --- a/Cubical/ZCohomology/Groups/Prelims.agda +++ b/Cubical/ZCohomology/Groups/Prelims.agda @@ -21,6 +21,13 @@ open import Cubical.Data.Int renaming (_+_ to _+ℤ_; +-comm to +ℤ-comm ; +-as open import Cubical.Data.Nat open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec) + +infixr 33 _⋄_ + +_⋄_ : _ +_⋄_ = compIso + + S¹map : hLevelTrunc 3 S¹ → S¹ S¹map = trRec isGroupoidS¹ (idfun _) @@ -137,31 +144,11 @@ Iso.rightInv S1→K2≡K2×K1 (a , b) = (b +ₖ 0ₖ ≡⟨ rUnitₖ b ⟩ b ∎) - moveabout : {x : hLevelTrunc 4 (S₊ 2)} (p q : x ≡ 0ₖ) (mid : 0ₖ ≡ 0ₖ) - → sym q ∙ (p ∙ mid ∙ sym p) ∙ q ≡ mid - moveabout p q mid = (assoc (sym q) _ _ - ∙∙ cong (_∙ q) (assoc (sym q) p (mid ∙ sym p)) - ∙∙ sym (assoc (sym q ∙ p) (mid ∙ sym p) q)) - ∙∙ (cong ((sym q ∙ p) ∙_) (sym (assoc mid (sym p) q)) - ∙∙ cong (λ x → (sym q ∙ p) ∙ (mid ∙ x)) (sym (symDistr (sym q) p)) - ∙∙ cong ((sym q ∙ p)∙_) (isCommΩK 2 mid _)) - ∙∙ (assoc _ _ _ - ∙∙ cong (_∙ mid) (rCancel (sym q ∙ p)) - ∙∙ sym (lUnit mid)) - - congHelper : ∀ {ℓ} {A : Type ℓ} {a1 : A} (p : a1 ≡ a1) (f : A → A) (id : (λ x → x) ≡ f) - → cong f p ≡ sym (funExt⁻ id a1) ∙ p ∙ funExt⁻ id a1 - congHelper {a1 = a1} p f id = - (λ i → lUnit (rUnit (cong f p) i) i) - ∙ (λ i → (λ j → id ((~ i) ∨ (~ j)) a1) ∙ cong (id (~ i)) p ∙ λ j → id (~ i ∨ j) a1) - - - congHelper2 : (p : 0ₖ ≡ 0ₖ) (f : coHomK 2 → coHomK 2) (id : (λ x → x) ≡ f) - → (q : (f 0ₖ) ≡ 0ₖ) + congHelper2 : (p : 0ₖ ≡ 0ₖ) (f : coHomK 2 → coHomK 2) (id : (λ x → x) ≡ f) → (q : (f 0ₖ) ≡ 0ₖ) → (sym q) ∙ cong f p ∙ q ≡ p - congHelper2 p f id q = - cong (λ x → sym q ∙ x ∙ q) (congHelper p f id) ∙ - moveabout (sym (funExt⁻ id ∣ north ∣)) q p + congHelper2 p f = J (λ f _ → (q : (f 0ₖ) ≡ 0ₖ) → (sym q) ∙ cong f p ∙ q ≡ p) + λ q → (cong (sym q ∙_) (isCommΩK 2 p _) ∙∙ assoc _ _ _ ∙∙ cong (_∙ p) (lCancel q)) + ∙ sym (lUnit p) Iso.leftInv S1→K2≡K2×K1 a = funExt λ {north → rUnitₖ (a north) @@ -232,12 +219,23 @@ Iso.leftInv S1→K2≡K2×K1 a = (b +ₖ (-ₖ a) +ₖ a ≡⟨ cong (b +ₖ_) (lCancelₖ a) ⟩ b +ₖ 0ₖ ∎) - helperFun : ∀ {ℓ} {A : Type ℓ} {a b : A} (p p' : a ≡ b) (q q' : b ≡ b) (r : a ≡ a) + helperFun : ∀ {ℓ} {A : Type ℓ} {a b : A} (p' p : a ≡ b) (q q' : b ≡ b) (r : a ≡ a) -- (p p' : a ≡ b) (q q' : b ≡ b) (r : a ≡ a) → ((p q : a ≡ a) → p ∙ q ≡ q ∙ p) → q ≡ q' → PathP (λ i → p' (~ i) ≡ p' (~ i)) q' r → p ∙ q ∙ sym p ≡ r - helperFun p p' q q' r comm qid dep = + helperFun {a = a} {- p p' q q' r comm qid dep -} = + J (λ b p' → (p : a ≡ b) → (q q' : b ≡ b) → (r : a ≡ a) -- (p p' : a ≡ b) (q q' : b ≡ b) (r : a ≡ a) + → ((p q : a ≡ a) → p ∙ q ≡ q ∙ p) + → q ≡ q' + → PathP (λ i → p' (~ i) ≡ p' (~ i)) q' r + → p ∙ q ∙ sym p ≡ r) + λ p q q' r comm id id2 → (cong (λ x → p ∙ x ∙ (sym p)) (id ∙ id2) + ∙∙ cong (p ∙_) (comm _ _) + ∙∙ assoc _ _ _) + ∙∙ cong (_∙ r) (rCancel _) + ∙∙ sym (lUnit r) + {- p ∙ q ∙ sym p ≡⟨ cong (λ x → p ∙ x ∙ sym p) qid ⟩≡⟨ cong (λ x → p ∙ x ∙ sym p) (λ i → lUnit (rUnit q' i) i) ⟩ p ∙ (refl ∙ q' ∙ refl) ∙ sym p ≡⟨ cong (λ x → p ∙ x ∙ sym p) (λ i → (λ j → p' (~ i ∨ ~ j)) ∙ dep i ∙ λ j → p' (~ i ∨ j)) @@ -249,15 +247,15 @@ Iso.leftInv S1→K2≡K2×K1 a = (p ∙ sym p') ∙ r ∙ sym (p ∙ sym p') ≡⟨ cong ((p ∙ sym p') ∙_) (comm r (sym (p ∙ sym p'))) ⟩≡⟨ assoc (p ∙ sym p') (sym (p ∙ sym p')) r ⟩ ((p ∙ sym p') ∙ sym (p ∙ sym p')) ∙ r ≡⟨ cong (_∙ r) (rCancel (p ∙ sym p')) ⟩≡⟨ sym (lUnit r) ⟩ - r ∎ + r ∎ -} together : Path (_ ≡ _) (cong (a north +ₖ_) (sym (rCancelₖ (a north))) ∙ ((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north))) ∙ cong (a north +ₖ_) (rCancelₖ (a north))) (cong(_+ₖ 0ₖ) (cong a (merid south) ∙ cong a (sym (merid north)))) together = - helperFun (cong (a north +ₖ_) (sym (rCancelₖ (a north)))) - (λ i → pathHelper (a north) (a north) (~ i)) + helperFun (λ i → pathHelper (a north) (a north) (~ i)) + (cong (a north +ₖ_) (sym (rCancelₖ (a north)))) ((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north))) (((λ j → a north +ₖ a (merid south j) +ₖ (-ₖ (a north))) ∙ λ j → a north +ₖ a (merid north (~ j)) +ₖ (-ₖ (a north)))) (cong(_+ₖ 0ₖ) (cong a (merid south) ∙ cong a (sym (merid north)))) diff --git a/Cubical/ZCohomology/Groups/Sn.agda b/Cubical/ZCohomology/Groups/Sn.agda index 4b4a21996..466211e9c 100644 --- a/Cubical/ZCohomology/Groups/Sn.agda +++ b/Cubical/ZCohomology/Groups/Sn.agda @@ -33,11 +33,13 @@ open import Cubical.Data.Unit open import Cubical.Structures.Group +infixr 31 _□_ +_□_ : _ +_□_ = compGroupEquiv open GroupEquiv open vSES - Sn-connected : (n : ℕ) (x : S₊ (suc n)) → ∥ north ≡ x ∥₁ Sn-connected n = suspToPropRec north (λ _ → propTruncIsProp) ∣ refl ∣₁ @@ -255,7 +257,7 @@ H¹-S²≅0 = -- we begin with some functions translating between H¹(S₊ 1) and ∥ S¹ → S¹ ∥₀. The latter type is easy to characterise, -- by (S¹ → S¹) ≃ S¹ × ℤ (see Cubical.ZCohomology.Groups.Prelims). Truncating this leaves only ℤ, since S¹ is connected. --- The translation mention above uses the basechange function. We use basechange-lemma (Cubical.ZCohomology.Groups.Prelims) to prove the basechange2⁻ preserves +-- The translation mentioned above uses the basechange function. We use basechange-lemma (Cubical.ZCohomology.Groups.Prelims) to prove the basechange2⁻ preserves -- path composition (in a more general sense than what is proved in basechange2⁻-morph) -- We can now give the group equivalence. The first bit is just a big composition of our previously defined translations and is pretty uninteresting. diff --git a/Cubical/ZCohomology/Groups/Unit.agda b/Cubical/ZCohomology/Groups/Unit.agda index 2ab01ec7a..14913286f 100644 --- a/Cubical/ZCohomology/Groups/Unit.agda +++ b/Cubical/ZCohomology/Groups/Unit.agda @@ -34,7 +34,7 @@ H⁰-Unit≅ℤ = GrIsoToGrEquiv H⁰-Unit≅ℤ' {- Hⁿ(Unit) for n ≥ 1 -} isContrHⁿ-Unit : (n : ℕ) → isContr (coHom (suc n) Unit) -isContrHⁿ-Unit n = subst isContr (λ i → ∥ UnitToTypeId (coHomK (suc n)) (~ i) ∥₂) (helper' n) +isContrHⁿ-Unit n = subst isContr (λ i → ∥ UnitToTypePath (coHomK (suc n)) (~ i) ∥₂) (helper' n) where helper' : (n : ℕ) → isContr (∥ coHomK (suc n) ∥₂) helper' n = @@ -63,8 +63,8 @@ Hⁿ-Unit≅0 n = GrIsoToGrEquiv (Hⁿ-Unit≅0' n) private Hⁿ-contrTypeIso : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → isContr A → Iso (coHom (suc n) A) (coHom (suc n) Unit) - Hⁿ-contrTypeIso n contr = compIso (setTruncIso (ContrToTypeIso contr)) - (setTruncIso (invIso (ContrToTypeIso isContrUnit))) + Hⁿ-contrTypeIso n contr = compIso (setTruncIso (isContr→Iso2 contr)) + (setTruncIso (invIso (isContr→Iso2 isContrUnit))) Hⁿ-contrType≅0' : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → isContr A → GroupIso (coHomGr (suc n) A) trivialGroup diff --git a/Cubical/ZCohomology/MayerVietorisUnreduced.agda b/Cubical/ZCohomology/MayerVietorisUnreduced.agda index db27289d1..6ad27651f 100644 --- a/Cubical/ZCohomology/MayerVietorisUnreduced.agda +++ b/Cubical/ZCohomology/MayerVietorisUnreduced.agda @@ -15,7 +15,7 @@ open import Cubical.HITs.Pushout open import Cubical.HITs.Sn open import Cubical.HITs.Susp open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2) -open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₋₁ ; ∣_∣ to ∣_∣₋₁) +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim ; elim2 to pElim2 ; ∥_∥ to ∥_∥₁ ; ∣_∣ to ∣_∣₁) open import Cubical.HITs.Nullification open import Cubical.Data.Nat open import Cubical.Data.Prod hiding (_×_) @@ -136,7 +136,7 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : (λ p1 → pRec propTruncIsProp λ p2 → ∣ ∣ (λ c → ΩKn+1→Kn (sym (cong (λ F → F (f c)) p1) ∙∙ cong a (push c) ∙∙ cong (λ F → F (g c)) p2)) ∣₂ - , cong ∣_∣₂ (funExt (λ δ → helper n a p1 p2 δ)) ∣₋₁) + , cong ∣_∣₂ (funExt (λ δ → helper n a p1 p2 δ)) ∣₁) (Iso.fun (PathIdTrunc₀Iso) (cong fst p)) (Iso.fun (PathIdTrunc₀Iso) (cong snd p)) where @@ -214,7 +214,7 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : (λ Fa → sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) λ Fb p → pRec propTruncIsProp (λ q → ∣ ∣ helpFun n Fa Fb (funExt⁻ q) ∣₂ - , anotherHelper n Fa Fb q ∣₋₁) + , anotherHelper n Fa Fb q ∣₁) (helper n Fa Fb p)) Fa Fb @@ -222,7 +222,7 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : where helper : (n : ℕ) (Fa : A → coHomK n) (Fb : B → coHomK n) → (fun (Δ n)) (∣ Fa ∣₂ , ∣ Fb ∣₂) ≡ 0ₕ - → ∥ (Path (_ → _) (λ c → Fa (f c)) (λ c → Fb (g c))) ∥₋₁ + → ∥ (Path (_ → _) (λ c → Fa (f c)) (λ c → Fb (g c))) ∥₁ helper zero Fa Fb p = Iso.fun (PathIdTrunc₀Iso) ((sym (rUnitₕ (coHomFun zero f ∣ Fa ∣₂)) ∙∙ (λ i → coHomFun zero f ∣ Fa ∣₂ +ₕ (lCancelₕ (coHomFun zero g ∣ Fb ∣₂) (~ i))) @@ -257,7 +257,7 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) λ Fc p → pRec propTruncIsProp (λ p → ∣ (∣ (λ a → ΩKn+1→Kn (cong (λ f → f (inl a)) p)) ∣₂ , ∣ (λ b → ΩKn+1→Kn (cong (λ f → f (inr b)) p)) ∣₂) , - Iso.inv (PathIdTrunc₀Iso) ∣ funExt (λ c → helper2 n Fc p c) ∣₋₁ ∣₋₁) + Iso.inv (PathIdTrunc₀Iso) ∣ funExt (λ c → helper2 n Fc p c) ∣₁ ∣₁) (Iso.fun (PathIdTrunc₀Iso) p) where diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index 57bfc284e..e11371613 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -19,7 +19,7 @@ open import Cubical.HITs.Susp open import Cubical.HITs.Wedge open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2 ; setTruncIsSet to §) open import Cubical.HITs.Nullification -open import Cubical.Data.Int hiding (_+_) +open import Cubical.Data.Int renaming (_+_ to _ℤ+_) open import Cubical.Data.Nat open import Cubical.HITs.Truncation renaming (elim to trElim ; map to trMap ; rec to trRec ; elim3 to trElim3) open import Cubical.Homotopy.Loopspace @@ -27,6 +27,8 @@ open import Cubical.Homotopy.Connected open import Cubical.Homotopy.Freudenthal open import Cubical.HITs.SmashProduct.Base open import Cubical.Structures.Group +open import Cubical.Structures.Semigroup +open import Cubical.Structures.Monoid open import Cubical.Foundations.Equiv.HalfAdjoint @@ -109,9 +111,8 @@ Kn→ΩKn+10ₖ (suc n) = (λ i → cong ∣_∣ (rCancel (merid north) i)) -- c ∙∙ (λ i → ΩKn+1→Kn (Kn→ΩKn+10ₖ n (~ i))) ∙∙ Iso.leftInv (Iso3-Kn-ΩKn+1 n) 0ₖ -abstract - +ₖ→∙ : (n : ℕ) (a b : coHomK n) → Kn→ΩKn+1 n (a +ₖ b) ≡ Kn→ΩKn+1 n a ∙ Kn→ΩKn+1 n b - +ₖ→∙ n a b = Iso.rightInv (Iso3-Kn-ΩKn+1 n) (Kn→ΩKn+1 n a ∙ Kn→ΩKn+1 n b) ++ₖ→∙ : (n : ℕ) (a b : coHomK n) → Kn→ΩKn+1 n (a +ₖ b) ≡ Kn→ΩKn+1 n a ∙ Kn→ΩKn+1 n b ++ₖ→∙ n a b = Iso.rightInv (Iso3-Kn-ΩKn+1 n) (Kn→ΩKn+1 n a ∙ Kn→ΩKn+1 n b) lUnitₖ : {n : ℕ} (x : coHomK n) → 0ₖ +ₖ x ≡ x lUnitₖ {n = zero} x = cong ΩKn+1→Kn (sym (lUnit (Kn→ΩKn+1 zero x))) ∙ @@ -270,8 +271,6 @@ rUnitlUnit0 {n = suc n} = 0ₕ∙ zero = ∣ (λ _ → 0ₖ) , refl ∣₂ 0ₕ∙ (suc n) = ∣ (λ _ → 0ₖ) , refl ∣₂ -open import Cubical.Structures.Semigroup -open import Cubical.Structures.Monoid open IsSemigroup open IsMonoid open Group @@ -335,7 +334,6 @@ isCommΩK-based (suc n) x p q = (sym (substComposite (λ x → x) (isoToPath ((congIso (invIso (congIsoHelper x))))) (λ i → rCancelₖ x i ≡ rCancelₖ x i) q)) -open import Cubical.Data.Int renaming (_+_ to _ℤ+_) addLemma : (a b : Int) → a +ₖ b ≡ (a ℤ+ b) addLemma a b = (cong ΩKn+1→Kn (sym (congFunct ∣_∣ (looper a) (looper b))) ∙∙ cong₂ (λ x y → ΩKn+1→Kn (cong ∣_∣ (x ∙ y))) (looper≡looper2 a) (looper≡looper2 b) From 2d13fafdaa15fcbbf1d041103930e46c2df8cdcf Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Thu, 16 Jul 2020 18:31:20 +0200 Subject: [PATCH 55/89] anders+whitespace --- Cubical/Data/Empty/Base.agda | 3 --- Cubical/Homotopy/Loopspace.agda | 3 ++- Cubical/Structures/AbGroup.agda | 2 ++ Cubical/Structures/Group/Base.agda | 13 ------------- 4 files changed, 4 insertions(+), 17 deletions(-) diff --git a/Cubical/Data/Empty/Base.agda b/Cubical/Data/Empty/Base.agda index e55bc3860..43cef0613 100644 --- a/Cubical/Data/Empty/Base.agda +++ b/Cubical/Data/Empty/Base.agda @@ -10,6 +10,3 @@ rec () elim : ∀ {ℓ} {A : ⊥ → Type ℓ} → (x : ⊥) → A x elim () - -nope : ∀ {ℓ} (A : Type ℓ) → Type ℓ -nope A = A → ⊥ diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda index c6e99c623..dc3642c61 100644 --- a/Cubical/Homotopy/Loopspace.agda +++ b/Cubical/Homotopy/Loopspace.agda @@ -27,10 +27,11 @@ open import Cubical.HITs.SetTruncation generalEH : {ℓ : Level} {A : Type ℓ} {a b c : A} {p q : a ≡ b} {r s : b ≡ c} (α : p ≡ q) (β : r ≡ s) → (cong (_∙ r) α) ∙ (cong (q ∙_) β) ≡ (cong (p ∙_) β) ∙ (cong (_∙ s) α) -generalEH {p = p} {r = r} α β j i = +generalEH {p = p} {r = r} α β j i = hcomp (λ k → λ { (i = i0) → p ∙ r ; (i = i1) → α (k ∨ ~ j) ∙ β (k ∨ j) }) (α (~ j ∧ i) ∙ β (j ∧ i)) + Eckmann-Hilton : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (α β : typ ((Ω^ (2 + n)) A)) → α ∙ β ≡ β ∙ α Eckmann-Hilton {A = A} n α β i = diff --git a/Cubical/Structures/AbGroup.agda b/Cubical/Structures/AbGroup.agda index 57989a28e..a68d2e3d2 100644 --- a/Cubical/Structures/AbGroup.agda +++ b/Cubical/Structures/AbGroup.agda @@ -135,6 +135,8 @@ module AbGroupΣTheory {ℓ} where IsGroup.inverse (IsAbGroup.isGroup (isAbGroup (helper2 a i))) = inverse a IsAbGroup.comm (isAbGroup (helper2 a i)) = comm a + -- TODO: investigate why we need all of the qualified projections + abGroupUnivalentStr : UnivalentStr AbGroupStructure AbGroupEquivStr abGroupUnivalentStr = axiomsUnivalentStr _ isPropAbGroupAxioms rawGroupUnivalentStr diff --git a/Cubical/Structures/Group/Base.agda b/Cubical/Structures/Group/Base.agda index 28980a554..721b314c6 100644 --- a/Cubical/Structures/Group/Base.agda +++ b/Cubical/Structures/Group/Base.agda @@ -240,16 +240,3 @@ is-set (isSemigroup (isMonoid (isGroup intGroup))) = isSetInt assoc (isSemigroup (isMonoid (isGroup intGroup))) = +-assoc identity (isMonoid (isGroup intGroup)) x = refl , (+-comm (pos 0) x) inverse (isGroup intGroup) x = +-comm x (pos 0 -Int x) ∙ minusPlus x 0 , (minusPlus x 0) - - -dirProd' : Group₀ → Group₀ → Group₀ -dirProd' G H = - makeGroup (0g G , 0g H) - (λ x y → (_+_ G) (fst x) (fst y) , (_+_ H) (snd x) (snd y)) - (λ x → ((- G) (fst x)) , (- H) (snd x)) - (isOfHLevelΣ 2 (isSetCarrier G) (λ _ → isSetCarrier H)) - (λ x y z → ΣPathP ((assoc G _ _ _) , assoc H _ _ _)) - (λ x → ΣPathP ((Group.rid G (fst x)) , (Group.rid H (snd x)))) - (λ x → ΣPathP ((Group.lid G (fst x)) , (Group.lid H (snd x)))) - (λ x → ΣPathP ((Group.invr G _) , (Group.invr H _))) - λ x → ΣPathP ((Group.invl G _) , (Group.invl H _)) From fbf33d10a86dae5ca8539263786304726417959d Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Fri, 24 Jul 2020 13:23:20 +0200 Subject: [PATCH 56/89] fixes --- Cubical/Structures/Group/Base.agda | 4 ++-- Cubical/Structures/Monoid.agda | 1 - Cubical/ZCohomology/Groups/Torus.agda | 10 ++++++++-- Cubical/ZCohomology/Properties.agda | 1 + 4 files changed, 11 insertions(+), 5 deletions(-) diff --git a/Cubical/Structures/Group/Base.agda b/Cubical/Structures/Group/Base.agda index 721b314c6..10ff817b6 100644 --- a/Cubical/Structures/Group/Base.agda +++ b/Cubical/Structures/Group/Base.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --no-import-sorts --safe #-} +{-# OPTIONS --cubical --no-import-sorts --no-syntactic-equality --safe #-} module Cubical.Structures.Group.Base where open import Cubical.Foundations.Prelude @@ -201,7 +201,7 @@ _+_ (η-Group _ _ r _ _ i) = r i isGroup (η-Group _ _ _ _ t i) = t i isSetCarrier : ∀ {ℓ} → (G : Group {ℓ}) → isSet ⟨ G ⟩ -isSetCarrier G = IsSemigroup.is-set (IsMonoid.isSemigroup (isMonoid (isGroup G))) +isSetCarrier G = IsSemigroup.is-set (IsMonoid.isSemigroup (isMonoid G)) open IsMonoid open IsSemigroup diff --git a/Cubical/Structures/Monoid.agda b/Cubical/Structures/Monoid.agda index 455c60d04..ad22f0891 100644 --- a/Cubical/Structures/Monoid.agda +++ b/Cubical/Structures/Monoid.agda @@ -24,7 +24,6 @@ private ℓ : Level record IsMonoid {A : Type ℓ} (ε : A) (_·_ : A → A → A) : Type ℓ where - no-eta-equality constructor ismonoid field diff --git a/Cubical/ZCohomology/Groups/Torus.agda b/Cubical/ZCohomology/Groups/Torus.agda index 5d651a2f0..316249a3c 100644 --- a/Cubical/ZCohomology/Groups/Torus.agda +++ b/Cubical/ZCohomology/Groups/Torus.agda @@ -33,8 +33,6 @@ open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim2 t open import Cubical.HITs.Nullification open import Cubical.HITs.Truncation renaming (elim to trElim ; elim2 to trElim2 ; map to trMap ; rec to trRec) - - --------- H⁰(T²) ------------ H⁰-T²≅0 : GroupEquiv (coHomGr 0 (S₊ 1 × S₊ 1)) intGroup H⁰-T²≅0 = @@ -173,3 +171,11 @@ H²-T²≅ℤ = invGroupEquiv ℤ≅H²-T² ∙∙ λ i → assoc (lUnitₖ 0ₖ ∙ sym (rUnitₖ 0ₖ)) (symDistr ((lUnitₖ 0ₖ)) (sym (rUnitₖ 0ₖ)) (~ i)) (cong (_+ₖ 0ₖ) a) i) ∙∙ cong (_∙ (cong (_+ₖ 0ₖ) a)) (rCancel (lUnitₖ 0ₖ ∙ sym (rUnitₖ 0ₖ))) ∙∙ sym (lUnit (cong (_+ₖ 0ₖ) a)) +private + toℤ₂ : coHom 2 (S₊ 1 × S₊ 1) → Int + toℤ₂ = fst (GroupEquiv.eq H²-T²≅ℤ) + + fromℤ₂ : Int → coHom 2 (S₊ 1 × S₊ 1) + fromℤ₂ = invEq (GroupEquiv.eq H²-T²≅ℤ) + + diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index e11371613..40dc2ac7f 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -215,6 +215,7 @@ commₕ {n = n} = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) -- Proof that rUnitₖ and lUnitₖ agree on 0ₖ. Needed for Mayer-Vietoris. + rUnitlUnit0 : {n : ℕ} → rUnitₖ {n = n} 0ₖ ≡ lUnitₖ 0ₖ rUnitlUnit0 {n = zero} = refl rUnitlUnit0 {n = suc n} = From f96e4214bd2a62ef26a148aef92b4806d76f8d94 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 27 Jul 2020 02:27:03 +0200 Subject: [PATCH 57/89] h1-sn --- Cubical/Homotopy/Loopspace.agda | 13 +--- Cubical/Structures/Group/Algebra.agda | 63 +++++++++++++++++++ Cubical/ZCohomology/Groups/Sn.agda | 62 ++++++++++++++---- Cubical/ZCohomology/Groups/Torus.agda | 22 +++++-- .../ZCohomology/Groups/WedgeOfSpheres.agda | 21 ++++++- 5 files changed, 153 insertions(+), 28 deletions(-) diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda index dc3642c61..3e5b2a13a 100644 --- a/Cubical/Homotopy/Loopspace.agda +++ b/Cubical/Homotopy/Loopspace.agda @@ -25,21 +25,14 @@ open import Cubical.HITs.SetTruncation Ω→ : ∀ {ℓA ℓB} {A : Pointed ℓA} {B : Pointed ℓB} (f : A →∙ B) → (Ω A →∙ Ω B) Ω→ (f , f∙) = (λ p → (sym f∙ ∙ cong f p) ∙ f∙) , cong (λ q → q ∙ f∙) (sym (rUnit (sym f∙))) ∙ lCancel f∙ -generalEH : {ℓ : Level} {A : Type ℓ} {a b c : A} {p q : a ≡ b} {r s : b ≡ c} (α : p ≡ q) (β : r ≡ s) - → (cong (_∙ r) α) ∙ (cong (q ∙_) β) ≡ (cong (p ∙_) β) ∙ (cong (_∙ s) α) -generalEH {p = p} {r = r} α β j i = - hcomp (λ k → λ { (i = i0) → p ∙ r - ; (i = i1) → α (k ∨ ~ j) ∙ β (k ∨ j) }) - (α (~ j ∧ i) ∙ β (j ∧ i)) - Eckmann-Hilton : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (α β : typ ((Ω^ (2 + n)) A)) → α ∙ β ≡ β ∙ α Eckmann-Hilton {A = A} n α β i = - comp (λ k → rUnit (snd ((Ω^ (1 + n)) A)) (~ k) ≡ rUnit (snd ((Ω^ (1 + n)) A)) (~ k)) - -- note : rUnit refl := lUnit refl + comp (λ k → rUnit (snd ((Ω^ (1 + n)) A)) (~ k) ≡ rUnit (snd ((Ω^ (1 + n)) A)) (~ k)) -- note : rUnit refl := lUnit refl (λ k → λ { (i = i0) → (cong (λ x → rUnit x (~ k)) α) ∙ cong (λ x → lUnit x (~ k)) β ; (i = i1) → (cong (λ x → lUnit x (~ k)) β) ∙ cong (λ x → rUnit x (~ k)) α}) - (generalEH α β i) + ((λ j → α (j ∧ ~ i) ∙ β (j ∧ i)) ∙ λ j → α (~ i ∨ j) ∙ β (i ∨ j)) + {- Homotopy group version -} π-comp : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → ∥ typ ((Ω^ (suc n)) A) ∥₂ diff --git a/Cubical/Structures/Group/Algebra.agda b/Cubical/Structures/Group/Algebra.agda index 1a4a18cc8..f890ab94b 100644 --- a/Cubical/Structures/Group/Algebra.agda +++ b/Cubical/Structures/Group/Algebra.agda @@ -222,3 +222,66 @@ rUnitGroupIso = contrGroup≅trivialGroup : {G : Group {ℓ}} → isContr ⟨ G ⟩ → GroupEquiv G trivialGroup GroupEquiv.eq (contrGroup≅trivialGroup contr) = isContr→≃Unit contr GroupEquiv.isHom (contrGroup≅trivialGroup contr) _ _ = refl + +open GroupEquiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Equiv +open import Cubical.Structures.Monoid hiding (⟨_⟩) +open import Cubical.Structures.Semigroup hiding (⟨_⟩) +open Monoid +open IsMonoid +open IsGroup +open Semigroup +open IsSemigroup + +GroupUa' : (G H : Group {ℓ}) → (p : ⟨ G ⟩ ≡ ⟨ H ⟩) + → PathP (λ i → p i) (0g G) (0g H) + → PathP (λ i → p i → p i → p i) (_+_ G) (_+_ H) + → PathP (λ i → p i → p i) (- G) (- H) + → G ≡ H +Carrier (GroupUa' G H p p0 p+ p- i) = p i +0g (GroupUa' G H p p0 p+ p- i) = p0 i +_+_ (GroupUa' G H p p0 p+ p- i) = p+ i +- GroupUa' G H p p0 p+ p- i = p- i +is-set (isSemigroup (isMonoid (isGroup (GroupUa' G H p p0 p+ p- i)))) = + isOfHLevel→isOfHLevelDep 1 {B = λ x → isSet x} (λ _ → isPropIsSet) (is-set G) (is-set H) p i +IsSemigroup.assoc (isSemigroup (isMonoid (isGroup (GroupUa' G H p p0 p+ p- i)))) = + hcomp (λ k → λ {(i = i0) → IsSemigroup.assoc (isSemigroup G) + ; (i = i1) → isPropΠ3 (λ _ _ _ → is-set H _ _) + (transport (λ j → ((x y z : p j) → p+ j x (p+ j y z) ≡ p+ j (p+ j x y) z)) + (IsSemigroup.assoc (isSemigroup G))) (IsSemigroup.assoc (isSemigroup H)) k }) + (transp (λ j → ((x y z : p (i ∧ j)) → p+ (i ∧ j) x (p+ (i ∧ j) y z) ≡ p+ (i ∧ j) (p+ (i ∧ j) x y) z)) (~ i) + (IsSemigroup.assoc (isSemigroup G))) +identity (isMonoid (isGroup (GroupUa' G H p p0 p+ p- i))) = + hcomp (λ k → λ {(i = i0) → identity G + ; (i = i1) → isPropΠ (λ _ → isOfHLevelΣ 1 (is-set H _ _) λ _ → is-set H _ _) + (transport (λ j → (x : p j) → (p+ j x (p0 j) ≡ x) × (p+ j (p0 j) x ≡ x)) + (identity G)) + (identity H) k } ) + (transp (λ j → (x : p (i ∧ j)) → (p+ (i ∧ j) x (p0 (i ∧ j)) ≡ x) × (p+ (i ∧ j) (p0 (i ∧ j)) x ≡ x)) (~ i) + (identity G)) +inverse (isGroup (GroupUa' G H p p0 p+ p- i)) = + hcomp (λ k → λ {(i = i0) → inverse G + ; (i = i1) → isPropΠ (λ _ → isOfHLevelΣ 1 (is-set H _ _) λ _ → is-set H _ _) + (transport (λ i → (x : p i) → (p+ i x (p- i x) ≡ p0 i) × (p+ i (p- i x) x ≡ p0 i)) (inverse G)) + (inverse H) k } ) + (transp (λ j → (x : p (i ∧ j)) → (p+ (i ∧ j) x (p- (i ∧ j) x) ≡ p0 (i ∧ j)) × (p+ (i ∧ j) (p- (i ∧ j) x) x ≡ p0 (i ∧ j))) (~ i) (inverse G)) + +GroupUa : (G H : Group {ℓ}) → GroupEquiv G H → G ≡ H +GroupUa G H e = + GroupUa' G H + (ua (eq e)) + (λ i → hcomp (λ k → λ {(i = i0) → 0g G + ; (i = i1) → (transportRefl (fst (eq e) (0g G)) + ∙ morph0→0 G H (grouphom (fst (eq e)) (isHom e))) k}) + (transp (λ j → ua (eq e) (i ∧ j)) (~ i) (0g G))) + (λ i a b → hcomp (λ k → λ {(i = i0) → _+_ G a b + ; (i = i1) → ((λ z → transportRefl (fst (eq e) ((_+_ G (invEq (eq e) (transportRefl a z)) (invEq (eq e) (transportRefl b z))))) z) + ∙∙ isHom e (invEq (eq e) a) (invEq (eq e) b) + ∙∙ cong₂ (_+_ H) (retEq (eq e) a) (retEq (eq e) b)) k}) + (transp (λ j → ua (eq e) (i ∧ j)) (~ i) (_+_ G (transp (λ j → ua (eq e) (i ∧ (~ j))) (~ i) a) (transp (λ j → ua (eq e) (i ∧ (~ j))) (~ i) b)))) + λ i a → hcomp (λ k → λ {(i = i0) → (- G) a + ; (i = i1) → ((λ z → transportRefl (fst (eq e) ((- G) (invEq (eq e) (transportRefl a z)))) z) + ∙∙ morphMinus G H (grouphom (fst (eq e)) (isHom e)) (invEq (eq e) a) + ∙∙ cong (- H) (retEq (eq e) a)) k}) + (transp (λ j → ua (eq e) (i ∧ j)) (~ i) ((- G) (transp (λ j → ua (eq e) (i ∧ (~ j))) (~ i) a))) diff --git a/Cubical/ZCohomology/Groups/Sn.agda b/Cubical/ZCohomology/Groups/Sn.agda index 466211e9c..fe03356cd 100644 --- a/Cubical/ZCohomology/Groups/Sn.agda +++ b/Cubical/ZCohomology/Groups/Sn.agda @@ -224,18 +224,17 @@ H²-S¹≅0 = --------------- H¹(S²) -------------------------------------------- -H¹-S²≅0 : GroupEquiv (coHomGr 1 (S₊ 2)) trivialGroup -H¹-S²≅0 = - coHomPushout≅coHomSn 1 1 - □ BijectionIsoToGroupEquiv - (bij-iso (I.i 1) - helper - λ x → ∣ 0ₕ , isOfHLevelSuc 0 (isOfHLevelΣ 0 (isContrHⁿ-Unit 0) (λ _ → isContrHⁿ-Unit 0)) _ x ∣₁) - □ dirProdEquiv (Hⁿ-Unit≅0 0) (Hⁿ-Unit≅0 0) - □ lUnitGroupIso +H¹-Sⁿ≅0 : (n : ℕ) → GroupEquiv (coHomGr 1 (S₊ (2 + n))) trivialGroup +H¹-Sⁿ≅0 n = coHomPushout≅coHomSn (1 + n) 1 + □ BijectionIsoToGroupEquiv + (bij-iso (I.i 1) + helper + λ x → ∣ 0ₕ , isOfHLevelSuc 0 (isOfHLevelΣ 0 (isContrHⁿ-Unit zero) (λ _ → isContrHⁿ-Unit zero)) _ x ∣₁) + □ dirProdEquiv (Hⁿ-Unit≅0 zero) (Hⁿ-Unit≅0 zero) + □ lUnitGroupIso where - module I = MV Unit Unit (S₊ 1) (λ _ → tt) (λ _ → tt) - surj-helper : (x : ⟨ coHomGr 0 (S₊ 1) ⟩) → isInIm _ _ (I.Δ 0) x + module I = MV Unit Unit (S₊ (1 + n)) (λ _ → tt) (λ _ → tt) + surj-helper : (x : ⟨ coHomGr 0 (S₊ _) ⟩) → isInIm _ _ (I.Δ 0) x surj-helper = sElim (λ _ → isOfHLevelSuc 1 propTruncIsProp) λ f → ∣ (∣ (λ _ → f north) ∣₂ , 0ₕ) @@ -250,6 +249,47 @@ H¹-S²≅0 = λ a id → sym id ∙ I.Im-Δ⊂Ker-d 0 ∣ a ∣₂ (surj-helper _)) (I.Ker-i⊂Im-d 0 ∣ x ∣₂ inker) +{- +open import Cubical.Data.Sum +open import Cubical.Data.Empty renaming (rec to ⊥-rec) +Hᵐ-Sⁿ≅0 : (n m : ℕ) → (Σ[ x ∈ ℕ ] n + (suc x) ≡ m) ⊎ (Σ[ x ∈ ℕ ] m + (suc x) ≡ n) + → GroupEquiv (coHomGr (suc m) (S₊ (suc n))) trivialGroup +Hᵐ-Sⁿ≅0 n zero (inl (dif , id)) = ⊥-rec (snotz (sym (+-suc n dif) ∙ id)) +Hᵐ-Sⁿ≅0 zero (suc m) (inl (dif , id)) = {!!} + where + module I = MV Unit Unit (S₊ 1) (λ _ → tt) (λ _ → tt) +Hᵐ-Sⁿ≅0 (suc n) (suc m) (inl (dif , id)) = + coHomPushout≅coHomSn (suc n) (2 + m) + □ invGroupEquiv (vSES→GroupEquiv _ _ + (ses (isOfHLevelSuc 0 (isOfHLevelΣ 0 (isContrHⁿ-Unit m) λ _ → isContrHⁿ-Unit m)) + (isOfHLevelSuc 0 (isOfHLevelΣ 0 (isContrHⁿ-Unit (suc m)) λ _ → isContrHⁿ-Unit (suc m))) + (I.Δ (suc m)) + (I.i (suc (suc m))) + (I.d (suc m)) + (I.Ker-d⊂Im-Δ (suc m)) + (I.Ker-i⊂Im-d (suc m)))) + □ Hᵐ-Sⁿ≅0 n m (inl (dif , (cong predℕ id))) + + where + module I = MV Unit Unit (S₊ (suc n)) (λ _ → tt) (λ _ → tt) +Hᵐ-Sⁿ≅0 n zero (inr (dif , id)) = transport (λ i → GroupEquiv (coHomGr 1 (S₊ (suc (id i)))) trivialGroup) + (H¹-Sⁿ≅0 dif) +Hᵐ-Sⁿ≅0 n (suc m) (inr (dif , id)) = + coHomPushout≅coHomSn n (2 + m) + □ invGroupEquiv + (vSES→GroupEquiv _ _ + (ses (isOfHLevelSuc 0 (isOfHLevelΣ 0 (isContrHⁿ-Unit m) λ _ → isContrHⁿ-Unit m)) + (isOfHLevelSuc 0 (isOfHLevelΣ 0 (isContrHⁿ-Unit (suc m)) λ _ → isContrHⁿ-Unit (suc m))) + (I.Δ (suc m)) + (I.i (suc (suc m))) + (I.d (suc m)) + (I.Ker-d⊂Im-Δ (suc m)) + (I.Ker-i⊂Im-d (suc m)))) + □ transport (λ i → GroupEquiv (coHomGr (suc m) (S₊ ((id) i))) trivialGroup) + (Hᵐ-Sⁿ≅0 (m + suc dif) m (inr (dif , refl))) + where + module I = MV Unit Unit (S₊ n) (λ _ → tt) (λ _ → tt) +-} --------- Direct proof of H¹(S¹) ≅ ℤ without Mayer-Vietoris ------- diff --git a/Cubical/ZCohomology/Groups/Torus.agda b/Cubical/ZCohomology/Groups/Torus.agda index 316249a3c..929dc4f81 100644 --- a/Cubical/ZCohomology/Groups/Torus.agda +++ b/Cubical/ZCohomology/Groups/Torus.agda @@ -34,8 +34,8 @@ open import Cubical.HITs.Nullification open import Cubical.HITs.Truncation renaming (elim to trElim ; elim2 to trElim2 ; map to trMap ; rec to trRec) --------- H⁰(T²) ------------ -H⁰-T²≅0 : GroupEquiv (coHomGr 0 (S₊ 1 × S₊ 1)) intGroup -H⁰-T²≅0 = +H⁰-T²≅ℤ : GroupEquiv (coHomGr 0 (S₊ 1 × S₊ 1)) intGroup +H⁰-T²≅ℤ = H⁰-connected (north , north) λ (a , b) → pRec propTruncIsProp (λ id1 → pRec propTruncIsProp @@ -172,10 +172,20 @@ H²-T²≅ℤ = invGroupEquiv ℤ≅H²-T² ∙∙ cong (_∙ (cong (_+ₖ 0ₖ) a)) (rCancel (lUnitₖ 0ₖ ∙ sym (rUnitₖ 0ₖ))) ∙∙ sym (lUnit (cong (_+ₖ 0ₖ) a)) private - toℤ₂ : coHom 2 (S₊ 1 × S₊ 1) → Int - toℤ₂ = fst (GroupEquiv.eq H²-T²≅ℤ) + to₂ : coHom 2 (S₊ 1 × S₊ 1) → Int + to₂ = fst (GroupEquiv.eq H²-T²≅ℤ) - fromℤ₂ : Int → coHom 2 (S₊ 1 × S₊ 1) - fromℤ₂ = invEq (GroupEquiv.eq H²-T²≅ℤ) + from₂ : Int → coHom 2 (S₊ 1 × S₊ 1) + from₂ = invEq (GroupEquiv.eq H²-T²≅ℤ) + to₁ : coHom 1 (S₊ 1 × S₊ 1) → Int × Int + to₁ = fst (GroupEquiv.eq H¹-T²≅ℤ×ℤ) + from₁ : Int × Int → coHom 1 (S₊ 1 × S₊ 1) + from₁ = invEq (GroupEquiv.eq H¹-T²≅ℤ×ℤ) + + to₀ : coHom 0 (S₊ 1 × S₊ 1) → Int + to₀ = fst (GroupEquiv.eq H⁰-T²≅ℤ) + + from₀ : Int → coHom 0 (S₊ 1 × S₊ 1) + from₀ = invEq (GroupEquiv.eq H⁰-T²≅ℤ) diff --git a/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda b/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda index c9416010d..3a7203158 100644 --- a/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda +++ b/Cubical/ZCohomology/Groups/WedgeOfSpheres.agda @@ -43,7 +43,7 @@ H⁰-S²⋁S¹⋁S¹ = H⁰-connected (inl north) H¹-S²⋁S¹⋁S¹ : GroupEquiv (coHomGr 1 S²⋁S¹⋁S¹) (dirProd intGroup intGroup) H¹-S²⋁S¹⋁S¹ = Hⁿ-⋁ (S₊∙ 2) (S¹⋁S¹ , inl north) 0 - □ dirProdEquiv H¹-S²≅0 H¹-S¹⋁S¹ + □ dirProdEquiv (H¹-Sⁿ≅0 0) H¹-S¹⋁S¹ □ lUnitGroupIso ------------- H²(S²⋁S¹⋁S¹) --------- @@ -56,3 +56,22 @@ H²-S²⋁S¹⋁S¹ = □ dirProdEquiv H²-S¹≅0 H²-S¹≅0 □ rUnitGroupIso) □ rUnitGroupIso) + +private + open import Cubical.Data.Int + open import Cubical.Foundations.Equiv + open import Cubical.Data.Sigma + to₂ : coHom 2 S²⋁S¹⋁S¹ → Int + to₂ = fst (GroupEquiv.eq H²-S²⋁S¹⋁S¹) + from₂ : Int → coHom 2 S²⋁S¹⋁S¹ + from₂ = invEq (GroupEquiv.eq H²-S²⋁S¹⋁S¹) + + to₁ : coHom 1 S²⋁S¹⋁S¹ → Int × Int + to₁ = fst (GroupEquiv.eq H¹-S²⋁S¹⋁S¹) + from₁ : Int × Int → coHom 1 S²⋁S¹⋁S¹ + from₁ = invEq (GroupEquiv.eq H¹-S²⋁S¹⋁S¹) + + to₀ : coHom 0 S²⋁S¹⋁S¹ → Int + to₀ = fst (GroupEquiv.eq H⁰-S²⋁S¹⋁S¹) + from₀ : Int → coHom 0 S²⋁S¹⋁S¹ + from₀ = invEq (GroupEquiv.eq H⁰-S²⋁S¹⋁S¹) From d99305869f8c22b76aff7119b8543972f92089ad Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 4 Aug 2020 14:16:32 +0200 Subject: [PATCH 58/89] cleanup --- Cubical/Homotopy/Loopspace.agda | 1 - Cubical/ZCohomology/Groups/Prelims.agda | 1 - Cubical/ZCohomology/Groups/Sn.agda | 84 +++++------------ .../ZCohomology/MayerVietorisUnreduced.agda | 13 +-- Cubical/new.agda | 91 +++++++++++++++++++ 5 files changed, 117 insertions(+), 73 deletions(-) create mode 100644 Cubical/new.agda diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda index 3e5b2a13a..abd6a581b 100644 --- a/Cubical/Homotopy/Loopspace.agda +++ b/Cubical/Homotopy/Loopspace.agda @@ -33,7 +33,6 @@ Eckmann-Hilton {A = A} n α β i = ; (i = i1) → (cong (λ x → lUnit x (~ k)) β) ∙ cong (λ x → rUnit x (~ k)) α}) ((λ j → α (j ∧ ~ i) ∙ β (j ∧ i)) ∙ λ j → α (~ i ∨ j) ∙ β (i ∨ j)) - {- Homotopy group version -} π-comp : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → ∥ typ ((Ω^ (suc n)) A) ∥₂ → ∥ typ ((Ω^ (suc n)) A) ∥₂ → ∥ typ ((Ω^ (suc n)) A) ∥₂ diff --git a/Cubical/ZCohomology/Groups/Prelims.agda b/Cubical/ZCohomology/Groups/Prelims.agda index c2a3c7a93..b8a4cda98 100644 --- a/Cubical/ZCohomology/Groups/Prelims.agda +++ b/Cubical/ZCohomology/Groups/Prelims.agda @@ -122,7 +122,6 @@ S1→K₁≡S1×Int = helper2 ⋄ S¹→S¹≡S¹×Int ⋄ helper helper3 = trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ a → cong ∣_∣ (S1→S¹-retr a) - {- Proof that S¹→K2 is isomorphic to K2×K1 (as types). Needed for H²(T²) -} S1→K2≡K2×K1 : Iso (S₊ 1 → coHomK 2) (coHomK 2 × coHomK 1) Iso.fun S1→K2≡K2×K1 f = f north , ΩKn+1→Kn (sym (rCancelₖ (f north)) diff --git a/Cubical/ZCohomology/Groups/Sn.agda b/Cubical/ZCohomology/Groups/Sn.agda index fe03356cd..1e39cf4ed 100644 --- a/Cubical/ZCohomology/Groups/Sn.agda +++ b/Cubical/ZCohomology/Groups/Sn.agda @@ -222,7 +222,7 @@ H²-S¹≅0 = Ker-right⊂Im-ϕ vSES-helper = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 propTruncIsProp) -- doesn't terminate without elimination λ a → I.Ker-i⊂Im-d 1 ∣ a ∣₂ ---------------- H¹(S²) -------------------------------------------- +--------------- H¹(Sⁿ), n ≥ 1 -------------------------------------------- H¹-Sⁿ≅0 : (n : ℕ) → GroupEquiv (coHomGr 1 (S₊ (2 + n))) trivialGroup H¹-Sⁿ≅0 n = coHomPushout≅coHomSn (1 + n) 1 @@ -249,48 +249,6 @@ H¹-Sⁿ≅0 n = coHomPushout≅coHomSn (1 + n) 1 λ a id → sym id ∙ I.Im-Δ⊂Ker-d 0 ∣ a ∣₂ (surj-helper _)) (I.Ker-i⊂Im-d 0 ∣ x ∣₂ inker) -{- -open import Cubical.Data.Sum -open import Cubical.Data.Empty renaming (rec to ⊥-rec) -Hᵐ-Sⁿ≅0 : (n m : ℕ) → (Σ[ x ∈ ℕ ] n + (suc x) ≡ m) ⊎ (Σ[ x ∈ ℕ ] m + (suc x) ≡ n) - → GroupEquiv (coHomGr (suc m) (S₊ (suc n))) trivialGroup -Hᵐ-Sⁿ≅0 n zero (inl (dif , id)) = ⊥-rec (snotz (sym (+-suc n dif) ∙ id)) -Hᵐ-Sⁿ≅0 zero (suc m) (inl (dif , id)) = {!!} - where - module I = MV Unit Unit (S₊ 1) (λ _ → tt) (λ _ → tt) -Hᵐ-Sⁿ≅0 (suc n) (suc m) (inl (dif , id)) = - coHomPushout≅coHomSn (suc n) (2 + m) - □ invGroupEquiv (vSES→GroupEquiv _ _ - (ses (isOfHLevelSuc 0 (isOfHLevelΣ 0 (isContrHⁿ-Unit m) λ _ → isContrHⁿ-Unit m)) - (isOfHLevelSuc 0 (isOfHLevelΣ 0 (isContrHⁿ-Unit (suc m)) λ _ → isContrHⁿ-Unit (suc m))) - (I.Δ (suc m)) - (I.i (suc (suc m))) - (I.d (suc m)) - (I.Ker-d⊂Im-Δ (suc m)) - (I.Ker-i⊂Im-d (suc m)))) - □ Hᵐ-Sⁿ≅0 n m (inl (dif , (cong predℕ id))) - - where - module I = MV Unit Unit (S₊ (suc n)) (λ _ → tt) (λ _ → tt) -Hᵐ-Sⁿ≅0 n zero (inr (dif , id)) = transport (λ i → GroupEquiv (coHomGr 1 (S₊ (suc (id i)))) trivialGroup) - (H¹-Sⁿ≅0 dif) -Hᵐ-Sⁿ≅0 n (suc m) (inr (dif , id)) = - coHomPushout≅coHomSn n (2 + m) - □ invGroupEquiv - (vSES→GroupEquiv _ _ - (ses (isOfHLevelSuc 0 (isOfHLevelΣ 0 (isContrHⁿ-Unit m) λ _ → isContrHⁿ-Unit m)) - (isOfHLevelSuc 0 (isOfHLevelΣ 0 (isContrHⁿ-Unit (suc m)) λ _ → isContrHⁿ-Unit (suc m))) - (I.Δ (suc m)) - (I.i (suc (suc m))) - (I.d (suc m)) - (I.Ker-d⊂Im-Δ (suc m)) - (I.Ker-i⊂Im-d (suc m)))) - □ transport (λ i → GroupEquiv (coHomGr (suc m) (S₊ ((id) i))) trivialGroup) - (Hᵐ-Sⁿ≅0 (m + suc dif) m (inr (dif , refl))) - where - module I = MV Unit Unit (S₊ n) (λ _ → tt) (λ _ → tt) --} - --------- Direct proof of H¹(S¹) ≅ ℤ without Mayer-Vietoris ------- -- The strategy is to use the proof that ΩS¹ ≃ ℤ. Since we only have this for S¹ with the base/loop definition @@ -305,24 +263,28 @@ Hᵐ-Sⁿ≅0 n (suc m) (inr (dif , id)) = -- application of cong₂, which allows us to use basechange-lemma. coHom1S1≃ℤ : GroupEquiv (coHomGr 1 (S₊ 1)) intGroup -eq coHom1S1≃ℤ = - isoToEquiv - (iso (sRec isSetInt (λ f → snd (Iso.fun S¹→S¹≡S¹×Int (Iso.fun S1→S1≡S¹→S¹ f)))) -- fun - (λ a → ∣ Iso.inv S1→S1≡S¹→S¹ (Iso.inv S¹→S¹≡S¹×Int (base , a)) ∣₂) -- inv - (λ a → (λ i → snd (Iso.fun S¹→S¹≡S¹×Int (Iso.fun S1→S1≡S¹→S¹ (Iso.inv S1→S1≡S¹→S¹ (Iso.inv S¹→S¹≡S¹×Int (base , a)))))) -- rInv - ∙ (λ i → snd (Iso.fun S¹→S¹≡S¹×Int (Iso.rightInv S1→S1≡S¹→S¹ (Iso.inv S¹→S¹≡S¹×Int (base , a)) i))) - ∙ λ i → snd (Iso.rightInv S¹→S¹≡S¹×Int (base , a) i)) - ((sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) -- lInv - λ f → (λ i → ∣ Iso.inv S1→S1≡S¹→S¹ (Iso.inv S¹→S¹≡S¹×Int (base , snd (Iso.fun S¹→S¹≡S¹×Int (Iso.fun S1→S1≡S¹→S¹ f)))) ∣₂) - ∙∙ ((λ i → sRec setTruncIsSet - (λ x → ∣ Iso.inv S1→S1≡S¹→S¹ x ∣₂) - (sRec setTruncIsSet - (λ x → ∣ Iso.inv S¹→S¹≡S¹×Int (x , (snd (Iso.fun S¹→S¹≡S¹×Int (Iso.fun S1→S1≡S¹→S¹ f)))) ∣₂) - (Iso.inv PathIdTrunc₀Iso (isConnectedS¹ (fst (Iso.fun S¹→S¹≡S¹×Int (Iso.fun S1→S1≡S¹→S¹ f)))) i))) - ∙ (λ i → ∣ Iso.inv S1→S1≡S¹→S¹ (Iso.leftInv S¹→S¹≡S¹×Int (Iso.fun S1→S1≡S¹→S¹ f) i) ∣₂)) - ∙∙ (λ i → ∣ Iso.leftInv S1→S1≡S¹→S¹ f i ∣₂)))) +eq coHom1S1≃ℤ = isoToEquiv theIso + where + F = Iso.fun S¹→S¹≡S¹×Int + F⁻ = Iso.inv S¹→S¹≡S¹×Int + G = Iso.fun S1→S1≡S¹→S¹ + G⁻ = Iso.inv S1→S1≡S¹→S¹ + + theIso : Iso ⟨ coHomGr 1 (S₊ 1) ⟩ ⟨ intGroup ⟩ + Iso.fun theIso = sRec isSetInt (λ f → snd (F (G f))) + Iso.inv theIso a = ∣ G⁻ (F⁻ (base , a)) ∣₂ + Iso.rightInv theIso a = + (cong (snd ∘ F) (Iso.rightInv S1→S1≡S¹→S¹ (F⁻ (base , a))) + ∙ cong snd (Iso.rightInv S¹→S¹≡S¹×Int (base , a))) + Iso.leftInv theIso = + sElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ f → cong ((sRec setTruncIsSet (λ x → ∣ G⁻ x ∣₂)) + ∘ sRec setTruncIsSet λ x → ∣ F⁻ (x , (snd (F (G f)))) ∣₂) + (Iso.inv PathIdTrunc₀Iso (isConnectedS¹ (fst (F (G f))))) + ∙∙ cong (∣_∣₂ ∘ G⁻) (Iso.leftInv S¹→S¹≡S¹×Int (G f)) + ∙∙ cong ∣_∣₂ (Iso.leftInv S1→S1≡S¹→S¹ f) isHom coHom1S1≃ℤ = - sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) -- isMorph + sElim2 (λ _ _ → isOfHLevelPath 2 isSetInt _ _) λ f g → (λ i → winding (guy (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 base)) ∙ Kn→ΩKn+1 1 (g (S¹→S1 base)))) (λ i → pre-guy (ΩKn+1→Kn (Kn→ΩKn+1 1 (f (S¹→S1 (loop i))) ∙ Kn→ΩKn+1 1 (g (S¹→S1 (loop i)))))))) ∙∙ cong winding (helper (f (S¹→S1 base)) (g (S¹→S1 base)) f g refl refl) @@ -334,6 +296,8 @@ isHom coHom1S1≃ℤ = pre-guy = S¹map ∘ trMap S1→S¹ guy = basechange2⁻ ∘ pre-guy + + helper : (x y : coHomK 1) (f g : S₊ 1 → coHomK 1) → (f (S¹→S1 base)) ≡ x → (g (S¹→S1 base)) ≡ y diff --git a/Cubical/ZCohomology/MayerVietorisUnreduced.agda b/Cubical/ZCohomology/MayerVietorisUnreduced.agda index 6ad27651f..4443719d1 100644 --- a/Cubical/ZCohomology/MayerVietorisUnreduced.agda +++ b/Cubical/ZCohomology/MayerVietorisUnreduced.agda @@ -140,15 +140,6 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : (Iso.fun (PathIdTrunc₀Iso) (cong fst p)) (Iso.fun (PathIdTrunc₀Iso) (cong snd p)) where - pushFiller : (n : ℕ) (F : (Pushout f g) → hLevelTrunc (3 + n) (S₊ (suc n))) - (p1 : Path (_ → hLevelTrunc (3 + n) (S₊ (suc n))) (λ a₁ → F (inl a₁)) (λ _ → ∣ north ∣)) - (p2 : Path (_ → hLevelTrunc (3 + n) (S₊ (suc n))) (λ a₁ → F (inr a₁)) (λ _ → ∣ north ∣)) → - (c : C) → I → I → hLevelTrunc (3 + n) (S₊ (suc n)) - pushFiller n F p1 p2 c i j = - hcomp (λ k → λ { (i = i1) → F (push c j) - ; (j = i0) → p1 (~ i ∧ k) (f c) - ; (j = i1) → p2 (~ i ∧ k) (g c)}) - (F (push c j)) helper : (n : ℕ) (F : (Pushout f g) → hLevelTrunc (3 + n) (S₊ (suc n))) (p1 : Path (_ → hLevelTrunc (3 + n) (S₊ (suc n))) (λ a₁ → F (inl a₁)) (λ _ → ∣ north ∣)) (p2 : Path (_ → hLevelTrunc (3 + n) (S₊ (suc n))) (λ a₁ → F (inr a₁)) (λ _ → ∣ north ∣)) @@ -166,7 +157,7 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : ∙∙ cong F (push a) ∙∙ cong (λ F₁ → F₁ (g a)) p2) (~ k) i ; (j = i1) → F (push a i)}) - (pushFiller zero F p1 p2 a j i) + (doubleCompPath-filler (sym (cong (λ F → F (f a)) p1)) (cong F (push a)) (cong (λ F → F (g a)) p2) (~ j) i) helper (suc n) F p1 p2 (push a i) j = hcomp (λ k → λ { (i = i0) → p1 (~ j) (f a) ; (i = i1) → p2 (~ j) (g a) @@ -174,7 +165,7 @@ module MV {ℓ ℓ' ℓ''} (A : Type ℓ) (B : Type ℓ') (C : Type ℓ'') (f : ∙∙ cong F (push a) ∙∙ cong (λ F₁ → F₁ (g a)) p2) (~ k) i ; (j = i1) → F (push a i)}) - (pushFiller (suc n) F p1 p2 a j i) + (doubleCompPath-filler (sym (cong (λ F → F (f a)) p1)) (cong F (push a)) (cong (λ F → F (g a)) p2) (~ j) i) open GroupHom diff --git a/Cubical/new.agda b/Cubical/new.agda new file mode 100644 index 000000000..b66dc4bc0 --- /dev/null +++ b/Cubical/new.agda @@ -0,0 +1,91 @@ +{-# OPTIONS --cubical --no-import-sorts --safe #-} +module Cubical.new where + +open import Cubical.HITs.Rationals.HITQ.Base +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Relation.Nullary + +open import Cubical.Data.Int + +open import Cubical.Data.Nat as ℕ hiding (_*_ ; +-comm) renaming (_+_ to _+n_) +open import Cubical.Data.NatPlusOne +open import Cubical.Data.Sigma + +open import Cubical.Data.Sigma +open import Cubical.Data.Empty +_≤_ : Int → Int → Type₀ +_≤_ a b = Σ[ x ∈ ℕ ] a + pos x ≡ b + +almostHom : (f : Int → Int) → Type₀ +almostHom f = Σ[ (max , min) ∈ Int × Int ] ((a b : Int) → (f (a + b) - ((f b) - f a)) ≤ max) × ((a b : Int) → min ≤ (f (a + b) - ((f b) - f a))) + +_≡almost_ : (f g : Int → Int) → Type₀ +_≡almost_ f g = Σ[ (max , min) ∈ Int × Int ] ((a : Int) → ((f a - g a) ≤ max) × (min ≤ (f a - g a))) + +data R : Type₀ where + r : (f : Int → Int) → R + eq : (g h : (Int → Int)) → g ≡almost h → r g ≡ r h + squashR : (a b : R) → (p q : a ≡ b) → p ≡ q +open import Cubical.HITs.Susp.Base +open import Cubical.Foundations.HLevels +open import Cubical.HITs.SetTruncation + +data R→ (A : R → Type₀) : Type₀ where + t : (r : R) (a : A r) → (R→ A) + squash→ : (f g : Int → Int) → (f ≡almost g) → {!!} + +isSetGenR : isSet R +isSetGenR = squashR + +genR-ind : ∀{ℓ} {A : R → Type ℓ} → ((r : R) → isSet (A r)) → ((f : Int → Int) → A (r f)) → (r : R) → A r +genR-ind is-set ind (r f) = ind f +genR-ind is-set ind (eq g h p i) = {!ind h!} +genR-ind is-set ind (squashR x x₁ p q i i₁) = {!!} + +funcGrAdd : {!!} +funcGrAdd = {!!} + +-- funcGrAdd : genR → genR → R +-- funcGrAdd (r f) b = {!!} +-- funcGrAdd (eq g h x i) b = {!!} +-- funcGrAdd (higher g h p q i i₁) b = {!!} + +-- -- funcGrAdd : genR → genR → genR +-- -- funcGrAdd (r f) (r g) = r λ x → f x + g x +-- -- funcGrAdd (r G) (eq f g p i) = +-- -- eq (λ x → G x + f x) (λ x → G x + g x) (fst p , (λ a → (fst (fst (snd p a)) , cong (_+ pos (fst (fst (snd p a)))) {!!} ∙ snd (fst (snd p a))) , {!!})) i +-- -- funcGrAdd (eq f g x i) (r x₁) = {!!} +-- -- funcGrAdd (eq f g x i) (eq f₁ g₁ x₁ i₁) = {!!} + +-- -- intHelper : (a b c : Int) → (a + b) - (a + c) ≡ b - c +-- -- intHelper (pos n) b c = {!+-comm!} +-- -- intHelper (negsuc n) b c = {!!} +-- -- intHelper2 : (a b c : Int) → (b + a) - (c + a) ≡ b - c +-- -- intHelper2 a b c = cong₂ (_-_) (+-comm b a) (+-comm c a) ∙ intHelper a b c + + +-- -- _+_→set : ∀{ℓ} {A : genR → Type ℓ} → ((r : genR) → isSet (A r)) → ((f : Int → Int) → A (r f)) → (r : genR) → A r +-- -- (set + ind →set) (r f) = ind f +-- -- (set + ind →set) (eq g h x i) = {!x!} + +-- -- _+r_ : R → R → R +-- -- _+r_ = elim2 (λ _ _ → setTruncIsSet) {!!} +-- -- where +-- -- +2 : (a b : genR) → R +-- -- +2 (r f) (r g) = ∣ r (λ x → f x + g x) ∣₂ +-- -- +2 (r f) (eq g h p i) = ∣ eq (λ x → f x + g x) +-- -- (λ x → f x + h x) +-- -- (fst p +-- -- , λ a → (fst (fst (snd p a)) +-- -- , cong (_+ pos (fst (fst (snd p a)))) (intHelper (f a) (g a) (h a)) ∙ snd (fst (snd p a))) +-- -- , fst (snd (snd p a)) +-- -- , snd (snd (snd p a)) ∙ sym (intHelper (f a) (g a) (h a))) i ∣₂ +-- -- +2 (eq g h p i) (r f) = ∣ eq (λ x → g x + f x) +-- -- (λ x → h x + f x) +-- -- (fst p +-- -- , λ a → (fst (fst (snd p a)) +-- -- , cong (_+ pos (fst (fst (snd p a)))) (intHelper2 (f a) (g a) (h a)) ∙ snd (fst (snd p a))) +-- -- , fst (snd (snd p a)) +-- -- , snd (snd (snd p a)) ∙ sym (intHelper2 (f a) (g a) (h a))) i ∣₂ +-- -- +2 (eq f g p i) (eq f2 g2 p2 j) = {!!} -- setTruncIsSet ∣ eq (λ x → f x + f2 x) (λ x → g x + f2 x) ((fst (fst p) , snd (fst p)) , λ a → {!!}) i ∣₂ {!!} {!!} {!!} i j From da2d5292eb9c63dc2dff988ab4177039a04d2d33 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 4 Aug 2020 14:22:54 +0200 Subject: [PATCH 59/89] cleanup2 --- Cubical/Structures/Group/Algebra.agda | 63 ------------------- Cubical/ZCohomology/Groups/Sn.agda | 2 - Cubical/ZCohomology/Properties.agda | 1 - Cubical/new.agda | 91 --------------------------- 4 files changed, 157 deletions(-) delete mode 100644 Cubical/new.agda diff --git a/Cubical/Structures/Group/Algebra.agda b/Cubical/Structures/Group/Algebra.agda index f890ab94b..1a4a18cc8 100644 --- a/Cubical/Structures/Group/Algebra.agda +++ b/Cubical/Structures/Group/Algebra.agda @@ -222,66 +222,3 @@ rUnitGroupIso = contrGroup≅trivialGroup : {G : Group {ℓ}} → isContr ⟨ G ⟩ → GroupEquiv G trivialGroup GroupEquiv.eq (contrGroup≅trivialGroup contr) = isContr→≃Unit contr GroupEquiv.isHom (contrGroup≅trivialGroup contr) _ _ = refl - -open GroupEquiv -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.Equiv -open import Cubical.Structures.Monoid hiding (⟨_⟩) -open import Cubical.Structures.Semigroup hiding (⟨_⟩) -open Monoid -open IsMonoid -open IsGroup -open Semigroup -open IsSemigroup - -GroupUa' : (G H : Group {ℓ}) → (p : ⟨ G ⟩ ≡ ⟨ H ⟩) - → PathP (λ i → p i) (0g G) (0g H) - → PathP (λ i → p i → p i → p i) (_+_ G) (_+_ H) - → PathP (λ i → p i → p i) (- G) (- H) - → G ≡ H -Carrier (GroupUa' G H p p0 p+ p- i) = p i -0g (GroupUa' G H p p0 p+ p- i) = p0 i -_+_ (GroupUa' G H p p0 p+ p- i) = p+ i -- GroupUa' G H p p0 p+ p- i = p- i -is-set (isSemigroup (isMonoid (isGroup (GroupUa' G H p p0 p+ p- i)))) = - isOfHLevel→isOfHLevelDep 1 {B = λ x → isSet x} (λ _ → isPropIsSet) (is-set G) (is-set H) p i -IsSemigroup.assoc (isSemigroup (isMonoid (isGroup (GroupUa' G H p p0 p+ p- i)))) = - hcomp (λ k → λ {(i = i0) → IsSemigroup.assoc (isSemigroup G) - ; (i = i1) → isPropΠ3 (λ _ _ _ → is-set H _ _) - (transport (λ j → ((x y z : p j) → p+ j x (p+ j y z) ≡ p+ j (p+ j x y) z)) - (IsSemigroup.assoc (isSemigroup G))) (IsSemigroup.assoc (isSemigroup H)) k }) - (transp (λ j → ((x y z : p (i ∧ j)) → p+ (i ∧ j) x (p+ (i ∧ j) y z) ≡ p+ (i ∧ j) (p+ (i ∧ j) x y) z)) (~ i) - (IsSemigroup.assoc (isSemigroup G))) -identity (isMonoid (isGroup (GroupUa' G H p p0 p+ p- i))) = - hcomp (λ k → λ {(i = i0) → identity G - ; (i = i1) → isPropΠ (λ _ → isOfHLevelΣ 1 (is-set H _ _) λ _ → is-set H _ _) - (transport (λ j → (x : p j) → (p+ j x (p0 j) ≡ x) × (p+ j (p0 j) x ≡ x)) - (identity G)) - (identity H) k } ) - (transp (λ j → (x : p (i ∧ j)) → (p+ (i ∧ j) x (p0 (i ∧ j)) ≡ x) × (p+ (i ∧ j) (p0 (i ∧ j)) x ≡ x)) (~ i) - (identity G)) -inverse (isGroup (GroupUa' G H p p0 p+ p- i)) = - hcomp (λ k → λ {(i = i0) → inverse G - ; (i = i1) → isPropΠ (λ _ → isOfHLevelΣ 1 (is-set H _ _) λ _ → is-set H _ _) - (transport (λ i → (x : p i) → (p+ i x (p- i x) ≡ p0 i) × (p+ i (p- i x) x ≡ p0 i)) (inverse G)) - (inverse H) k } ) - (transp (λ j → (x : p (i ∧ j)) → (p+ (i ∧ j) x (p- (i ∧ j) x) ≡ p0 (i ∧ j)) × (p+ (i ∧ j) (p- (i ∧ j) x) x ≡ p0 (i ∧ j))) (~ i) (inverse G)) - -GroupUa : (G H : Group {ℓ}) → GroupEquiv G H → G ≡ H -GroupUa G H e = - GroupUa' G H - (ua (eq e)) - (λ i → hcomp (λ k → λ {(i = i0) → 0g G - ; (i = i1) → (transportRefl (fst (eq e) (0g G)) - ∙ morph0→0 G H (grouphom (fst (eq e)) (isHom e))) k}) - (transp (λ j → ua (eq e) (i ∧ j)) (~ i) (0g G))) - (λ i a b → hcomp (λ k → λ {(i = i0) → _+_ G a b - ; (i = i1) → ((λ z → transportRefl (fst (eq e) ((_+_ G (invEq (eq e) (transportRefl a z)) (invEq (eq e) (transportRefl b z))))) z) - ∙∙ isHom e (invEq (eq e) a) (invEq (eq e) b) - ∙∙ cong₂ (_+_ H) (retEq (eq e) a) (retEq (eq e) b)) k}) - (transp (λ j → ua (eq e) (i ∧ j)) (~ i) (_+_ G (transp (λ j → ua (eq e) (i ∧ (~ j))) (~ i) a) (transp (λ j → ua (eq e) (i ∧ (~ j))) (~ i) b)))) - λ i a → hcomp (λ k → λ {(i = i0) → (- G) a - ; (i = i1) → ((λ z → transportRefl (fst (eq e) ((- G) (invEq (eq e) (transportRefl a z)))) z) - ∙∙ morphMinus G H (grouphom (fst (eq e)) (isHom e)) (invEq (eq e) a) - ∙∙ cong (- H) (retEq (eq e) a)) k}) - (transp (λ j → ua (eq e) (i ∧ j)) (~ i) ((- G) (transp (λ j → ua (eq e) (i ∧ (~ j))) (~ i) a))) diff --git a/Cubical/ZCohomology/Groups/Sn.agda b/Cubical/ZCohomology/Groups/Sn.agda index 1e39cf4ed..818ac7fa1 100644 --- a/Cubical/ZCohomology/Groups/Sn.agda +++ b/Cubical/ZCohomology/Groups/Sn.agda @@ -296,8 +296,6 @@ isHom coHom1S1≃ℤ = pre-guy = S¹map ∘ trMap S1→S¹ guy = basechange2⁻ ∘ pre-guy - - helper : (x y : coHomK 1) (f g : S₊ 1 → coHomK 1) → (f (S¹→S1 base)) ≡ x → (g (S¹→S1 base)) ≡ y diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index 40dc2ac7f..e11371613 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -215,7 +215,6 @@ commₕ {n = n} = sElim2 (λ _ _ → isOfHLevelPath 1 (§ _ _)) -- Proof that rUnitₖ and lUnitₖ agree on 0ₖ. Needed for Mayer-Vietoris. - rUnitlUnit0 : {n : ℕ} → rUnitₖ {n = n} 0ₖ ≡ lUnitₖ 0ₖ rUnitlUnit0 {n = zero} = refl rUnitlUnit0 {n = suc n} = diff --git a/Cubical/new.agda b/Cubical/new.agda deleted file mode 100644 index b66dc4bc0..000000000 --- a/Cubical/new.agda +++ /dev/null @@ -1,91 +0,0 @@ -{-# OPTIONS --cubical --no-import-sorts --safe #-} -module Cubical.new where - -open import Cubical.HITs.Rationals.HITQ.Base -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Function -open import Cubical.Relation.Nullary - -open import Cubical.Data.Int - -open import Cubical.Data.Nat as ℕ hiding (_*_ ; +-comm) renaming (_+_ to _+n_) -open import Cubical.Data.NatPlusOne -open import Cubical.Data.Sigma - -open import Cubical.Data.Sigma -open import Cubical.Data.Empty -_≤_ : Int → Int → Type₀ -_≤_ a b = Σ[ x ∈ ℕ ] a + pos x ≡ b - -almostHom : (f : Int → Int) → Type₀ -almostHom f = Σ[ (max , min) ∈ Int × Int ] ((a b : Int) → (f (a + b) - ((f b) - f a)) ≤ max) × ((a b : Int) → min ≤ (f (a + b) - ((f b) - f a))) - -_≡almost_ : (f g : Int → Int) → Type₀ -_≡almost_ f g = Σ[ (max , min) ∈ Int × Int ] ((a : Int) → ((f a - g a) ≤ max) × (min ≤ (f a - g a))) - -data R : Type₀ where - r : (f : Int → Int) → R - eq : (g h : (Int → Int)) → g ≡almost h → r g ≡ r h - squashR : (a b : R) → (p q : a ≡ b) → p ≡ q -open import Cubical.HITs.Susp.Base -open import Cubical.Foundations.HLevels -open import Cubical.HITs.SetTruncation - -data R→ (A : R → Type₀) : Type₀ where - t : (r : R) (a : A r) → (R→ A) - squash→ : (f g : Int → Int) → (f ≡almost g) → {!!} - -isSetGenR : isSet R -isSetGenR = squashR - -genR-ind : ∀{ℓ} {A : R → Type ℓ} → ((r : R) → isSet (A r)) → ((f : Int → Int) → A (r f)) → (r : R) → A r -genR-ind is-set ind (r f) = ind f -genR-ind is-set ind (eq g h p i) = {!ind h!} -genR-ind is-set ind (squashR x x₁ p q i i₁) = {!!} - -funcGrAdd : {!!} -funcGrAdd = {!!} - --- funcGrAdd : genR → genR → R --- funcGrAdd (r f) b = {!!} --- funcGrAdd (eq g h x i) b = {!!} --- funcGrAdd (higher g h p q i i₁) b = {!!} - --- -- funcGrAdd : genR → genR → genR --- -- funcGrAdd (r f) (r g) = r λ x → f x + g x --- -- funcGrAdd (r G) (eq f g p i) = --- -- eq (λ x → G x + f x) (λ x → G x + g x) (fst p , (λ a → (fst (fst (snd p a)) , cong (_+ pos (fst (fst (snd p a)))) {!!} ∙ snd (fst (snd p a))) , {!!})) i --- -- funcGrAdd (eq f g x i) (r x₁) = {!!} --- -- funcGrAdd (eq f g x i) (eq f₁ g₁ x₁ i₁) = {!!} - --- -- intHelper : (a b c : Int) → (a + b) - (a + c) ≡ b - c --- -- intHelper (pos n) b c = {!+-comm!} --- -- intHelper (negsuc n) b c = {!!} --- -- intHelper2 : (a b c : Int) → (b + a) - (c + a) ≡ b - c --- -- intHelper2 a b c = cong₂ (_-_) (+-comm b a) (+-comm c a) ∙ intHelper a b c - - --- -- _+_→set : ∀{ℓ} {A : genR → Type ℓ} → ((r : genR) → isSet (A r)) → ((f : Int → Int) → A (r f)) → (r : genR) → A r --- -- (set + ind →set) (r f) = ind f --- -- (set + ind →set) (eq g h x i) = {!x!} - --- -- _+r_ : R → R → R --- -- _+r_ = elim2 (λ _ _ → setTruncIsSet) {!!} --- -- where --- -- +2 : (a b : genR) → R --- -- +2 (r f) (r g) = ∣ r (λ x → f x + g x) ∣₂ --- -- +2 (r f) (eq g h p i) = ∣ eq (λ x → f x + g x) --- -- (λ x → f x + h x) --- -- (fst p --- -- , λ a → (fst (fst (snd p a)) --- -- , cong (_+ pos (fst (fst (snd p a)))) (intHelper (f a) (g a) (h a)) ∙ snd (fst (snd p a))) --- -- , fst (snd (snd p a)) --- -- , snd (snd (snd p a)) ∙ sym (intHelper (f a) (g a) (h a))) i ∣₂ --- -- +2 (eq g h p i) (r f) = ∣ eq (λ x → g x + f x) --- -- (λ x → h x + f x) --- -- (fst p --- -- , λ a → (fst (fst (snd p a)) --- -- , cong (_+ pos (fst (fst (snd p a)))) (intHelper2 (f a) (g a) (h a)) ∙ snd (fst (snd p a))) --- -- , fst (snd (snd p a)) --- -- , snd (snd (snd p a)) ∙ sym (intHelper2 (f a) (g a) (h a))) i ∣₂ --- -- +2 (eq f g p i) (eq f2 g2 p2 j) = {!!} -- setTruncIsSet ∣ eq (λ x → f x + f2 x) (λ x → g x + f2 x) ((fst (fst p) , snd (fst p)) , λ a → {!!}) i ∣₂ {!!} {!!} {!!} i j From 8612eb5ee2db454da2e5492cc4cfdc014fc6253a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Axel=20Ljungstr=C3=B6m?= Date: Tue, 1 Sep 2020 17:12:51 +0200 Subject: [PATCH 60/89] cleanup --- Cubical/Homotopy/Loopspace.agda | 2 +- Cubical/ZCohomology/Groups/Sn.agda | 10 +++++----- 2 files changed, 6 insertions(+), 6 deletions(-) diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda index abd6a581b..e95cb5081 100644 --- a/Cubical/Homotopy/Loopspace.agda +++ b/Cubical/Homotopy/Loopspace.agda @@ -30,7 +30,7 @@ Eckmann-Hilton : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (α β : typ ((Ω^ (2 + n Eckmann-Hilton {A = A} n α β i = comp (λ k → rUnit (snd ((Ω^ (1 + n)) A)) (~ k) ≡ rUnit (snd ((Ω^ (1 + n)) A)) (~ k)) -- note : rUnit refl := lUnit refl (λ k → λ { (i = i0) → (cong (λ x → rUnit x (~ k)) α) ∙ cong (λ x → lUnit x (~ k)) β - ; (i = i1) → (cong (λ x → lUnit x (~ k)) β) ∙ cong (λ x → rUnit x (~ k)) α}) + ; (i = i1) → (cong (λ x → lUnit x (~ k)) β) ∙ cong (λ x → rUnit x (~ k)) α}) ((λ j → α (j ∧ ~ i) ∙ β (j ∧ i)) ∙ λ j → α (~ i ∨ j) ∙ β (i ∨ j)) {- Homotopy group version -} diff --git a/Cubical/ZCohomology/Groups/Sn.agda b/Cubical/ZCohomology/Groups/Sn.agda index 8a74a7204..8f00b16f5 100644 --- a/Cubical/ZCohomology/Groups/Sn.agda +++ b/Cubical/ZCohomology/Groups/Sn.agda @@ -245,9 +245,9 @@ H¹-Sⁿ≅0 n = coHomPushout≅coHomSn (1 + n) 1 helper = sElim (λ _ → isOfHLevelΠ 2 λ _ → isOfHLevelSuc 1 (setTruncIsSet _ _)) -- useless elimination speeds things up significantly λ x inker → pRec (setTruncIsSet _ _) - (sigmaElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) - λ a id → sym id ∙ I.Im-Δ⊂Ker-d 0 ∣ a ∣₂ (surj-helper _)) - (I.Ker-i⊂Im-d 0 ∣ x ∣₂ inker) + (sigmaElim (λ _ → isOfHLevelPath 2 setTruncIsSet _ _) + λ a id → sym id ∙ I.Im-Δ⊂Ker-d 0 ∣ a ∣₂ (surj-helper _)) + (I.Ker-i⊂Im-d 0 ∣ x ∣₂ inker) --------- Direct proof of H¹(S¹) ≅ ℤ without Mayer-Vietoris ------- @@ -319,8 +319,8 @@ isHom coHom1S1≃ℤ = (cong (Kn→ΩKn+1 1) reflg ∙ Kn→ΩKn+10ₖ 1)) ∙ λ j → guy (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 base)) j) (λ i → pre-guy (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (f (S¹→S1 (loop i))) j)) - ∙ guy (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 base)) j) - (λ i → pre-guy (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 (loop i))) j))) + ∙ guy (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 base)) j) + (λ i → pre-guy (Iso.leftInv (Iso3-Kn-ΩKn+1 1) (g (S¹→S1 (loop i))) j))) From 39f34309a9e6f77a308a30ef861ef62cf8a00a8e Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 6 Jul 2021 16:29:35 +0200 Subject: [PATCH 61/89] Benchmarks updated --- .../Experiments/ZCohomology/Benchmarks.agda | 210 ++++++++++++++---- 1 file changed, 168 insertions(+), 42 deletions(-) diff --git a/Cubical/Experiments/ZCohomology/Benchmarks.agda b/Cubical/Experiments/ZCohomology/Benchmarks.agda index 1378e360f..f1e403259 100644 --- a/Cubical/Experiments/ZCohomology/Benchmarks.agda +++ b/Cubical/Experiments/ZCohomology/Benchmarks.agda @@ -28,17 +28,27 @@ open import Cubical.HITs.Sn open import Cubical.Algebra.Group hiding (ℤ ; Bool) open import Cubical.ZCohomology.Base open import Cubical.ZCohomology.Properties -open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.GroupStructure hiding (_+ₕ_) renaming (_+'ₕ_ to _+ₕ_) +-- _+'ₕ_ is just (λ x y → (x +ₕ 0ₕ) +ₕ (x +ₕ 0ₕ)) +-- For technical reason, this gives nicer reductions and computes better open import Cubical.ZCohomology.Groups.Sn open import Cubical.ZCohomology.Groups.Wedge open import Cubical.ZCohomology.Groups.Torus open import Cubical.ZCohomology.Groups.KleinBottle open import Cubical.ZCohomology.Groups.WedgeOfSpheres open import Cubical.ZCohomology.Groups.RP2 +open import Cubical.ZCohomology.Groups.CP2 open import Cubical.Data.Sigma open import Cubical.HITs.KleinBottle open import Cubical.HITs.RPn.Base +open import Cubical.HITs.SetTruncation +open import Cubical.HITs.Pushout +open import Cubical.HITs.Hopf +open import Cubical.HITs.Truncation +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 + open IsGroupHom open Iso @@ -52,25 +62,25 @@ module S1-tests where ϕ⁻¹ : ℤ → coHom 1 (S₊ 1) ϕ⁻¹ = inv (fst (Hⁿ-Sⁿ≅ℤ 0)) - test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- 30ms + test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- <10ms test₁ = refl test₂ : ϕ (ϕ⁻¹ 1) ≡ 1 -- <10ms test₂ = refl - test₃ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- <10ms + test₃ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- <10ms test₃ = refl - test₄ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 1) ≡ 1 -- 10ms + test₄ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 1) ≡ 1 -- 12ms test₄ = refl - test₅ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 0) ≡ 1 -- 11ms + test₅ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 0) ≡ 1 -- 13ms test₅ = refl - test₆ : ϕ (ϕ⁻¹ -3 +ₕ ϕ⁻¹ 4) ≡ 1 -- 29ms + test₆ : ϕ (ϕ⁻¹ -3 +ₕ ϕ⁻¹ 4) ≡ 1 -- 37ms test₆ = refl - test₇ : ϕ (ϕ⁻¹ -5 +ₕ ϕ⁻¹ -2) ≡ -7 -- 28ms + test₇ : ϕ (ϕ⁻¹ -5 +ₕ ϕ⁻¹ -2) ≡ -7 -- 38ms test₇ = refl -- S² @@ -85,26 +95,72 @@ module S2-tests where test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- 13ms test₁ = refl - test₂ : ϕ (ϕ⁻¹ 1) ≡ 1 -- 16ms + test₂ : ϕ (ϕ⁻¹ 1) ≡ 1 -- 17ms test₂ = refl - test₃ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- 278ms + test₃ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- 1,152ms test₃ = refl - test₄ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 1) ≡ 1 -- 290ms + test₄ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 1) ≡ 1 -- 1,235ms test₄ = refl -{- - test₅ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 0) ≡ 1 -- nope + test₅ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 0) ≡ 1 -- 1,208ms test₅ = refl - test₆ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 1) ≡ 2 -- nope + test₆ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 1) ≡ 2 -- 1,153ms test₆ = refl - test₇ : ϕ (ϕ⁻¹ 2 +ₕ ϕ⁻¹ 4) ≡ 6 -- nope + test₇ : ϕ (ϕ⁻¹ 2 +ₕ ϕ⁻¹ 4) ≡ 6 -- 1,365ms test₇ = refl + +-- S³ +module S3-tests where + + ϕ : coHom 3 (S₊ 3) → ℤ + ϕ = fun (fst (Hⁿ-Sⁿ≅ℤ 2)) + + ϕ⁻¹ : ℤ → coHom 3 (S₊ 3) + ϕ⁻¹ = inv (fst (Hⁿ-Sⁿ≅ℤ 2)) + + test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- 228ms + test₁ = refl + + test₂ : ϕ (ϕ⁻¹ 1) ≡ 1 -- 231ms + test₂ = refl + + test₃ : ϕ (ϕ⁻¹ 3) ≡ 3 -- 325ms + test₃ = refl + +{- + test₄ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- nope + test₄ = refl + + test₅ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 1) ≡ 1 -- nope + test₅ = refl -} +-- S⁴ +module S4-tests where + + ϕ : coHom 4 (S₊ 4) → ℤ + ϕ = fun (fst (Hⁿ-Sⁿ≅ℤ 3)) + + ϕ⁻¹ : ℤ → coHom 4 (S₊ 4) + ϕ⁻¹ = inv (fst (Hⁿ-Sⁿ≅ℤ 3)) + +{- _+ₕ_ Fails already here... + test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- nope + test₁ = refl +-} + +-- ϕ can handle 0ₕ fast + test₂ : ϕ (0ₕ _) ≡ 0 -- < 10ms + test₂ = refl + +{- It fails to map the generator to 1, however. + test₂ : ϕ (∣ ∣_∣ ∣₂) ≡ 1 -- nope + test₂ = refl +-} module S1∨S1∨S2-tests₁ where -- everything fast @@ -114,19 +170,19 @@ module S1∨S1∨S2-tests₁ where -- everything fast ϕ⁻¹ : ℤ × ℤ → coHom 1 S²⋁S¹⋁S¹ ϕ⁻¹ = inv (fst H¹-S²⋁S¹⋁S¹) - test₁ : ϕ (ϕ⁻¹ (0 , 0)) ≡ (0 , 0) -- <10ms + test₁ : ϕ (ϕ⁻¹ (0 , 0)) ≡ (0 , 0) -- 11ms test₁ = refl - test₂ : ϕ (ϕ⁻¹ (3 , 1)) ≡ (3 , 1) -- 21ms + test₂ : ϕ (ϕ⁻¹ (3 , 1)) ≡ (3 , 1) -- 23ms test₂ = refl - test₃ : ϕ (ϕ⁻¹ (0 , 0) +ₕ ϕ⁻¹ (0 , 0)) ≡ (0 , 0) -- 15ms + test₃ : ϕ (ϕ⁻¹ (0 , 0) +ₕ ϕ⁻¹ (0 , 0)) ≡ (0 , 0) -- 19ms test₃ = refl - test₄ : ϕ (ϕ⁻¹ (0 , 1) +ₕ ϕ⁻¹ (1 , 0)) ≡ (1 , 1) -- 21ms + test₄ : ϕ (ϕ⁻¹ (0 , 1) +ₕ ϕ⁻¹ (1 , 0)) ≡ (1 , 1) -- 26ms test₄ = refl - test₅ : ϕ (ϕ⁻¹ (3 , 2) +ₕ ϕ⁻¹ (-1 , 5)) ≡ (2 , 7) -- 47ms + test₅ : ϕ (ϕ⁻¹ (3 , 2) +ₕ ϕ⁻¹ (-1 , 5)) ≡ (2 , 7) -- 62ms test₅ = refl @@ -138,15 +194,27 @@ module S1∨S1∨S2-tests₂ where ϕ⁻¹ : ℤ → coHom 2 S²⋁S¹⋁S¹ ϕ⁻¹ = inv (fst H²-S²⋁S¹⋁S¹) - test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- 157ms + test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- 106ms test₁ = refl - test₂ : ϕ (ϕ⁻¹ 3) ≡ 3 -- 119ms + test₂ : ϕ (ϕ⁻¹ 3) ≡ 3 -- 125ms test₂ = refl - test₃ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- 1,820ms + test₃ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- 9,689ms test₃ = refl + test₄ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 1) ≡ 1 -- 9,235ms + test₄ = refl + + test₅ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 0) ≡ 1 -- 9,748ms + test₅ = refl + + test₆ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 1) ≡ 2 -- 9,136ms + test₆ = refl + + test₇ : ϕ (ϕ⁻¹ 2 +ₕ ϕ⁻¹ 4) ≡ 6 -- 9,557ms + test₇ = refl + module Torus-test₁ where -- fast @@ -156,19 +224,19 @@ module Torus-test₁ where -- fast ϕ⁻¹ : ℤ × ℤ → coHom 1 (S₊ 1 × S₊ 1) ϕ⁻¹ = inv (fst H¹-T²≅ℤ×ℤ) - test₁ : ϕ (ϕ⁻¹ (0 , 0)) ≡ (0 , 0) -- <10ms + test₁ : ϕ (ϕ⁻¹ (0 , 0)) ≡ (0 , 0) -- 11ms test₁ = refl - test₂ : ϕ (ϕ⁻¹ (3 , 1)) ≡ (3 , 1) -- 18ms + test₂ : ϕ (ϕ⁻¹ (3 , 1)) ≡ (3 , 1) -- 17ms test₂ = refl - test₃ : ϕ (ϕ⁻¹ (0 , 0) +ₕ ϕ⁻¹ (0 , 0)) ≡ (0 , 0) -- 15ms + test₃ : ϕ (ϕ⁻¹ (0 , 0) +ₕ ϕ⁻¹ (0 , 0)) ≡ (0 , 0) -- 19ms test₃ = refl - test₄ : ϕ (ϕ⁻¹ (0 , 1) +ₕ ϕ⁻¹ (1 , 0)) ≡ (1 , 1) -- 20ms + test₄ : ϕ (ϕ⁻¹ (0 , 1) +ₕ ϕ⁻¹ (1 , 0)) ≡ (1 , 1) -- 26ms test₄ = refl - test₅ : ϕ (ϕ⁻¹ (-3 , 2) +ₕ ϕ⁻¹ (-1 , 5)) ≡ (-4 , 7) -- 44ms + test₅ : ϕ (ϕ⁻¹ (-3 , 2) +ₕ ϕ⁻¹ (-1 , 5)) ≡ (-4 , 7) -- 61ms test₅ = refl @@ -180,22 +248,27 @@ module Torus-test₂ where ϕ⁻¹ : ℤ → coHom 2 (S₊ 1 × S₊ 1) ϕ⁻¹ = inv (fst H²-T²≅ℤ) - test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- 121ms + test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- 136sm test₁ = refl - test₂ : ϕ (ϕ⁻¹ 3) ≡ 3 -- 142ms + test₂ : ϕ (ϕ⁻¹ 3) ≡ 3 -- 154ms test₂ = refl - test₃ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- 3,674ms + test₃ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- 12,790ms test₃ = refl - test₄ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 1) ≡ 1 -- 3,772ms + test₄ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 1) ≡ 1 -- 12,366ms test₄ = refl -{- - test₅ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 0) ≡ 1 -- nope + test₅ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 0) ≡ 1 -- 12,257ms test₅ = refl --} + + test₆ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 1) ≡ 2 -- 13,092ms + test₆ = refl + + test₇ : ϕ (ϕ⁻¹ 2 +ₕ ϕ⁻¹ 4) ≡ 6 -- 12,528ms + test₇ = refl + module Klein-test₁ where -- fast @@ -208,26 +281,26 @@ module Klein-test₁ where -- fast test₁ : ϕ (ϕ⁻¹ 0) ≡ 0 -- <10ms test₁ = refl - test₂ : ϕ (ϕ⁻¹ 3) ≡ 3 -- 12ms + test₂ : ϕ (ϕ⁻¹ 3) ≡ 3 -- 13ms test₂ = refl - test₃ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- <10ms + test₃ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- 10ms test₃ = refl - test₄ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 1) ≡ 1 -- 11ms + test₄ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 1) ≡ 1 -- 14ms test₄ = refl - test₅ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 0) ≡ 1 -- 12ms + test₅ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 0) ≡ 1 -- 14ms test₅ = refl - test₆ : ϕ (ϕ⁻¹ -3 +ₕ ϕ⁻¹ 4) ≡ 1 -- 29ms + test₆ : ϕ (ϕ⁻¹ -3 +ₕ ϕ⁻¹ 4) ≡ 1 -- 38ms test₆ = refl - test₇ : ϕ (ϕ⁻¹ -5 +ₕ ϕ⁻¹ -2) ≡ -7 -- 29ms + test₇ : ϕ (ϕ⁻¹ -5 +ₕ ϕ⁻¹ -2) ≡ -7 -- 38ms test₇ = refl -- The example in the paper: - test : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 2) ≡ 3 -- 15ms + test : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 2) ≡ 3 -- 22ms test = refl @@ -243,6 +316,7 @@ module Klein-test₂ where test₀ = refl -} + module RP2-test₂ where ϕ : coHom 2 RP² → Bool ϕ = fun (fst H²-RP²≅Bool) @@ -250,10 +324,62 @@ module RP2-test₂ where ϕ⁻¹ : Bool → coHom 2 RP² ϕ⁻¹ = inv (fst H²-RP²≅Bool) - test₀ : ϕ (0ₕ _) ≡ true -- 1,210ms (unlike for Klein, this works) + test₀ : ϕ (0ₕ _) ≡ true -- 1,328ms (unlike for Klein, this works) test₀ = refl {- test₁ : ϕ (ϕ⁻¹ true) ≡ true -- nope test₁ = refl -} + + +module CP2-test₂ where + ϕ : coHom 2 CP² → ℤ + ϕ = fun (fst H²CP²≅ℤ) + + ϕ⁻¹ : ℤ → coHom 2 CP² + ϕ⁻¹ = inv (fst H²CP²≅ℤ) + + -- For explicitly constructed elmenet g : H²CP², ϕ works well + test₀ : ϕ (0ₕ _) ≡ 0 -- <10ms + test₀ = refl + + generator : coHom 2 CP² + generator = ∣ (λ { (inl x) → ∣ x ∣ ; (inr x) → 0ₖ _ ; (push a i) → p a i}) ∣₂ + where + ind : (B : TotalHopf → Type) → ((x : _) → isOfHLevel 3 (B x)) → B (north , base) → (x : _) → B x + ind = + transport (λ i → (B : isoToPath IsoS³TotalHopf i → Type) + → ((x : _) → isOfHLevel 3 (B x)) + → B (transp (λ j → isoToPath IsoS³TotalHopf (i ∨ ~ j)) i (north , base)) → (x : _) → B x) + λ B hLev ind → sphereElim _ (λ _ → hLev _) ind + + p : (a : TotalHopf) → ∣ fst a ∣ ≡ 0ₖ 2 + p = ind _ (λ _ → isOfHLevelTrunc 4 _ _) refl + + test₁ : ϕ generator ≡ 1 -- 24ms + test₁ = refl + + + -- For _+ₕ_ too + test₂ : ϕ (ϕ⁻¹ 0 +ₕ ϕ⁻¹ 0) ≡ 0 -- 1,343ms + test₂ = refl + + test₃ : ϕ (ϕ⁻¹ 1 +ₕ ϕ⁻¹ 1) ≡ 2 -- 1,302ms + test₃ = refl + + test₄ : ϕ (ϕ⁻¹ 2 +ₕ ϕ⁻¹ 2) ≡ 4 -- 1,410ms + test₄ = refl + + +module CP2-test₄ where + ϕ : coHom 4 CP² → ℤ + ϕ = fun (fst H⁴CP²≅ℤ) + + ϕ⁻¹ : ℤ → coHom 4 CP² + ϕ⁻¹ = inv (fst H⁴CP²≅ℤ) + +{- + test₀ : ϕ (0ₕ _) ≡ true -- fails already here... + test₀ = refl +-} From f6bdb86f4747fb45722f351281fca9c314c3af9a Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 30 Aug 2021 02:23:05 +0200 Subject: [PATCH 62/89] stuff --- Cubical/Foundations/Pointed/Homogeneous.agda | 43 + Cubical/ZCohomology/Groups/CP2.agda | 913 ++++++++++++++---- Cubical/ZCohomology/Gysin.agda | 765 +++++++++++++++ .../RingStructure/GradedCommutativity.agda | 45 +- .../ZCohomology/RingStructure/RingLaws.agda | 4 + 5 files changed, 1564 insertions(+), 206 deletions(-) create mode 100644 Cubical/ZCohomology/Gysin.agda diff --git a/Cubical/Foundations/Pointed/Homogeneous.agda b/Cubical/Foundations/Pointed/Homogeneous.agda index 5315e2685..7bb614d26 100644 --- a/Cubical/Foundations/Pointed/Homogeneous.agda +++ b/Cubical/Foundations/Pointed/Homogeneous.agda @@ -46,6 +46,49 @@ isHomogeneous {ℓ} (A , x) = ∀ y → Path (Pointed ℓ) (A , x) (A , y) }) (sym (h (pt B∙)) ∙ h ((sym f₀ ∙∙ funExt⁻ p a₀ ∙∙ g₀) i)) +→∙Homogeneous≡Path : ∀ {ℓ ℓ'} {A∙ : Pointed ℓ} {B∙ : Pointed ℓ'} {f∙ g∙ : A∙ →∙ B∙} + (h : isHomogeneous B∙) → (p q : f∙ ≡ g∙) → cong fst p ≡ cong fst q → p ≡ q +→∙Homogeneous≡Path {A∙ = A∙@(A , a₀)} {B∙@(B , b)} {f∙@(f , f₀)} {g∙@(g , g₀)} h p q r = + transport (λ k + → PathP (λ i + → PathP (λ j → (A , a₀) →∙ newPath-refl p q r i j (~ k)) + (f , f₀) (g , g₀)) p q) + (badPath p q r) + where + newPath : (p q : f∙ ≡ g∙) (r : cong fst p ≡ cong fst q) + → Square (refl {x = b}) refl refl refl + newPath p q r i j = + hcomp (λ k → λ {(i = i0) → cong snd p j k + ; (i = i1) → cong snd q j k + ; (j = i0) → f₀ k + ; (j = i1) → g₀ k}) + (r i j a₀) + + newPath-refl : (p q : f∙ ≡ g∙) (r : cong fst p ≡ cong fst q) + → PathP (λ i → (PathP (λ j → B∙ ≡ (B , newPath p q r i j))) refl refl) refl refl + newPath-refl p q r i j k = + hcomp (λ w → λ { (i = i0) → lCancel (h b) w k + ; (i = i1) → lCancel (h b) w k + ; (j = i0) → lCancel (h b) w k + ; (j = i1) → lCancel (h b) w k + ; (k = i0) → lCancel (h b) w k + ; (k = i1) → B , newPath p q r i j}) + ((sym (h b) ∙ h (newPath p q r i j)) k) + + badPath : (p q : f∙ ≡ g∙) (r : cong fst p ≡ cong fst q) + → PathP (λ i → + PathP (λ j → A∙ →∙ (B , newPath p q r i j)) + (f , f₀) (g , g₀)) + p q + fst (badPath p q r i j) = r i j + snd (badPath p q s i j) k = + hcomp (λ r → λ { (i = i0) → snd (p j) (r ∧ k) + ; (i = i1) → snd (q j) (r ∧ k) + ; (j = i0) → f₀ (k ∧ r) + ; (j = i1) → g₀ (k ∧ r) + ; (k = i0) → s i j a₀}) + (s i j a₀) + isHomogeneousPi : ∀ {ℓ ℓ'} {A : Type ℓ} {B∙ : A → Pointed ℓ'} → (∀ a → isHomogeneous (B∙ a)) → isHomogeneous (Πᵘ∙ A B∙) isHomogeneousPi h f i .fst = ∀ a → typ (h a (f a) i) diff --git a/Cubical/ZCohomology/Groups/CP2.agda b/Cubical/ZCohomology/Groups/CP2.agda index f7a7f0483..039c4f174 100644 --- a/Cubical/ZCohomology/Groups/CP2.agda +++ b/Cubical/ZCohomology/Groups/CP2.agda @@ -9,6 +9,7 @@ open import Cubical.ZCohomology.MayerVietorisUnreduced open import Cubical.ZCohomology.Groups.Unit open import Cubical.ZCohomology.Groups.Sn open import Cubical.ZCohomology.RingStructure.CupProduct +open import Cubical.ZCohomology.RingStructure.RingLaws open import Cubical.Foundations.HLevels open import Cubical.Foundations.Function @@ -26,7 +27,7 @@ open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) open import Cubical.Data.Nat.Order open import Cubical.Data.Unit open import Cubical.Algebra.Group - renaming (ℤ to ℤGroup ; Unit to UnitGroup) + renaming (ℤ to ℤGroup ; Unit to UnitGroup) hiding (Bool) open import Cubical.HITs.Pushout open import Cubical.HITs.S1 @@ -35,213 +36,727 @@ open import Cubical.HITs.Susp open import Cubical.HITs.Hopf open import Cubical.HITs.Join open import Cubical.HITs.SetTruncation - renaming (rec to sRec ; elim to sElim ; elim2 to sElim2 ; map to sMap) + renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap) open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim2 to pElim2 ; ∣_∣ to ∣_∣₁ ; map to pMap) open import Cubical.HITs.Truncation + renaming (elim to trElim ; rec2 to trRec2 ; rec to trRec) open import Cubical.Relation.Nullary open IsGroupHom open Iso -CP² : Type -CP² = Pushout {A = TotalHopf} fst λ _ → tt - -characFunSpaceCP² : ∀ {ℓ} {A : Type ℓ} - → Iso (CP² → A) (Σ[ x ∈ A ] Σ[ f ∈ (S₊ (suc (suc zero)) → A) ] - ((y : TotalHopf) → f (fst y) ≡ x)) -fun characFunSpaceCP² f = (f (inr tt)) , ((f ∘ inl ) , (λ a → cong f (push a))) -inv characFunSpaceCP² (a , f , p) (inl x) = f x -inv characFunSpaceCP² (a , f , p) (inr x) = a -inv characFunSpaceCP² (a , f , p) (push b i) = p b i -rightInv characFunSpaceCP² _ = refl -leftInv characFunSpaceCP² _ = - funExt λ { (inl x) → refl - ; (inr x) → refl - ; (push a i) → refl} - -H⁰CP²≅ℤ : GroupIso (coHomGr 0 CP²) ℤGroup -H⁰CP²≅ℤ = - H⁰-connected (inr tt) - (PushoutToProp (λ _ → squash) - (sphereElim _ (λ _ → isOfHLevelSuc 1 squash) - ∣ sym (push (north , base)) ∣₁) - λ _ → ∣ refl ∣₁) - -module M = MV (S₊ 2) Unit TotalHopf fst (λ _ → tt) - -H²CP²≅ℤ : GroupIso (coHomGr 2 CP²) ℤGroup -H²CP²≅ℤ = compGroupIso (BijectionIso→GroupIso bij) - (compGroupIso (invGroupIso trivIso) (Hⁿ-Sⁿ≅ℤ 1)) - where - isContrH¹TotalHopf : isContr (coHom 1 TotalHopf) - isContrH¹TotalHopf = - isOfHLevelRetractFromIso 0 (setTruncIso (domIso (invIso (IsoS³TotalHopf)))) - (isOfHLevelRetractFromIso 0 ((fst (H¹-Sⁿ≅0 1))) isContrUnit) - - isContrH²TotalHopf : isContr (coHom 2 TotalHopf) - isContrH²TotalHopf = - isOfHLevelRetractFromIso 0 (setTruncIso (domIso (invIso (IsoS³TotalHopf)))) - ((isOfHLevelRetractFromIso 0 - (fst (Hⁿ-Sᵐ≅0 1 2 λ p → snotz (sym (cong predℕ p)))) isContrUnit)) - - trivIso : GroupIso (coHomGr 2 (Susp S¹)) (×coHomGr 2 (Susp S¹) Unit) - fun (fst trivIso) x = x , 0ₕ _ - inv (fst trivIso) = fst - rightInv (fst trivIso) (x , p) = - ΣPathP (refl , isContr→isProp (isContrHⁿ-Unit 1) _ _) - leftInv (fst trivIso) x = refl - snd trivIso = makeIsGroupHom λ _ _ → refl - - bij : BijectionIso (coHomGr 2 CP²) (×coHomGr 2 (Susp S¹) Unit) - BijectionIso.fun bij = M.i 2 - BijectionIso.inj bij x p = - pRec (squash₂ _ _) - (uncurry (λ z q - → sym q - ∙∙ cong (fst (M.d 1)) (isContr→isProp isContrH¹TotalHopf z (0ₕ _)) - ∙∙ (M.d 1) .snd .pres1)) - (M.Ker-i⊂Im-d 1 x p) - where - help : isInIm (M.d 1) x - help = M.Ker-i⊂Im-d 1 x p - BijectionIso.surj bij y = - M.Ker-Δ⊂Im-i 2 y (isContr→isProp isContrH²TotalHopf _ _) - -H⁴CP²≅ℤ : GroupIso (coHomGr 4 CP²) ℤGroup -H⁴CP²≅ℤ = compGroupIso (invGroupIso (BijectionIso→GroupIso bij)) - (compGroupIso help (Hⁿ-Sⁿ≅ℤ 2)) + + +diag : ∀ {ℓ} {A : Type ℓ} {x y z w : A} (p : x ≡ y) (r : z ≡ w) (q : y ≡ z) (s : x ≡ w) (P : PathP (λ i → p i ≡ r (~ i)) s q) + → (λ i → P i i) ≡ p ∙ q +diag {x = x} {y = y} {z = z} {w = w} p r q s P i j = + hcomp (λ k → λ {(i = i0) → P j (j ∧ k) + ; (j = i0) → x + ; (j = i1) → q k}) + (p j) + {- + ———— Boundary —————————————————————————————————————————————— +i = i0 ⊢ P j j +i = i1 ⊢ (p ∙ q) j +j = i0 ⊢ x +j = i1 ⊢ z + -} + + +testagain : (g : S₊ 3 → S₊ 2) (p : Σ[ f ∈ (S₊ 2 → coHomK 4) ] ((x : S₊ 3) → f (g x) ≡ 0ₖ 4)) + → (s : (λ x → 0ₖ 4) ≡ fst p) + → p ≡ ((λ _ → 0ₖ 4) , λ x → funExt⁻ s (g x) ∙ (snd p x)) +testagain g (f , k) s = help f s k where - help : GroupIso (coHomGr 3 TotalHopf) (coHomGr 3 (S₊ 3)) - help = isoType→isoCohom 3 (invIso IsoS³TotalHopf) - - bij : BijectionIso (coHomGr 3 TotalHopf) (coHomGr 4 CP²) - BijectionIso.fun bij = M.d 3 - BijectionIso.inj bij x p = - pRec (squash₂ _ _) - (uncurry (λ z q → - sym q - ∙∙ cong (M.Δ 3 .fst) - (isOfHLevelΣ 1 (isContr→isProp - (isOfHLevelRetractFromIso 0 - (fst (Hⁿ-Sᵐ≅0 2 1 λ p → snotz (cong predℕ p))) isContrUnit)) - (λ _ → isContr→isProp (isContrHⁿ-Unit 2)) - z (0ₕ _ , 0ₕ _)) - ∙∙ M.Δ 3 .snd .pres1)) - (M.Ker-d⊂Im-Δ _ x p) - BijectionIso.surj bij y = - M.Ker-i⊂Im-d 3 y (isOfHLevelΣ 1 - (isContr→isProp (isOfHLevelRetractFromIso 0 - (fst (Hⁿ-Sᵐ≅0 3 1 λ p → snotz (cong predℕ p))) isContrUnit)) - (λ _ → isContr→isProp (isContrHⁿ-Unit _)) _ _) - - --- Characterisations of the trivial groups - -private - elim-TotalHopf : (B : TotalHopf → Type) - → ((x : _) → isOfHLevel 3 (B x)) → B (north , base) - → (x : _) → B x - elim-TotalHopf = - transport (λ i → (B : isoToPath IsoS³TotalHopf i → Type) - → ((x : _) → isOfHLevel 3 (B x)) - → B (transp (λ j → isoToPath IsoS³TotalHopf (i ∨ ~ j)) i (north , base)) → (x : _) → B x) - λ B hLev elim-TotalHopf → sphereElim _ (λ _ → hLev _) elim-TotalHopf - -H¹-CP²≅0 : GroupIso (coHomGr 1 CP²) UnitGroup -H¹-CP²≅0 = - contrGroupIsoUnit - (isOfHLevelRetractFromIso 0 (setTruncIso characFunSpaceCP²) - (isOfHLevelRetractFromIso 0 lem₂ lem₃)) + help : (f : (S₊ 2 → coHomK 4)) (s : (λ x → 0ₖ 4) ≡ f) (k : (x : S₊ 3) → f (g x) ≡ 0ₖ 4) + → Path (Σ[ f ∈ (S₊ 2 → coHomK 4) ] ((x : S₊ 3) → f (g x) ≡ 0ₖ 4)) (f , k) + (((λ _ → 0ₖ 4) , λ x → funExt⁻ s (g x) ∙ (k x))) + help f = + J (λ f s → (k : (x : S₊ 3) → f (g x) ≡ 0ₖ 4) + → Path (Σ[ f ∈ (S₊ 2 → coHomK 4) ] ((x : S₊ 3) → f (g x) ≡ 0ₖ 4)) (f , k) + (((λ _ → 0ₖ 4) , λ x → funExt⁻ s (g x) ∙ (k x)))) + λ k → ΣPathP (refl , (funExt (λ x → lUnit (k x)))) + +open import Cubical.Homotopy.Loopspace +open import Cubical.HITs.SmashProduct + +hopfy : ∀ {ℓ} {A : Pointed ℓ} + → (Ω^ 2) A .fst → (Ω^ 3) A .fst +hopfy α l i = + hcomp + (λ m → λ + { (i = i0) → refl + ; (i = i1) → λ j → side l m j + ; (l = i0) → α (i ∧ ~ m) + ; (l = i1) → α (~ i ∨ m) + }) + (λ j → top l i j) where - lem₁ : (f : (Susp S¹ → coHomK 1)) → ∥ (λ _ → 0ₖ _) ≡ f ∥ - lem₁ f = pMap (λ p → p) - (Iso.fun PathIdTrunc₀Iso (isOfHLevelRetractFromIso 1 - (fst (Hⁿ-Sᵐ≅0 0 1 (λ p → snotz (sym p)))) isPropUnit (0ₕ _) ∣ f ∣₂)) - - lem₂ : Iso ∥ (Σ[ x ∈ coHomK 1 ] ( Σ[ f ∈ (Susp S¹ → coHomK 1) ] ((y : TotalHopf) → f (fst y) ≡ x))) ∥₂ - ∥ (Σ[ f ∈ (Susp S¹ → coHomK 1) ] ((y : TotalHopf) → f (fst y) ≡ 0ₖ 1)) ∥₂ - fun lem₂ = sMap (uncurry λ x → uncurry λ f p → (λ y → (-ₖ x) +ₖ f y) , λ y → cong ((-ₖ x) +ₖ_) (p y) ∙ lCancelₖ _ x) - inv lem₂ = sMap λ p → 0ₖ _ , p - rightInv lem₂ = - sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - λ {(f , p) → cong ∣_∣₂ (ΣPathP ((funExt (λ x → lUnitₖ _ (f x))) - , (funExt (λ y → sym (rUnit (λ i → (-ₖ 0ₖ 1) +ₖ p y i))) - ◁ λ j y i → lUnitₖ _ (p y i) j)))} - leftInv lem₂ = - sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - (uncurry (coHomK-elim _ (λ _ → isPropΠ (λ _ → squash₂ _ _)) - (uncurry λ f p → cong ∣_∣₂ (ΣPathP (refl , (ΣPathP ((funExt (λ x → lUnitₖ _ (f x))) - , ((funExt (λ y → sym (rUnit (λ i → (-ₖ 0ₖ 1) +ₖ p y i))) - ◁ λ j y i → lUnitₖ _ (p y i) j))))))))) - - lem₃ : isContr _ - fst lem₃ = ∣ (λ _ → 0ₖ 1) , (λ _ → refl) ∣₂ - snd lem₃ = - sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - (uncurry λ f → pRec (isPropΠ (λ _ → squash₂ _ _)) - (J (λ f _ → (y : (y₁ : TotalHopf) → f (fst y₁) ≡ 0ₖ 1) → - ∣ (λ _ → 0ₖ 1) , (λ _ _ → 0ₖ 1) ∣₂ ≡ ∣ f , y ∣₂) - (λ y → cong ∣_∣₂ (ΣPathP ((funExt (λ z → sym (y (north , base)))) , toPathP (s y))))) - (lem₁ f)) - - where - s : (y : TotalHopf → 0ₖ 1 ≡ 0ₖ 1) - → transport (λ i → (_ : TotalHopf) → y (north , base) (~ i) ≡ ∣ base ∣) - (λ _ _ → 0ₖ 1) ≡ y - s y = funExt (elim-TotalHopf _ (λ _ → isOfHLevelPath 3 (isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) _ _) - λ k → transp (λ i → y (north , base) (~ i ∧ ~ k) ≡ ∣ base ∣) k - λ j → y (north , base) (~ k ∨ j)) - -Hⁿ-CP²≅0-higher : (n : ℕ) → ¬ (n ≡ 1) → GroupIso (coHomGr (3 +ℕ n) CP²) UnitGroup -Hⁿ-CP²≅0-higher n p = contrGroupIsoUnit ((0ₕ _) , (λ x → sym (main x))) + a₀ = α i0 i0 + + top : (l i j : I) → _ + top l i j = + hcomp + (λ k → λ + { (i = i0) → a₀ + ; (j = i0) → a₀ + ; (j = i1) → α (~ (i ∧ l)) k + ; (l = i0) → α i j + ; (l = i1) → α (~ i) (j ∧ k) + }) + (α (i ∧ ~ l) j) + + side : (l m j : I) → _ + side l m j = + hcomp + (λ k → λ + { (m = i1) → a₀ + ; (j = i0) → a₀ + ; (j = i1) → α (~ (m ∨ l)) k + ; (l = i0) → α (~ m) (j ∧ k) + ; (l = i1) → α m j + }) + (α (m ∨ ~ l) j) + +EH-funct : ∀ {ℓ} (p q : typ ((Ω^ 2) (S₊∙ 2))) → EH 0 (p ∙ q) (sym (p ∙ q)) ≡ {!EH 0 p (sym p)!} +EH-funct = {!!} + +zz : (p q r : typ (Ω (S₊∙ 2))) (b : r ≡ r) (t : p ∙ sym p ≡ sym p ∙ p) (l : r ≡ p) (r : r ≡ q) (P : PathP (λ i → {!!} ≡ {!p i!}) {!!} {!!}) + → Path (Square {!!} {!!} {!!} {!!}) (λ i j → P (i ∨ j) i j) {!!} +zz p q = {!!} + +Iso1 : Iso ∥ typ ((Ω^ 3) (S₊∙ 2)) ∥₂ ∥ typ ((Ω^ 2) (S₊∙ 2)) ∥₂ +fun Iso1 = {!!} +inv Iso1 = sMap hopfy +rightInv Iso1 = sElim {!!} λ p → cong ∣_∣₂ {!!} -- (l p) where - h : GroupHom (coHomGr (2 +ℕ n) TotalHopf) (coHomGr (3 +ℕ n) CP²) - h = M.d (2 +ℕ n) - - propᵣ : isProp (fst (×coHomGr (3 +ℕ n) (Susp S¹) Unit)) - propᵣ = - isPropΣ - (isOfHLevelRetractFromIso 1 - (fst (Hⁿ-Sᵐ≅0 (2 +ℕ n) 1 λ p → ⊥-rec (snotz (cong predℕ p)))) isPropUnit) - λ _ → isContr→isProp (isContrHⁿ-Unit _) - - propₗ : isProp (coHom (2 +ℕ n) TotalHopf) - propₗ = subst (λ x → isProp (coHom (2 +ℕ n) x)) (isoToPath IsoS³TotalHopf) - (isOfHLevelRetractFromIso 1 - (fst (Hⁿ-Sᵐ≅0 (suc n) 2 λ q → p (cong predℕ q))) isPropUnit) - - inIm : (x : coHom (3 +ℕ n) CP²) → isInIm h x - inIm x = M.Ker-i⊂Im-d (2 +ℕ n) x (propᵣ _ _) - - main : (x : coHom (3 +ℕ n) CP²) → x ≡ 0ₕ _ - main x = - pRec (squash₂ _ _) - (uncurry (λ f p → sym p ∙∙ cong (h .fst) (propₗ f (0ₕ _)) ∙∙ pres1 (snd h))) - (inIm x) - --- All trivial groups: -Hⁿ-CP²≅0 : (n : ℕ) → ¬ suc n ≡ 2 → ¬ suc n ≡ 4 - → GroupIso (coHomGr (suc n) CP²) UnitGroup -Hⁿ-CP²≅0 zero p q = H¹-CP²≅0 -Hⁿ-CP²≅0 (suc zero) p q = ⊥-rec (p refl) -Hⁿ-CP²≅0 (suc (suc zero)) p q = Hⁿ-CP²≅0-higher 0 λ p → snotz (sym p) -Hⁿ-CP²≅0 (suc (suc (suc zero))) p q = ⊥-rec (q refl) -Hⁿ-CP²≅0 (suc (suc (suc (suc n)))) p q = - Hⁿ-CP²≅0-higher (suc (suc n)) - λ p → snotz (cong predℕ p) - --- Another Brunerie number -ℤ→HⁿCP²→ℤ : ℤ → ℤ -ℤ→HⁿCP²→ℤ x = - Iso.fun (fst H⁴CP²≅ℤ) - (Iso.inv (fst H²CP²≅ℤ) x ⌣ Iso.inv (fst H²CP²≅ℤ) x) - -brunerie2 : ℤ -brunerie2 = ℤ→HⁿCP²→ℤ 1 - -{- -|brunerie2|≡1 : abs (ℤ→HⁿCP²→ℤ 1) ≡ 1 -|brunerie2|≡1 = refl --} + l : (p : _) → {!!} + l p k i j = + hcomp {!λ r → {(i = i0) → ? ; (i = i1) → ? ; (j = i0) → ? ; (j = i1) → ? ; (k = i1) → ? ; (k = i0) → ?}!} + {!!} +leftInv Iso1 = + {!k = i0 ⊢ ((λ i₁ → rCancel p (~ i₁)) ∙∙ EH 0 p (p ⁻¹) ∙∙ lCancel p) + (~ i ∨ j) (~ i ∧ j) j +k = i1 ⊢ p i j +i = i0 ⊢ snd (Ω (S₊∙ 2)) j +i = i1 ⊢ snd (Ω (S₊∙ 2)) j +j = i0 ⊢ snd (S₊∙ 2) +j = i1 ⊢ snd (S₊∙ 2)!} + +surf' : typ ((Ω^ 2) (S₊∙ 2)) +surf' i j = hcomp (λ k → λ { (i = i0) → north + ; (i = i1) → north + ; (j = i0) → rCancel (merid base) k i + ; (j = i1) → rCancel (merid base) k i}) + ((merid (loop j) ∙ sym (merid base)) i) + +test : ℤ +test = sRec isSetℤ (λ p → Iso.inv (Iso-Kn-ΩKn+1 0) (cong (Iso.inv (Iso-Kn-ΩKn+1 1)) (cong (cong ∣_∣ₕ) p))) + (Iso.fun Iso1 ∣ hopfy surf' ∣₂) + +-- fun1 : ∥ typ ((Ω^ 3) (S₊∙ 2)) ∥₂ → coHom 2 (Smash (S₊∙ 1) (S₊∙ 1)) +-- fun1 = {!!} + + +-- open import Cubical.Data.Bool +-- S¹* : Type +-- S¹* = Susp Bool + +-- S³* : Type +-- S³* = Susp (Susp (Susp Bool)) + +-- open import Cubical.HITs.S3 +-- open import Cubical.HITs.Join +-- S¹*S¹→S³ : join S¹ S¹ → S₊ 2 +-- S¹*S¹→S³ (inl x) = north +-- S¹*S¹→S³ (inr x) = north +-- S¹*S¹→S³ (push base b i) = north +-- S¹*S¹→S³ (push (loop j) base i) = north +-- S¹*S¹→S³ (push (loop i) (loop j) k) = (sym (rCancel surf') ∙∙ EH 0 surf' (sym surf') ∙∙ lCancel surf') k i j + +-- c : Iso S³* TotalHopf +-- fun c north = north , base +-- fun c south = north , base +-- fun c (merid a i) = h1 a i , S³≡S³ a i +-- where +-- h1 : (a : Susp (Susp Bool)) → Path (S₊ 2) north north +-- h1 north = refl +-- h1 south = refl +-- h1 (merid north i) j = (surf' ∙ sym (surf')) i j +-- h1 (merid south i) j = (sym surf' ∙ surf') i j +-- h1 (merid (merid false k) i) j = (rCancel surf' ∙ sym (lCancel surf')) k i j +-- h1 (merid (merid true k) i) j = EH 0 surf' (sym surf') k i j + +-- S³≡S³ : (a : Susp (Susp Bool)) → PathP (λ i → HopfSuspS¹ (h1 a i)) base base +-- S³≡S³ north = refl +-- S³≡S³ south = refl +-- S³≡S³ (merid a i) = {!!} -- help a i +-- where +-- sl : {!(j : I) → !} +-- sl = {!!} +-- l : (j : I) (x : Susp (Susp Bool)) (p : north ≡ x) → Path (PathP {!!} {!!} {!!}) (λ j → transp (λ i₂ → HopfSuspS¹ (h1 (p i₂) j)) i0 base) {!!} +-- l = {!!} + +-- s : (a : _) → (λ j → transp (λ i₂ → HopfSuspS¹ (h1 (merid a i₂) j)) i0 base) ≡ refl +-- s x = {!!} +-- help : (a : _) → SquareP (λ i j → HopfSuspS¹ (h1 (merid a i) j)) refl refl refl refl +-- help a = toPathP (transportRefl _ ∙ s a) + +-- h : (a : S₊ 2) → Path TotalHopf (north , base) (north , base) +-- h a = cong₂ _,_ refl {!!} +-- where +-- help : (a : S₊ 2) → PathP (λ i → Glue S¹ (Border base i)) base base +-- help = {!!} +-- inv c = {!!} +-- rightInv c = {!!} +-- leftInv c = {!!} + + +-- hej : S₊ 2 → Type₀ +-- hej north = S¹ +-- hej south = S¹ +-- hej (merid a i) = {!!} + +-- sauto : S¹ → I → S¹ +-- sauto base i = base +-- sauto (loop j) i = +-- hcomp (λ k → λ {(i = i0) → loop (j ∨ k) +-- ; (i = i1) → base +-- ; (j = i0) → loop (i ∨ k) +-- ; (j = i1) → base}) +-- (loop (i ∨ j)) + +-- open import Cubical.HITs.S1 renaming (_·_ to _*_) +-- S₊2→Type : S₊ 2 → Type +-- S₊2→Type north = typ (Ω (S₊∙ 2)) +-- S₊2→Type south = typ (Ω (S₊∙ 2)) +-- S₊2→Type (merid a i) = l i +-- where +-- h : (a : S¹) → isEquiv λ p → p ∙ merid (sauto a i) ∙ sym (merid base) +-- h = sphereElim 0 (λ _ → isPropIsEquiv _) (subst isEquiv (λ j p → ((cong (p ∙_) (rCancel (merid base))) ∙ sym (rUnit p)) (~ j)) +-- (idEquiv _ .snd)) + +-- l : typ (Ω (S₊∙ 2)) ≡ typ (Ω (S₊∙ 2)) +-- l = ua ((λ p → p ∙ merid (sauto a i) ∙ sym (merid base)) , h a) -- ua (a *_ , (rotIsEquiv a)) + +-- t2 : hLevelTrunc 3 ((typ ((Ω^ 2) (S₊∙ 2)))) → S¹ +-- t2 = {!!} + +-- test : ℤ +-- test = Iso.fun (fst (Hⁿ-Sⁿ≅ℤ 1)) ∣ {!!} ∣₂ + +-- S2-add : hLevelTrunc 5 (S₊ 2) → hLevelTrunc 5 (S₊ 2) → hLevelTrunc 4 (S₊ 2) +-- S2-add = trRec2 {!!} (wedgeconFun 1 1 (λ _ _ → isOfHLevelTrunc {!!}) {!!} {!!} {!!}) + +-- CP² : Type +-- CP² = Pushout {A = TotalHopf} fst λ _ → tt + +-- characFunSpaceCP² : ∀ {ℓ} {A : Type ℓ} +-- → Iso (CP² → A) (Σ[ x ∈ A ] Σ[ f ∈ (S₊ (suc (suc zero)) → A) ] +-- ((y : TotalHopf) → f (fst y) ≡ x)) +-- fun characFunSpaceCP² f = (f (inr tt)) , ((f ∘ inl ) , (λ a → cong f (push a))) +-- inv characFunSpaceCP² (a , f , p) (inl x) = f x +-- inv characFunSpaceCP² (a , f , p) (inr x) = a +-- inv characFunSpaceCP² (a , f , p) (push b i) = p b i +-- rightInv characFunSpaceCP² _ = refl +-- leftInv characFunSpaceCP² _ = +-- funExt λ { (inl x) → refl +-- ; (inr x) → refl +-- ; (push a i) → refl} + +-- H⁰CP²≅ℤ : GroupIso (coHomGr 0 CP²) ℤGroup +-- H⁰CP²≅ℤ = +-- H⁰-connected (inr tt) +-- (PushoutToProp (λ _ → squash) +-- (sphereElim _ (λ _ → isOfHLevelSuc 1 squash) +-- ∣ sym (push (north , base)) ∣₁) +-- λ _ → ∣ refl ∣₁) + +-- module M = MV (S₊ 2) Unit TotalHopf fst (λ _ → tt) + +-- H²CP²≅ℤ : GroupIso (coHomGr 2 CP²) ℤGroup +-- H²CP²≅ℤ = compGroupIso (BijectionIso→GroupIso bij) +-- (compGroupIso (invGroupIso trivIso) (Hⁿ-Sⁿ≅ℤ 1)) +-- where +-- isContrH¹TotalHopf : isContr (coHom 1 TotalHopf) +-- isContrH¹TotalHopf = +-- isOfHLevelRetractFromIso 0 (setTruncIso (domIso (invIso (IsoS³TotalHopf)))) +-- (isOfHLevelRetractFromIso 0 ((fst (H¹-Sⁿ≅0 1))) isContrUnit) + +-- isContrH²TotalHopf : isContr (coHom 2 TotalHopf) +-- isContrH²TotalHopf = +-- isOfHLevelRetractFromIso 0 (setTruncIso (domIso (invIso (IsoS³TotalHopf)))) +-- ((isOfHLevelRetractFromIso 0 +-- (fst (Hⁿ-Sᵐ≅0 1 2 λ p → snotz (sym (cong predℕ p)))) isContrUnit)) + +-- trivIso : GroupIso (coHomGr 2 (Susp S¹)) (×coHomGr 2 (Susp S¹) Unit) +-- fun (fst trivIso) x = x , 0ₕ _ +-- inv (fst trivIso) = fst +-- rightInv (fst trivIso) (x , p) = +-- ΣPathP (refl , isContr→isProp (isContrHⁿ-Unit 1) _ _) +-- leftInv (fst trivIso) x = refl +-- snd trivIso = makeIsGroupHom λ _ _ → refl + +-- bij : BijectionIso (coHomGr 2 CP²) (×coHomGr 2 (Susp S¹) Unit) +-- BijectionIso.fun bij = M.i 2 +-- BijectionIso.inj bij x p = +-- pRec (squash₂ _ _) +-- (uncurry (λ z q +-- → sym q +-- ∙∙ cong (fst (M.d 1)) (isContr→isProp isContrH¹TotalHopf z (0ₕ _)) +-- ∙∙ (M.d 1) .snd .pres1)) +-- (M.Ker-i⊂Im-d 1 x p) +-- where +-- help : isInIm (M.d 1) x +-- help = M.Ker-i⊂Im-d 1 x p +-- BijectionIso.surj bij y = +-- M.Ker-Δ⊂Im-i 2 y (isContr→isProp isContrH²TotalHopf _ _) + +-- H⁴CP²≅ℤ : GroupIso (coHomGr 4 CP²) ℤGroup +-- H⁴CP²≅ℤ = compGroupIso (invGroupIso (BijectionIso→GroupIso bij)) +-- (compGroupIso help (Hⁿ-Sⁿ≅ℤ 2)) +-- where +-- help : GroupIso (coHomGr 3 TotalHopf) (coHomGr 3 (S₊ 3)) +-- help = isoType→isoCohom 3 (invIso IsoS³TotalHopf) + +-- bij : BijectionIso (coHomGr 3 TotalHopf) (coHomGr 4 CP²) +-- BijectionIso.fun bij = M.d 3 +-- BijectionIso.inj bij x p = +-- pRec (squash₂ _ _) +-- (uncurry (λ z q → +-- sym q +-- ∙∙ cong (M.Δ 3 .fst) +-- (isOfHLevelΣ 1 (isContr→isProp +-- (isOfHLevelRetractFromIso 0 +-- (fst (Hⁿ-Sᵐ≅0 2 1 λ p → snotz (cong predℕ p))) isContrUnit)) +-- (λ _ → isContr→isProp (isContrHⁿ-Unit 2)) +-- z (0ₕ _ , 0ₕ _)) +-- ∙∙ M.Δ 3 .snd .pres1)) +-- (M.Ker-d⊂Im-Δ _ x p) +-- BijectionIso.surj bij y = +-- M.Ker-i⊂Im-d 3 y (isOfHLevelΣ 1 +-- (isContr→isProp (isOfHLevelRetractFromIso 0 +-- (fst (Hⁿ-Sᵐ≅0 3 1 λ p → snotz (cong predℕ p))) isContrUnit)) +-- (λ _ → isContr→isProp (isContrHⁿ-Unit _)) _ _) + + +-- -- Characterisations of the trivial groups + +-- private +-- elim-TotalHopf : (B : TotalHopf → Type) +-- → ((x : _) → isOfHLevel 3 (B x)) → B (north , base) +-- → (x : _) → B x +-- elim-TotalHopf = +-- transport (λ i → (B : isoToPath IsoS³TotalHopf i → Type) +-- → ((x : _) → isOfHLevel 3 (B x)) +-- → B (transp (λ j → isoToPath IsoS³TotalHopf (i ∨ ~ j)) i (north , base)) → (x : _) → B x) +-- λ B hLev elim-TotalHopf → sphereElim _ (λ _ → hLev _) elim-TotalHopf + +-- H¹-CP²≅0 : GroupIso (coHomGr 1 CP²) UnitGroup +-- H¹-CP²≅0 = +-- contrGroupIsoUnit +-- (isOfHLevelRetractFromIso 0 (setTruncIso characFunSpaceCP²) +-- (isOfHLevelRetractFromIso 0 lem₂ lem₃)) +-- where +-- lem₁ : (f : (Susp S¹ → coHomK 1)) → ∥ (λ _ → 0ₖ _) ≡ f ∥ +-- lem₁ f = pMap (λ p → p) +-- (Iso.fun PathIdTrunc₀Iso (isOfHLevelRetractFromIso 1 +-- (fst (Hⁿ-Sᵐ≅0 0 1 (λ p → snotz (sym p)))) isPropUnit (0ₕ _) ∣ f ∣₂)) + +-- lem₂ : Iso ∥ (Σ[ x ∈ coHomK 1 ] ( Σ[ f ∈ (Susp S¹ → coHomK 1) ] ((y : TotalHopf) → f (fst y) ≡ x))) ∥₂ +-- ∥ (Σ[ f ∈ (Susp S¹ → coHomK 1) ] ((y : TotalHopf) → f (fst y) ≡ 0ₖ 1)) ∥₂ +-- fun lem₂ = sMap (uncurry λ x → uncurry λ f p → (λ y → (-ₖ x) +ₖ f y) , λ y → cong ((-ₖ x) +ₖ_) (p y) ∙ lCancelₖ _ x) +-- inv lem₂ = sMap λ p → 0ₖ _ , p +-- rightInv lem₂ = +-- sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) +-- λ {(f , p) → cong ∣_∣₂ (ΣPathP ((funExt (λ x → lUnitₖ _ (f x))) +-- , (funExt (λ y → sym (rUnit (λ i → (-ₖ 0ₖ 1) +ₖ p y i))) +-- ◁ λ j y i → lUnitₖ _ (p y i) j)))} +-- leftInv lem₂ = +-- sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) +-- (uncurry (coHomK-elim _ (λ _ → isPropΠ (λ _ → squash₂ _ _)) +-- (uncurry λ f p → cong ∣_∣₂ (ΣPathP (refl , (ΣPathP ((funExt (λ x → lUnitₖ _ (f x))) +-- , ((funExt (λ y → sym (rUnit (λ i → (-ₖ 0ₖ 1) +ₖ p y i))) +-- ◁ λ j y i → lUnitₖ _ (p y i) j))))))))) + +-- lem₃ : isContr _ +-- fst lem₃ = ∣ (λ _ → 0ₖ 1) , (λ _ → refl) ∣₂ +-- snd lem₃ = +-- sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) +-- (uncurry λ f → pRec (isPropΠ (λ _ → squash₂ _ _)) +-- (J (λ f _ → (y : (y₁ : TotalHopf) → f (fst y₁) ≡ 0ₖ 1) → +-- ∣ (λ _ → 0ₖ 1) , (λ _ _ → 0ₖ 1) ∣₂ ≡ ∣ f , y ∣₂) +-- (λ y → cong ∣_∣₂ (ΣPathP ((funExt (λ z → sym (y (north , base)))) , toPathP (s y))))) +-- (lem₁ f)) + +-- where +-- s : (y : TotalHopf → 0ₖ 1 ≡ 0ₖ 1) +-- → transport (λ i → (_ : TotalHopf) → y (north , base) (~ i) ≡ ∣ base ∣) +-- (λ _ _ → 0ₖ 1) ≡ y +-- s y = funExt (elim-TotalHopf _ (λ _ → isOfHLevelPath 3 (isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) _ _) +-- λ k → transp (λ i → y (north , base) (~ i ∧ ~ k) ≡ ∣ base ∣) k +-- λ j → y (north , base) (~ k ∨ j)) + +-- Hⁿ-CP²≅0-higher : (n : ℕ) → ¬ (n ≡ 1) → GroupIso (coHomGr (3 +ℕ n) CP²) UnitGroup +-- Hⁿ-CP²≅0-higher n p = contrGroupIsoUnit ((0ₕ _) , (λ x → sym (main x))) +-- where +-- h : GroupHom (coHomGr (2 +ℕ n) TotalHopf) (coHomGr (3 +ℕ n) CP²) +-- h = M.d (2 +ℕ n) + +-- propᵣ : isProp (fst (×coHomGr (3 +ℕ n) (Susp S¹) Unit)) +-- propᵣ = +-- isPropΣ +-- (isOfHLevelRetractFromIso 1 +-- (fst (Hⁿ-Sᵐ≅0 (2 +ℕ n) 1 λ p → ⊥-rec (snotz (cong predℕ p)))) isPropUnit) +-- λ _ → isContr→isProp (isContrHⁿ-Unit _) + +-- propₗ : isProp (coHom (2 +ℕ n) TotalHopf) +-- propₗ = subst (λ x → isProp (coHom (2 +ℕ n) x)) (isoToPath IsoS³TotalHopf) +-- (isOfHLevelRetractFromIso 1 +-- (fst (Hⁿ-Sᵐ≅0 (suc n) 2 λ q → p (cong predℕ q))) isPropUnit) + +-- inIm : (x : coHom (3 +ℕ n) CP²) → isInIm h x +-- inIm x = M.Ker-i⊂Im-d (2 +ℕ n) x (propᵣ _ _) + +-- main : (x : coHom (3 +ℕ n) CP²) → x ≡ 0ₕ _ +-- main x = +-- pRec (squash₂ _ _) +-- (uncurry (λ f p → sym p ∙∙ cong (h .fst) (propₗ f (0ₕ _)) ∙∙ pres1 (snd h))) +-- (inIm x) + +-- -- All trivial groups: +-- Hⁿ-CP²≅0 : (n : ℕ) → ¬ suc n ≡ 2 → ¬ suc n ≡ 4 +-- → GroupIso (coHomGr (suc n) CP²) UnitGroup +-- Hⁿ-CP²≅0 zero p q = H¹-CP²≅0 +-- Hⁿ-CP²≅0 (suc zero) p q = ⊥-rec (p refl) +-- Hⁿ-CP²≅0 (suc (suc zero)) p q = Hⁿ-CP²≅0-higher 0 λ p → snotz (sym p) +-- Hⁿ-CP²≅0 (suc (suc (suc zero))) p q = ⊥-rec (q refl) +-- Hⁿ-CP²≅0 (suc (suc (suc (suc n)))) p q = +-- Hⁿ-CP²≅0-higher (suc (suc n)) +-- λ p → snotz (cong predℕ p) + +-- -- Another Brunerie number +-- ℤ→HⁿCP²→ℤ : ℤ → ℤ +-- ℤ→HⁿCP²→ℤ x = +-- Iso.fun (fst H⁴CP²≅ℤ) +-- (Iso.inv (fst H²CP²≅ℤ) x ⌣ Iso.inv (fst H²CP²≅ℤ) x) + +-- brunerie2 : ℤ +-- brunerie2 = ℤ→HⁿCP²→ℤ 1 + +-- {- +-- |brunerie2|≡1 : abs (ℤ→HⁿCP²→ℤ 1) ≡ 1 +-- |brunerie2|≡1 = refl +-- -} + +-- t : ∀ (a : TotalHopf) → ∣ fst a ∣ ≡ 0ₖ 2 +-- t = {!!} + +-- one : coHom 2 CP² +-- one = ∣ (λ { (inl x) → ∣ x ∣ +-- ; (inr x) → 0ₖ _ +-- ; (push a i) → t a i}) ∣₂ + + +-- CP²→CP⁴ : (x : coHom 2 CP²) → coHom 2 CP² → coHom 4 CP² +-- CP²→CP⁴ x f = f ⌣ x + +-- zz : Iso.fun (fst H⁴CP²≅ℤ) (CP²→CP⁴ one one) ≡ 1 +-- zz = {!pres1!} ∙∙ {!!} ∙∙ {!!} +-- open import Cubical.ZCohomology.Properties +-- CP4→CP2 : (x : coHom 2 CP²) → coHom 4 CP² → coHom 2 CP² +-- CP4→CP2 = sRec2 squash₂ +-- λ e f → ∣ (λ { (inl x) → {!!} +-- ; (inr x) → {!!} +-- ; (push a i) → {!!}}) ∣₂ + + +-- CP' : (f : S₊∙ 3 →∙ S₊∙ 2) → Type₀ +-- CP' f = Pushout {A = S₊ 3} (fst f) λ _ → tt + +-- open import Cubical.Homotopy.Loopspace + +-- contrH2 : (f : S₊ 2 → coHomK 4) → ∥ (λ x → 0ₖ 4) ≡ f ∥₂ +-- contrH2 = {!!} + +-- CP4→CP2'' : (f : S₊∙ 3 →∙ S₊∙ 2) → ∥ (Σ[ g ∈ (S₊ 2 → coHomK 4) ] ((a : S₊ 3) → g (fst f a) ≡ 0ₖ 4)) ∥₂ +-- → ∥ ((S₊∙ 3 →∙ (Ω (coHomK-ptd 4)))) ∥₂ +-- CP4→CP2'' (f , p) = +-- sRec squash₂ +-- (uncurry λ g → +-- sRec (isSetΠ (λ _ → squash₂)) (J (λ g _ → (y : (a : Susp (Susp S¹)) → g (f a) ≡ 0ₖ 4) +-- → ∥ ((S₊∙ 3 →∙ (Ω (coHomK-ptd 4)))) ∥₂) +-- λ p → ∣ (λ a → p a ∙ sym (p north)) , rCancel (p north) ∣₂) +-- (contrH2 g)) + +-- CP²' = CP' (fst ∘ Iso.fun (IsoS³TotalHopf) , refl) + +-- gen0 : CP²' → coHomK 2 +-- gen0 = {!!} + +-- gen1 : CP²' → coHomK 4 +-- gen1 (inl x) = 0ₖ 4 +-- gen1 (inr x) = 0ₖ 4 +-- gen1 (push a i) = Kn→ΩKn+1 3 ∣ a ∣ i + + +-- hehehe : S₊ 3 → S₊ 2 +-- hehehe north = north +-- hehehe south = north +-- hehehe (merid north i) = north +-- hehehe (merid south i) = north +-- hehehe (merid (merid base i₁) i) = north +-- hehehe (merid (merid (loop k) j) i) = +-- (sym (rCancel surf') ∙∙ EH 0 surf' (sym surf') ∙∙ lCancel surf') i j k + +-- wtf : (x : S₊ 2) → Σ[ y ∈ S₊ 3 ] hehehe y ≡ x +-- wtf north = north , refl +-- wtf south = north , merid base +-- wtf (merid base i) = north , λ j → (merid base (i ∧ j)) +-- wtf (merid (loop j) i) = sq i j , {!!} +-- where +-- sq : Square {A = S₊ 3} (refl {x = north}) refl refl refl +-- sq i j = hcomp (λ k → λ {(i = i0) → north +-- ; (i = i1) → merid north (~ k) +-- ; (j = i0) → merid north (~ k ∧ i) +-- ; (j = i1) → merid {!!} (i ∧ ~ k)}) +-- (merid (merid (loop j) j) i) + + +-- -- sl : S₊ 2 → Type +-- -- sl north = S¹ +-- -- sl south = S¹ +-- -- sl (merid base i) = S¹ +-- -- sl (merid (loop j) i) = h i j +-- -- where +-- -- h : Square (refl {x = S¹}) (λ _ → S¹) (λ _ → S¹) λ _ → S¹ +-- -- h i j = Glue {!HopfSuspS¹ (merid base i)!} +-- -- λ { (i = i0) → S¹ , {!!} +-- -- ; (i = i1) → {!!} +-- -- ; (j = i0) → {!!} +-- -- ; (j = i1) → {!!}} + +-- -- CP7 : Type₀ +-- -- CP7 = Pushout hehehe λ _ → tt + +-- -- indMap : typ ((Ω^ 4) (coHomK-ptd 4)) → CP7 → coHomK 4 +-- -- indMap P (inl x) = 0ₖ 4 +-- -- indMap P (inr x) = 0ₖ 4 +-- -- indMap P (push north i) = 0ₖ 4 +-- -- indMap P (push south i) = 0ₖ 4 +-- -- indMap P (push (merid north j) i) = 0ₖ 4 +-- -- indMap P (push (merid south j) i) = 0ₖ 4 +-- -- indMap P (push (merid (merid base i₁) j) i) = 0ₖ 4 +-- -- indMap P (push (merid (merid (loop l) k) j) i) = P i j k l + +-- -- ss : Iso ∥ (CP7 → coHomK 4) ∥₂ ∥ (Σ[ x ∈ coHomK 4 ] Σ[ f ∈ (S₊ 2 → coHomK 4) ] ((a : S₊ 3) → f (hehehe a) ≡ x)) ∥₂ +-- -- fun ss = sMap λ f → (f (inr tt)) , ((f ∘ inl) , (λ a → cong f (push a))) +-- -- inv ss = sMap λ {(pt , f , p) → λ { (inl x) → f x ; (inr x) → pt ; (push a i) → p a i}} +-- -- rightInv ss = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) (λ x → refl) +-- -- leftInv ss = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) +-- -- λ f → cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i) → refl}) + +-- -- ss2 : Iso (∥ (Σ[ x ∈ coHomK 4 ] Σ[ f ∈ (S₊ 2 → coHomK 4) ] ((a : S₊ 3) → f (hehehe a) ≡ x)) ∥₂) +-- -- ∥ (Σ[ f ∈ (S₊ 2 → coHomK 4) ] ((a : S₊ 3) → f (hehehe a) ≡ 0ₖ 4)) ∥₂ +-- -- fun ss2 = +-- -- sRec squash₂ (uncurry (coHomK-elim 3 (λ _ → isOfHLevelΠ 4 λ _ → isOfHLevelSuc 3 (isOfHLevelSuc 2 squash₂)) +-- -- ∣_∣₂)) +-- -- inv ss2 = sMap λ p → (0ₖ 4 , p) +-- -- rightInv ss2 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl +-- -- leftInv ss2 = +-- -- sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) +-- -- (uncurry (coHomK-elim 3 +-- -- (λ _ → isOfHLevelΠ 4 +-- -- λ _ → isOfHLevelPath 4 +-- -- (isOfHLevelSuc 3 (isOfHLevelSuc 2 squash₂)) _ _) λ _ → refl)) +-- -- open import Cubical.Homotopy.Connected + +-- -- truncLem : (f : (S₊ 2 → coHomK 4)) → ∥ (λ x → 0ₖ 4) ≡ f ∥₂ +-- -- truncLem f = +-- -- sRec squash₂ +-- -- (λ hs → sRec squash₂ +-- -- (λ h → sRec squash₂ (λ m-b → sRec squash₂ (λ m-l → ∣ funExt (λ { north → h +-- -- ; south → hs +-- -- ; (merid base i) → m-b i +-- -- ; (merid (loop j) i) k → m-l j i k} ) ∣₂) +-- -- (final (f north) (f south) h hs (cong f (merid base)) m-b (λ j i → f (merid (loop j) i)))) +-- -- (h-m-b (f north) (f south) h hs (cong f (merid base)))) +-- -- (h (f north))) (h (f south)) +-- -- where +-- -- h : (x : coHomK 4) → ∥ 0ₖ 4 ≡ x ∥₂ +-- -- h = coHomK-elim _ (λ _ → isOfHLevelSuc 3 (isOfHLevelSuc 2 squash₂)) ∣ refl ∣₂ + +-- -- stupid : isConnected (1 +ℕ 4) (coHomK 4) +-- -- stupid = isConnectedKn 3 + +-- -- h-m-b : (x y : coHomK 4) → (p : 0ₖ 4 ≡ x) (q : 0ₖ 4 ≡ y) (z : x ≡ y) → ∥ PathP (λ i → 0ₖ 4 ≡ z i) p q ∥₂ +-- -- h-m-b x y p q z = trElim {B = λ _ → ∥ PathP (λ i → 0ₖ 4 ≡ z i) p q ∥₂} (λ _ → squash₂) ∣_∣₂ +-- -- (isConnectedPathP 2 {A = λ i → 0ₖ 4 ≡ z i} (isConnectedPath 3 (isConnectedSubtr 4 1 stupid) _ _) p q .fst) + +-- -- final : (x y : coHomK 4) → (p : 0ₖ 4 ≡ x) (q : 0ₖ 4 ≡ y) (z : x ≡ y) (w : PathP (λ i → 0ₖ 4 ≡ z i) p q) (r : z ≡ z) +-- -- → ∥ PathP (λ j → PathP (λ i → 0ₖ 4 ≡ r j i) p q) w w ∥₂ +-- -- final x y p q z w r = +-- -- trRec {B = ∥ PathP (λ j → PathP (λ i → 0ₖ 4 ≡ r j i) p q) w w ∥₂} +-- -- squash₂ +-- -- ∣_∣₂ +-- -- (isConnectedPathP 2 {A = (λ j → PathP (λ i → 0ₖ 4 ≡ r j i) p q) } +-- -- (isConnectedPathP 3 (isConnectedPath 4 (isConnectedKn 3) _ _) _ _) w w .fst) + +-- -- truncLemHLev : (f : (S₊ 2 → coHomK 4)) → isContr ∥ (λ x → 0ₖ 4) ≡ f ∥₂ +-- -- truncLemHLev f = +-- -- isOfHLevelRetractFromIso 0 +-- -- (compIso setTruncTrunc2Iso (invIso (PathIdTruncIso 2))) +-- -- (isOfHLevelPath 0 (∣ (λ _ → 0ₖ 4) ∣ , (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ f → sRec (isOfHLevelTrunc 3 _ _) (cong ∣_∣) (truncLem f))) _ _) + +-- -- FF : (f : (S₊ 2 → coHomK 4)) → ∥ (λ _ → 0ₖ 4) ≡ f ∥₂ → ((a : S₊ 3) → f (hehehe a) ≡ 0ₖ 4) → ∥ (S₊ 3 → 0ₖ 4 ≡ 0ₖ 4) ∥₂ +-- -- FF f = sRec (isSetΠ (λ _ → squash₂)) λ id p → ∣ (λ x → funExt⁻ id (hehehe x) ∙ (p x)) ∣₂ + +-- -- sll : (S₊ 3 → 0ₖ 4 ≡ 0ₖ 4) → CP7 → coHomK 4 +-- -- sll f (inl x) = 0ₖ 4 +-- -- sll f (inr x) = 0ₖ 4 +-- -- sll f (push a i) = f a i + +-- -- ssl* : ∥ (S₊ 3 → 0ₖ 4 ≡ 0ₖ 4) ∥₂ → ∥ (CP7 → coHomK 4) ∥₂ +-- -- ssl* = sMap sll + +-- -- ss3 : Iso ∥ (Σ[ f ∈ (S₊ 2 → coHomK 4) ] ((a : S₊ 3) → f (hehehe a) ≡ 0ₖ 4)) ∥₂ ∥ (S₊ 3 → 0ₖ 4 ≡ 0ₖ 4) ∥₂ +-- -- fun ss3 = sRec squash₂ (uncurry (λ f → FF f (truncLem f))) +-- -- inv ss3 = sMap λ f → (λ _ → 0ₖ 4) , f +-- -- rightInv ss3 = sElim {!!} λ f → {!λ i → FF (λ _ → 0ₖ 4) (h i) f!} ∙ {!!} +-- -- where +-- -- h : truncLem (λ _ → 0ₖ _) ≡ ∣ refl ∣₂ +-- -- h = {!cong ∣_∣₂ ?!} +-- -- leftInv ss3 = {!!} + +-- -- s : ∀ (f : CP7 → coHomK 4) → Σ[ P ∈ (S₊∙ 3 →∙ Ω (coHomK-ptd 4)) ] +-- -- ∣ f ∣₂ ≡ ∣ sll (fst P) ∣₂ +-- -- s f = (transport (λ i → Σ[ P ∈ (S₊∙ 3 →∙ Ω (coHomK-ptd 4)) ] +-- -- Iso.leftInv myIso ∣ f ∣₂ i ≡ ∣ sll (fst P) ∣₂) +-- -- ((sRec l (λ f → (λ x → f x ∙ sym (f north)) , rCancel (f north)) (Iso.fun myIso ∣ f ∣₂)) , +-- -- mainLem (Iso.fun myIso ∣ f ∣₂))) +-- -- where +-- -- myIso : Iso _ _ +-- -- myIso = compIso ss (compIso ss2 ss3) + +-- -- ll : (f : Susp (Susp S¹) → 0ₖ 4 ≡ 0ₖ 4) → ∥ f north ≡ refl ∥₂ +-- -- ll f = trRec squash₂ ∣_∣₂ +-- -- (isConnectedPath 2 (isConnectedSubtr 3 1 (isConnectedPathKn 3 (0ₖ 4) (0ₖ 4))) (f north) refl .fst) + +-- -- l : isSet (S₊∙ 3 →∙ Ω (coHomK-ptd 4)) +-- -- l = subst isSet (λ i → S₊∙ 3 →∙ Kn≃ΩKn+1∙ {n = 3} i) +-- -- (isOfHLevel↑∙' 0 2) + +-- -- kaha : (f : _) → inv myIso ∣ f ∣₂ ≡ ∣ sll f ∣₂ +-- -- kaha f = cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i) → refl}) + +-- -- mainLem : (f : _) → inv myIso f ≡ +-- -- ∣ sll (fst (sRec l (λ f₁ → (λ x → f₁ x ∙ (λ i → f₁ north (~ i))) , rCancel (f₁ (snd (S₊∙ 3)))) +-- -- f)) ∣₂ +-- -- mainLem = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) +-- -- λ f → sRec (isOfHLevelPath 2 squash₂ _ _) +-- -- (λ p → kaha f ∙ cong ∣_∣₂ (cong sll ((λ i x → rUnit (f x) i) ∙ λ i x → f x ∙ sym (p (~ i))))) +-- -- (ll f) + +-- -- genCP7 : coHom 4 CP7 +-- -- genCP7 = ∣ (λ { (inl x) → 0ₖ 4 +-- -- ; (inr x) → 0ₖ 4 +-- -- ; (push a i) → Kn→ΩKn+1 3 ∣ a ∣ i}) ∣₂ + +-- -- heheid : (a : _) → ∣ hehehe a ∣ ≡ 0ₖ (suc (suc zero)) +-- -- heheid = sphereElim _ (λ _ → isOfHLevelTrunc 4 _ _) refl + +-- -- gen2CP7 : coHom 4 CP7 +-- -- gen2CP7 = ∣ (λ { (inl x) → _⌣ₖ_ {n = 2} {m = 2} ∣ x ∣ ∣ x ∣ +-- -- ; (inr x) → 0ₖ 4 +-- -- ; (push a i) → heheid a i ⌣ₖ heheid a i}) ∣₂ + + +-- -- testS : gen2CP7 ≡ genCP7 +-- -- testS = +-- -- cong ∣_∣₂ +-- -- (funExt λ { (inl north) → refl +-- -- ; (inl south) → refl +-- -- ; (inl (merid a i)) j → h a j i +-- -- ; (inr x) → {!!} +-- -- ; (push a i) → {!!}}) +-- -- where +-- -- h : (a : S¹) → cong₂ (_⌣ₖ_ {n = 2} {m = 2}) (cong ∣_∣ₕ (merid a)) (cong ∣_∣ₕ (merid a)) ≡ λ i → 0ₖ 4 +-- -- h = {!!} + +-- -- gen2 : coHom 2 CP7 +-- -- gen2 = ∣ (λ { (inl x) → ∣ x ∣ ; (inr x) → 0ₖ 2 ; (push a i) → heheid a i}) ∣₂ + +-- -- open import Cubical.HITs.S1 renaming (_·_ to _*_) +-- -- hopfSurj' : (x : S₊ 2) → HopfSuspS¹ x +-- -- hopfSurj' north = base +-- -- hopfSurj' south = base +-- -- hopfSurj' (merid a i) = {!!} +-- -- where +-- -- k : (a : _) → PathP (λ i → Glue S¹ (Border a i)) base base +-- -- k a = toPathP ((λ i → a * base) ∙ {!!}) + +-- -- -- hopfSurj : (x : S₊ 2) → Σ[ s ∈ S₊ 3 ] fst (JoinS¹S¹→TotalHopf (S³→joinS¹S¹ (inv IsoS³S3 s))) ≡ x +-- -- -- hopfSurj x = (Iso.inv (IsoS³TotalHopf) (x , l x)) , cong fst (Iso.rightInv (IsoS³TotalHopf) (x , l x)) + +-- -- -- gen2 : (p : (a : S₊ 3) → ∣ fst (JoinS¹S¹→TotalHopf (S³→joinS¹S¹ (inv IsoS³S3 a))) ∣ ≡ 0ₖ 2) +-- -- -- → CP²' → coHomK 4 +-- -- -- gen2 p (inl x) = _⌣ₖ_ {n = 2} {m = 2} ∣ x ∣ ∣ x ∣ +-- -- -- gen2 p (inr x) = 0ₖ _ +-- -- -- gen2 p (push a i) = p a i ⌣ₖ p a i + +-- -- -- gen≡ : (p : _) (a : _) → gen1 a ≡ gen2 p a +-- -- -- gen≡ p (inl x) = (λ i → p (fst (hopfSurj x)) (~ i) ⌣ₖ p (fst (hopfSurj x)) (~ i)) +-- -- -- ∙ λ i → _⌣ₖ_ {n = 2} {m = 2} ∣ hopfSurj x .snd i ∣ ∣ hopfSurj x .snd i ∣ +-- -- -- gen≡ p (inr x) = refl +-- -- -- gen≡ p (push a i) = {!!} + + +-- -- -- -- CP4→CP2''' : (f : S₊∙ 3 →∙ S₊∙ 2) → ∥ (S₊∙ 3 →∙ (Ω (coHomK-ptd 4))) ∥₂ → coHom 2 (CP' f) +-- -- -- -- CP4→CP2''' f = sRec squash₂ λ f → ∣ (λ { (inl x) → ΩKn+1→Kn 2 (cong (ΩKn+1→Kn 3) (h x f)) +-- -- -- -- ; (inr x) → 0ₖ 2 +-- -- -- -- ; (push a i) → {!!}}) ∣₂ +-- -- -- -- where +-- -- -- -- h : (x : S₊ 2) (f : ((S₊∙ 3 →∙ (Ω (coHomK-ptd 4))))) → typ ((Ω^ 2) (coHomK-ptd 4)) +-- -- -- -- h x f i j = +-- -- -- -- hcomp (λ k → λ { (i = i0) → 0ₖ 4 +-- -- -- -- ; (i = i1) → 0ₖ 4 +-- -- -- -- ; (j = i0) → snd f k i +-- -- -- -- ; (j = i1) → snd f k i}) +-- -- -- -- (fst f ((merid x ∙ sym (merid north)) j) i) + +-- -- -- -- CP4→CP2' : (f : S₊∙ 3 →∙ S₊∙ 2) (x : coHom 2 (CP' f)) → coHom 4 (CP' f) → coHom 2 (CP' f) +-- -- -- -- CP4→CP2' f = +-- -- -- -- sRec2 squash₂ +-- -- -- -- λ e g → ∣ (λ { (inl x) → ΩKn+1→Kn 2 ({!g!} ∙∙ cong (ΩKn+1→Kn 3) {!λ i j → g (push (merid x i)) j!} ∙∙ {!!}) +-- -- -- -- ; (inr x) → {!!} +-- -- -- -- ; (push a i) → {!!}}) ∣₂ +-- -- -- -- where +-- -- -- -- h : (g : CP' f → coHomK 4) → g (inr tt) ≡ 0ₖ _ → g ∘ inl ≡ (λ x → 0ₖ _) → S₊ 2 → typ ((Ω^ 2) (coHomK-ptd 4)) +-- -- -- -- h g gpd idf x i j = +-- -- -- -- hcomp (λ k → λ { (i = i0) → rCancelₖ 4 (g (push north j)) k +-- -- -- -- ; (i = i1) → rCancelₖ 4 (g (push south j)) k +-- -- -- -- ; (j = i0) → {!idf k ((fst f (merid x i))) -ₖ idf k ((fst f (merid north i)))!} -- idf k ((fst f (merid x i))) -ₖ gpd k +-- -- -- -- ; (j = i1) → {!!}}) +-- -- -- -- (g (push (merid x i) j) -ₖ g (push (merid north i) j)) + + +-- -- contrfstIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} → (s : isContr A) → Iso (Σ A B) (B (s .fst)) +-- -- fun (contrfstIso {B = B} s) (a , b) = subst B (sym (snd s a)) b +-- -- inv (contrfstIso s) b = (fst s) , b +-- -- rightInv (contrfstIso {B = B} s) b = (λ i → subst B (sym (h i)) b) ∙ transportRefl b +-- -- where +-- -- h : snd s (fst s) ≡ refl +-- -- h = isProp→isSet (isContr→isProp s) _ _ _ _ +-- -- leftInv (contrfstIso{B = B} s) (a , b) = +-- -- ΣPathP ((snd s a) , λ i → transp (λ j → B (snd s a (i ∨ ~ j))) i b) + +-- -- miniFib : hLevelTrunc 3 (S₊ 2 → coHomK 4) → Type +-- -- miniFib x = trRec {B = TypeOfHLevel ℓ-zero 2} (isOfHLevelTypeOfHLevel 2) (λ f → (∥ ((x : S₊ 3) → f (hehehe x) ≡ 0ₖ 4) ∥₂) , squash₂) x .fst + +-- -- l3 : isContr (hLevelTrunc 3 (S₊ 2 → coHomK 4)) +-- -- l3 = {!!} + +-- -- l7 : {!!} +-- -- l7 = {!!} diff --git a/Cubical/ZCohomology/Gysin.agda b/Cubical/ZCohomology/Gysin.agda new file mode 100644 index 000000000..c6495485b --- /dev/null +++ b/Cubical/ZCohomology/Gysin.agda @@ -0,0 +1,765 @@ +{-# OPTIONS --safe --experimental-lossy-unification --no-forcing #-} +module Cubical.ZCohomology.Gysin where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Groups.Connected +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.MayerVietorisUnreduced +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.RingStructure.CupProduct +open import Cubical.ZCohomology.RingStructure.RingLaws +open import Cubical.ZCohomology.RingStructure.GradedCommutativity + +open import Cubical.Functions.Embedding + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Pointed.Homogeneous + +open import Cubical.Data.Empty renaming (rec to ⊥-rec) +open import Cubical.Data.Sigma +open import Cubical.Data.Int hiding (_+'_) +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Unit +open import Cubical.Data.Bool +open import Cubical.Algebra.Group + renaming (ℤ to ℤGroup ; Unit to UnitGroup) hiding (Bool) + +open import Cubical.Homotopy.Connected +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 +open import Cubical.HITs.Truncation + renaming (rec to trRec ; elim to trElim ; elim2 to trElim2) +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; elim to sElim ; elim2 to sElim2 ; map to sMap) +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; elim to pElim) + +open import Cubical.Algebra.AbGroup + +open import Cubical.Homotopy.Loopspace + +module _ {ℓ : Level} {A : Pointed ℓ} {n : ℕ} + where + funTyp : Type _ + funTyp = A →∙ coHomK-ptd n + + _++_ : funTyp → funTyp → funTyp + fst (f ++ g) x = fst f x +ₖ fst g x + snd (f ++ g) = cong₂ _+ₖ_ (snd f) (snd g) ∙ rUnitₖ _ (0ₖ _) + +addAgree : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (x y : funTyp {A = A} {n = n}) + → Path (fst (coHomRedGrDir n A)) + ∣ x ++ y ∣₂ + (∣ x ∣₂ +ₕ∙ ∣ y ∣₂) +addAgree {A = A} zero f g = + cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) refl) +addAgree {A = A} (suc zero) f g = + cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) refl) +addAgree {A = A} (suc (suc n)) f g = + cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) refl) + +ss : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Type ℓ''} + → isHomogeneous B + → (x y : (A →∙ B)) (f : C → x ≡ y) + → isEquiv (cong fst ∘ f) + → isEquiv f +ss homb x y f e = + isoToIsEquiv + (iso _ + (λ p → invEq (_ , e) (cong fst p)) + (λ p → →∙Homogeneous≡Path homb _ _ (secEq (_ , e) (cong fst p))) + (retEq (_ , e))) + +Whitehead : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → (f : A → B) + → isEmbedding f + → (∀ (b : B) → ∃[ a ∈ A ] f a ≡ b) + → isEquiv f +equiv-proof (Whitehead f emb p) b = + pRec isPropIsContr + (λ fib → fib , isEmbedding→hasPropFibers emb b fib) + (p b) + +Int-ind : ∀ {ℓ} (P : ℤ → Type ℓ) + → P (pos zero) → P (pos 1) + → P (negsuc zero) + → ((x y : ℤ) → P x → P y → P (x + y)) → (x : ℤ) → P x +Int-ind P z one min ind (pos zero) = z +Int-ind P z one min ind (pos (suc zero)) = one +Int-ind P z one min ind (pos (suc (suc n))) = + ind (pos (suc n)) 1 (Int-ind P z one min ind (pos (suc n))) one +Int-ind P z one min ind (negsuc zero) = min +Int-ind P z one min ind (negsuc (suc n)) = + ind (negsuc n) (negsuc zero) (Int-ind P z one min ind (negsuc n)) min + +genFunSpace : (n : ℕ) → S₊∙ n →∙ coHomK-ptd n +fst (genFunSpace zero) false = 1 +fst (genFunSpace zero) true = 0 +snd (genFunSpace zero) = refl +fst (genFunSpace (suc n)) = ∣_∣ +snd (genFunSpace (suc zero)) = refl +snd (genFunSpace (suc (suc n))) = refl + +module _ where + open import Cubical.Algebra.Monoid + open import Cubical.Algebra.Semigroup + open GroupStr + open IsGroup + open IsMonoid + open IsSemigroup + πS : (n : ℕ) → Group ℓ-zero + fst (πS n) = S₊∙ n →∙ coHomK-ptd n + 1g (snd (πS n)) = (λ _ → 0ₖ n) , refl + GroupStr._·_ (snd (πS n)) = + λ f g → (λ x → fst f x +ₖ fst g x) + , cong₂ _+ₖ_ (snd f) (snd g) ∙ rUnitₖ n (0ₖ n) + inv (snd (πS n)) f = (λ x → -ₖ fst f x) , cong -ₖ_ (snd f) ∙ -0ₖ {n = n} + is-set (isSemigroup (isMonoid (isGroup (snd (πS zero))))) = + isOfHLevelΣ 2 (isSetΠ (λ _ → isSetℤ)) + λ _ → isOfHLevelPath 2 isSetℤ _ _ + is-set (isSemigroup (isMonoid (isGroup (snd (πS (suc n)))))) = isOfHLevel↑∙' 0 n + IsSemigroup.assoc (isSemigroup (isMonoid (isGroup (snd (πS n))))) x y z = + →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ w → assocₖ n (fst x w) (fst y w) (fst z w)) + fst (identity (isMonoid (isGroup (snd (πS n)))) (f , p)) = + →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → rUnitₖ n (f x)) + snd (identity (isMonoid (isGroup (snd (πS n)))) (f , p)) = + →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → lUnitₖ n (f x)) + fst (inverse (isGroup (snd (πS n))) (f , p)) = + →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → rCancelₖ n (f x)) + snd (inverse (isGroup (snd (πS n))) (f , p)) = + →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → lCancelₖ n (f x)) + + isSetπS : (n : ℕ) → isSet (S₊∙ n →∙ coHomK-ptd n) + isSetπS = λ n → is-set (snd (πS n)) + +K : (n : ℕ) → GroupIso (coHomRedGrDir n (S₊∙ n)) (πS n) +fst (K n) = setTruncIdempotentIso (isSetπS n) +snd (K zero) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 (isSetπS 0) _ _) + λ f g → →∙Homogeneous≡ (isHomogeneousKn 0) refl) +snd (K (suc zero)) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 (isSetπS 1) _ _) + λ f g → →∙Homogeneous≡ (isHomogeneousKn 1) refl) +snd (K (suc (suc n))) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 (isSetπS _) _ _) + λ f g → →∙Homogeneous≡ (isHomogeneousKn _) refl) + +t : ∀ {ℓ} {A : Pointed ℓ} → Iso ((Bool , true) →∙ A) (typ A) +Iso.fun t f = fst f false +fst (Iso.inv t a) false = a +fst (Iso.inv (t {A = A}) a) true = pt A +snd (Iso.inv t a) = refl +Iso.rightInv t a = refl +Iso.leftInv t (f , p) = + ΣPathP ((funExt (λ { false → refl ; true → sym p})) , λ i j → p (~ i ∨ j)) + +S1' : (n : ℕ) → GroupIso (πS n) ℤGroup +fst (S1' zero) = t +snd (S1' zero) = makeIsGroupHom λ _ _ → refl +S1' (suc n) = + compGroupIso + (invGroupIso (K (suc n))) + (compGroupIso + (GroupEquiv→GroupIso (coHomGr≅coHomRedGr n (S₊∙ (suc n)))) + (Hⁿ-Sⁿ≅ℤ n)) + +S1 : (n : ℕ) → Iso (S₊∙ (suc n) →∙ coHomK-ptd (suc n)) ℤ +S1 n = compIso (invIso (setTruncIdempotentIso (isOfHLevel↑∙' 0 n))) + (compIso (equivToIso (fst (coHomGr≅coHomRedGr n (S₊∙ (suc n))))) + (fst (Hⁿ-Sⁿ≅ℤ n))) + +connFunSpace : (n i : ℕ) → (x y : S₊∙ n →∙ coHomK-ptd (suc i +' n)) → ∥ x ≡ y ∥ +connFunSpace n i f g = Iso.fun PathIdTrunc₀Iso (isContr→isProp (lem n) ∣ f ∣₂ ∣ g ∣₂) + where + open import Cubical.Relation.Nullary + natLem : (n i : ℕ) → ¬ (suc (i +ℕ n) ≡ n) + natLem zero i = snotz + natLem (suc n) i p = natLem n i (sym (+-suc i n) ∙ (cong predℕ p)) + + lem : (n : ℕ) → isContr ∥ (S₊∙ n →∙ coHomK-ptd (suc i +' n)) ∥₂ + fst (lem zero) = 0ₕ∙ (suc i) + snd (lem zero) = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → pRec (squash₂ _ _) + (λ id → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn (suc i)) + (funExt λ { false → sym id ; true → sym (snd f)}))) + (help (f .fst false)) + where + help : (x : coHomK (suc i)) → ∥ x ≡ 0ₖ _ ∥ + help = coHomK-elim _ (λ _ → isProp→isOfHLevelSuc i squash) ∣ refl ∣ + lem (suc n) = + isOfHLevelRetractFromIso 0 + (compIso (equivToIso (fst (coHomGr≅coHomRedGr (suc (i +ℕ n)) (S₊∙ (suc n))))) + (fst (Hⁿ-Sᵐ≅0 (suc (i +ℕ n)) n (natLem n i)))) + isContrUnit + +S1Pres1 : (n : ℕ) → Iso.fun (fst (S1' (suc n))) (∣_∣ , refl) ≡ pos 1 +S1Pres1 zero = refl +S1Pres1 (suc n) = cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ n))) (lem n) ∙ S1Pres1 n + where + lem : (n : ℕ) → Iso.fun (fst (suspensionAx-Sn n n)) ∣ ∣_∣ ∣₂ ≡ ∣ ∣_∣ ∣₂ + lem zero = cong ∣_∣₂ (funExt λ x → transportRefl (∣ x ∣ +ₖ (0ₖ 1)) ∙ rUnitₖ 1 ∣ x ∣) + lem (suc n) = cong ∣_∣₂ + (funExt λ x → (λ i → transportRefl ((ΩKn+1→Kn (suc (suc n)) + (transp (λ j → 0ₖ (suc (suc (suc n))) ≡ ∣ merid north (~ j ∧ ~ i) ∣) i + (λ z → ∣ compPath-filler (merid (transportRefl (transportRefl x i) i)) (sym (merid north)) i z + ∣)))) i) + ∙ Iso.leftInv (Iso-Kn-ΩKn+1 (suc (suc n))) ∣ x ∣) + +S1Pres1← : (n : ℕ) → Iso.inv (fst (S1' (suc n))) (pos 1) ≡ (∣_∣ , refl) +S1Pres1← n = sym (cong (Iso.inv (fst (S1' (suc n)))) (S1Pres1 n)) ∙ Iso.leftInv (fst (S1' (suc n))) (∣_∣ , refl) + +IsoFunSpace : (n : ℕ) → Iso (S₊∙ n →∙ coHomK-ptd n) ℤ +IsoFunSpace n = fst (S1' n) + +module g-base where + g : (n : ℕ) (i : ℕ) → coHomK i → (S₊∙ n →∙ coHomK-ptd (i +' n)) + fst (g n i x) y = x ⌣ₖ (genFunSpace n) .fst y + snd (g n i x) = cong (x ⌣ₖ_) ((genFunSpace n) .snd) ∙ ⌣ₖ-0ₖ i n x + + G : (n : ℕ) (i : ℕ) → coHomK i → (S₊∙ n →∙ coHomK-ptd (n +' i)) + fst (G n i x) y = (genFunSpace n) .fst y ⌣ₖ x + snd (G n i x) = cong (_⌣ₖ x) ((genFunSpace n) .snd) ∙ 0ₖ-⌣ₖ n i x + + eq1 : (n : ℕ) (i : ℕ) → (S₊∙ n →∙ coHomK-ptd (i +' n)) ≃ (S₊∙ n →∙ coHomK-ptd (i +' n)) + eq1 n i = isoToEquiv (iso F F FF FF) + where + lem : (i n : ℕ) → (-ₖ^ i · n) (snd (coHomK-ptd (i +' n))) ≡ 0ₖ _ + lem zero zero = refl + lem zero (suc zero) = refl + lem zero (suc (suc n)) = refl + lem (suc zero) zero = refl + lem (suc (suc i)) zero = refl + lem (suc i) (suc n) = refl + + F : S₊∙ n →∙ coHomK-ptd (i +' n) → S₊∙ n →∙ coHomK-ptd (i +' n) + fst (F f) x = (-ₖ^ i · n) (fst f x) + snd (F f) = cong (-ₖ^ i · n) (snd f) ∙ lem i n + + FF : (x : _) → F (F x) ≡ x + FF x = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ y → -ₖ-gen² i n _ _ (fst x y)) + + rCancel'' : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) → sym p ∙∙ refl ∙∙ p ≡ refl + rCancel'' p = (λ j → (λ i → p (~ i ∨ j)) ∙∙ refl ∙∙ λ i → p (i ∨ j)) ∙ sym (rUnit refl) + + transpPres0 : ∀ {k m : ℕ} (p : k ≡ m) → subst coHomK p (0ₖ k) ≡ 0ₖ m + transpPres0 {k = k} = J (λ m p → subst coHomK p (0ₖ k) ≡ 0ₖ m) (transportRefl _) + + eq2 : (n : ℕ) (i : ℕ) → (S₊∙ n →∙ coHomK-ptd (n +' i)) ≃ (S₊∙ n →∙ coHomK-ptd (i +' n)) + eq2 n i = + isoToEquiv (iso (λ f → (λ x → subst coHomK (+'-comm n i) (fst f x)) , + cong (subst coHomK (+'-comm n i)) (snd f) ∙ transpPres0 (+'-comm n i)) + (λ f → (λ x → subst coHomK (sym (+'-comm n i)) (fst f x)) + , (cong (subst coHomK (sym (+'-comm n i))) (snd f) ∙ transpPres0 (sym (+'-comm n i)))) + (λ f → →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ x → transportTransport⁻ _ (f .fst x))) + λ f → →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ x → transportTransport⁻ _ (f .fst x))) + + g≡ : (n : ℕ) (i : ℕ) → g n i ≡ λ x → fst (compEquiv (eq2 n i) (eq1 n i)) ((G n i) x) + g≡ n i = funExt (λ f → →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ y → gradedComm-⌣ₖ _ _ f (genFunSpace n .fst y))) + + glIso3-fun : (n m : ℕ) → + (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) + → (S₊ n → coHomK m) + glIso3-fun zero m (f , p) false = ΩKn+1→Kn _ (sym p ∙∙ cong f loop ∙∙ p) + glIso3-fun zero m (f , p) true = 0ₖ _ + glIso3-fun (suc n) m (f , p) x = + ΩKn+1→Kn _ (sym p ∙∙ cong f (merid x ∙ sym (merid (ptSn _))) ∙∙ p) + + glIso3-fun∙ : (n m : ℕ) → (f : _) → glIso3-fun n m f (ptSn _) ≡ 0ₖ m + glIso3-fun∙ zero m = λ _ → refl + glIso3-fun∙ (suc n) m (f , p) = + cong (ΩKn+1→Kn m) (cong (sym p ∙∙_∙∙ p) (cong (cong f) (rCancel (merid (ptSn _))))) + ∙∙ cong (ΩKn+1→Kn m) (rCancel'' p) + ∙∙ ΩKn+1→Kn-refl m + + + glIso3-inv : (n m : ℕ) → (S₊∙ n →∙ coHomK-ptd m) → (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) + fst (glIso3-inv zero m (f , p)) base = 0ₖ _ + fst (glIso3-inv zero m (f , p)) (loop i) = Kn→ΩKn+1 m (f false) i + snd (glIso3-inv zero m (f , p)) = refl + fst (glIso3-inv (suc n) m (f , p)) north = 0ₖ _ + fst (glIso3-inv (suc n) m (f , p)) south = 0ₖ _ + fst (glIso3-inv (suc n) m (f , p)) (merid a i) = Kn→ΩKn+1 m (f a) i + snd (glIso3-inv (suc n) m (f , p)) = refl + + glIso3 : (n m : ℕ) → + Iso (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) + (S₊∙ n →∙ coHomK-ptd m) + Iso.fun (glIso3 n m) f = (glIso3-fun n m f) , (glIso3-fun∙ n m f) + Iso.inv (glIso3 n m) = glIso3-inv n m + Iso.rightInv (glIso3 zero m) (f , p) = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { false → cong (ΩKn+1→Kn m) (sym (rUnit _)) ∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f false) + ; true → sym p}) + Iso.rightInv (glIso3 (suc n) m) (f , p) = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ x → (λ i → ΩKn+1→Kn m (sym (rUnit (cong-∙ (glIso3-inv (suc n) m (f , p) .fst) (merid x) (sym (merid (ptSn _))) i)) i)) + ∙∙ cong (ΩKn+1→Kn m) (cong (Kn→ΩKn+1 m (f x) ∙_) (cong sym (cong (Kn→ΩKn+1 m) p ∙ Kn→ΩKn+10ₖ m)) ∙ sym (rUnit _)) + ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f x)) + Iso.leftInv (glIso3 zero m) (f , p) = →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { base → sym p + ; (loop i) j → h j i}) + where + h : PathP (λ i → p (~ i) ≡ p (~ i)) (Kn→ΩKn+1 m (ΩKn+1→Kn m (sym p ∙∙ cong f loop ∙∙ p))) (cong f loop) + h = Iso.rightInv (Iso-Kn-ΩKn+1 _) _ + ◁ λ i → doubleCompPath-filler (sym p) (cong f loop) p (~ i) + Iso.leftInv (glIso3 (suc n) m) (f , p) = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { north → sym p + ; south → sym p ∙ cong f (merid (ptSn _)) + ; (merid a i) j → h a j i}) + where + h : (a : S₊ (suc n)) → PathP (λ i → p (~ i) ≡ (sym p ∙ cong f (merid (ptSn (suc n)))) i) + (Kn→ΩKn+1 m (ΩKn+1→Kn m (sym p ∙∙ cong f (merid a ∙ sym (merid (ptSn _))) ∙∙ p))) + (cong f (merid a)) + h a = Iso.rightInv (Iso-Kn-ΩKn+1 _) _ + ◁ λ i j → hcomp (λ k → λ { (i = i1) → (f (merid a j)) + ; (j = i0) → p (k ∧ ~ i) + ; (j = i1) → compPath-filler' (sym p) (cong f (merid (ptSn _))) k i }) + (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) + + glIsoInvHom : (n m : ℕ) (x y : coHomK n) (z : S₊ (suc m)) + → Iso.inv (glIso3 _ _) (G m n (x +ₖ y)) .fst z + ≡ Iso.inv (glIso3 _ _) (G m n x) .fst z + +ₖ Iso.inv (glIso3 _ _) (G m n y) .fst z + glIsoInvHom zero zero x y base = refl + glIsoInvHom (suc n) zero x y base = refl + glIsoInvHom zero zero x y (loop i) j = h j i + where + h : (cong (Iso.inv (glIso3 _ _) (G zero zero (x + y)) .fst) loop) + ≡ cong₂ _+ₖ_ (cong (Iso.inv (glIso3 _ _) (G zero zero x) .fst) loop) + (cong (Iso.inv (glIso3 _ _) (G zero zero y) .fst) loop) + h = Kn→ΩKn+1-hom 0 x y + ∙ ∙≡+₁ (cong (Iso.inv (glIso3 _ _) (G zero zero x) .fst) loop) + (cong (Iso.inv (glIso3 _ _) (G zero zero y) .fst) loop) + glIsoInvHom (suc n) zero x y (loop i) j = h j i + where + h : Kn→ΩKn+1 (suc n) ((pos (suc zero)) ·₀ (x +ₖ y)) + ≡ cong₂ _+ₖ_ (cong (Iso.inv (glIso3 zero (zero +' suc n)) (G zero (suc n) x) .fst) loop) + (cong (Iso.inv (glIso3 zero (zero +' suc n)) (G zero (suc n) y) .fst) loop) + h = cong (Kn→ΩKn+1 (suc n)) (lUnit⌣ₖ _ (x +ₖ y)) + ∙∙ Kn→ΩKn+1-hom (suc n) x y + ∙∙ (λ i → ∙≡+₂ n (Kn→ΩKn+1 (suc n) (lUnit⌣ₖ _ x (~ i))) + (Kn→ΩKn+1 (suc n) (lUnit⌣ₖ _ y (~ i))) i) + glIsoInvHom zero (suc m) x y north = refl + glIsoInvHom zero (suc m) x y south = refl + glIsoInvHom zero (suc m) x y (merid a i) j = h j i + where + h : Kn→ΩKn+1 (suc m) (_⌣ₖ_ {n = suc m} {m = zero} ∣ a ∣ (x + y)) + ≡ cong₂ _+ₖ_ (Kn→ΩKn+1 (suc m) (_⌣ₖ_ {n = suc m} {m = zero} ∣ a ∣ x)) + (Kn→ΩKn+1 (suc m) (_⌣ₖ_ {n = suc m} {m = zero} ∣ a ∣ y)) + h = cong (Kn→ΩKn+1 (suc m)) (leftDistr-⌣ₖ (suc m) 0 ∣ a ∣ x y) + ∙∙ Kn→ΩKn+1-hom (suc m) _ _ + ∙∙ ∙≡+₂ _ _ _ + glIsoInvHom (suc n) (suc m) x y north = refl + glIsoInvHom (suc n) (suc m) x y south = refl + glIsoInvHom (suc n) (suc m) x y (merid a i) j = h j i + where + h : Kn→ΩKn+1 (suc (suc (m +ℕ n))) (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ (x +ₖ y)) + ≡ cong₂ _+ₖ_ (Kn→ΩKn+1 (suc (suc (m +ℕ n))) (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ x)) + (Kn→ΩKn+1 (suc (suc (m +ℕ n))) (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ y)) + h = cong (Kn→ΩKn+1 (suc (suc (m +ℕ n)))) + (leftDistr-⌣ₖ (suc m) (suc n) ∣ a ∣ x y) + ∙∙ Kn→ΩKn+1-hom _ _ _ + ∙∙ ∙≡+₂ _ _ _ + + +'-suc : (n m : ℕ) → suc n +' m ≡ suc (n +' m) + +'-suc zero zero = refl + +'-suc (suc n) zero = refl + +'-suc zero (suc m) = refl + +'-suc (suc n) (suc m) = refl + + LEM : (i n : ℕ) (x : coHomK i) + → Path _ (G (suc n) i x) + (subst (λ x → S₊∙ (suc n) →∙ coHomK-ptd x) + (sym (+'-suc n i)) + ((Iso.inv (glIso3 n _)) (G n i x))) + LEM zero zero x = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ z → (λ i → x ·₀ ∣ z ∣) ∙ h x z ∙ sym (transportRefl _)) + where + h : (x : ℤ) (z : S¹) → _·₀_ {n = 1} x ∣ z ∣ ≡ fst (Iso.inv (glIso3 0 zero) (G zero zero x)) z + h = Int-ind _ (λ { base → refl ; (loop i) j → Kn→ΩKn+10ₖ zero (~ j) i}) + (λ { base → refl ; (loop i) j → rUnit (cong ∣_∣ₕ (lUnit loop j)) j i}) + (λ { base → refl ; (loop i) j → rUnit (cong ∣_∣ₕ (sym loop)) j i}) + λ x y inx iny z + → rightDistr-⌣ₖ 0 1 x y ∣ z ∣ + ∙∙ cong₂ _+ₖ_ (inx z) (iny z) + ∙∙ sym (glIsoInvHom zero zero x y z) + LEM (suc i) zero x = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ z → h z + ∙ sym (transportRefl + ((Iso.inv (glIso3 zero (suc i)) (G zero (suc i) x)) .fst z))) + where + h : (z : S₊ 1) → _ ≡ Iso.inv (glIso3 zero (suc i)) (G zero (suc i) x) .fst z + h base = refl + h (loop k) j = Kn→ΩKn+1 (suc i) (lUnit⌣ₖ _ x (~ j)) k + LEM zero (suc n) x = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ z → h x z ∙ λ i → transportRefl (Iso.inv (glIso3 (suc n) (suc n)) (G (suc n) zero x)) (~ i) .fst z) + where + +merid : rUnitₖ (suc (suc n)) ∣ south ∣ ≡ cong ∣_∣ₕ (merid (ptSn _)) + +merid = sym (lUnit _) + ∙ cong (cong ∣_∣ₕ) + λ j i → transp (λ _ → S₊ (suc (suc n))) (i ∨ j) (merid (ptSn (suc n)) i) + open import Cubical.Foundations.Path + + pp : (a : S₊ (suc n)) → PathP (λ i → ∣ merid a i ∣ₕ ≡ Kn→ΩKn+1 (suc n) (∣ a ∣ +ₖ (0ₖ _)) i) + refl (sym (rUnitₖ (suc (suc n)) ∣ south ∣)) + pp a = flipSquare ((λ i j → ∣ compPath-filler (merid a) (sym (merid (ptSn _))) i j ∣ₕ ) + ▷ cong (Kn→ΩKn+1 (suc n)) (sym (rUnitₖ (suc n) ∣ a ∣ₕ))) + ▷ sym (cong sym +merid) + + pp2 : (a : S₊ (suc n)) → (λ i → -ₖ ∣ merid a i ∣) + ≡ Kn→ΩKn+1 (suc n) (-ₖ ∣ a ∣) + pp2 a = cong (cong ∣_∣ₕ) (sym (symDistr (merid a) (sym (merid (ptSn (suc n)))))) + ∙ sym (Kn→ΩKn+1-ₖ (suc n) ∣ a ∣) + + h : (x : ℤ) (z : S₊ (suc (suc n))) + → _⌣ₖ_ {n = suc (suc n)} {m = 0} ∣ z ∣ x + ≡ Iso.inv (glIso3 (suc n) (suc n)) (G (suc n) zero x) .fst z + h = Int-ind _ + (λ { north → refl ; south → refl ; (merid a i) j → Kn→ΩKn+10ₖ (suc n) (~ j) i}) + (λ { north → refl ; south → refl + ; (merid a i) j → hcomp (λ k → λ { (i = i0) → ∣ north ∣ + ; (i = i1) → rUnitₖ (suc (suc n)) ∣ south ∣ (~ k) + ; (j = i0) → rUnitₖ (suc (suc n)) ∣ merid a i ∣ (~ k) + ; (j = i1) → pp a i k}) + ∣ merid a i ∣ₕ}) + (λ { north → refl + ; south → refl + ; (merid a i) j → pp2 a j i}) + λ x y indx indy z → leftDistr-⌣ₖ _ _ ∣ z ∣ x y + ∙ cong₂ _+ₖ_ (indx z) (indy z) + ∙ sym (glIsoInvHom _ _ _ _ _) + LEM (suc i) (suc n) x = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ z → h z + ∙ λ j → transportRefl ((Iso.inv (glIso3 (suc n) (suc (suc (n +ℕ i)))) + (G (suc n) (suc i) x))) (~ j) .fst z) + where + h : (z : S₊ (suc (suc n))) → _ + h north = refl + h south = refl + h (merid a i) = refl + + isEquivGzero : (i : ℕ) → isEquiv (G zero i) + isEquivGzero i = + isoToIsEquiv + (iso _ (λ f → fst f false) + (λ {(f , p) → →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { false → rUnitₖ _ (f false) ; true → sym p})}) + (lUnit⌣ₖ _)) + + isEquivG : (n i : ℕ) → isEquiv (G n i) + isEquivG zero i = + isoToIsEquiv + (iso _ (λ f → fst f false) + (λ {(f , p) → →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { false → rUnitₖ _ (f false) ; true → sym p})}) + (lUnit⌣ₖ _)) + isEquivG (suc n) i = + subst isEquiv (sym (funExt (LEM i n))) + (compEquiv (compEquiv (G n i , isEquivG n i) + (isoToEquiv (invIso (glIso3 n (n +' i))))) + (pathToEquiv (λ j → S₊∙ (suc n) →∙ coHomK-ptd (+'-suc n i (~ j)))) .snd) + + + isEquiv-g : (n i : ℕ) → isEquiv (g n i) + isEquiv-g n i = + subst isEquiv (sym (g≡ n i)) + (compEquiv (G n i , isEquivG n i) (compEquiv (eq2 n i) (eq1 n i)) .snd) + +module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) + (conB : (x y : typ B) → ∥ x ≡ y ∥) + (n : ℕ) (Q-is : Iso (typ (Q (pt B))) (S₊ n)) + (Q-is-ptd : Iso.fun Q-is (pt (Q (pt B))) ≡ snd (S₊∙ n)) + (c : (b : typ B) → (Q b →∙ coHomK-ptd n)) + (c-pt : c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) where + + g : (b : typ B) (i : ℕ) → coHomK i → (Q b →∙ coHomK-ptd (i +' n)) + fst (g b i x) y = x ⌣ₖ c b .fst y + snd (g b i x) = cong (x ⌣ₖ_) (c b .snd) ∙ ⌣ₖ-0ₖ i n x + + g' : (b : typ B) (i : ℕ) → coHomK i → coHomK i → (Q b →∙ coHomK-ptd (i +' n)) + fst (g' b i x y) z = x ⌣ₖ c b .fst z +ₖ y ⌣ₖ c b .fst z + snd (g' b i x y) = cong₂ _+ₖ_ (cong (x ⌣ₖ_) (c b .snd) ∙ ⌣ₖ-0ₖ i n x) + (cong (y ⌣ₖ_) (c b .snd) ∙ ⌣ₖ-0ₖ i n y) ∙ rUnitₖ _ (0ₖ _) + + g-hom : (b : typ B) (i : ℕ) → (x y : coHomK i) → g b i (x +ₖ y) ≡ ((g b i x) ++ (g b i y)) + g-hom b i x y = →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ z → rightDistr-⌣ₖ i n x y (c b .fst z)) + + gPathP' : (i : ℕ) → PathP (λ j → coHomK i → (isoToPath Q-is j , ua-gluePath (isoToEquiv Q-is) (Q-is-ptd) j) + →∙ coHomK-ptd (i +' n)) (g (pt B) i) (g-base.g n i) + gPathP' i = + toPathP + (funExt + λ x → →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ y → (λ i → transportRefl (transportRefl x i ⌣ₖ c (pt B) .fst (Iso.inv Q-is (transportRefl y i))) i) + ∙ cong (x ⌣ₖ_) (funExt⁻ c-pt (Iso.inv Q-is y) ∙ cong (genFunSpace n .fst) (Iso.rightInv Q-is y)))) + + g-base : (i : ℕ) → isEquiv (g (pt B) i) + g-base i = transport (λ j → isEquiv (gPathP' i (~ j))) (g-base.isEquiv-g n i) + + g-equiv : (b : typ B) (i : ℕ) → isEquiv (g b i) + g-equiv b i = + pRec (isPropIsEquiv _) + (J (λ b _ → isEquiv (g b i)) + (g-base i)) + (conB (pt B) b) + +module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) + (conB : (x y : typ B) → ∥ x ≡ y ∥₂) + (n : ℕ) (Q-is : Iso (typ (Q (pt B))) (S₊ n)) + (Q-is-ptd : Iso.fun Q-is (pt (Q (pt B))) ≡ snd (S₊∙ n)) where + + is-setQ→K : (b : typ B) → isSet (Q b →∙ coHomK-ptd n) + is-setQ→K b = + sRec (isProp→isOfHLevelSuc 1 isPropIsSet) + (J (λ b _ → isSet (Q b →∙ coHomK-ptd n)) + (subst isSet (cong (_→∙ coHomK-ptd n) + (ua∙ (isoToEquiv (invIso Q-is)) (cong (Iso.inv Q-is) (sym Q-is-ptd) ∙ Iso.leftInv Q-is _))) + (isOfHLevelRetractFromIso 2 (fst (S1' n)) isSetℤ))) + (conB (pt B) b) + + c* : Σ[ c ∈ ((b : typ B) → (Q b →∙ coHomK-ptd n)) ] + (c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) + fst c* b = + sRec (is-setQ→K b) + (J (λ b _ → Q b →∙ coHomK-ptd n) + ((λ x → genFunSpace n .fst (Iso.fun Q-is x)) + , cong (genFunSpace n .fst) Q-is-ptd ∙ genFunSpace n .snd)) (conB (pt B) b) + snd c* = + funExt λ x → (λ i → sRec (is-setQ→K (pt B)) + (J (λ b _ → Q b →∙ coHomK-ptd n) + ((λ x₁ → genFunSpace n .fst (Iso.fun Q-is x₁)) , + (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) + (isPropPath (conB (pt B) (pt B)) ∣ refl ∣₂ i) .fst x) + ∙ (λ i → transportRefl (genFunSpace n .fst (Iso.fun Q-is (transportRefl x i))) i) + where + isConnB : isConnected 3 (typ B) + fst isConnB = ∣ pt B ∣ + snd isConnB = + trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → sRec (isOfHLevelTrunc 3 _ _) (cong ∣_∣ₕ) (conB (pt B) a) + + isPropPath : isProp (∥ pt B ≡ pt B ∥₂) + isPropPath = + isOfHLevelRetractFromIso 1 setTruncTrunc2Iso + (isContr→isProp (isConnectedPath _ isConnB (pt B) (pt B))) + +module _ {ℓ ℓ'} (B : Pointed ℓ) (P : typ B → Type ℓ') where + E : Type _ + E = Σ (typ B) P + + Ẽ : Type _ + Ẽ = Pushout {A = E} (λ _ → tt) fst + + i : (n : ℕ) → (P-base : Iso (P (pt B)) (S₊ n)) → S₊ (suc n) → Ẽ + i zero P-base base = inr (pt B) + i zero P-base (loop j) = (sym (push (pt B , Iso.inv P-base false)) + ∙ push ((pt B) , Iso.inv P-base true)) j + i (suc n) P-base north = inl tt + i (suc n) P-base south = inr (pt B) + i (suc n) P-base (merid a i₁) = push (pt B , Iso.inv P-base a) i₁ + + Q : typ B → Pointed ℓ' + Q x = Susp (P x) , north + + F : Type _ + F = Σ (typ B) λ x → Q x .fst + + F̃ : Type _ + F̃ = Pushout {A = typ B} {C = F} (λ _ → tt) λ b → b , north + + invFE : Ẽ → F̃ + invFE (inl x) = inl tt + invFE (inr x) = inr (x , south) + invFE (push (x , a) i₁) = ((push x) ∙ λ i → inr (x , merid a i)) i₁ + + funFE : F̃ → Ẽ + funFE (inl x) = inl tt + funFE (inr (x , north)) = inl tt + funFE (inr (x , south)) = inr x + funFE (inr (x , merid a i₁)) = push (x , a) i₁ + funFE (push a i₁) = inl tt + + IsoFE : Iso F̃ Ẽ + Iso.fun IsoFE = funFE + Iso.inv IsoFE = invFE + Iso.rightInv IsoFE (inl x) = refl + Iso.rightInv IsoFE (inr x) = refl + Iso.rightInv IsoFE (push (x , a) i₁) k = h k i₁ + where + h : cong funFE (((push x) ∙ λ i → inr (x , merid a i))) + ≡ push (x , a) + h = congFunct funFE (push x) (λ i → inr (x , merid a i)) + ∙ sym (lUnit (push (x , a))) + Iso.leftInv IsoFE (inl x) = refl + Iso.leftInv IsoFE (inr (x , north)) = push x + Iso.leftInv IsoFE (inr (x , south)) = refl + Iso.leftInv IsoFE (inr (x , merid a i)) j = + compPath-filler' (push x) (λ i₁ → inr (x , merid a i₁)) (~ j) i + Iso.leftInv IsoFE (push a i₁) k = push a (i₁ ∧ k) + + + funDir1 : ∀ {ℓ} {A : Pointed ℓ} → (F̃ , inl tt) →∙ A → (b : typ B) → Q b →∙ A + fst (funDir1 {A = A , a} (f , p) b) north = f (inr (b , north)) + fst (funDir1 {A = A , a} (f , p) b) south = f (inr (b , south)) + fst (funDir1 {A = A , a} (f , p) b) (merid a₁ i₁) = f (inr (b , merid a₁ i₁)) + snd (funDir1 {A = A , a} (f , p) b) = sym (cong f (push b)) ∙ p + + funDir2 : ∀ {ℓ} {A : Pointed ℓ} → ((b : typ B) → Q b →∙ A) → (F̃ , inl tt) →∙ A + fst (funDir2 {A = A , a} f) (inl x) = a + fst (funDir2 {A = A , a} f) (inr (x , north)) = f x .fst north + fst (funDir2 {A = A , a} f) (inr (x , south)) = f x .fst south + fst (funDir2 {A = A , a} f) (inr (x , merid a₁ i₁)) = f x .fst (merid a₁ i₁) + fst (funDir2 {A = A , a} f) (push a₁ i₁) = snd (f a₁) (~ i₁) + snd (funDir2 {A = A , a} f) = refl + + funDir2-hom : (n : ℕ) → (f g : ((b : typ B) → Q b →∙ coHomK-ptd n)) + → funDir2 (λ b → f b ++ g b) ≡ (funDir2 f ++ funDir2 g) + funDir2-hom n f g = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { (inl x) → sym (rUnitₖ _ (0ₖ _)) + ; (inr (x , north)) → refl + ; (inr (x , south)) → refl + ; (inr (x , merid a i₁)) → refl + ; (push a j) i → compPath-filler (cong₂ _+ₖ_ (snd (f a)) (snd (g a))) + (rUnitₖ n (0ₖ n)) (~ i) (~ j)}) + + funDirSect : ∀ {ℓ} {A : Pointed ℓ} → (x : (b : typ B) → Q b →∙ A) (b : typ B) (q : Q b .fst) + → funDir1 (funDir2 x) b .fst q ≡ x b .fst q + funDirSect f b north = refl + funDirSect f b south = refl + funDirSect f b (merid a i₁) = refl + + funDirRetr : ∀ {ℓ} {A : Pointed ℓ} (f : F̃ → typ A) (p : _) + → (x : F̃) → fst (funDir2 {A = A} (funDir1 (f , p))) x ≡ f x + funDirRetr f p (inl x) = sym p + funDirRetr f p (inr (x , north)) = refl + funDirRetr f p (inr (x , south)) = refl + funDirRetr f p (inr (x , merid a i₁)) = refl + funDirRetr f p (push a i₁) j = compPath-filler (sym (cong f (push a))) p (~ j) (~ i₁) + + Iso2 : ∀ {ℓ} {A : Pointed ℓ} + → Iso ((F̃ , inl tt) →∙ A) + ((b : typ B) → Q b →∙ A) + Iso.fun (Iso2 {A = A , a}) = funDir1 + Iso.inv (Iso2 {A = A , a}) = funDir2 + Iso.rightInv (Iso2 {A = A , a}) f = + funExt λ b → ΣPathP (funExt (funDirSect f b) + , sym (rUnit (snd (f b)))) + Iso.leftInv (Iso2 {A = A , a}) (f , p) = + ΣPathP ((funExt (funDirRetr f p)) + , λ i j → p (~ i ∨ j)) + + ι : (k : ℕ) → Iso ((b : typ B) → Q b →∙ coHomK-ptd k) + ((Ẽ , inl tt) →∙ coHomK-ptd k) + ι k = compIso (invIso Iso2) h + where + h : Iso ((F̃ , inl tt) →∙ coHomK-ptd k) + ((Ẽ , inl tt) →∙ coHomK-ptd k) + Iso.fun h G = (λ x → G .fst (Iso.inv IsoFE x)) + , (snd G) + Iso.inv h G = (λ x → G .fst (Iso.fun IsoFE x)) + , (snd G) + Iso.rightInv h G = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ x → cong (G .fst) (Iso.rightInv IsoFE x)) + Iso.leftInv h G = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ x → cong (G .fst) (Iso.leftInv IsoFE x)) + + ι-hom : (k : ℕ) → (f g : ((b : typ B) → Q b →∙ coHomK-ptd k)) + → Iso.fun (ι k) (λ b → f b ++ g b) + ≡ (Iso.fun (ι k) f ++ Iso.fun (ι k) g) + ι-hom k f g = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ x → funExt⁻ (cong fst (funDir2-hom _ f g)) (invFE x)) + +codomainIsoDep : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : A → Type ℓ'} {C : A → Type ℓ''} + → ((a : A) → Iso (B a) (C a)) + → Iso ((a : A) → B a) ((a : A) → C a) +Iso.fun (codomainIsoDep is) f a = Iso.fun (is a) (f a) +Iso.inv (codomainIsoDep is) f a = Iso.inv (is a) (f a) +Iso.rightInv (codomainIsoDep is) f = funExt λ a → Iso.rightInv (is a) (f a) +Iso.leftInv (codomainIsoDep is) f = funExt λ a → Iso.leftInv (is a) (f a) + +module Thom {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) + (conB : (x y : typ B) → ∥ x ≡ y ∥) + (n : ℕ) (Q-is : Iso (typ (Q B P (pt B))) (S₊ n)) + (Q-is-ptd : Iso.fun Q-is (pt (Q B P (pt B))) ≡ snd (S₊∙ n)) + (c : (b : typ B) → (Q B P b →∙ coHomK-ptd n)) + (c-pt : c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) where + + ϕ : (i : ℕ) → GroupEquiv (coHomGr i (typ B)) + (coHomRedGrDir (i +' n) (Ẽ B P , inl tt)) + fst (ϕ i) = + isoToEquiv + (setTruncIso + (compIso + (codomainIsoDep + λ b → equivToIso ((g B (Q B P) conB n Q-is Q-is-ptd c c-pt b i) + , g-equiv B (Q B P) conB n Q-is Q-is-ptd c c-pt b i)) + (ι B P (i +' n)))) + snd (ϕ i) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ F G → cong ∣_∣₂ (cong (Iso.fun (ι B P (i +' n))) + (funExt (λ a → g-hom B (Q B P) conB n Q-is Q-is-ptd c c-pt a i (F a) (G a))) + ∙ ι-hom B P (i +' n) _ _) + ∙ addAgree (i +' n) _ _) + +module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) + (conB : (x y : typ B) → ∥ x ≡ y ∥₂) + (n : ℕ) (Q-is : Iso (typ (Q B P (pt B))) (S₊ n)) + (Q-is-ptd : Iso.fun Q-is (pt (Q B P (pt B))) ≡ snd (S₊∙ n)) where + 0-connB : (x y : typ B) → ∥ x ≡ y ∥ + 0-connB x y = sRec (isProp→isSet squash) (∥_∥.∣_∣) (conB x y) + + c = (c* B (Q B P) conB n Q-is Q-is-ptd .fst) + c-ptd = (c* B (Q B P) conB n Q-is Q-is-ptd .snd) + + module w = Thom B P 0-connB n Q-is Q-is-ptd c c-ptd + + ϕ = w.ϕ + + E' = E B P + + E'̃ = Ẽ B P + + p : E' → typ B + p = fst + + e : coHom n (typ B) + e = ∣ (λ b → c b .fst south) ∣₂ + + ⌣-hom : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr (i +' n) (typ B)) + fst (⌣-hom i) x = x ⌣ e + snd (⌣-hom i) = + makeIsGroupHom λ f g → rightDistr-⌣ _ _ f g e + + p-hom : (i : ℕ) → GroupHom (coHomGr (i +' n) (typ B)) (coHomGr (i +' n) E') + p-hom i = coHomMorph _ p + + diff --git a/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda b/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda index 059b5af70..126228a53 100644 --- a/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda +++ b/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda @@ -39,11 +39,10 @@ private variable ℓ : Level -private - natTranspLem : ∀ {A B : ℕ → Type} {n m : ℕ} (a : A n) - (f : (n : ℕ) → (a : A n) → B n) (p : n ≡ m) - → f m (subst A p a) ≡ subst B p (f n a) - natTranspLem {A = A} {B = B} a f p = sym (substCommSlice A B f p a) +natTranspLem : ∀ {A B : ℕ → Type} {n m : ℕ} (a : A n) + (f : (n : ℕ) → (a : A n) → B n) (p : n ≡ m) + → f m (subst A p a) ≡ subst B p (f n a) +natTranspLem {A = A} {B = B} a f p = sym (substCommSlice A B f p a) +'-comm : (n m : ℕ) → n +' m ≡ m +' n +'-comm n m = +'≡+ n m ∙∙ +-comm n m ∙∙ sym (+'≡+ m n) @@ -97,8 +96,7 @@ private -- -ₖⁿ̇*ᵐ -ₖ^_·_ : {k : ℕ} (n m : ℕ) → coHomK k → coHomK k --ₖ^_·_ {k = zero} n m = -ₖ-helper {k = zero} n m (evenOrOdd n) (evenOrOdd m) --ₖ^_·_ {k = suc k} n m = trMap (-ₖ-helper n m (evenOrOdd n) (evenOrOdd m)) +-ₖ^_·_ {k = k} n m = -ₖ-gen n m (evenOrOdd n) (evenOrOdd m) -- cohomology version -ₕ^_·_ : {k : ℕ} {A : Type ℓ} (n m : ℕ) → coHom k A → coHom k A @@ -153,6 +151,39 @@ private ; south → cong ∣_∣ₕ (merid (ptSn _)) ; (merid a i) k → ∣ compPath-filler (merid a) (sym (merid (ptSn _))) (~ k) i ∣ₕ} + +open import Cubical.Foundations.Isomorphism +-ₖ-gen² : {k : ℕ} (n m : ℕ) + (p : isEvenT n ⊎ isOddT n) + (q : isEvenT m ⊎ isOddT m) + → (x : coHomK k) → -ₖ-gen n m p q (-ₖ-gen n m p q x) ≡ x +-ₖ-gen² {k = zero} n m (inl x₁) q x = refl +-ₖ-gen² {k = zero} n m (inr x₁) (inl x₂) x = refl +-ₖ-gen² {k = zero} n m (inr x₁) (inr x₂) x = + cong (pos 0 -_) (-AntiComm (pos 0) x) + ∙∙ -AntiComm (pos 0) (-ℤ (x - pos 0)) + ∙∙ h x + where + h : (x : _) → -ℤ (-ℤ (x - pos 0) - pos 0) ≡ x + h (pos zero) = refl + h (pos (suc n)) = refl + h (negsuc n) = refl +-ₖ-gen² {k = suc k} n m (inl x₁) q x i = + -ₖ-gen-inl-left n m x₁ q (-ₖ-gen-inl-left n m x₁ q x i) i +-ₖ-gen² {k = suc k} n m (inr x₁) (inl x₂) x i = + -ₖ-gen-inl-right n m (inr x₁) x₂ (-ₖ-gen-inl-right n m (inr x₁) x₂ x i) i +-ₖ-gen² {k = suc k} n m (inr x₁) (inr x₂) x = + (λ i → -ₖ-gen-inr≡-ₖ n m x₁ x₂ (-ₖ-gen-inr≡-ₖ n m x₁ x₂ x i) i) ∙ -ₖ^2 x + +-ₖ-genIso : {k : ℕ} (n m : ℕ) + (p : isEvenT n ⊎ isOddT n) + (q : isEvenT m ⊎ isOddT m) + → Iso (coHomK k) (coHomK k) +Iso.fun (-ₖ-genIso {k = k} n m p q) = -ₖ-gen n m p q +Iso.inv (-ₖ-genIso {k = k} n m p q) = -ₖ-gen n m p q +Iso.rightInv (-ₖ-genIso {k = k} n m p q) = -ₖ-gen² n m p q +Iso.leftInv (-ₖ-genIso {k = k} n m p q) = -ₖ-gen² n m p q + -- action of cong on -ₖⁿ̇*ᵐ cong-ₖ-gen-inr : {k : ℕ} (n m : ℕ) (p : _) (q : _) (P : Path (coHomK (2 + k)) (0ₖ _) (0ₖ _)) → cong (-ₖ-gen n m (inr p) (inr q)) P ≡ sym P diff --git a/Cubical/ZCohomology/RingStructure/RingLaws.agda b/Cubical/ZCohomology/RingStructure/RingLaws.agda index 99b719f6f..58d715584 100644 --- a/Cubical/ZCohomology/RingStructure/RingLaws.agda +++ b/Cubical/ZCohomology/RingStructure/RingLaws.agda @@ -637,6 +637,10 @@ private n+'0 zero = refl n+'0 (suc n) = refl +lUnit⌣ₖ : (n : ℕ) (x : coHomK n) → _⌣ₖ_ {n = 0} (pos 1) x ≡ x +lUnit⌣ₖ zero = λ _ → refl +lUnit⌣ₖ (suc n) x = rUnitₖ _ x + lUnit⌣ : ∀ {ℓ} {A : Type ℓ} (n : ℕ) (x : coHom n A) → x ⌣ 1⌣ ≡ subst (λ n → coHom n A) (sym (n+'0 n)) x lUnit⌣ zero = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) From fc075e155f2c018fa4ec0744baccf1527f97be26 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Wed, 1 Sep 2021 16:59:44 +0200 Subject: [PATCH 63/89] stuff --- Cubical/Data/Sigma/Properties.agda | 7 + Cubical/Homotopy/Hopf.agda | 400 ++++++++++++++ Cubical/ZCohomology/Gysin.agda | 495 +++++++++++++++++- .../RingStructure/GradedCommutativity.agda | 2 +- 4 files changed, 899 insertions(+), 5 deletions(-) create mode 100644 Cubical/Homotopy/Hopf.agda diff --git a/Cubical/Data/Sigma/Properties.agda b/Cubical/Data/Sigma/Properties.agda index 16af06c24..0679e8a88 100644 --- a/Cubical/Data/Sigma/Properties.agda +++ b/Cubical/Data/Sigma/Properties.agda @@ -358,3 +358,10 @@ module _ {A : Type ℓ} {B : A → Type ℓ'} {C : ∀ a → B a → Type ℓ''} Iso.leftInv curryIso f = refl unquoteDecl curryEquiv = declStrictIsoToEquiv curryEquiv curryIso + +ΣΣ : Iso (Σ[ x ∈ A ] (Σ[ y ∈ B x ] (C x y))) + (Σ[ x ∈ Σ A B ] C (fst x) (snd x)) +fun ΣΣ (x , y , p) = (x , y) , p +inv ΣΣ ((x , y) , p) = x , y , p +rightInv ΣΣ _ = refl +leftInv ΣΣ _ = refl diff --git a/Cubical/Homotopy/Hopf.agda b/Cubical/Homotopy/Hopf.agda new file mode 100644 index 000000000..5db7125a1 --- /dev/null +++ b/Cubical/Homotopy/Hopf.agda @@ -0,0 +1,400 @@ +{-# OPTIONS --safe #-} +module Cubical.Homotopy.Hopf where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Groups.Connected +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.MayerVietorisUnreduced +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.RingStructure.CupProduct +open import Cubical.ZCohomology.RingStructure.RingLaws +open import Cubical.ZCohomology.RingStructure.GradedCommutativity + +open import Cubical.Functions.Embedding + +open import Cubical.Data.Fin + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Pointed.Homogeneous + +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Empty renaming (rec to ⊥-rec) +open import Cubical.Data.Sigma +open import Cubical.Data.Int hiding (_+'_) +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Unit +open import Cubical.Data.Bool +open import Cubical.Algebra.Group + renaming (ℤ to ℤGroup ; Unit to UnitGroup) hiding (Bool) + +open import Cubical.HITs.Pushout.Flattening +open import Cubical.Homotopy.Connected +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 renaming (_·_ to _*_) +open import Cubical.HITs.Truncation + renaming (rec to trRec ; elim to trElim ; elim2 to trElim2) +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; elim to sElim ; elim2 to sElim2 ; map to sMap) +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; elim to pElim) + +open import Cubical.Algebra.AbGroup + +open import Cubical.Homotopy.Loopspace + +open import Cubical.HITs.Join + + +retEq≡secEq : ∀ {ℓ} {A B : Type ℓ} (e : A ≃ B) + → (x : _) → secEq e (e .fst x) ≡ cong (e .fst) (retEq e x) +retEq≡secEq {A = A} = + EquivJ (λ B e → (x : _) → secEq e (e .fst x) ≡ cong (e .fst) (retEq e x) ) + λ _ → refl + +record HSpace {ℓ : Level} (A : Pointed ℓ) : Type ℓ where + constructor HSp + field + μ : typ A → typ A → typ A + μₗ : (x : typ A) → μ (pt A) x ≡ x + μᵣ : (x : typ A) → μ x (pt A) ≡ x + μₗᵣ : μₗ (pt A) ≡ μᵣ (pt A) + +record assocHSpace {ℓ : Level} {A : Pointed ℓ} (e : HSpace A) : Type ℓ where + constructor assocHSp + field + μ-assoc : (x y z : _) → HSpace.μ e (HSpace.μ e x y) z + ≡ HSpace.μ e x (HSpace.μ e y z) + + μ-assoc-filler : (y z : typ A) + → PathP (λ i → HSpace.μ e (HSpace.μₗ e y i) z ≡ HSpace.μₗ e (HSpace.μ e y z) i) + (μ-assoc (pt A) y z) refl + +S1-HSpace : HSpace (S₊∙ 1) +HSpace.μ S1-HSpace = _*_ +HSpace.μₗ S1-HSpace x = refl +HSpace.μᵣ S1-HSpace base = refl +HSpace.μᵣ S1-HSpace (loop i) = refl +HSpace.μₗᵣ S1-HSpace x = refl + +S1-AssocHSpace : assocHSpace S1-HSpace +assocHSpace.μ-assoc S1-AssocHSpace base _ _ = refl +assocHSpace.μ-assoc S1-AssocHSpace (loop i) x y j = h x y j i + where + h : (x y : S₊ 1) → cong (_* y) (rotLoop x) ≡ rotLoop (x * y) + h = wedgeconFun _ _ (λ _ _ → isOfHLevelPath 2 (isGroupoidS¹ _ _) _ _) + (λ x → refl) + (λ { base → refl ; (loop i) → refl}) + refl +assocHSpace.μ-assoc-filler S1-AssocHSpace _ _ = refl + +module hopfBase {ℓ : Level} {A : Pointed ℓ} {e : HSpace A} (e-ass : assocHSpace e) (conA : ((x y : typ A) → ∥ x ≡ y ∥)) where + isEquiv-μ : (x : typ A) → isEquiv (λ z → (HSpace.μ e z x)) + isEquiv-μ x = pRec (isPropIsEquiv _) + (J (λ x _ → isEquiv (λ z → HSpace.μ e z x)) + (subst isEquiv (funExt (λ z → sym (HSpace.μᵣ e z))) (idIsEquiv (typ A)))) + (conA (pt A) x) + + μ-eq : (x : typ A) → typ A ≃ typ A + μ-eq x = (λ z → HSpace.μ e z x) , (isEquiv-μ x) + + isEquiv-μ' : (x : typ A) → isEquiv (HSpace.μ e x) + isEquiv-μ' x = + pRec (isPropIsEquiv _) + (J (λ x _ → isEquiv (HSpace.μ e x)) + (subst isEquiv (funExt (λ x → sym (HSpace.μₗ e x))) (idIsEquiv (typ A)))) + (conA (pt A) x) + + μ-eq' : (x : typ A) → typ A ≃ typ A + μ-eq' x = HSpace.μ e x , isEquiv-μ' x + + Hopf : Susp (typ A) → Type ℓ + Hopf north = typ A + Hopf south = typ A + Hopf (merid a i₁) = ua (μ-eq a) i₁ + + TotalSpaceHopf' : Type _ + TotalSpaceHopf' = Pushout {A = typ A × typ A} fst λ x → HSpace.μ e (fst x) (snd x) + + TotalSpaceHopf'→TotalSpace : TotalSpaceHopf' → Σ[ x ∈ Susp (typ A) ] Hopf x + TotalSpaceHopf'→TotalSpace (inl x) = north , x + TotalSpaceHopf'→TotalSpace (inr x) = south , x + TotalSpaceHopf'→TotalSpace (push (x , y) i₁) = merid y i₁ , ua-gluePt (μ-eq y) i₁ x + + joinIso₁ : Iso (TotalSpaceHopf') (join (typ A) (typ A)) + joinIso₁ = iso F G s r + where + F : TotalSpaceHopf' → join (typ A) (typ A) + F (inl x) = inl x + F (inr x) = inr x + F (push (a , x) i) = push a (HSpace.μ e a x) i + + G : join (typ A) (typ A) → TotalSpaceHopf' + G (inl x) = inl x + G (inr x) = inr x + G (push a b i) = (push (a , invEq (μ-eq' a) b) ∙ cong inr (secEq (μ-eq' a) b)) i + + s : section F G + s (inl x) = refl + s (inr x) = refl + s (push a b i) j = + hcomp (λ k → λ { (i = i0) → inl a + ; (i = i1) → inr (secEq (μ-eq' a) b (j ∨ k)) + ; (j = i0) → F (compPath-filler (push (a , invEq (μ-eq' a) b)) (cong inr (secEq (μ-eq' a) b)) k i) + ; (j = i1) → push a b i}) + (hcomp (λ k → λ { (i = i0) → inl a + ; (i = i1) → inr (secEq (μ-eq' a) b (~ k ∨ j)) + ; (j = i0) → push a (secEq (μ-eq' a) b (~ k)) i + ; (j = i1) → push a b i}) + (push a b i)) + + r : retract F G + r (inl x) = refl + r (inr x) = refl + r (push (x , y) i) j = + hcomp (λ k → λ { (i = i0) → inl x + ; (i = i1) → inr (HSpace.μ e x y) + ; (j = i0) → (push (x , invEq (μ-eq' x) (HSpace.μ e x y)) + ∙ (λ i₁ → inr (retEq≡secEq (μ-eq' x) y (~ k) i₁))) i + ; (j = i1) → push (x , y) i}) + (hcomp (λ k → λ { (i = i0) → inl x + ; (i = i1) → inr (HSpace.μ e x (retEq (μ-eq' x) y k)) + ; (j = i1) → push (x , retEq (μ-eq' x) y k) i}) + ((push (x , invEq (μ-eq' x) (HSpace.μ e x y))) i)) + + isEquivTotalSpaceHopf'→TotalSpace : isEquiv TotalSpaceHopf'→TotalSpace + isEquivTotalSpaceHopf'→TotalSpace = + isoToIsEquiv (iso _ inv inv* retr) + where + inv : _ → _ + inv (north , y) = inl y + inv (south , y) = inr y + inv (merid a i , y) = + hcomp (λ k → λ { (i = i0) → push (y , a) (~ k) + ; (i = i1) → inr y}) + (inr (ua-unglue (μ-eq a) i y)) + where + + pp : PathP (λ i → ua (μ-eq a) i → TotalSpaceHopf') inl inr + pp = ua→ {e = μ-eq a} {B = λ _ → TotalSpaceHopf'} λ b → push (b , a) + + inv* : (x : _) → TotalSpaceHopf'→TotalSpace (inv x) ≡ x + inv* (north , x) = refl + inv* (south , x) = refl + inv* (merid a i , y) j = + hcomp (λ k → λ { (i = i0) → merid a (~ k ∧ ~ j) , ua-gluePt (μ-eq a) (~ k ∧ ~ j) y + ; (i = i1) → south , y + ; (j = i0) → TotalSpaceHopf'→TotalSpace + (hfill (λ k → λ { (i = i0) → push (y , a) (~ k) + ; (i = i1) → inr y}) + (inS (inr (ua-unglue (μ-eq a) i y))) + k) + ; (j = i1) → merid a i , y}) + ((merid a (i ∨ ~ j)) , s (μ-eq a) i j y) + where + s : ∀ {ℓ} {A B : Type ℓ} (e : A ≃ B) → + PathP (λ i → PathP (λ j → (y : ua e i) → ua e (i ∨ ~ j)) + (λ y → ua-unglue e i y) + λ y → y) + (λ j y → ua-gluePt e (~ j) y) + refl + s {A = A} {B = B} = EquivJ (λ B e → PathP (λ i → PathP (λ j → (y : ua e i) → ua e (i ∨ ~ j)) + (λ y → ua-unglue e i y) + λ y → y) + (λ j y → ua-gluePt e (~ j) y) + refl) + h + where + h : _ + h i j a = ua-gluePt (idEquiv B) (i ∨ ~ j) (ua-unglue (idEquiv B) i a) + + + retr : retract TotalSpaceHopf'→TotalSpace inv + retr (inl x) = refl + retr (inr x) = refl + retr (push (x , y) i) j = + hcomp (λ k → λ { (i = i0) → push (x , y) (~ k) + ; (i = i1) → inr (HSpace.μ e x y) + ; (j = i1) → push (x , y) (i ∨ ~ k)}) + (inr (HSpace.μ e x y)) + + IsoTotalSpaceJoin : Iso (Σ[ x ∈ Susp (typ A) ] Hopf x) (join (typ A) (typ A)) + IsoTotalSpaceJoin = + compIso (equivToIso (invEquiv (_ , isEquivTotalSpaceHopf'→TotalSpace))) + joinIso₁ + + induced : TotalSpaceHopf' → Susp (typ A) + induced = fst ∘ TotalSpaceHopf'→TotalSpace + + lem : (x y z : typ A) → (i j : I) → ua (μ-eq y) i + lem x y z i j = + fill (λ k → ua (μ-eq y) i) + (λ j → λ { (i = i0) → HSpace.μ e z x + ; (i = i1) → assocHSpace.μ-assoc e-ass z x y j}) + (inS (ua-gluePt (μ-eq y) i (HSpace.μ e z x))) + j + + v : (x : TotalSpaceHopf') → typ A ≃ Hopf (induced x) + v (inl x) = μ-eq x + v (inr x) = μ-eq x + v (push (x , y) i₁) = pp x y i₁ + where + pp : (x y : _) → PathP (λ i → typ A ≃ ua (μ-eq y) i) (μ-eq x) (μ-eq (HSpace.μ e x y)) + pp x y = ΣPathP (P , help) + where + P : PathP (λ z → typ A → ua (μ-eq y) z) (fst (μ-eq x)) + (fst (μ-eq (HSpace.μ e x y))) + P i z = lem x y z i i1 + + abstract + help : PathP (λ i₂ → isEquiv (P i₂)) (snd (μ-eq x)) + (snd (μ-eq (HSpace.μ e x y))) + help = toPathP (isPropIsEquiv _ _ _) + + v' : (a : typ A) (x : TotalSpaceHopf') → Σ[ x ∈ Susp (typ A) ] Hopf x + v' a x = (induced x) , fst (v x) a + + v'-equiv : (a : typ A) → isEquiv (v' a) + v'-equiv a = pRec (isPropIsEquiv _) + (J (λ a _ → isEquiv (v' a)) + (subst isEquiv (sym main) + isEquivTotalSpaceHopf'→TotalSpace)) + (conA (pt A) a) + where + help1 : (x : _) → fst ((v' (pt A)) x) ≡ fst (TotalSpaceHopf'→TotalSpace x) + help1 (inl x) = refl + help1 (inr x) = refl + help1 (push a i) = refl + + help2 : (x : _) + → PathP (λ i → Hopf (help1 x i)) + (snd ((v' (pt A)) x)) + (snd (TotalSpaceHopf'→TotalSpace x)) + help2 (inl x) = HSpace.μₗ e x + help2 (inr x) = HSpace.μₗ e x + help2 (push (x , y) i) j = + hcomp (λ k → λ {(i = i0) → HSpace.μₗ e x j + ; (i = i1) → assocHSpace.μ-assoc-filler e-ass x y j k + ; (j = i0) → lem x y (pt A) i k + ; (j = i1) → ua-gluePt (μ-eq y) i x}) + (ua-gluePt (μ-eq y) i (HSpace.μₗ e x j)) + + main : v' (pt A) ≡ TotalSpaceHopf'→TotalSpace + main i x = (help1 x i) , (help2 x i) + + megaPush : Type _ + megaPush = Pushout {A = TotalSpaceHopf'} (λ _ → tt) induced + + P : megaPush → Type _ + P (inl x) = typ A + P (inr x) = Hopf x + P (push a i) = ua (v a) i + + totalSpaceMegaPush : Type _ + totalSpaceMegaPush = Σ[ x ∈ megaPush ] P x + + totalSpaceMegaPush' : Type _ + totalSpaceMegaPush' = + Pushout {A = typ A × TotalSpaceHopf'} + {C = Σ[ x ∈ Susp (typ A) ] Hopf x} + fst + λ x → v' (fst x) (snd x) + + z : Iso totalSpaceMegaPush totalSpaceMegaPush' + z = compIso whe2 (compIso (equivToIso fl.flatten) whe) + where + module fl = FlatteningLemma ((λ _ → tt)) induced (λ x → typ A) Hopf v + + whe : Iso (Pushout fl.Σf fl.Σg) totalSpaceMegaPush' + Iso.fun whe (inl x) = inl (snd x) + Iso.fun whe (inr x) = inr x + Iso.fun whe (push a i) = push ((snd a) , (fst a)) i + Iso.inv whe (inl x) = inl (tt , x) + Iso.inv whe (inr x) = inr x + Iso.inv whe (push a i) = push (snd a , fst a) i + Iso.rightInv whe (inl x) = refl + Iso.rightInv whe (inr x) = refl + Iso.rightInv whe (push a i) = refl + Iso.leftInv whe (inl x) = refl + Iso.leftInv whe (inr x) = refl + Iso.leftInv whe (push a i) = refl + + whe2 : Iso totalSpaceMegaPush (Σ (Pushout (λ _ → tt) induced) fl.E) + Iso.fun whe2 (inl x , y) = inl x , y + Iso.fun whe2 (inr x , y) = inr x , y + Iso.fun whe2 (push a i , y) = push a i , y + Iso.inv whe2 (inl x , y) = inl x , y + Iso.inv whe2 (inr x , y) = inr x , y + Iso.inv whe2 (push a i , y) = push a i , y + Iso.rightInv whe2 (inl x , snd₁) = refl + Iso.rightInv whe2 (inr x , snd₁) = refl + Iso.rightInv whe2 (push a i , snd₁) = refl + Iso.leftInv whe2 (inl x , snd₁) = refl + Iso.leftInv whe2 (inr x , snd₁) = refl + Iso.leftInv whe2 (push a i , snd₁) = refl + + F : totalSpaceMegaPush' + → (Pushout {A = typ A × Σ (Susp (typ A)) Hopf} + fst + snd) + F (inl x) = inl x + F (inr x) = inr x + F (push (x , y) i) = push (x , v' x y) i + + G : (Pushout {A = typ A × Σ (Susp (typ A)) Hopf} + fst + snd) + → totalSpaceMegaPush' + G (inl x) = inl x + G (inr x) = inr x + G (push (x , y) i) = + hcomp (λ k → λ { (i = i0) → inl x + ; (i = i1) → inr (secEq (_ , v'-equiv x) y k)}) + (push (x , invEq (_ , v'-equiv x) y) i) + + zz : Iso totalSpaceMegaPush' + (Pushout {A = typ A × Σ (Susp (typ A)) Hopf} + fst + snd) + Iso.fun zz = F + Iso.inv zz = G + Iso.rightInv zz (inl x) = refl + Iso.rightInv zz (inr x) = refl + Iso.rightInv zz (push (x , y) i) j = + hcomp (λ k → λ { (i = i0) → inl x + ; (i = i1) → inr (secEq (_ , v'-equiv x) y k) + ; (j = i0) → F ( + hfill (λ k → λ { (i = i0) → inl x + ; (i = i1) → inr (secEq (_ , v'-equiv x) y k)}) + (inS (push (x , invEq (_ , v'-equiv x) y) i)) k) + ; (j = i1) → push (x , (secEq (_ , v'-equiv x) y k)) i}) + (push (x , (secEq (_ , v'-equiv x) y i0)) i) + Iso.leftInv zz (inl x) = refl + Iso.leftInv zz (inr x) = refl + Iso.leftInv zz (push (x , y) i) j = + hcomp (λ k → λ { (i = i0) → inl x + ; (i = i1) → inr (secEq (v' x , v'-equiv x) (v' x y) (j ∨ k)) + ; (j = i1) → push (x , y) i}) + (hcomp (λ k → λ { (i = i0) → inl x + ; (i = i1) → inr (retEq≡secEq (v' x , v'-equiv x) y (~ k) j) + ; (j = i1) → push (x , y) i + ; (j = i0) → push (x , invEq (v' x , v'-equiv x) (v' x y)) i}) + (push (x , retEq (v' x , v'-equiv x) y j) i)) + + IsoJoin₁ : Iso totalSpaceMegaPush (join (typ A) (join (typ A) (typ A))) + IsoJoin₁ = + compIso z (compIso zz (compIso (equivToIso (joinPushout≃join _ _)) + (pathToIso (cong (join (typ A)) (isoToPath IsoTotalSpaceJoin))))) diff --git a/Cubical/ZCohomology/Gysin.agda b/Cubical/ZCohomology/Gysin.agda index c6495485b..124798ccc 100644 --- a/Cubical/ZCohomology/Gysin.agda +++ b/Cubical/ZCohomology/Gysin.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --safe --experimental-lossy-unification --no-forcing #-} +{-# OPTIONS --safe --experimental-lossy-unification #-} module Cubical.ZCohomology.Gysin where open import Cubical.ZCohomology.Base @@ -14,6 +14,8 @@ open import Cubical.ZCohomology.RingStructure.GradedCommutativity open import Cubical.Functions.Embedding +open import Cubical.Data.Fin + open import Cubical.Foundations.HLevels open import Cubical.Foundations.Transport open import Cubical.Foundations.Function @@ -25,6 +27,8 @@ open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Equiv open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.Foundations.Univalence + open import Cubical.Data.Empty renaming (rec to ⊥-rec) open import Cubical.Data.Sigma open import Cubical.Data.Int hiding (_+'_) @@ -35,15 +39,17 @@ open import Cubical.Data.Bool open import Cubical.Algebra.Group renaming (ℤ to ℤGroup ; Unit to UnitGroup) hiding (Bool) +open import Cubical.HITs.Pushout.Flattening open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.EilenbergSteenrod open import Cubical.HITs.Pushout open import Cubical.HITs.Sn open import Cubical.HITs.Susp -open import Cubical.HITs.S1 +open import Cubical.HITs.S1 renaming (_·_ to _*_) open import Cubical.HITs.Truncation renaming (rec to trRec ; elim to trElim ; elim2 to trElim2) open import Cubical.HITs.SetTruncation - renaming (rec to sRec ; elim to sElim ; elim2 to sElim2 ; map to sMap) + renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap) open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim) @@ -51,6 +57,120 @@ open import Cubical.Algebra.AbGroup open import Cubical.Homotopy.Loopspace +open import Cubical.HITs.Join + +open import Cubical.Homotopy.Hopf + +characFunSpaceS¹ : ∀ {ℓ} {A : Type ℓ} → + Iso (S₊ 1 → A) (Σ[ x ∈ A ] x ≡ x) +Iso.fun characFunSpaceS¹ f = f base , cong f loop +Iso.inv characFunSpaceS¹ (x , p) base = x +Iso.inv characFunSpaceS¹ (x , p) (loop i) = p i +Iso.rightInv characFunSpaceS¹ _ = refl +Iso.leftInv characFunSpaceS¹ f = funExt λ { base → refl ; (loop i) → refl} + +characFunSpace : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + → Iso (join A B → C) + (Σ[ f ∈ (A → C) ] Σ[ g ∈ (B → C) ] + ((a : A) (b : B) → f a ≡ g b)) +Iso.fun characFunSpace f = (f ∘ inl) , ((f ∘ inr) , (λ a b → cong f (push a b))) +Iso.inv characFunSpace (f , g , p) (inl x) = f x +Iso.inv characFunSpace (f , g , p) (inr x) = g x +Iso.inv characFunSpace (f , g , p) (push a b i) = p a b i +Iso.rightInv characFunSpace (f , g , p) = refl +Iso.leftInv characFunSpace f = + funExt λ { (inl x) → refl ; (inr x) → refl ; (push a b i) → refl} + +coHomS¹-ish : (n : ℕ) → Type _ +coHomS¹-ish n = hLevelTrunc 3 (S₊ 1 → coHomK (3 +ℕ n)) + +fib : (n : ℕ) → coHomS¹-ish n → Type _ +fib n x = + trRec {B = TypeOfHLevel ℓ-zero 2} (isOfHLevelTypeOfHLevel 2) + (λ f → ∥ (Σ[ g ∈ (S₊ 3 → coHomK (3 +ℕ n)) ] + ((a : S₊ 1) (b : S₊ 3) → f a ≡ g b)) ∥₂ , squash₂) x .fst + +contrFstΣ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} + → (e : isContr A) + → Iso (Σ A B) (B (fst e)) +Iso.fun (contrFstΣ {B = B} e) (a , b) = subst B (sym (snd e a)) b +Iso.inv (contrFstΣ {B = B} e) b = (fst e) , b +Iso.rightInv (contrFstΣ {B = B} e) b = cong (λ x → subst B x b) h ∙ transportRefl b + where + h : sym (snd e (fst e)) ≡ refl + h = isProp→isSet (isContr→isProp e) _ _ _ _ +Iso.leftInv (contrFstΣ {B = B} e) b = + ΣPathP ((snd e (fst b)) + , (λ j → transp (λ i → B (snd e (fst b) (~ i ∨ j))) j (snd b))) + +Iso1 : (n : ℕ) → Iso (coHom (3 +ℕ n) (join S¹ (S₊ 3))) ∥ Σ (coHomS¹-ish n) (fib n) ∥₂ +Iso1 n = compIso (setTruncIso characFunSpace) Iso2 + where + Iso2 : Iso _ ∥ Σ (coHomS¹-ish n) (fib n) ∥₂ + Iso.fun Iso2 = sMap λ F → ∣ fst F ∣ , ∣ (fst (snd F)) , (snd (snd F)) ∣₂ + Iso.inv Iso2 = + sRec squash₂ + (uncurry (trElim (λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelSuc 2 squash₂) + λ f → sRec squash₂ λ p → ∣ f , p ∣₂)) + Iso.rightInv Iso2 = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + (uncurry (trElim (λ _ → isOfHLevelΠ 3 λ _ → isProp→isOfHLevelSuc 2 (squash₂ _ _)) + λ f → sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl )) + Iso.leftInv Iso2 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl + +isContrcoHomS¹-ish : (n : ℕ) → isContr (coHomS¹-ish n) +isContrcoHomS¹-ish n = isOfHLevelRetractFromIso 0 (mapCompIso characFunSpaceS¹) isContrEnd + where + isContrEnd : isContr (hLevelTrunc 3 (Σ[ x ∈ coHomK (3 +ℕ n) ] x ≡ x)) + fst isContrEnd = ∣ 0ₖ _ , refl ∣ + snd isContrEnd = + trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + (uncurry (coHomK-elim _ + (λ _ → isOfHLevelΠ (suc (suc (suc n))) + λ _ → isOfHLevelPlus' {n = n} 3 (isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _)) + (λ p → (trRec (isOfHLevelPlus' {n = n} 3 (isOfHLevelTrunc 3) _ _) + (λ p i → ∣ (0ₖ (3 +ℕ n) , p i) ∣) + (Iso.fun (PathIdTruncIso _) + (isContr→isProp + (isConnectedPath _ (isConnectedKn (2 +ℕ n)) (0ₖ _) (0ₖ _)) ∣ refl ∣ ∣ p ∣)))))) + +Iso2' : (n : ℕ) → Iso (∥ Σ (coHomS¹-ish n) (fib n) ∥₂) (fib n ∣ (λ _ → 0ₖ _) ∣) +Iso2' n = compIso (setTruncIso (contrFstΣ centerChange)) + (setTruncIdempotentIso squash₂) + where + centerChange : isContr (coHomS¹-ish n) + fst centerChange = ∣ (λ _ → 0ₖ _) ∣ + snd centerChange y = isContr→isProp (isContrcoHomS¹-ish n) _ _ + +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Relation.Nullary +ll5 : (n : ℕ) → ¬ (n ≡ 2) → isContr (fib n ∣ (λ _ → 0ₖ _) ∣) +ll5 n p = + isOfHLevelRetractFromIso 0 + (compIso + (setTruncIso + (compIso (Σ-cong-iso-snd λ f → + (compIso characFunSpaceS¹ (invIso (Σ-cong-iso-fst (iso funExt⁻ funExt (λ _ → refl) λ _ → refl))))) + (compIso ΣΣ (contrFstΣ (isContrSingl _))))) + (compIso (setTruncIso (iso funExt⁻ funExt (λ _ → refl) λ _ → refl)) + (compIso (setTruncIso (codomainIso (congIso (invIso (Iso-Kn-ΩKn+1 _))))) + ((compIso (invIso (fst (coHom≅coHomΩ _ (S₊ _)))) + (fst (Hⁿ-Sᵐ≅0 n 2 p))))))) + isContrUnit + +record ExactSeqℤ {ℓ : Level} (G : ℤ → Group ℓ) : Type ℓ where + field + maps : ∀ n → GroupHom (G n) (G (sucℤ n)) + im⊂ker : ∀ n → ∀ g → isInIm (maps n) g → isInKer (maps (sucℤ n)) g + ker⊂im : ∀ n → ∀ g → isInKer (maps (sucℤ n)) g → isInIm (maps n) g + +record ExactSeqℕ {ℓ : Level} (G : ℕ → Group ℓ) : Type ℓ where + field + maps : ∀ n → GroupHom (G n) (G (suc n)) + im⊂ker : ∀ n → ∀ g → isInIm (maps n) g → isInKer (maps (suc n)) g + ker⊂im : ∀ n → ∀ g → isInKer (maps (suc n)) g → isInIm (maps n) g + + module _ {ℓ : Level} {A : Pointed ℓ} {n : ℕ} where funTyp : Type _ @@ -734,6 +854,7 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) (conB : (x y : typ B) → ∥ x ≡ y ∥₂) (n : ℕ) (Q-is : Iso (typ (Q B P (pt B))) (S₊ n)) (Q-is-ptd : Iso.fun Q-is (pt (Q B P (pt B))) ≡ snd (S₊∙ n)) where + 0-connB : (x y : typ B) → ∥ x ≡ y ∥ 0-connB x y = sRec (isProp→isSet squash) (∥_∥.∣_∣) (conB x y) @@ -759,7 +880,373 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) snd (⌣-hom i) = makeIsGroupHom λ f g → rightDistr-⌣ _ _ f g e - p-hom : (i : ℕ) → GroupHom (coHomGr (i +' n) (typ B)) (coHomGr (i +' n) E') + p-hom : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr i E') p-hom i = coHomMorph _ p + E-susp : (i : ℕ) → GroupHom (coHomGr i E') (coHomRedGrDir (suc i) (E'̃ , inl tt)) + fst (E-susp i) = sMap λ f → (λ { (inl x) → 0ₖ _ + ; (inr x) → 0ₖ _ + ; (push a j) → Kn→ΩKn+1 _ (f a) j}) , refl + snd (E-susp zero) = + makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f g → + cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn 1) + (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a j) k → (Kn→ΩKn+1-hom zero (f a) (g a) + ∙ ∙≡+₁ (Kn→ΩKn+1 _ (f a)) (Kn→ΩKn+1 _ (g a))) k j}))) + snd (E-susp (suc i)) = + makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f g → + cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a j) k → (Kn→ΩKn+1-hom _ (f a) (g a) + ∙ ∙≡+₂ _ (Kn→ΩKn+1 _ (f a)) (Kn→ΩKn+1 _ (g a))) k j}))) + + module cofibSeq where + j* : (i : ℕ) → GroupHom (coHomRedGrDir i (E'̃ , (inl tt))) (coHomGr i (typ B)) + fst (j* i) = sMap λ f → λ x → fst f (inr x) + snd (j* zero) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) + snd (j* (suc zero)) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) + snd (j* (suc (suc i₁))) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) + + Im-j⊂Ker-p : (i : ℕ) (x : _) → isInIm (j* i) x → isInKer (p-hom i) x + Im-j⊂Ker-p i = + sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ f → pRec (squash₂ _ _) + (uncurry (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ g P → subst (isInKer (p-hom i)) P + (cong ∣_∣₂ (funExt λ x → cong (g .fst) (sym (push x)) ∙ g .snd)))) + + Ker-p⊂Im-j : (i : ℕ) (x : _) → isInKer (p-hom i) x → isInIm (j* i) x + Ker-p⊂Im-j i = + sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) + λ f ker → pRec squash + (λ id → ∣ ∣ (λ { (inl x) → 0ₖ _ + ; (inr x) → f x + ; (push a i₁) → funExt⁻ id a (~ i₁)}) , refl ∣₂ , refl ∣) + (Iso.fun PathIdTrunc₀Iso ker) + + Im-p⊂Ker-Susp : (i : ℕ) (x : _) → isInIm (p-hom i) x → isInKer (E-susp i) x + Im-p⊂Ker-Susp i = + sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ f → pRec (squash₂ _ _) + (uncurry (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ g y → subst (isInKer (E-susp i)) y + (cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { (inl x) → refl + ; (inr x) → sym (Kn→ΩKn+1 _ (g x)) + ; (push a j) k → Kn→ΩKn+1 i (g (fst a)) (~ k ∧ j)}))))) + open import Cubical.Foundations.Path + Ker-Susp⊂Im-p : (i : ℕ) (x : _) → isInKer (E-susp i) x → isInIm (p-hom i) x + Ker-Susp⊂Im-p i = + sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) + λ f ker → pRec squash + (λ id → ∣ ∣ (λ x → ΩKn+1→Kn i (sym (funExt⁻ (cong fst id) (inr x)))) ∣₂ + , cong ∣_∣₂ (funExt (λ { (a , b) → + cong (ΩKn+1→Kn i) (lUnit _ ∙ cong (_∙ sym (funExt⁻ (λ i₁ → fst (id i₁)) (inr a))) (sym (flipSquare (cong snd id)) + ∙∙ (PathP→compPathR (cong (funExt⁻ (cong fst id)) (push (a , b)))) + ∙∙ assoc _ _ _ + ∙ sym (rUnit _)) + ∙ (sym (assoc _ _ _) + ∙∙ cong (Kn→ΩKn+1 i (f (a , b)) ∙_) (rCancel _) + ∙∙ sym (rUnit _))) + ∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f (a , b))})) ∣) + (Iso.fun PathIdTrunc₀Iso ker) + + Im-Susp⊂Ker-j : (i : ℕ) (x : _) → isInIm (E-susp i) x → isInKer (cofibSeq.j* (suc i)) x + Im-Susp⊂Ker-j i = + sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ g → pRec (squash₂ _ _) + (uncurry (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ f id → pRec (squash₂ _ _) + (λ P → subst (isInKer (cofibSeq.j* (suc i))) (cong ∣_∣₂ P) + (cong ∣_∣₂ refl)) + (Iso.fun PathIdTrunc₀Iso id))) + + Ker-j⊂Im-Susp : (i : ℕ) (x : _) → isInKer (cofibSeq.j* (suc i)) x → isInIm (E-susp i) x + Ker-j⊂Im-Susp i = + sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) + λ f ker + → pRec squash + (λ p → ∣ ∣ (λ x → ΩKn+1→Kn _ (sym (snd f) ∙∙ cong (fst f) (push x) ∙∙ funExt⁻ p (fst x))) ∣₂ + , cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) + (funExt (λ { (inl x) → sym (snd f) + ; (inr x) → sym (funExt⁻ p x) + ; (push a j) k → h f p a k j}))) ∣) + (Iso.fun PathIdTrunc₀Iso ker) + where + h : (f : (E'̃ , inl tt) →∙ coHomK-ptd (suc i)) + → (p : (fst f ∘ inr) ≡ (λ _ → 0ₖ (suc i))) + → (a : E B P) + → PathP (λ i → snd f (~ i) ≡ p (~ i) (fst a)) + (Kn→ΩKn+1 i (ΩKn+1→Kn i (sym (snd f) ∙∙ cong (fst f) (push a) ∙∙ funExt⁻ p (fst a)))) + (cong (fst f) (push a)) + h f p a = Iso.rightInv (Iso-Kn-ΩKn+1 i) _ + ◁ λ i j → doubleCompPath-filler (sym (snd f)) (cong (fst f) (push a)) (funExt⁻ p (fst a)) (~ i) j + + ϕ∘j : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr (i +' n) (typ B)) + ϕ∘j i = compGroupHom (fst (fst (ϕ i)) , snd (ϕ i)) (cofibSeq.j* (i +' n)) + + +'-suc : (i n : ℕ) → (suc i +' n) ≡ suc (i +' n) + +'-suc zero zero = refl + +'-suc (suc i₁) zero = refl + +'-suc zero (suc n) = refl + +'-suc (suc i₁) (suc n) = refl + + private + h : ∀ {ℓ ℓ'} {n m : ℕ} (G : ℕ → Group ℓ) (H : Group ℓ') (p : n ≡ m) + → GroupEquiv (G n) H + → GroupEquiv (G m) H + h {n = n} G H = + J (λ m _ → GroupEquiv (G n) H → GroupEquiv (G m) H) + λ p → p + + h-ret : ∀ {ℓ ℓ'} {n m : ℕ} (G : ℕ → Group ℓ) (H : Group ℓ') + → (e : GroupEquiv (G n) H) + → (p : n ≡ m) + → (x : G m .fst) → invEq (fst e) (fst (fst (h G H p e)) x) ≡ subst (λ x → G x .fst) (sym p) x + h-ret G H e = + J (λ m p → ((x : G m .fst) → invEq (fst e) (fst (fst (h G H p e)) x) ≡ subst (λ x → G x .fst) (sym p) x)) + λ x → cong (invEq (fst e) ) + (λ i → transportRefl (transportRefl (fst (fst e) (transportRefl (transportRefl x i) i)) i) i) + ∙∙ retEq (fst e) x + ∙∙ sym (transportRefl _) + + isEquivϕ' : (i : ℕ) → GroupEquiv (coHomRedGrDir (suc (i +' n)) (E'̃ , inl tt)) + (coHomGr (suc i) (typ B)) + isEquivϕ' i = (h (λ n → coHomRedGrDir n (E'̃ , inl tt)) _ (+'-suc i n) + (invGroupEquiv (ϕ (suc i)))) + + ϕ' : (i : ℕ) → GroupHom (coHomRedGrDir (suc (i +' n)) (E'̃ , inl tt)) + (coHomGr (suc i) (typ B)) + ϕ' i = fst (fst (isEquivϕ' i)) , snd (isEquivϕ' i) + + susp∘ϕ : (i : ℕ) → GroupHom (coHomGr (i +' n) E') (coHomGr (suc i) (typ B)) + susp∘ϕ i = compGroupHom (E-susp (i +' n)) (ϕ' i) + + Im-ϕ∘j⊂Ker-p : (i : ℕ) (x : _) → isInIm (ϕ∘j i) x → isInKer (p-hom _) x + Im-ϕ∘j⊂Ker-p i x p = + cofibSeq.Im-j⊂Ker-p _ x + (pRec squash (uncurry (λ f p → ∣ fst (fst (ϕ _)) f , p ∣)) p) + + Ker-p⊂Im-ϕ∘j : (i : ℕ) (x : _) → isInKer (p-hom _) x → isInIm (ϕ∘j i) x + Ker-p⊂Im-ϕ∘j i x p = + pRec squash (uncurry (λ f p → ∣ (invEq (fst (ϕ _)) f) + , (cong (fst (cofibSeq.j* _)) (secEq (fst (ϕ _)) f) ∙ p) ∣)) + (cofibSeq.Ker-p⊂Im-j _ x p) + + + Im-p⊂KerSusp∘ϕ : (i : ℕ) (x : _) → isInIm (p-hom _) x → isInKer (susp∘ϕ i) x + Im-p⊂KerSusp∘ϕ i x p = cong (fst (ϕ' _)) (Im-p⊂Ker-Susp _ x p) ∙ IsGroupHom.pres1 (snd (ϕ' _)) + + KerSusp∘ϕ⊂Im-p : (i : ℕ) (x : _) → isInKer (susp∘ϕ i) x → isInIm (p-hom _) x + KerSusp∘ϕ⊂Im-p i x p = + Ker-Susp⊂Im-p _ x (sym (retEq (fst (isEquivϕ' _)) _) + ∙ (cong (invEq (fst (isEquivϕ' _))) p + ∙ IsGroupHom.pres1 (snd (invGroupEquiv (isEquivϕ' _))))) + + Im-Susp∘ϕ⊂Ker-ϕ∘j : (i : ℕ) → (x : _) → isInIm (susp∘ϕ i) x → isInKer (ϕ∘j (suc i)) x + Im-Susp∘ϕ⊂Ker-ϕ∘j i x = + pRec (squash₂ _ _) + (uncurry λ f → J (λ x p → isInKer (ϕ∘j (suc i)) x) + ((λ i → fst (cofibSeq.j* _) (fst (fst (ϕ _)) (fst (ϕ' _) (fst (E-susp _) f)))) + ∙∙ cong (fst (cofibSeq.j* _)) + ((h-ret (λ n → coHomRedGrDir n (E'̃ , inl tt)) _ + (invGroupEquiv (ϕ (suc i))) (+'-suc i n)) (fst (E-susp _) f)) + ∙∙ (natTranspLem _ (λ n → fst (cofibSeq.j* n)) (sym (+'-suc i n)) + ∙ cong (subst (λ z → coHomGr z (typ B) .fst) (sym (+'-suc i n))) + (Im-Susp⊂Ker-j _ (fst (E-susp (i +' n)) f) ∣ f , refl ∣) + ∙ tLem i n))) + where + tLem : (i n : ℕ) → subst (λ z → coHomGr z (typ B) .fst) (sym (+'-suc i n)) + (0ₕ _) ≡ 0ₕ _ + tLem zero zero = refl + tLem zero (suc n) = refl + tLem (suc i₁) zero = refl + tLem (suc i₁) (suc n) = refl + + + Ker-ϕ∘j⊂Im-Susp∘ϕ : (i : ℕ) (x : _) + → isInKer (ϕ∘j (suc i)) x → isInIm (susp∘ϕ i) x + Ker-ϕ∘j⊂Im-Susp∘ϕ i x p = + pRec squash + (uncurry (λ f p → ∣ f , cong (fst (fst (isEquivϕ' i))) p ∙ secEq (fst (isEquivϕ' _)) x ∣)) + (Ker-j⊂Im-Susp _ (invEq (fst (isEquivϕ' _)) x) + ((cong (cofibSeq.j* (suc (i +' n)) .fst ) lem2 + ∙ natTranspLem _ (λ n → cofibSeq.j* n .fst) (+'-suc i n)) + ∙∙ cong (transport (λ j → (coHomGr (+'-suc i n j) (typ B) .fst))) p + ∙∙ h2 i n)) + where + lem2 : invEq (fst (isEquivϕ' i)) x ≡ transport (λ j → coHomRed (+'-suc i n j) (E'̃ , inl tt)) (fst (fst (ϕ _)) x) + lem2 = cong (transport (λ j → coHomRed (+'-suc i n j) (E'̃ , inl tt))) + (transportRefl _ ∙ cong (fst (fst (ϕ _))) + λ i → transportRefl (transportRefl x i) i) + + h2 : (i n : ℕ) → transport (λ j → coHomGr (+'-suc i n j) (typ B) .fst) + (GroupStr.1g (coHomGr (suc i +' n) (typ B) .snd)) ≡ 0ₕ _ + h2 zero zero = refl + h2 zero (suc n) = refl + h2 (suc i₁) zero = refl + h2 (suc i₁) (suc n) = refl + + ϕ∘j≡ : (i : ℕ) → ϕ∘j i ≡ ⌣-hom i + ϕ∘j≡ i = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ _ → refl)) + +open import Cubical.Experiments.Brunerie +open import Cubical.HITs.Hopf + +open import Cubical.HITs.Join + +module fibS1 = hopfBase S1-AssocHSpace (sphereElim2 _ (λ _ _ → squash) ∣ refl ∣) + +S¹Hopf : S₊ 2 → Type +S¹Hopf = fibS1.Hopf + +TotalHopf' : Type _ +TotalHopf' = Σ (S₊ 2) S¹Hopf + +CP2 : Type _ +CP2 = fibS1.megaPush + +fibr : CP2 → Type _ +fibr = fibS1.P + +hopf : join S¹ S¹ → S₊ 2 +hopf x = fst (JoinS¹S¹→TotalHopf x) + +E* : Type _ +E* = fibS1.totalSpaceMegaPush + +IsoE' : Iso E* (join S¹ (join S¹ S¹)) +IsoE' = fibS1.IsoJoin₁ + +IsoE2 : (join S¹ (join S¹ S¹)) ≡ join S¹ (S₊ 3) +IsoE2 = cong (join S¹) (sym S³≡joinS¹S¹ ∙ isoToPath IsoS³S3) + +CP' : Type _ +CP' = Pushout (λ _ → tt) hopf + +conCP2 : (x y : CP2) → ∥ x ≡ y ∥₂ +conCP2 x y = sRec2 squash₂ (λ p q → ∣ p ∙ sym q ∣₂) (conCP2' x) (conCP2' y) + where + conCP2' : (x : CP2) → ∥ x ≡ inl tt ∥₂ + conCP2' (inl x) = ∣ refl ∣₂ + conCP2' (inr x) = sphereElim 1 {A = λ x → ∥ inr x ≡ inl tt ∥₂} (λ _ → squash₂) ∣ sym (push (inl base)) ∣₂ x + conCP2' (push a i₁) = ll a i₁ + where + h2 : ∀ {ℓ} {A : fibS1.TotalSpaceHopf' → Type ℓ} → ((a : _) → isProp (A a)) + → A (inl base) + → ((a : fibS1.TotalSpaceHopf') → A a) + h2 {A = A} p b = + PushoutToProp p (sphereElim 0 (λ _ → p _) b) + (sphereElim 0 (λ _ → p _) (subst A (push (base , base)) b)) + + ll : (a : fibS1.TotalSpaceHopf') → PathP (λ i → ∥ Path CP2 (push a i) (inl tt) ∥₂) (conCP2' (inl tt)) (conCP2' (inr (fibS1.induced a))) + ll = h2 (λ _ → isOfHLevelPathP' 1 squash₂ _ _) λ j → ∣ (λ i → push (inl base) (~ i ∧ j)) ∣₂ + +module GysinS1 = Gysin (CP2 , inl tt) fibr conCP2 2 idIso refl + +PushoutReplaceBase : + ∀ {ℓ ℓ' ℓ''} {A B : Type ℓ} {C : Type ℓ'} {D : Type ℓ''} {f : A → C} {g : A → D} + (e : B ≃ A) → Pushout (f ∘ fst e) (g ∘ fst e) ≡ Pushout f g +PushoutReplaceBase {f = f} {g = g} = + EquivJ (λ _ e → Pushout (f ∘ fst e) (g ∘ fst e) ≡ Pushout f g) + refl + +isContrH³E : isContr (coHom 3 (GysinS1.E')) +isContrH³E = + subst isContr + (sym (isoToPath (compIso (Iso1 0) (Iso2' 0))) + ∙ cong (coHom 3) (sym (isoToPath IsoE' ∙ IsoE2))) + (ll5 0 (snotz ∘ sym)) + +isContrH⁵E : isContr (coHom 4 (GysinS1.E')) +isContrH⁵E = + subst isContr + (sym (isoToPath (compIso (Iso1 1) (Iso2' 1))) + ∙ cong (coHom 4) (sym (isoToPath IsoE' ∙ IsoE2))) + (ll5 1 λ p → snotz (sym (cong predℕ p))) + +HopfA*A : ∀ {ℓ} {A : Type ℓ} + → (join A A → Type ℓ) + → {!GysinS1.susp∘ϕ 1!} +HopfA*A = {!!} + +S₊2→S₊2 : CP' → Type ℓ-zero +S₊2→S₊2 (inl x) = {!!} +S₊2→S₊2 (inr x) = {!x!} +S₊2→S₊2 (push a i₁) = {!!} + + where + {- + fill (λ k → ua (μ-eq y) i) + (λ k → λ {(i = i0) → HSpace.μ e (pt A) x + ; (i = i1) → assocHSpace.μ-assoc e-ass (pt A) x y k}) + (inS ((ua-gluePt (μ-eq y) i (HSpace.μ e (pt A) x)))) + j -} + +{- +Goal: snd (v' (pt A) (push a i₁)) ≡ + ua-gluePt (μ-eq (snd a)) i₁ (fst a) +———— Boundary —————————————————————————————————————————————— +i₁ = i0 ⊢ HSpace.μₗ e (fst a) +i₁ = i1 ⊢ HSpace.μₗ e (HSpace.μ e (fst a) (snd a)) +-} + + -- help : (x : _) → (v' (pt A)) x ≡ TotalSpaceHopf'→TotalSpace x + -- help (inl x) = ΣPathP (refl , HSpace.μₗ e x) + -- help (inr x) = ΣPathP (refl , (HSpace.μₗ e x)) + -- help (push (x , y) i) j = + -- comp (λ _ → Σ (Susp (typ A)) Hopf) + -- (λ k → λ {(i = i0) → merid y i , HSpace.μₗ e x j + -- ; (i = i1) → merid y i , assocHSpace.μ-assoc-filler e-ass x y j k + -- ; (j = i0) → merid y i , hfill + -- (λ j → λ { (i = i0) → HSpace.μ e (pt A) x + -- ; (i = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j + -- }) + -- (inS (ua-gluePt (μ-eq y) i (HSpace.μ e (pt A) x))) + -- k + -- ; (j = i1) → merid y i , ua-gluePt (μ-eq y) i x}) + -- (merid y i , ua-gluePt (μ-eq y) i (HSpace.μₗ e x j)) + -- where + -- open import Cubical.Foundations.Path + + -- PPΣ : ∀ {ℓ} {A : Type ℓ} {f : A ≃ A} (p : f ≡ f) → {!!} + -- PPΣ = {!!} + + -- V : PathP (λ i₁ → hcomp + -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x + -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j + -- }) + -- (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x)) ≡ + -- ua-gluePt (μ-eq y) i₁ x) + -- (HSpace.μₗ e x) + -- (HSpace.μₗ e (HSpace.μ e x y)) -- (HSpace.μₗ e (HSpace.μ e (fst a) (snd a))) + -- V = transport (λ z → {!PathP (λ i₁ → hfill + -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x + -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j + -- }) + -- (inS (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x))) z ≡ ua-gluePt (μ-eq y) i₁ x) + -- ? ?!}) + -- {!hfill + -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x + -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j + -- }) + -- (inS (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x))) ?!} -- toPathP ({!!} ∙∙ {!!} ∙∙ {!!}) -- toPathP (flipSquare {!!}) -- hcomp {!!} {!!} + + -- P : Pushout {A = TotalSpaceHopf'} (λ _ → tt) induced → Type _ + -- P (inl x) = typ A + -- P (inr x) = Hopf x + -- P (push a i₁) = ua (v a) i₁ diff --git a/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda b/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda index 126228a53..8b20487de 100644 --- a/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda +++ b/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda @@ -39,7 +39,7 @@ private variable ℓ : Level -natTranspLem : ∀ {A B : ℕ → Type} {n m : ℕ} (a : A n) +natTranspLem : ∀ {ℓ} {A B : ℕ → Type ℓ} {n m : ℕ} (a : A n) (f : (n : ℕ) → (a : A n) → B n) (p : n ≡ m) → f m (subst A p a) ≡ subst B p (f n a) natTranspLem {A = A} {B = B} a f p = sym (substCommSlice A B f p a) From 57990e7c3313cfccf3c5e0de860d6b93fcb8f879 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 6 Sep 2021 17:03:27 +0200 Subject: [PATCH 64/89] stuff --- Cubical/Homotopy/Hopf.agda | 1 - Cubical/ZCohomology/Gysin.agda | 522 ++++++++++++++++++++++++++++++++- 2 files changed, 507 insertions(+), 16 deletions(-) diff --git a/Cubical/Homotopy/Hopf.agda b/Cubical/Homotopy/Hopf.agda index 5db7125a1..1e30fd1ec 100644 --- a/Cubical/Homotopy/Hopf.agda +++ b/Cubical/Homotopy/Hopf.agda @@ -58,7 +58,6 @@ open import Cubical.Homotopy.Loopspace open import Cubical.HITs.Join - retEq≡secEq : ∀ {ℓ} {A B : Type ℓ} (e : A ≃ B) → (x : _) → secEq e (e .fst x) ≡ cong (e .fst) (retEq e x) retEq≡secEq {A = A} = diff --git a/Cubical/ZCohomology/Gysin.agda b/Cubical/ZCohomology/Gysin.agda index 124798ccc..43d8946c7 100644 --- a/Cubical/ZCohomology/Gysin.agda +++ b/Cubical/ZCohomology/Gysin.agda @@ -32,7 +32,7 @@ open import Cubical.Foundations.Univalence open import Cubical.Data.Empty renaming (rec to ⊥-rec) open import Cubical.Data.Sigma open import Cubical.Data.Int hiding (_+'_) -open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_ ; _·_ to _·ℕ_) open import Cubical.Data.Nat.Order open import Cubical.Data.Unit open import Cubical.Data.Bool @@ -61,6 +61,186 @@ open import Cubical.HITs.Join open import Cubical.Homotopy.Hopf +SES→Iso : ∀ {ℓ ℓ'} {L R : Group ℓ-zero} + → {G : Group ℓ} {H : Group ℓ'} + → UnitGroup ≡ L + → UnitGroup ≡ R + → (lhom : GroupHom L G) (midhom : GroupHom G H) (rhom : GroupHom H R) + → ((x : _) → isInKer midhom x → isInIm lhom x) + → ((x : _) → isInKer rhom x → isInIm midhom x) + → isEquiv (fst midhom) +SES→Iso {R = R} {G = G} {H = H} = + J (λ L _ → UnitGroup ≡ R → + (lhom : GroupHom L G) (midhom : GroupHom G H) + (rhom : GroupHom H R) → + ((x : fst G) → isInKer midhom x → isInIm lhom x) → + ((x : fst H) → isInKer rhom x → isInIm midhom x) → + isEquiv (fst midhom)) + ((J (λ R _ → (lhom : GroupHom UnitGroup G) (midhom : GroupHom G H) + (rhom : GroupHom H R) → + ((x : fst G) → isInKer midhom x → isInIm lhom x) → + ((x : _) → isInKer rhom x → isInIm midhom x) → + isEquiv (fst midhom)) + main)) + where + main : (lhom : GroupHom UnitGroup G) (midhom : GroupHom G H) + (rhom : GroupHom H UnitGroup) → + ((x : fst G) → isInKer midhom x → isInIm lhom x) → + ((x : fst H) → isInKer rhom x → isInIm midhom x) → + isEquiv (fst midhom) + main lhom midhom rhom lexact rexact = + BijectionIsoToGroupEquiv {G = G} {H = H} + bijIso' .fst .snd + where + bijIso' : BijectionIso G H + BijectionIso.fun bijIso' = midhom + BijectionIso.inj bijIso' x inker = + pRec (GroupStr.is-set (snd G) _ _) + (λ s → sym (snd s) ∙ IsGroupHom.pres1 (snd lhom)) (lexact _ inker) + BijectionIso.surj bijIso' x = rexact x refl + +pres- : (e : GroupHom ℤGroup ℤGroup) → (a : ℤ) → fst e (- a) ≡ - fst e a +pres- e a = +Comm (fst e (- a)) (pos 0) + ∙ cong (_+ (fst e (- a))) (sym (l (fst e a)) ∙ +Comm (fst e a) (- fst e a)) + ∙∙ sym (+Assoc _ _ _) + ∙∙ cong (- (fst e a) +_) (pp ∙ l (fst e a)) + where + l : (a : ℤ) → a + (- a) ≡ 0 + l (pos zero) = refl + l (pos (suc zero)) = refl + l (pos (suc (suc n))) = predℤ+negsuc n (pos (suc (suc n))) ∙ l (pos (suc n)) + l (negsuc zero) = refl + l (negsuc (suc n)) = (cong sucℤ (sucℤ+pos n (negsuc (suc n)))) ∙ l (negsuc n) + + pp : fst e a + fst e (- a) ≡ fst e a + (- fst e a) + pp = sym (IsGroupHom.pres· (snd e) a (- a)) + ∙∙ cong (fst e) (l a) + ∙∙ (IsGroupHom.pres1 (snd e) + ∙ sym (l _)) + +GroupHomPres· : (e : GroupHom ℤGroup ℤGroup) → (a b : ℤ) → fst e (a · b) ≡ a · fst e b +GroupHomPres· e (pos zero) b = IsGroupHom.pres1 (snd e) +GroupHomPres· e (pos (suc n)) b = + IsGroupHom.pres· (snd e) b (pos n · b) ∙ cong (fst e b +_) (GroupHomPres· e (pos n) b) +GroupHomPres· e (negsuc zero) b = pres- e b +GroupHomPres· e (negsuc (suc n)) b = + IsGroupHom.pres· (snd e) (- b) (negsuc n · b) + ∙ cong₂ _+_ (pres- e b) (GroupHomPres· e (negsuc n) b) + +open import Cubical.Data.Sum +open import Cubical.Relation.Nullary +l2 : (n m : ℕ) → Σ[ a ∈ ℕ ] (negsuc n · pos (suc m)) ≡ negsuc a +l2 n zero = n , ·Comm (negsuc n) (pos 1) +l2 n (suc m) = h _ _ .fst , + (·Comm (negsuc n) (pos (suc (suc m))) + ∙∙ cong (negsuc n +_) (·Comm (pos (suc m)) (negsuc n) ∙ (l2 n m .snd)) + ∙∙ h _ _ .snd) + where + h : (x y : ℕ) → Σ[ a ∈ ℕ ] negsuc x + negsuc y ≡ negsuc a + h zero zero = 1 , refl + h zero (suc y) = (suc (suc y)) , +Comm (negsuc zero) (negsuc (suc y)) + h (suc x) zero = (suc (suc x)) , refl + h (suc x) (suc y) = (h (suc (suc x)) y .fst) , (predℤ+negsuc y (negsuc (suc x)) ∙ snd ((h (suc (suc x))) y)) + +zz : (n x : ℕ) → Σ[ a ∈ ℕ ] ((pos (suc x)) · pos (suc (suc n)) ≡ pos (suc (suc a))) +zz n zero = n , refl +zz n (suc x) = h _ _ (zz n x) + where + h : (x : ℤ) (n : ℕ) → Σ[ a ∈ ℕ ] (x ≡ pos (suc (suc a))) + → Σ[ a ∈ ℕ ] pos n + x ≡ pos (suc (suc a)) + h x n (a , p) = n +ℕ a + , cong (pos n +_) p ∙ cong sucℤ (sucℤ+pos a (pos n)) + ∙ sucℤ+pos a (pos (suc n)) ∙ (sym (pos+ (suc (suc n)) a)) + +lem22 : (n : ℕ) (x : ℤ) → ¬ pos 1 ≡ x · pos (suc (suc n)) +lem22 n (pos zero) p = snotz (injPos p) +lem22 n (pos (suc n₁)) p = snotz (injPos (sym (cong predℤ (snd (zz n n₁))) ∙ sym (cong predℤ p))) +lem22 n (negsuc n₁) p = posNotnegsuc _ _ (p ∙ l2 _ _ .snd) + + +grr : (e : GroupEquiv ℤGroup ℤGroup) (x : ℤ) → (fst (fst e) 1) ≡ x → abs (fst (fst e) 1) ≡ 1 +grr e (pos zero) p = + ⊥-rec (snotz (injPos (sym (retEq (fst e) 1) + ∙∙ (cong (fst (fst (invGroupEquiv e))) p) + ∙∙ IsGroupHom.pres1 (snd (invGroupEquiv e))))) +grr e (pos (suc zero)) p = cong abs p +grr e (pos (suc (suc n))) p = ⊥-rec (lem22 _ _ (h ∙ ·Comm (pos (suc (suc n))) (invEq (fst e) 1))) + where + + h : pos 1 ≡ _ + h = sym (retEq (fst e) 1) + ∙∙ (cong (fst (fst (invGroupEquiv e))) (p ∙ ·Comm 1 (pos (suc (suc n))))) + ∙∙ GroupHomPres· (_ , snd (invGroupEquiv e)) (pos (suc (suc n))) 1 +grr e (negsuc zero) p = cong abs p +grr e (negsuc (suc n)) p = ⊥-rec (lem22 _ _ l33) + where + h : fst (fst e) (negsuc zero) ≡ pos (suc (suc n)) + h = pres- (_ , snd e) (pos 1) ∙ cong -_ p + + compGroup : GroupEquiv ℤGroup ℤGroup + fst compGroup = isoToEquiv (iso -_ -_ -Involutive -Involutive) + snd compGroup = makeIsGroupHom -Dist+ + + compHom : GroupEquiv ℤGroup ℤGroup + compHom = compGroupEquiv compGroup e + + l32 : fst (fst compHom) (pos 1) ≡ pos (suc (suc n)) + l32 = h + + l33 : pos 1 ≡ invEq (fst compHom) (pos 1) · pos (suc (suc n)) + l33 = sym (retEq (fst compHom) (pos 1)) + ∙∙ cong (invEq (fst compHom)) l32 + ∙∙ (cong (invEq (fst compHom)) (·Comm (pos 1) (pos (suc (suc n)))) + ∙ GroupHomPres· (_ , (snd (invGroupEquiv compHom))) (pos (suc (suc n))) (pos 1) + ∙ ·Comm (pos (suc (suc n))) (invEq (fst compHom) (pos 1))) + +abs→⊎ : (x : ℤ) (n : ℕ) → abs x ≡ n → (x ≡ pos n) ⊎ (x ≡ - pos n) +abs→⊎ x n = J (λ n _ → (x ≡ pos n) ⊎ (x ≡ - pos n)) (help x) + where + help : (x : ℤ) → (x ≡ pos (abs x)) ⊎ (x ≡ - pos (abs x)) + help (pos n) = inl refl + help (negsuc n) = inr refl + +JGroupEquiv : ∀ {ℓ ℓ'} {G : Group ℓ} (P : (H : Group ℓ) → GroupEquiv G H → Type ℓ') + → P G idGroupEquiv + → ∀ {H} e → P H e +JGroupEquiv {G = G} P p {H} e = + transport (λ i → P (GroupPath G H .fst e i) + (transp (λ j → GroupEquiv G (GroupPath G H .fst e (i ∨ ~ j))) i e)) + (subst (P G) (sym l) p) + where + l : (transp (λ j → GroupEquiv G (GroupPath G H .fst e (~ j))) i0 e) ≡ idGroupEquiv + l = Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (Σ≡Prop (λ _ → isPropIsEquiv _) + (funExt λ x → (λ i → fst (fst (fst e .snd .equiv-proof + (transportRefl (fst (fst e) (transportRefl x i)) i)))) + ∙ retEq (fst e) x)) + +sesIsoPresGen : + ∀ (G : Group ℓ-zero) + → (iso : GroupEquiv ℤGroup G) + → (H : Group ℓ-zero) + → (iso2 : GroupEquiv ℤGroup H) + → (e : fst G) + → invEq (fst iso) e ≡ 1 + → (hom : GroupEquiv G H) + → abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1 +sesIsoPresGen G = JGroupEquiv (λ G iso → + (H : Group ℓ-zero) + (iso2 : GroupEquiv ℤGroup H) + → (e : fst G) + → invEq (fst iso) e ≡ 1 + → (hom : GroupEquiv G H) + → abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1) + λ H → JGroupEquiv (λ H iso2 → (e : ℤ) → e ≡ pos 1 → + (hom : GroupEquiv ℤGroup H) → + abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1) + λ e p hom → cong (abs ∘ (fst (fst hom))) p ∙ autoPres1 hom + where + autoPres1 : (e : GroupEquiv ℤGroup ℤGroup) + → abs (fst (fst e) 1) ≡ 1 + autoPres1 e = grr e _ refl + characFunSpaceS¹ : ∀ {ℓ} {A : Type ℓ} → Iso (S₊ 1 → A) (Σ[ x ∈ A ] x ≡ x) Iso.fun characFunSpaceS¹ f = f base , cong f loop @@ -661,6 +841,18 @@ module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) (isOfHLevelRetractFromIso 2 (fst (S1' n)) isSetℤ))) (conB (pt B) b) + + isConnB : isConnected 3 (typ B) + fst isConnB = ∣ pt B ∣ + snd isConnB = + trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → sRec (isOfHLevelTrunc 3 _ _) (cong ∣_∣ₕ) (conB (pt B) a) + + isPropPath : isProp (∥ pt B ≡ pt B ∥₂) + isPropPath = + isOfHLevelRetractFromIso 1 setTruncTrunc2Iso + (isContr→isProp (isConnectedPath _ isConnB (pt B) (pt B))) + c* : Σ[ c ∈ ((b : typ B) → (Q b →∙ coHomK-ptd n)) ] (c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) fst c* b = @@ -675,17 +867,35 @@ module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) (isPropPath (conB (pt B) (pt B)) ∣ refl ∣₂ i) .fst x) ∙ (λ i → transportRefl (genFunSpace n .fst (Iso.fun Q-is (transportRefl x i))) i) - where - isConnB : isConnected 3 (typ B) - fst isConnB = ∣ pt B ∣ - snd isConnB = - trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) - λ a → sRec (isOfHLevelTrunc 3 _ _) (cong ∣_∣ₕ) (conB (pt B) a) - isPropPath : isProp (∥ pt B ≡ pt B ∥₂) - isPropPath = - isOfHLevelRetractFromIso 1 setTruncTrunc2Iso - (isContr→isProp (isConnectedPath _ isConnB (pt B) (pt B))) + p-help : {b : fst B} (p : pt B ≡ b) → (subst (fst ∘ Q) (sym p) (snd (Q b))) ≡ (snd (Q (pt B))) + p-help {b = b} = + J (λ b p → subst (fst ∘ Q) (sym p) (snd (Q b)) ≡ snd (Q (pt B))) (transportRefl _) + + cEquiv : (b : fst B) (p : ∥ pt B ≡ b ∥₂) + → (c* .fst b) + ≡ sRec (is-setQ→K b) + (λ pp → (λ qb → genFunSpace n .fst (Iso.fun Q-is (subst (fst ∘ Q) (sym pp) qb))) + , cong (genFunSpace n .fst ∘ Iso.fun Q-is) (p-help pp) + ∙ ((λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) p + cEquiv b = + sElim (λ _ → isOfHLevelPath 2 (is-setQ→K b) _ _) + (J (λ b a → c* .fst b ≡ + sRec (is-setQ→K b) (λ pp → + (λ qb → + genFunSpace n .fst (Iso.fun Q-is (subst (fst ∘ Q) (sym pp) qb))) + , + cong (genFunSpace n .fst ∘ Iso.fun Q-is) (p-help pp) ∙ + (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd) + ∣ a ∣₂) + ((λ i → sRec (is-setQ→K (pt B)) + (J (λ b₁ _ → Q b₁ →∙ coHomK-ptd n) + ((λ x → genFunSpace n .fst (Iso.fun Q-is x)) , + (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) + (isPropPath (conB (pt B) (pt B)) ∣ refl ∣₂ i)) + ∙ →∙Homogeneous≡ (isHomogeneousKn n) + (transportRefl ((λ x → genFunSpace n .fst (Iso.fun Q-is x))) + ∙ funExt λ x → cong (genFunSpace n .fst ∘ Iso.fun Q-is) (sym (transportRefl x))))) module _ {ℓ ℓ'} (B : Pointed ℓ) (P : typ B → Type ℓ') where E : Type _ @@ -860,6 +1070,7 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) c = (c* B (Q B P) conB n Q-is Q-is-ptd .fst) c-ptd = (c* B (Q B P) conB n Q-is Q-is-ptd .snd) + cEq = cEquiv B (Q B P) conB n Q-is Q-is-ptd module w = Thom B P 0-connB n Q-is Q-is-ptd c c-ptd @@ -1135,6 +1346,9 @@ IsoE' = fibS1.IsoJoin₁ IsoE2 : (join S¹ (join S¹ S¹)) ≡ join S¹ (S₊ 3) IsoE2 = cong (join S¹) (sym S³≡joinS¹S¹ ∙ isoToPath IsoS³S3) +IsoTotalHopf' : Iso fibS1.TotalSpaceHopf' (join S¹ S¹) +IsoTotalHopf' = fibS1.joinIso₁ + CP' : Type _ CP' = Pushout (λ _ → tt) hopf @@ -1172,20 +1386,298 @@ isContrH³E = ∙ cong (coHom 3) (sym (isoToPath IsoE' ∙ IsoE2))) (ll5 0 (snotz ∘ sym)) -isContrH⁵E : isContr (coHom 4 (GysinS1.E')) -isContrH⁵E = +isContrH⁴E : isContr (coHom 4 (GysinS1.E')) +isContrH⁴E = subst isContr (sym (isoToPath (compIso (Iso1 1) (Iso2' 1))) ∙ cong (coHom 4) (sym (isoToPath IsoE' ∙ IsoE2))) (ll5 1 λ p → snotz (sym (cong predℕ p))) +genH2 = GysinS1.e + +S³→Groupoid : ∀ {ℓ} (P : join S¹ S¹ → Type ℓ) + → ((x : _) → isGroupoid (P x)) + → P (inl base) + → (x : _) → P x +S³→Groupoid P grp b = + transport (λ i → (x : (sym (isoToPath S³IsojoinS¹S¹) ∙ isoToPath IsoS³S3) (~ i)) + → P (transp (λ j → (sym (isoToPath S³IsojoinS¹S¹) ∙ isoToPath IsoS³S3) (~ i ∧ ~ j)) i x)) + (sphereElim _ (λ _ → grp _) b) + +TotalHopf→Gpd : ∀ {ℓ} (P : fibS1.TotalSpaceHopf' → Type ℓ) + → ((x : _) → isGroupoid (P x)) + → P (inl base) + → (x : _) → P x +TotalHopf→Gpd P grp b = + transport (λ i → (x : sym (isoToPath IsoTotalHopf') i) + → P (transp (λ j → isoToPath IsoTotalHopf' (~ i ∧ ~ j)) i x)) + (S³→Groupoid _ (λ _ → grp _) b) + +TotalHopf→Gpd' : ∀ {ℓ} (P : Σ (S₊ 2) fibS1.Hopf → Type ℓ) + → ((x : _) → isGroupoid (P x)) + → P (north , base) + → (x : _) → P x +TotalHopf→Gpd' P grp b = + transport (λ i → (x : ua (_ , fibS1.isEquivTotalSpaceHopf'→TotalSpace) i) + → P (transp (λ j → ua (_ , fibS1.isEquivTotalSpaceHopf'→TotalSpace) (i ∨ j)) i x)) + (TotalHopf→Gpd _ (λ _ → grp _) b) + +CP2→Groupoid : ∀ {ℓ} {P : CP2 → Type ℓ} + → ((x : _) → is2Groupoid (P x)) + → (b : P (inl tt)) + → (s2 : ((x : _) → P (inr x))) + → PathP (λ i → P (push (inl base) i)) b (s2 north) + → (x : _) → P x +CP2→Groupoid {P = P} grp b s2 pp (inl x) = b +CP2→Groupoid {P = P} grp b s2 pp (inr x) = s2 x +CP2→Groupoid {P = P} grp b s2 pp (push a i₁) = h23 a i₁ + where + h23 : (a : fibS1.TotalSpaceHopf') → PathP (λ i → P (push a i)) b (s2 _) + h23 = TotalHopf→Gpd _ (λ _ → isOfHLevelPathP' 3 (grp _) _ _) pp + +mm : (S₊ 2 → coHomK 2) → (CP2 → coHomK 2) +mm f = λ { (inl x) → 0ₖ _ + ; (inr x) → f x -ₖ f north + ; (push a i) → TotalHopf→Gpd (λ x → 0ₖ 2 ≡ f (fibS1.TotalSpaceHopf'→TotalSpace x .fst) -ₖ f north) + (λ _ → isOfHLevelTrunc 4 _ _) + (sym (rCancelₖ 2 (f north))) + a i} + +e-inv : coHomGr 2 (S₊ 2) .fst → coHomGr 2 CP2 .fst +e-inv = sMap mm + +cancel : (f : CP2 → coHomK 2) → ∥ mm (f ∘ inr) ≡ f ∥ +cancel f = + pRec squash + (λ p → pRec squash + (λ f → ∣ funExt f ∣) + (zz2 p)) + (h1 (f (inl tt))) + where + h1 : (x : coHomK 2) → ∥ x ≡ 0ₖ 2 ∥ + h1 = coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash) ∣ refl ∣ + + zz2 : (p : f (inl tt) ≡ 0ₖ 2) → ∥ ((x : CP2) → mm (f ∘ inr) x ≡ f x) ∥ + zz2 p = trRec squash (λ pp → + ∣ CP2→Groupoid (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) + (sym p) + (λ x → (λ i → f (inr x) -ₖ f (push (inl base) (~ i))) + ∙∙ (λ i → f (inr x) -ₖ p i) + ∙∙ rUnitₖ 2 (f (inr x))) + pp ∣) + l + where + l : hLevelTrunc 1 (PathP ((λ i₁ → mm (f ∘ inr) (push (inl base) i₁) ≡ f (push (inl base) i₁))) + (sym p) + (((λ i₁ → f (inr north) -ₖ f (push (inl base) (~ i₁))) ∙∙ + (λ i₁ → f (inr north) -ₖ p i₁) ∙∙ rUnitₖ 2 (f (inr north))))) + l = isConnectedPathP 1 (isConnectedPath 2 (isConnectedKn 1) _ _) _ _ .fst + +e' : GroupIso (coHomGr 2 CP2) (coHomGr 2 (S₊ 2)) +Iso.fun (fst e') = sMap (_∘ inr) +Iso.inv (fst e') = e-inv +Iso.rightInv (fst e') = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → trRec {B = Iso.fun (fst e') (Iso.inv (fst e') ∣ f ∣₂) ≡ ∣ f ∣₂} + (isOfHLevelPath 2 squash₂ _ _) + (λ p → cong ∣_∣₂ (funExt λ x → cong (λ y → (f x) -ₖ y) p ∙ rUnitₖ 2 (f x))) + (Iso.fun (PathIdTruncIso _) (isContr→isProp (isConnectedKn 1) ∣ f north ∣ ∣ 0ₖ 2 ∣)) +Iso.leftInv (fst e') = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → pRec (squash₂ _ _) (cong ∣_∣₂) (cancel f) +snd e' = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ f g → refl) + +grrg : GroupIso (coHomGr 2 CP2) ℤGroup +grrg = compGroupIso e' (Hⁿ-Sⁿ≅ℤ 1) + +⌣hom : GroupEquiv (coHomGr 2 CP2) (coHomGr 4 CP2) +fst (fst ⌣hom) = GysinS1.⌣-hom 2 .fst +snd (fst ⌣hom) = subst isEquiv (cong fst (GysinS1.ϕ∘j≡ 2)) h23 + where + h23 : isEquiv (GysinS1.ϕ∘j 2 .fst) + h23 = + SES→Iso + (GroupPath _ _ .fst (invGroupEquiv + (isContr→≃Unit isContrH³E + , makeIsGroupHom λ _ _ → refl))) + (GroupPath _ _ .fst (invGroupEquiv + (isContr→≃Unit isContrH⁴E + , makeIsGroupHom λ _ _ → refl))) + (GysinS1.susp∘ϕ 1) + (GysinS1.ϕ∘j 2) + (GysinS1.p-hom 4) + (GysinS1.Ker-ϕ∘j⊂Im-Susp∘ϕ _) + (GysinS1.Ker-p⊂Im-ϕ∘j _) +snd ⌣hom = GysinS1.⌣-hom 2 .snd + +⌣hom2 : GroupEquiv (coHomGr 0 CP2) (coHomGr 2 CP2) +fst (fst ⌣hom2) = GysinS1.⌣-hom 0 .fst +snd (fst ⌣hom2) = subst isEquiv (cong fst (GysinS1.ϕ∘j≡ 0)) h23 + where + cr : GroupHom (coHomGr 1 (S₊ 2)) (coHomGr 0 CP2) + fst cr = + sMap λ f → λ { (inl x) → 0 + ; (inr x) → ΩKn+1→Kn 0 ({!!} ∙∙ cong f {!merid x!} ∙∙ {!!}) + ; (push a i₁) → {!!}} + snd cr = {!!} + h23 : isEquiv (GysinS1.ϕ∘j 0 .fst) + h23 = + SES→Iso + {!!} + {!!} + cr + (GysinS1.ϕ∘j 0) + (GysinS1.p-hom 2) + (sElim {!!} (λ f inker → pRec {!!} (λ g → {!funExt⁻ g (inl tt)!}) (Iso.fun PathIdTrunc₀Iso inker))) + (GysinS1.Ker-p⊂Im-ϕ∘j _) +snd ⌣hom2 = {!!} + +isGenerator≃ℤ : ∀ {ℓ} (G : Group ℓ) (g : fst G) + → Type _ +isGenerator≃ℤ G g = + Σ[ e ∈ GroupIso G ℤGroup ] abs (Iso.fun (fst e) g) ≡ 1 + +⌣-pres1 : ∀ {ℓ} (G : Type ℓ) (n : ℕ) → Type _ +⌣-pres1 G n = + (z : Σ[ x ∈ coHomGr n G .fst ] isGenerator≃ℤ (coHomGr n G) x) + → isGenerator≃ℤ (coHomGr (n +' n) G) (z .fst ⌣ z .fst) + + +genr : coHom 2 CP2 +genr = ∣ CP2→Groupoid (λ _ → isOfHLevelTrunc 4) + (0ₖ _) + ∣_∣ + refl ∣₂ + +≡CP2 : (f g : CP2 → coHomK 2) → ∥ f ∘ inr ≡ g ∘ inr ∥ → Path (coHom 2 CP2) ∣ f ∣₂ ∣ g ∣₂ +≡CP2 f g = pRec (squash₂ _ _) + (λ p → pRec (squash₂ _ _) (λ id → trRec (squash₂ _ _) + (λ pp → cong ∣_∣₂ + (funExt (CP2→Groupoid (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) + id + (funExt⁻ p) + (compPathR→PathP pp)))) + (pp2 (f (inl tt)) (g (inl tt)) id + (cong f (push (inl base)) ∙ (funExt⁻ p north) ∙ cong g (sym (push (inl base)))))) + (conn (f (inl tt)) (g (inl tt)))) + + where + conn : (x y : coHomK 2) → ∥ x ≡ y ∥ + conn = coHomK-elim _ (λ _ → isSetΠ λ _ → isOfHLevelSuc 1 squash) + (coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash) ∣ refl ∣) + + pp2 : (x y : coHomK 2) (p q : x ≡ y) → hLevelTrunc 1 (p ≡ q) + pp2 x y = λ p q → Iso.fun (PathIdTruncIso _) (isContr→isProp (isConnectedPath _ (isConnectedKn 1) x y) ∣ p ∣ ∣ q ∣) + +s2 : ℤ +s2 = Iso.fun (fst grrg) (genr) + +-* : S¹ → S¹ +-* x = {!!} + +rUnit* : (x : S¹) → x * base ≡ x +rUnit* base = refl +rUnit* (loop i₁) = refl + + +mmInv : (x : S¹) → (merid (invLooper x) ∙ sym (merid base)) ≡ sym (merid x ∙ sym (merid base)) +mmInv x = {!!} + +meridP : (a x : S¹) → Path (Path (coHomK 2) _ _) (cong ∣_∣ₕ (merid (a * x) ∙ sym (merid base))) + ((cong ∣_∣ₕ (merid a ∙ sym (merid base))) ∙ (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) +meridP = wedgeconFun _ _ (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) + (λ x → {!!}) + {!!} + {!!} + +commS1 : (a x : S¹) → a * x ≡ x * a +commS1 = wedgeconFun _ _ (λ _ _ → isGroupoidS¹ _ _) + (sym ∘ rUnit*) + rUnit* + refl + +s233 : (a x : S¹) → (invEq (fibS1.μ-eq a) x) * a ≡ (invLooper a * x) * a +s233 a x = secEq (fibS1.μ-eq a) x ∙∙ cong (_* x) (lemmie a) ∙∙ assocHSpace.μ-assoc S1-AssocHSpace a (invLooper a) x ∙ commS1 _ _ + where + lemmie : (x : S¹) → ptSn 1 ≡ x * (invLooper x) + lemmie base = refl + lemmie (loop i) j = + hcomp (λ r → λ {(i = i0) → base ; (i = i1) → base ; (j = i0) → base}) + base + +ss23 : (a x : S¹) → invEq (fibS1.μ-eq a) x ≡ invLooper a * x +ss23 a x = sym (retEq (fibS1.μ-eq a) (invEq (fibS1.μ-eq a) x)) + ∙∙ cong (invEq (fibS1.μ-eq a)) (s233 a x) + ∙∙ retEq (fibS1.μ-eq a) (invLooper a * x) + +ll : GysinS1.e ≡ genr +ll = ≡CP2 _ _ ∣ funExt (λ x → funExt⁻ (cong fst (ll32 x)) south) ∣ + where + merid* : (a b : S¹) → Path (Path (coHomK 2) _ _) {!!} {!!} + merid* = {!!} + + kr : (x : Σ (S₊ 2) fibS1.Hopf) → Path (hLevelTrunc 4 _) ∣ fst x ∣ ∣ north ∣ + kr = uncurry λ { north → λ y → cong ∣_∣ₕ (merid base ∙ sym (merid y)) + ; south → λ y → cong ∣_∣ₕ (sym (merid y)) + ; (merid a i) → lem a i} + where + lem : (a : S¹) → PathP (λ i → (y : fibS1.Hopf (merid a i)) → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a i ∣ ∣ north ∣) + (λ y → cong ∣_∣ₕ (merid base ∙ sym (merid y))) + λ y → cong ∣_∣ₕ (sym (merid y)) + lem a = toPathP (funExt λ x → cong (transport ((λ i₁ → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a i₁ ∣ ∣ north ∣))) + ((λ i → (λ z → cong ∣_∣ₕ (merid base + ∙ sym (merid (transport (λ j → uaInvEquiv (fibS1.μ-eq a) (~ i) j) x))) z)) + ∙ λ i → cong ∣_∣ₕ (merid base + ∙ sym (merid (transportRefl (invEq (fibS1.μ-eq a) x) i)))) + ∙∙ (λ i → transp (λ i₁ → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a (i₁ ∨ i) ∣ ∣ north ∣) i + (compPath-filler' (cong ∣_∣ₕ (sym (merid a))) + (cong ∣_∣ₕ (merid base ∙ sym (merid (ss23 a x i)))) i)) + ∙∙ {!!} + ∙∙ {!!} + ∙∙ {!!}) + + psst : (x : S₊ 2) → Q (CP2 , inl tt) fibr (inr x) →∙ coHomK-ptd 2 + fst (psst x) north = ∣ north ∣ + fst (psst x) south = ∣ x ∣ + fst (psst x) (merid a i₁) = kr (x , a) (~ i₁) + snd (psst x) = refl + + ll32-fst : GysinS1.c (inr north) .fst ≡ psst north .fst + ll32-fst = cong fst (GysinS1.cEq (inr north) ∣ push (inl base) ∣₂) + ∙ (funExt l) + where + l : (qb : _) → + ∣ (subst (fst ∘ Q (CP2 , inl tt) fibr) (sym (push (inl base))) qb) ∣ + ≡ psst north .fst qb + l north = refl + l south = cong ∣_∣ₕ (sym (merid base)) + l (merid a i) j = + hcomp (λ k → λ { (i = i0) → ∣ merid a (~ k ∧ j) ∣ + ; (i = i1) → ∣ merid base (~ j) ∣ + ; (j = i0) → ∣ transportRefl (merid a i) (~ k) ∣ + ; (j = i1) → ∣ compPath-filler (merid base) (sym (merid a)) k (~ i) ∣ₕ}) + (hcomp (λ k → λ { (i = i0) → ∣ merid a (j ∨ ~ k) ∣ + ; (i = i1) → ∣ merid base (~ j ∨ ~ k) ∣ + ; (j = i0) → ∣ merid a (~ k ∨ i) ∣ + ; (j = i1) → ∣ merid base (~ i ∨ ~ k) ∣ₕ}) + ∣ south ∣) + + is-setHepl : (x : S₊ 2) → isSet (Q (CP2 , inl tt) fibr (inr x) →∙ coHomK-ptd 2) + is-setHepl = sphereElim _ (λ _ → isProp→isOfHLevelSuc 1 (isPropIsOfHLevel 2)) + (isOfHLevel↑∙' 0 1) + + ll32 : (x : S₊ 2) → (GysinS1.c (inr x) ≡ psst x) + ll32 = sphereElim _ (λ x → isOfHLevelPath 2 (is-setHepl x) _ _) + (→∙Homogeneous≡ (isHomogeneousKn _) ll32-fst) HopfA*A : ∀ {ℓ} {A : Type ℓ} → (join A A → Type ℓ) - → {!GysinS1.susp∘ϕ 1!} + → {!(funExt⁻ (λ i₁ → fst (ll32 north i₁)) south)!} HopfA*A = {!!} S₊2→S₊2 : CP' → Type ℓ-zero -S₊2→S₊2 (inl x) = {!!} +S₊2→S₊2 (inl x) = {!hopfBase.joinIso₁!} S₊2→S₊2 (inr x) = {!x!} S₊2→S₊2 (push a i₁) = {!!} From 44326735af34ab4360df6fbb58a1619b23d2c63e Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Wed, 15 Sep 2021 00:57:58 +0200 Subject: [PATCH 65/89] stuff --- Cubical/ZCohomology/Gysin.agda | 3749 +++++++++++++++------------ Cubical/ZCohomology/Properties.agda | 1490 ++++++----- 2 files changed, 2959 insertions(+), 2280 deletions(-) diff --git a/Cubical/ZCohomology/Gysin.agda b/Cubical/ZCohomology/Gysin.agda index 43d8946c7..e7bb4babe 100644 --- a/Cubical/ZCohomology/Gysin.agda +++ b/Cubical/ZCohomology/Gysin.agda @@ -29,6 +29,8 @@ open import Cubical.Foundations.Pointed.Homogeneous open import Cubical.Foundations.Univalence +open import Cubical.Relation.Nullary + open import Cubical.Data.Empty renaming (rec to ⊥-rec) open import Cubical.Data.Sigma open import Cubical.Data.Int hiding (_+'_) @@ -61,1684 +63,2071 @@ open import Cubical.HITs.Join open import Cubical.Homotopy.Hopf -SES→Iso : ∀ {ℓ ℓ'} {L R : Group ℓ-zero} - → {G : Group ℓ} {H : Group ℓ'} - → UnitGroup ≡ L - → UnitGroup ≡ R - → (lhom : GroupHom L G) (midhom : GroupHom G H) (rhom : GroupHom H R) - → ((x : _) → isInKer midhom x → isInIm lhom x) - → ((x : _) → isInKer rhom x → isInIm midhom x) - → isEquiv (fst midhom) -SES→Iso {R = R} {G = G} {H = H} = - J (λ L _ → UnitGroup ≡ R → - (lhom : GroupHom L G) (midhom : GroupHom G H) - (rhom : GroupHom H R) → - ((x : fst G) → isInKer midhom x → isInIm lhom x) → - ((x : fst H) → isInKer rhom x → isInIm midhom x) → - isEquiv (fst midhom)) - ((J (λ R _ → (lhom : GroupHom UnitGroup G) (midhom : GroupHom G H) - (rhom : GroupHom H R) → - ((x : fst G) → isInKer midhom x → isInIm lhom x) → - ((x : _) → isInKer rhom x → isInIm midhom x) → - isEquiv (fst midhom)) - main)) - where - main : (lhom : GroupHom UnitGroup G) (midhom : GroupHom G H) - (rhom : GroupHom H UnitGroup) → - ((x : fst G) → isInKer midhom x → isInIm lhom x) → - ((x : fst H) → isInKer rhom x → isInIm midhom x) → - isEquiv (fst midhom) - main lhom midhom rhom lexact rexact = - BijectionIsoToGroupEquiv {G = G} {H = H} - bijIso' .fst .snd - where - bijIso' : BijectionIso G H - BijectionIso.fun bijIso' = midhom - BijectionIso.inj bijIso' x inker = - pRec (GroupStr.is-set (snd G) _ _) - (λ s → sym (snd s) ∙ IsGroupHom.pres1 (snd lhom)) (lexact _ inker) - BijectionIso.surj bijIso' x = rexact x refl - -pres- : (e : GroupHom ℤGroup ℤGroup) → (a : ℤ) → fst e (- a) ≡ - fst e a -pres- e a = +Comm (fst e (- a)) (pos 0) - ∙ cong (_+ (fst e (- a))) (sym (l (fst e a)) ∙ +Comm (fst e a) (- fst e a)) - ∙∙ sym (+Assoc _ _ _) - ∙∙ cong (- (fst e a) +_) (pp ∙ l (fst e a)) - where - l : (a : ℤ) → a + (- a) ≡ 0 - l (pos zero) = refl - l (pos (suc zero)) = refl - l (pos (suc (suc n))) = predℤ+negsuc n (pos (suc (suc n))) ∙ l (pos (suc n)) - l (negsuc zero) = refl - l (negsuc (suc n)) = (cong sucℤ (sucℤ+pos n (negsuc (suc n)))) ∙ l (negsuc n) - - pp : fst e a + fst e (- a) ≡ fst e a + (- fst e a) - pp = sym (IsGroupHom.pres· (snd e) a (- a)) - ∙∙ cong (fst e) (l a) - ∙∙ (IsGroupHom.pres1 (snd e) - ∙ sym (l _)) - -GroupHomPres· : (e : GroupHom ℤGroup ℤGroup) → (a b : ℤ) → fst e (a · b) ≡ a · fst e b -GroupHomPres· e (pos zero) b = IsGroupHom.pres1 (snd e) -GroupHomPres· e (pos (suc n)) b = - IsGroupHom.pres· (snd e) b (pos n · b) ∙ cong (fst e b +_) (GroupHomPres· e (pos n) b) -GroupHomPres· e (negsuc zero) b = pres- e b -GroupHomPres· e (negsuc (suc n)) b = - IsGroupHom.pres· (snd e) (- b) (negsuc n · b) - ∙ cong₂ _+_ (pres- e b) (GroupHomPres· e (negsuc n) b) - -open import Cubical.Data.Sum -open import Cubical.Relation.Nullary -l2 : (n m : ℕ) → Σ[ a ∈ ℕ ] (negsuc n · pos (suc m)) ≡ negsuc a -l2 n zero = n , ·Comm (negsuc n) (pos 1) -l2 n (suc m) = h _ _ .fst , - (·Comm (negsuc n) (pos (suc (suc m))) - ∙∙ cong (negsuc n +_) (·Comm (pos (suc m)) (negsuc n) ∙ (l2 n m .snd)) - ∙∙ h _ _ .snd) - where - h : (x y : ℕ) → Σ[ a ∈ ℕ ] negsuc x + negsuc y ≡ negsuc a - h zero zero = 1 , refl - h zero (suc y) = (suc (suc y)) , +Comm (negsuc zero) (negsuc (suc y)) - h (suc x) zero = (suc (suc x)) , refl - h (suc x) (suc y) = (h (suc (suc x)) y .fst) , (predℤ+negsuc y (negsuc (suc x)) ∙ snd ((h (suc (suc x))) y)) - -zz : (n x : ℕ) → Σ[ a ∈ ℕ ] ((pos (suc x)) · pos (suc (suc n)) ≡ pos (suc (suc a))) -zz n zero = n , refl -zz n (suc x) = h _ _ (zz n x) - where - h : (x : ℤ) (n : ℕ) → Σ[ a ∈ ℕ ] (x ≡ pos (suc (suc a))) - → Σ[ a ∈ ℕ ] pos n + x ≡ pos (suc (suc a)) - h x n (a , p) = n +ℕ a - , cong (pos n +_) p ∙ cong sucℤ (sucℤ+pos a (pos n)) - ∙ sucℤ+pos a (pos (suc n)) ∙ (sym (pos+ (suc (suc n)) a)) - -lem22 : (n : ℕ) (x : ℤ) → ¬ pos 1 ≡ x · pos (suc (suc n)) -lem22 n (pos zero) p = snotz (injPos p) -lem22 n (pos (suc n₁)) p = snotz (injPos (sym (cong predℤ (snd (zz n n₁))) ∙ sym (cong predℤ p))) -lem22 n (negsuc n₁) p = posNotnegsuc _ _ (p ∙ l2 _ _ .snd) - - -grr : (e : GroupEquiv ℤGroup ℤGroup) (x : ℤ) → (fst (fst e) 1) ≡ x → abs (fst (fst e) 1) ≡ 1 -grr e (pos zero) p = - ⊥-rec (snotz (injPos (sym (retEq (fst e) 1) - ∙∙ (cong (fst (fst (invGroupEquiv e))) p) - ∙∙ IsGroupHom.pres1 (snd (invGroupEquiv e))))) -grr e (pos (suc zero)) p = cong abs p -grr e (pos (suc (suc n))) p = ⊥-rec (lem22 _ _ (h ∙ ·Comm (pos (suc (suc n))) (invEq (fst e) 1))) - where - - h : pos 1 ≡ _ - h = sym (retEq (fst e) 1) - ∙∙ (cong (fst (fst (invGroupEquiv e))) (p ∙ ·Comm 1 (pos (suc (suc n))))) - ∙∙ GroupHomPres· (_ , snd (invGroupEquiv e)) (pos (suc (suc n))) 1 -grr e (negsuc zero) p = cong abs p -grr e (negsuc (suc n)) p = ⊥-rec (lem22 _ _ l33) - where - h : fst (fst e) (negsuc zero) ≡ pos (suc (suc n)) - h = pres- (_ , snd e) (pos 1) ∙ cong -_ p - - compGroup : GroupEquiv ℤGroup ℤGroup - fst compGroup = isoToEquiv (iso -_ -_ -Involutive -Involutive) - snd compGroup = makeIsGroupHom -Dist+ - - compHom : GroupEquiv ℤGroup ℤGroup - compHom = compGroupEquiv compGroup e - - l32 : fst (fst compHom) (pos 1) ≡ pos (suc (suc n)) - l32 = h - - l33 : pos 1 ≡ invEq (fst compHom) (pos 1) · pos (suc (suc n)) - l33 = sym (retEq (fst compHom) (pos 1)) - ∙∙ cong (invEq (fst compHom)) l32 - ∙∙ (cong (invEq (fst compHom)) (·Comm (pos 1) (pos (suc (suc n)))) - ∙ GroupHomPres· (_ , (snd (invGroupEquiv compHom))) (pos (suc (suc n))) (pos 1) - ∙ ·Comm (pos (suc (suc n))) (invEq (fst compHom) (pos 1))) - -abs→⊎ : (x : ℤ) (n : ℕ) → abs x ≡ n → (x ≡ pos n) ⊎ (x ≡ - pos n) -abs→⊎ x n = J (λ n _ → (x ≡ pos n) ⊎ (x ≡ - pos n)) (help x) - where - help : (x : ℤ) → (x ≡ pos (abs x)) ⊎ (x ≡ - pos (abs x)) - help (pos n) = inl refl - help (negsuc n) = inr refl - -JGroupEquiv : ∀ {ℓ ℓ'} {G : Group ℓ} (P : (H : Group ℓ) → GroupEquiv G H → Type ℓ') - → P G idGroupEquiv - → ∀ {H} e → P H e -JGroupEquiv {G = G} P p {H} e = - transport (λ i → P (GroupPath G H .fst e i) - (transp (λ j → GroupEquiv G (GroupPath G H .fst e (i ∨ ~ j))) i e)) - (subst (P G) (sym l) p) - where - l : (transp (λ j → GroupEquiv G (GroupPath G H .fst e (~ j))) i0 e) ≡ idGroupEquiv - l = Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (Σ≡Prop (λ _ → isPropIsEquiv _) - (funExt λ x → (λ i → fst (fst (fst e .snd .equiv-proof - (transportRefl (fst (fst e) (transportRefl x i)) i)))) - ∙ retEq (fst e) x)) - -sesIsoPresGen : - ∀ (G : Group ℓ-zero) - → (iso : GroupEquiv ℤGroup G) - → (H : Group ℓ-zero) - → (iso2 : GroupEquiv ℤGroup H) - → (e : fst G) - → invEq (fst iso) e ≡ 1 - → (hom : GroupEquiv G H) - → abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1 -sesIsoPresGen G = JGroupEquiv (λ G iso → - (H : Group ℓ-zero) - (iso2 : GroupEquiv ℤGroup H) - → (e : fst G) - → invEq (fst iso) e ≡ 1 - → (hom : GroupEquiv G H) - → abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1) - λ H → JGroupEquiv (λ H iso2 → (e : ℤ) → e ≡ pos 1 → - (hom : GroupEquiv ℤGroup H) → - abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1) - λ e p hom → cong (abs ∘ (fst (fst hom))) p ∙ autoPres1 hom - where - autoPres1 : (e : GroupEquiv ℤGroup ℤGroup) - → abs (fst (fst e) 1) ≡ 1 - autoPres1 e = grr e _ refl - -characFunSpaceS¹ : ∀ {ℓ} {A : Type ℓ} → - Iso (S₊ 1 → A) (Σ[ x ∈ A ] x ≡ x) -Iso.fun characFunSpaceS¹ f = f base , cong f loop -Iso.inv characFunSpaceS¹ (x , p) base = x -Iso.inv characFunSpaceS¹ (x , p) (loop i) = p i -Iso.rightInv characFunSpaceS¹ _ = refl -Iso.leftInv characFunSpaceS¹ f = funExt λ { base → refl ; (loop i) → refl} - -characFunSpace : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} - → Iso (join A B → C) - (Σ[ f ∈ (A → C) ] Σ[ g ∈ (B → C) ] - ((a : A) (b : B) → f a ≡ g b)) -Iso.fun characFunSpace f = (f ∘ inl) , ((f ∘ inr) , (λ a b → cong f (push a b))) -Iso.inv characFunSpace (f , g , p) (inl x) = f x -Iso.inv characFunSpace (f , g , p) (inr x) = g x -Iso.inv characFunSpace (f , g , p) (push a b i) = p a b i -Iso.rightInv characFunSpace (f , g , p) = refl -Iso.leftInv characFunSpace f = - funExt λ { (inl x) → refl ; (inr x) → refl ; (push a b i) → refl} - -coHomS¹-ish : (n : ℕ) → Type _ -coHomS¹-ish n = hLevelTrunc 3 (S₊ 1 → coHomK (3 +ℕ n)) - -fib : (n : ℕ) → coHomS¹-ish n → Type _ -fib n x = - trRec {B = TypeOfHLevel ℓ-zero 2} (isOfHLevelTypeOfHLevel 2) - (λ f → ∥ (Σ[ g ∈ (S₊ 3 → coHomK (3 +ℕ n)) ] - ((a : S₊ 1) (b : S₊ 3) → f a ≡ g b)) ∥₂ , squash₂) x .fst - -contrFstΣ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} - → (e : isContr A) - → Iso (Σ A B) (B (fst e)) -Iso.fun (contrFstΣ {B = B} e) (a , b) = subst B (sym (snd e a)) b -Iso.inv (contrFstΣ {B = B} e) b = (fst e) , b -Iso.rightInv (contrFstΣ {B = B} e) b = cong (λ x → subst B x b) h ∙ transportRefl b - where - h : sym (snd e (fst e)) ≡ refl - h = isProp→isSet (isContr→isProp e) _ _ _ _ -Iso.leftInv (contrFstΣ {B = B} e) b = - ΣPathP ((snd e (fst b)) - , (λ j → transp (λ i → B (snd e (fst b) (~ i ∨ j))) j (snd b))) - -Iso1 : (n : ℕ) → Iso (coHom (3 +ℕ n) (join S¹ (S₊ 3))) ∥ Σ (coHomS¹-ish n) (fib n) ∥₂ -Iso1 n = compIso (setTruncIso characFunSpace) Iso2 - where - Iso2 : Iso _ ∥ Σ (coHomS¹-ish n) (fib n) ∥₂ - Iso.fun Iso2 = sMap λ F → ∣ fst F ∣ , ∣ (fst (snd F)) , (snd (snd F)) ∣₂ - Iso.inv Iso2 = - sRec squash₂ - (uncurry (trElim (λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelSuc 2 squash₂) - λ f → sRec squash₂ λ p → ∣ f , p ∣₂)) - Iso.rightInv Iso2 = - sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - (uncurry (trElim (λ _ → isOfHLevelΠ 3 λ _ → isProp→isOfHLevelSuc 2 (squash₂ _ _)) - λ f → sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl )) - Iso.leftInv Iso2 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl - -isContrcoHomS¹-ish : (n : ℕ) → isContr (coHomS¹-ish n) -isContrcoHomS¹-ish n = isOfHLevelRetractFromIso 0 (mapCompIso characFunSpaceS¹) isContrEnd - where - isContrEnd : isContr (hLevelTrunc 3 (Σ[ x ∈ coHomK (3 +ℕ n) ] x ≡ x)) - fst isContrEnd = ∣ 0ₖ _ , refl ∣ - snd isContrEnd = - trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) - (uncurry (coHomK-elim _ - (λ _ → isOfHLevelΠ (suc (suc (suc n))) - λ _ → isOfHLevelPlus' {n = n} 3 (isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _)) - (λ p → (trRec (isOfHLevelPlus' {n = n} 3 (isOfHLevelTrunc 3) _ _) - (λ p i → ∣ (0ₖ (3 +ℕ n) , p i) ∣) - (Iso.fun (PathIdTruncIso _) - (isContr→isProp - (isConnectedPath _ (isConnectedKn (2 +ℕ n)) (0ₖ _) (0ₖ _)) ∣ refl ∣ ∣ p ∣)))))) - -Iso2' : (n : ℕ) → Iso (∥ Σ (coHomS¹-ish n) (fib n) ∥₂) (fib n ∣ (λ _ → 0ₖ _) ∣) -Iso2' n = compIso (setTruncIso (contrFstΣ centerChange)) - (setTruncIdempotentIso squash₂) - where - centerChange : isContr (coHomS¹-ish n) - fst centerChange = ∣ (λ _ → 0ₖ _) ∣ - snd centerChange y = isContr→isProp (isContrcoHomS¹-ish n) _ _ - -open import Cubical.Foundations.Equiv.HalfAdjoint -open import Cubical.Relation.Nullary -ll5 : (n : ℕ) → ¬ (n ≡ 2) → isContr (fib n ∣ (λ _ → 0ₖ _) ∣) -ll5 n p = - isOfHLevelRetractFromIso 0 - (compIso - (setTruncIso - (compIso (Σ-cong-iso-snd λ f → - (compIso characFunSpaceS¹ (invIso (Σ-cong-iso-fst (iso funExt⁻ funExt (λ _ → refl) λ _ → refl))))) - (compIso ΣΣ (contrFstΣ (isContrSingl _))))) - (compIso (setTruncIso (iso funExt⁻ funExt (λ _ → refl) λ _ → refl)) - (compIso (setTruncIso (codomainIso (congIso (invIso (Iso-Kn-ΩKn+1 _))))) - ((compIso (invIso (fst (coHom≅coHomΩ _ (S₊ _)))) - (fst (Hⁿ-Sᵐ≅0 n 2 p))))))) - isContrUnit - -record ExactSeqℤ {ℓ : Level} (G : ℤ → Group ℓ) : Type ℓ where - field - maps : ∀ n → GroupHom (G n) (G (sucℤ n)) - im⊂ker : ∀ n → ∀ g → isInIm (maps n) g → isInKer (maps (sucℤ n)) g - ker⊂im : ∀ n → ∀ g → isInKer (maps (sucℤ n)) g → isInIm (maps n) g - -record ExactSeqℕ {ℓ : Level} (G : ℕ → Group ℓ) : Type ℓ where - field - maps : ∀ n → GroupHom (G n) (G (suc n)) - im⊂ker : ∀ n → ∀ g → isInIm (maps n) g → isInKer (maps (suc n)) g - ker⊂im : ∀ n → ∀ g → isInKer (maps (suc n)) g → isInIm (maps n) g - - -module _ {ℓ : Level} {A : Pointed ℓ} {n : ℕ} - where - funTyp : Type _ - funTyp = A →∙ coHomK-ptd n - - _++_ : funTyp → funTyp → funTyp - fst (f ++ g) x = fst f x +ₖ fst g x - snd (f ++ g) = cong₂ _+ₖ_ (snd f) (snd g) ∙ rUnitₖ _ (0ₖ _) - -addAgree : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (x y : funTyp {A = A} {n = n}) - → Path (fst (coHomRedGrDir n A)) - ∣ x ++ y ∣₂ - (∣ x ∣₂ +ₕ∙ ∣ y ∣₂) -addAgree {A = A} zero f g = - cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) refl) -addAgree {A = A} (suc zero) f g = - cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) refl) -addAgree {A = A} (suc (suc n)) f g = - cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) refl) - -ss : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Type ℓ''} - → isHomogeneous B - → (x y : (A →∙ B)) (f : C → x ≡ y) - → isEquiv (cong fst ∘ f) - → isEquiv f -ss homb x y f e = - isoToIsEquiv - (iso _ - (λ p → invEq (_ , e) (cong fst p)) - (λ p → →∙Homogeneous≡Path homb _ _ (secEq (_ , e) (cong fst p))) - (retEq (_ , e))) - -Whitehead : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} - → (f : A → B) - → isEmbedding f - → (∀ (b : B) → ∃[ a ∈ A ] f a ≡ b) - → isEquiv f -equiv-proof (Whitehead f emb p) b = - pRec isPropIsContr - (λ fib → fib , isEmbedding→hasPropFibers emb b fib) - (p b) - -Int-ind : ∀ {ℓ} (P : ℤ → Type ℓ) - → P (pos zero) → P (pos 1) - → P (negsuc zero) - → ((x y : ℤ) → P x → P y → P (x + y)) → (x : ℤ) → P x -Int-ind P z one min ind (pos zero) = z -Int-ind P z one min ind (pos (suc zero)) = one -Int-ind P z one min ind (pos (suc (suc n))) = - ind (pos (suc n)) 1 (Int-ind P z one min ind (pos (suc n))) one -Int-ind P z one min ind (negsuc zero) = min -Int-ind P z one min ind (negsuc (suc n)) = - ind (negsuc n) (negsuc zero) (Int-ind P z one min ind (negsuc n)) min - -genFunSpace : (n : ℕ) → S₊∙ n →∙ coHomK-ptd n -fst (genFunSpace zero) false = 1 -fst (genFunSpace zero) true = 0 -snd (genFunSpace zero) = refl -fst (genFunSpace (suc n)) = ∣_∣ -snd (genFunSpace (suc zero)) = refl -snd (genFunSpace (suc (suc n))) = refl - -module _ where - open import Cubical.Algebra.Monoid - open import Cubical.Algebra.Semigroup - open GroupStr - open IsGroup - open IsMonoid - open IsSemigroup - πS : (n : ℕ) → Group ℓ-zero - fst (πS n) = S₊∙ n →∙ coHomK-ptd n - 1g (snd (πS n)) = (λ _ → 0ₖ n) , refl - GroupStr._·_ (snd (πS n)) = - λ f g → (λ x → fst f x +ₖ fst g x) - , cong₂ _+ₖ_ (snd f) (snd g) ∙ rUnitₖ n (0ₖ n) - inv (snd (πS n)) f = (λ x → -ₖ fst f x) , cong -ₖ_ (snd f) ∙ -0ₖ {n = n} - is-set (isSemigroup (isMonoid (isGroup (snd (πS zero))))) = - isOfHLevelΣ 2 (isSetΠ (λ _ → isSetℤ)) - λ _ → isOfHLevelPath 2 isSetℤ _ _ - is-set (isSemigroup (isMonoid (isGroup (snd (πS (suc n)))))) = isOfHLevel↑∙' 0 n - IsSemigroup.assoc (isSemigroup (isMonoid (isGroup (snd (πS n))))) x y z = - →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ w → assocₖ n (fst x w) (fst y w) (fst z w)) - fst (identity (isMonoid (isGroup (snd (πS n)))) (f , p)) = - →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → rUnitₖ n (f x)) - snd (identity (isMonoid (isGroup (snd (πS n)))) (f , p)) = - →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → lUnitₖ n (f x)) - fst (inverse (isGroup (snd (πS n))) (f , p)) = - →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → rCancelₖ n (f x)) - snd (inverse (isGroup (snd (πS n))) (f , p)) = - →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → lCancelₖ n (f x)) - - isSetπS : (n : ℕ) → isSet (S₊∙ n →∙ coHomK-ptd n) - isSetπS = λ n → is-set (snd (πS n)) - -K : (n : ℕ) → GroupIso (coHomRedGrDir n (S₊∙ n)) (πS n) -fst (K n) = setTruncIdempotentIso (isSetπS n) -snd (K zero) = - makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 (isSetπS 0) _ _) - λ f g → →∙Homogeneous≡ (isHomogeneousKn 0) refl) -snd (K (suc zero)) = - makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 (isSetπS 1) _ _) - λ f g → →∙Homogeneous≡ (isHomogeneousKn 1) refl) -snd (K (suc (suc n))) = - makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 (isSetπS _) _ _) - λ f g → →∙Homogeneous≡ (isHomogeneousKn _) refl) - -t : ∀ {ℓ} {A : Pointed ℓ} → Iso ((Bool , true) →∙ A) (typ A) -Iso.fun t f = fst f false -fst (Iso.inv t a) false = a -fst (Iso.inv (t {A = A}) a) true = pt A -snd (Iso.inv t a) = refl -Iso.rightInv t a = refl -Iso.leftInv t (f , p) = - ΣPathP ((funExt (λ { false → refl ; true → sym p})) , λ i j → p (~ i ∨ j)) - -S1' : (n : ℕ) → GroupIso (πS n) ℤGroup -fst (S1' zero) = t -snd (S1' zero) = makeIsGroupHom λ _ _ → refl -S1' (suc n) = - compGroupIso - (invGroupIso (K (suc n))) - (compGroupIso - (GroupEquiv→GroupIso (coHomGr≅coHomRedGr n (S₊∙ (suc n)))) - (Hⁿ-Sⁿ≅ℤ n)) - -S1 : (n : ℕ) → Iso (S₊∙ (suc n) →∙ coHomK-ptd (suc n)) ℤ -S1 n = compIso (invIso (setTruncIdempotentIso (isOfHLevel↑∙' 0 n))) - (compIso (equivToIso (fst (coHomGr≅coHomRedGr n (S₊∙ (suc n))))) - (fst (Hⁿ-Sⁿ≅ℤ n))) - -connFunSpace : (n i : ℕ) → (x y : S₊∙ n →∙ coHomK-ptd (suc i +' n)) → ∥ x ≡ y ∥ -connFunSpace n i f g = Iso.fun PathIdTrunc₀Iso (isContr→isProp (lem n) ∣ f ∣₂ ∣ g ∣₂) - where - open import Cubical.Relation.Nullary - natLem : (n i : ℕ) → ¬ (suc (i +ℕ n) ≡ n) - natLem zero i = snotz - natLem (suc n) i p = natLem n i (sym (+-suc i n) ∙ (cong predℕ p)) - - lem : (n : ℕ) → isContr ∥ (S₊∙ n →∙ coHomK-ptd (suc i +' n)) ∥₂ - fst (lem zero) = 0ₕ∙ (suc i) - snd (lem zero) = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - λ f → pRec (squash₂ _ _) - (λ id → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn (suc i)) - (funExt λ { false → sym id ; true → sym (snd f)}))) - (help (f .fst false)) - where - help : (x : coHomK (suc i)) → ∥ x ≡ 0ₖ _ ∥ - help = coHomK-elim _ (λ _ → isProp→isOfHLevelSuc i squash) ∣ refl ∣ - lem (suc n) = - isOfHLevelRetractFromIso 0 - (compIso (equivToIso (fst (coHomGr≅coHomRedGr (suc (i +ℕ n)) (S₊∙ (suc n))))) - (fst (Hⁿ-Sᵐ≅0 (suc (i +ℕ n)) n (natLem n i)))) - isContrUnit - -S1Pres1 : (n : ℕ) → Iso.fun (fst (S1' (suc n))) (∣_∣ , refl) ≡ pos 1 -S1Pres1 zero = refl -S1Pres1 (suc n) = cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ n))) (lem n) ∙ S1Pres1 n - where - lem : (n : ℕ) → Iso.fun (fst (suspensionAx-Sn n n)) ∣ ∣_∣ ∣₂ ≡ ∣ ∣_∣ ∣₂ - lem zero = cong ∣_∣₂ (funExt λ x → transportRefl (∣ x ∣ +ₖ (0ₖ 1)) ∙ rUnitₖ 1 ∣ x ∣) - lem (suc n) = cong ∣_∣₂ - (funExt λ x → (λ i → transportRefl ((ΩKn+1→Kn (suc (suc n)) - (transp (λ j → 0ₖ (suc (suc (suc n))) ≡ ∣ merid north (~ j ∧ ~ i) ∣) i - (λ z → ∣ compPath-filler (merid (transportRefl (transportRefl x i) i)) (sym (merid north)) i z - ∣)))) i) - ∙ Iso.leftInv (Iso-Kn-ΩKn+1 (suc (suc n))) ∣ x ∣) - -S1Pres1← : (n : ℕ) → Iso.inv (fst (S1' (suc n))) (pos 1) ≡ (∣_∣ , refl) -S1Pres1← n = sym (cong (Iso.inv (fst (S1' (suc n)))) (S1Pres1 n)) ∙ Iso.leftInv (fst (S1' (suc n))) (∣_∣ , refl) - -IsoFunSpace : (n : ℕ) → Iso (S₊∙ n →∙ coHomK-ptd n) ℤ -IsoFunSpace n = fst (S1' n) - -module g-base where - g : (n : ℕ) (i : ℕ) → coHomK i → (S₊∙ n →∙ coHomK-ptd (i +' n)) - fst (g n i x) y = x ⌣ₖ (genFunSpace n) .fst y - snd (g n i x) = cong (x ⌣ₖ_) ((genFunSpace n) .snd) ∙ ⌣ₖ-0ₖ i n x - - G : (n : ℕ) (i : ℕ) → coHomK i → (S₊∙ n →∙ coHomK-ptd (n +' i)) - fst (G n i x) y = (genFunSpace n) .fst y ⌣ₖ x - snd (G n i x) = cong (_⌣ₖ x) ((genFunSpace n) .snd) ∙ 0ₖ-⌣ₖ n i x - - eq1 : (n : ℕ) (i : ℕ) → (S₊∙ n →∙ coHomK-ptd (i +' n)) ≃ (S₊∙ n →∙ coHomK-ptd (i +' n)) - eq1 n i = isoToEquiv (iso F F FF FF) - where - lem : (i n : ℕ) → (-ₖ^ i · n) (snd (coHomK-ptd (i +' n))) ≡ 0ₖ _ - lem zero zero = refl - lem zero (suc zero) = refl - lem zero (suc (suc n)) = refl - lem (suc zero) zero = refl - lem (suc (suc i)) zero = refl - lem (suc i) (suc n) = refl - - F : S₊∙ n →∙ coHomK-ptd (i +' n) → S₊∙ n →∙ coHomK-ptd (i +' n) - fst (F f) x = (-ₖ^ i · n) (fst f x) - snd (F f) = cong (-ₖ^ i · n) (snd f) ∙ lem i n - - FF : (x : _) → F (F x) ≡ x - FF x = - →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ y → -ₖ-gen² i n _ _ (fst x y)) - - rCancel'' : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) → sym p ∙∙ refl ∙∙ p ≡ refl - rCancel'' p = (λ j → (λ i → p (~ i ∨ j)) ∙∙ refl ∙∙ λ i → p (i ∨ j)) ∙ sym (rUnit refl) - - transpPres0 : ∀ {k m : ℕ} (p : k ≡ m) → subst coHomK p (0ₖ k) ≡ 0ₖ m - transpPres0 {k = k} = J (λ m p → subst coHomK p (0ₖ k) ≡ 0ₖ m) (transportRefl _) - - eq2 : (n : ℕ) (i : ℕ) → (S₊∙ n →∙ coHomK-ptd (n +' i)) ≃ (S₊∙ n →∙ coHomK-ptd (i +' n)) - eq2 n i = - isoToEquiv (iso (λ f → (λ x → subst coHomK (+'-comm n i) (fst f x)) , - cong (subst coHomK (+'-comm n i)) (snd f) ∙ transpPres0 (+'-comm n i)) - (λ f → (λ x → subst coHomK (sym (+'-comm n i)) (fst f x)) - , (cong (subst coHomK (sym (+'-comm n i))) (snd f) ∙ transpPres0 (sym (+'-comm n i)))) - (λ f → →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ x → transportTransport⁻ _ (f .fst x))) - λ f → →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ x → transportTransport⁻ _ (f .fst x))) - - g≡ : (n : ℕ) (i : ℕ) → g n i ≡ λ x → fst (compEquiv (eq2 n i) (eq1 n i)) ((G n i) x) - g≡ n i = funExt (λ f → →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ y → gradedComm-⌣ₖ _ _ f (genFunSpace n .fst y))) - - glIso3-fun : (n m : ℕ) → - (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) - → (S₊ n → coHomK m) - glIso3-fun zero m (f , p) false = ΩKn+1→Kn _ (sym p ∙∙ cong f loop ∙∙ p) - glIso3-fun zero m (f , p) true = 0ₖ _ - glIso3-fun (suc n) m (f , p) x = - ΩKn+1→Kn _ (sym p ∙∙ cong f (merid x ∙ sym (merid (ptSn _))) ∙∙ p) - - glIso3-fun∙ : (n m : ℕ) → (f : _) → glIso3-fun n m f (ptSn _) ≡ 0ₖ m - glIso3-fun∙ zero m = λ _ → refl - glIso3-fun∙ (suc n) m (f , p) = - cong (ΩKn+1→Kn m) (cong (sym p ∙∙_∙∙ p) (cong (cong f) (rCancel (merid (ptSn _))))) - ∙∙ cong (ΩKn+1→Kn m) (rCancel'' p) - ∙∙ ΩKn+1→Kn-refl m - - - glIso3-inv : (n m : ℕ) → (S₊∙ n →∙ coHomK-ptd m) → (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) - fst (glIso3-inv zero m (f , p)) base = 0ₖ _ - fst (glIso3-inv zero m (f , p)) (loop i) = Kn→ΩKn+1 m (f false) i - snd (glIso3-inv zero m (f , p)) = refl - fst (glIso3-inv (suc n) m (f , p)) north = 0ₖ _ - fst (glIso3-inv (suc n) m (f , p)) south = 0ₖ _ - fst (glIso3-inv (suc n) m (f , p)) (merid a i) = Kn→ΩKn+1 m (f a) i - snd (glIso3-inv (suc n) m (f , p)) = refl - - glIso3 : (n m : ℕ) → - Iso (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) - (S₊∙ n →∙ coHomK-ptd m) - Iso.fun (glIso3 n m) f = (glIso3-fun n m f) , (glIso3-fun∙ n m f) - Iso.inv (glIso3 n m) = glIso3-inv n m - Iso.rightInv (glIso3 zero m) (f , p) = - →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ { false → cong (ΩKn+1→Kn m) (sym (rUnit _)) ∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f false) - ; true → sym p}) - Iso.rightInv (glIso3 (suc n) m) (f , p) = - →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ x → (λ i → ΩKn+1→Kn m (sym (rUnit (cong-∙ (glIso3-inv (suc n) m (f , p) .fst) (merid x) (sym (merid (ptSn _))) i)) i)) - ∙∙ cong (ΩKn+1→Kn m) (cong (Kn→ΩKn+1 m (f x) ∙_) (cong sym (cong (Kn→ΩKn+1 m) p ∙ Kn→ΩKn+10ₖ m)) ∙ sym (rUnit _)) - ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f x)) - Iso.leftInv (glIso3 zero m) (f , p) = →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ { base → sym p - ; (loop i) j → h j i}) - where - h : PathP (λ i → p (~ i) ≡ p (~ i)) (Kn→ΩKn+1 m (ΩKn+1→Kn m (sym p ∙∙ cong f loop ∙∙ p))) (cong f loop) - h = Iso.rightInv (Iso-Kn-ΩKn+1 _) _ - ◁ λ i → doubleCompPath-filler (sym p) (cong f loop) p (~ i) - Iso.leftInv (glIso3 (suc n) m) (f , p) = - →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ { north → sym p - ; south → sym p ∙ cong f (merid (ptSn _)) - ; (merid a i) j → h a j i}) - where - h : (a : S₊ (suc n)) → PathP (λ i → p (~ i) ≡ (sym p ∙ cong f (merid (ptSn (suc n)))) i) - (Kn→ΩKn+1 m (ΩKn+1→Kn m (sym p ∙∙ cong f (merid a ∙ sym (merid (ptSn _))) ∙∙ p))) - (cong f (merid a)) - h a = Iso.rightInv (Iso-Kn-ΩKn+1 _) _ - ◁ λ i j → hcomp (λ k → λ { (i = i1) → (f (merid a j)) - ; (j = i0) → p (k ∧ ~ i) - ; (j = i1) → compPath-filler' (sym p) (cong f (merid (ptSn _))) k i }) - (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) - - glIsoInvHom : (n m : ℕ) (x y : coHomK n) (z : S₊ (suc m)) - → Iso.inv (glIso3 _ _) (G m n (x +ₖ y)) .fst z - ≡ Iso.inv (glIso3 _ _) (G m n x) .fst z - +ₖ Iso.inv (glIso3 _ _) (G m n y) .fst z - glIsoInvHom zero zero x y base = refl - glIsoInvHom (suc n) zero x y base = refl - glIsoInvHom zero zero x y (loop i) j = h j i - where - h : (cong (Iso.inv (glIso3 _ _) (G zero zero (x + y)) .fst) loop) - ≡ cong₂ _+ₖ_ (cong (Iso.inv (glIso3 _ _) (G zero zero x) .fst) loop) - (cong (Iso.inv (glIso3 _ _) (G zero zero y) .fst) loop) - h = Kn→ΩKn+1-hom 0 x y - ∙ ∙≡+₁ (cong (Iso.inv (glIso3 _ _) (G zero zero x) .fst) loop) - (cong (Iso.inv (glIso3 _ _) (G zero zero y) .fst) loop) - glIsoInvHom (suc n) zero x y (loop i) j = h j i - where - h : Kn→ΩKn+1 (suc n) ((pos (suc zero)) ·₀ (x +ₖ y)) - ≡ cong₂ _+ₖ_ (cong (Iso.inv (glIso3 zero (zero +' suc n)) (G zero (suc n) x) .fst) loop) - (cong (Iso.inv (glIso3 zero (zero +' suc n)) (G zero (suc n) y) .fst) loop) - h = cong (Kn→ΩKn+1 (suc n)) (lUnit⌣ₖ _ (x +ₖ y)) - ∙∙ Kn→ΩKn+1-hom (suc n) x y - ∙∙ (λ i → ∙≡+₂ n (Kn→ΩKn+1 (suc n) (lUnit⌣ₖ _ x (~ i))) - (Kn→ΩKn+1 (suc n) (lUnit⌣ₖ _ y (~ i))) i) - glIsoInvHom zero (suc m) x y north = refl - glIsoInvHom zero (suc m) x y south = refl - glIsoInvHom zero (suc m) x y (merid a i) j = h j i - where - h : Kn→ΩKn+1 (suc m) (_⌣ₖ_ {n = suc m} {m = zero} ∣ a ∣ (x + y)) - ≡ cong₂ _+ₖ_ (Kn→ΩKn+1 (suc m) (_⌣ₖ_ {n = suc m} {m = zero} ∣ a ∣ x)) - (Kn→ΩKn+1 (suc m) (_⌣ₖ_ {n = suc m} {m = zero} ∣ a ∣ y)) - h = cong (Kn→ΩKn+1 (suc m)) (leftDistr-⌣ₖ (suc m) 0 ∣ a ∣ x y) - ∙∙ Kn→ΩKn+1-hom (suc m) _ _ - ∙∙ ∙≡+₂ _ _ _ - glIsoInvHom (suc n) (suc m) x y north = refl - glIsoInvHom (suc n) (suc m) x y south = refl - glIsoInvHom (suc n) (suc m) x y (merid a i) j = h j i - where - h : Kn→ΩKn+1 (suc (suc (m +ℕ n))) (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ (x +ₖ y)) - ≡ cong₂ _+ₖ_ (Kn→ΩKn+1 (suc (suc (m +ℕ n))) (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ x)) - (Kn→ΩKn+1 (suc (suc (m +ℕ n))) (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ y)) - h = cong (Kn→ΩKn+1 (suc (suc (m +ℕ n)))) - (leftDistr-⌣ₖ (suc m) (suc n) ∣ a ∣ x y) - ∙∙ Kn→ΩKn+1-hom _ _ _ - ∙∙ ∙≡+₂ _ _ _ - - +'-suc : (n m : ℕ) → suc n +' m ≡ suc (n +' m) - +'-suc zero zero = refl - +'-suc (suc n) zero = refl - +'-suc zero (suc m) = refl - +'-suc (suc n) (suc m) = refl - - LEM : (i n : ℕ) (x : coHomK i) - → Path _ (G (suc n) i x) - (subst (λ x → S₊∙ (suc n) →∙ coHomK-ptd x) - (sym (+'-suc n i)) - ((Iso.inv (glIso3 n _)) (G n i x))) - LEM zero zero x = - →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ z → (λ i → x ·₀ ∣ z ∣) ∙ h x z ∙ sym (transportRefl _)) - where - h : (x : ℤ) (z : S¹) → _·₀_ {n = 1} x ∣ z ∣ ≡ fst (Iso.inv (glIso3 0 zero) (G zero zero x)) z - h = Int-ind _ (λ { base → refl ; (loop i) j → Kn→ΩKn+10ₖ zero (~ j) i}) - (λ { base → refl ; (loop i) j → rUnit (cong ∣_∣ₕ (lUnit loop j)) j i}) - (λ { base → refl ; (loop i) j → rUnit (cong ∣_∣ₕ (sym loop)) j i}) - λ x y inx iny z - → rightDistr-⌣ₖ 0 1 x y ∣ z ∣ - ∙∙ cong₂ _+ₖ_ (inx z) (iny z) - ∙∙ sym (glIsoInvHom zero zero x y z) - LEM (suc i) zero x = - →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ z → h z - ∙ sym (transportRefl - ((Iso.inv (glIso3 zero (suc i)) (G zero (suc i) x)) .fst z))) - where - h : (z : S₊ 1) → _ ≡ Iso.inv (glIso3 zero (suc i)) (G zero (suc i) x) .fst z - h base = refl - h (loop k) j = Kn→ΩKn+1 (suc i) (lUnit⌣ₖ _ x (~ j)) k - LEM zero (suc n) x = - →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ z → h x z ∙ λ i → transportRefl (Iso.inv (glIso3 (suc n) (suc n)) (G (suc n) zero x)) (~ i) .fst z) - where - +merid : rUnitₖ (suc (suc n)) ∣ south ∣ ≡ cong ∣_∣ₕ (merid (ptSn _)) - +merid = sym (lUnit _) - ∙ cong (cong ∣_∣ₕ) - λ j i → transp (λ _ → S₊ (suc (suc n))) (i ∨ j) (merid (ptSn (suc n)) i) - open import Cubical.Foundations.Path - - pp : (a : S₊ (suc n)) → PathP (λ i → ∣ merid a i ∣ₕ ≡ Kn→ΩKn+1 (suc n) (∣ a ∣ +ₖ (0ₖ _)) i) - refl (sym (rUnitₖ (suc (suc n)) ∣ south ∣)) - pp a = flipSquare ((λ i j → ∣ compPath-filler (merid a) (sym (merid (ptSn _))) i j ∣ₕ ) - ▷ cong (Kn→ΩKn+1 (suc n)) (sym (rUnitₖ (suc n) ∣ a ∣ₕ))) - ▷ sym (cong sym +merid) - - pp2 : (a : S₊ (suc n)) → (λ i → -ₖ ∣ merid a i ∣) - ≡ Kn→ΩKn+1 (suc n) (-ₖ ∣ a ∣) - pp2 a = cong (cong ∣_∣ₕ) (sym (symDistr (merid a) (sym (merid (ptSn (suc n)))))) - ∙ sym (Kn→ΩKn+1-ₖ (suc n) ∣ a ∣) - - h : (x : ℤ) (z : S₊ (suc (suc n))) - → _⌣ₖ_ {n = suc (suc n)} {m = 0} ∣ z ∣ x - ≡ Iso.inv (glIso3 (suc n) (suc n)) (G (suc n) zero x) .fst z - h = Int-ind _ - (λ { north → refl ; south → refl ; (merid a i) j → Kn→ΩKn+10ₖ (suc n) (~ j) i}) - (λ { north → refl ; south → refl - ; (merid a i) j → hcomp (λ k → λ { (i = i0) → ∣ north ∣ - ; (i = i1) → rUnitₖ (suc (suc n)) ∣ south ∣ (~ k) - ; (j = i0) → rUnitₖ (suc (suc n)) ∣ merid a i ∣ (~ k) - ; (j = i1) → pp a i k}) - ∣ merid a i ∣ₕ}) - (λ { north → refl - ; south → refl - ; (merid a i) j → pp2 a j i}) - λ x y indx indy z → leftDistr-⌣ₖ _ _ ∣ z ∣ x y - ∙ cong₂ _+ₖ_ (indx z) (indy z) - ∙ sym (glIsoInvHom _ _ _ _ _) - LEM (suc i) (suc n) x = - →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ z → h z - ∙ λ j → transportRefl ((Iso.inv (glIso3 (suc n) (suc (suc (n +ℕ i)))) - (G (suc n) (suc i) x))) (~ j) .fst z) - where - h : (z : S₊ (suc (suc n))) → _ - h north = refl - h south = refl - h (merid a i) = refl - - isEquivGzero : (i : ℕ) → isEquiv (G zero i) - isEquivGzero i = - isoToIsEquiv - (iso _ (λ f → fst f false) - (λ {(f , p) → →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ { false → rUnitₖ _ (f false) ; true → sym p})}) - (lUnit⌣ₖ _)) - - isEquivG : (n i : ℕ) → isEquiv (G n i) - isEquivG zero i = - isoToIsEquiv - (iso _ (λ f → fst f false) - (λ {(f , p) → →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ { false → rUnitₖ _ (f false) ; true → sym p})}) - (lUnit⌣ₖ _)) - isEquivG (suc n) i = - subst isEquiv (sym (funExt (LEM i n))) - (compEquiv (compEquiv (G n i , isEquivG n i) - (isoToEquiv (invIso (glIso3 n (n +' i))))) - (pathToEquiv (λ j → S₊∙ (suc n) →∙ coHomK-ptd (+'-suc n i (~ j)))) .snd) - - - isEquiv-g : (n i : ℕ) → isEquiv (g n i) - isEquiv-g n i = - subst isEquiv (sym (g≡ n i)) - (compEquiv (G n i , isEquivG n i) (compEquiv (eq2 n i) (eq1 n i)) .snd) - -module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) - (conB : (x y : typ B) → ∥ x ≡ y ∥) - (n : ℕ) (Q-is : Iso (typ (Q (pt B))) (S₊ n)) - (Q-is-ptd : Iso.fun Q-is (pt (Q (pt B))) ≡ snd (S₊∙ n)) - (c : (b : typ B) → (Q b →∙ coHomK-ptd n)) - (c-pt : c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) where - - g : (b : typ B) (i : ℕ) → coHomK i → (Q b →∙ coHomK-ptd (i +' n)) - fst (g b i x) y = x ⌣ₖ c b .fst y - snd (g b i x) = cong (x ⌣ₖ_) (c b .snd) ∙ ⌣ₖ-0ₖ i n x - - g' : (b : typ B) (i : ℕ) → coHomK i → coHomK i → (Q b →∙ coHomK-ptd (i +' n)) - fst (g' b i x y) z = x ⌣ₖ c b .fst z +ₖ y ⌣ₖ c b .fst z - snd (g' b i x y) = cong₂ _+ₖ_ (cong (x ⌣ₖ_) (c b .snd) ∙ ⌣ₖ-0ₖ i n x) - (cong (y ⌣ₖ_) (c b .snd) ∙ ⌣ₖ-0ₖ i n y) ∙ rUnitₖ _ (0ₖ _) - - g-hom : (b : typ B) (i : ℕ) → (x y : coHomK i) → g b i (x +ₖ y) ≡ ((g b i x) ++ (g b i y)) - g-hom b i x y = →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ z → rightDistr-⌣ₖ i n x y (c b .fst z)) - - gPathP' : (i : ℕ) → PathP (λ j → coHomK i → (isoToPath Q-is j , ua-gluePath (isoToEquiv Q-is) (Q-is-ptd) j) - →∙ coHomK-ptd (i +' n)) (g (pt B) i) (g-base.g n i) - gPathP' i = - toPathP - (funExt - λ x → →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ y → (λ i → transportRefl (transportRefl x i ⌣ₖ c (pt B) .fst (Iso.inv Q-is (transportRefl y i))) i) - ∙ cong (x ⌣ₖ_) (funExt⁻ c-pt (Iso.inv Q-is y) ∙ cong (genFunSpace n .fst) (Iso.rightInv Q-is y)))) - - g-base : (i : ℕ) → isEquiv (g (pt B) i) - g-base i = transport (λ j → isEquiv (gPathP' i (~ j))) (g-base.isEquiv-g n i) - - g-equiv : (b : typ B) (i : ℕ) → isEquiv (g b i) - g-equiv b i = - pRec (isPropIsEquiv _) - (J (λ b _ → isEquiv (g b i)) - (g-base i)) - (conB (pt B) b) - -module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) - (conB : (x y : typ B) → ∥ x ≡ y ∥₂) - (n : ℕ) (Q-is : Iso (typ (Q (pt B))) (S₊ n)) - (Q-is-ptd : Iso.fun Q-is (pt (Q (pt B))) ≡ snd (S₊∙ n)) where - - is-setQ→K : (b : typ B) → isSet (Q b →∙ coHomK-ptd n) - is-setQ→K b = - sRec (isProp→isOfHLevelSuc 1 isPropIsSet) - (J (λ b _ → isSet (Q b →∙ coHomK-ptd n)) - (subst isSet (cong (_→∙ coHomK-ptd n) - (ua∙ (isoToEquiv (invIso Q-is)) (cong (Iso.inv Q-is) (sym Q-is-ptd) ∙ Iso.leftInv Q-is _))) - (isOfHLevelRetractFromIso 2 (fst (S1' n)) isSetℤ))) - (conB (pt B) b) - - - isConnB : isConnected 3 (typ B) - fst isConnB = ∣ pt B ∣ - snd isConnB = - trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) - λ a → sRec (isOfHLevelTrunc 3 _ _) (cong ∣_∣ₕ) (conB (pt B) a) - - isPropPath : isProp (∥ pt B ≡ pt B ∥₂) - isPropPath = - isOfHLevelRetractFromIso 1 setTruncTrunc2Iso - (isContr→isProp (isConnectedPath _ isConnB (pt B) (pt B))) - - c* : Σ[ c ∈ ((b : typ B) → (Q b →∙ coHomK-ptd n)) ] - (c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) - fst c* b = - sRec (is-setQ→K b) - (J (λ b _ → Q b →∙ coHomK-ptd n) - ((λ x → genFunSpace n .fst (Iso.fun Q-is x)) - , cong (genFunSpace n .fst) Q-is-ptd ∙ genFunSpace n .snd)) (conB (pt B) b) - snd c* = - funExt λ x → (λ i → sRec (is-setQ→K (pt B)) - (J (λ b _ → Q b →∙ coHomK-ptd n) - ((λ x₁ → genFunSpace n .fst (Iso.fun Q-is x₁)) , - (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) - (isPropPath (conB (pt B) (pt B)) ∣ refl ∣₂ i) .fst x) - ∙ (λ i → transportRefl (genFunSpace n .fst (Iso.fun Q-is (transportRefl x i))) i) - - p-help : {b : fst B} (p : pt B ≡ b) → (subst (fst ∘ Q) (sym p) (snd (Q b))) ≡ (snd (Q (pt B))) - p-help {b = b} = - J (λ b p → subst (fst ∘ Q) (sym p) (snd (Q b)) ≡ snd (Q (pt B))) (transportRefl _) - - cEquiv : (b : fst B) (p : ∥ pt B ≡ b ∥₂) - → (c* .fst b) - ≡ sRec (is-setQ→K b) - (λ pp → (λ qb → genFunSpace n .fst (Iso.fun Q-is (subst (fst ∘ Q) (sym pp) qb))) - , cong (genFunSpace n .fst ∘ Iso.fun Q-is) (p-help pp) - ∙ ((λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) p - cEquiv b = - sElim (λ _ → isOfHLevelPath 2 (is-setQ→K b) _ _) - (J (λ b a → c* .fst b ≡ - sRec (is-setQ→K b) (λ pp → - (λ qb → - genFunSpace n .fst (Iso.fun Q-is (subst (fst ∘ Q) (sym pp) qb))) - , - cong (genFunSpace n .fst ∘ Iso.fun Q-is) (p-help pp) ∙ - (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd) - ∣ a ∣₂) - ((λ i → sRec (is-setQ→K (pt B)) - (J (λ b₁ _ → Q b₁ →∙ coHomK-ptd n) - ((λ x → genFunSpace n .fst (Iso.fun Q-is x)) , - (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) - (isPropPath (conB (pt B) (pt B)) ∣ refl ∣₂ i)) - ∙ →∙Homogeneous≡ (isHomogeneousKn n) - (transportRefl ((λ x → genFunSpace n .fst (Iso.fun Q-is x))) - ∙ funExt λ x → cong (genFunSpace n .fst ∘ Iso.fun Q-is) (sym (transportRefl x))))) - -module _ {ℓ ℓ'} (B : Pointed ℓ) (P : typ B → Type ℓ') where - E : Type _ - E = Σ (typ B) P - - Ẽ : Type _ - Ẽ = Pushout {A = E} (λ _ → tt) fst - - i : (n : ℕ) → (P-base : Iso (P (pt B)) (S₊ n)) → S₊ (suc n) → Ẽ - i zero P-base base = inr (pt B) - i zero P-base (loop j) = (sym (push (pt B , Iso.inv P-base false)) - ∙ push ((pt B) , Iso.inv P-base true)) j - i (suc n) P-base north = inl tt - i (suc n) P-base south = inr (pt B) - i (suc n) P-base (merid a i₁) = push (pt B , Iso.inv P-base a) i₁ - - Q : typ B → Pointed ℓ' - Q x = Susp (P x) , north - - F : Type _ - F = Σ (typ B) λ x → Q x .fst - - F̃ : Type _ - F̃ = Pushout {A = typ B} {C = F} (λ _ → tt) λ b → b , north - - invFE : Ẽ → F̃ - invFE (inl x) = inl tt - invFE (inr x) = inr (x , south) - invFE (push (x , a) i₁) = ((push x) ∙ λ i → inr (x , merid a i)) i₁ - - funFE : F̃ → Ẽ - funFE (inl x) = inl tt - funFE (inr (x , north)) = inl tt - funFE (inr (x , south)) = inr x - funFE (inr (x , merid a i₁)) = push (x , a) i₁ - funFE (push a i₁) = inl tt - - IsoFE : Iso F̃ Ẽ - Iso.fun IsoFE = funFE - Iso.inv IsoFE = invFE - Iso.rightInv IsoFE (inl x) = refl - Iso.rightInv IsoFE (inr x) = refl - Iso.rightInv IsoFE (push (x , a) i₁) k = h k i₁ - where - h : cong funFE (((push x) ∙ λ i → inr (x , merid a i))) - ≡ push (x , a) - h = congFunct funFE (push x) (λ i → inr (x , merid a i)) - ∙ sym (lUnit (push (x , a))) - Iso.leftInv IsoFE (inl x) = refl - Iso.leftInv IsoFE (inr (x , north)) = push x - Iso.leftInv IsoFE (inr (x , south)) = refl - Iso.leftInv IsoFE (inr (x , merid a i)) j = - compPath-filler' (push x) (λ i₁ → inr (x , merid a i₁)) (~ j) i - Iso.leftInv IsoFE (push a i₁) k = push a (i₁ ∧ k) - - - funDir1 : ∀ {ℓ} {A : Pointed ℓ} → (F̃ , inl tt) →∙ A → (b : typ B) → Q b →∙ A - fst (funDir1 {A = A , a} (f , p) b) north = f (inr (b , north)) - fst (funDir1 {A = A , a} (f , p) b) south = f (inr (b , south)) - fst (funDir1 {A = A , a} (f , p) b) (merid a₁ i₁) = f (inr (b , merid a₁ i₁)) - snd (funDir1 {A = A , a} (f , p) b) = sym (cong f (push b)) ∙ p - - funDir2 : ∀ {ℓ} {A : Pointed ℓ} → ((b : typ B) → Q b →∙ A) → (F̃ , inl tt) →∙ A - fst (funDir2 {A = A , a} f) (inl x) = a - fst (funDir2 {A = A , a} f) (inr (x , north)) = f x .fst north - fst (funDir2 {A = A , a} f) (inr (x , south)) = f x .fst south - fst (funDir2 {A = A , a} f) (inr (x , merid a₁ i₁)) = f x .fst (merid a₁ i₁) - fst (funDir2 {A = A , a} f) (push a₁ i₁) = snd (f a₁) (~ i₁) - snd (funDir2 {A = A , a} f) = refl - - funDir2-hom : (n : ℕ) → (f g : ((b : typ B) → Q b →∙ coHomK-ptd n)) - → funDir2 (λ b → f b ++ g b) ≡ (funDir2 f ++ funDir2 g) - funDir2-hom n f g = - →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ { (inl x) → sym (rUnitₖ _ (0ₖ _)) - ; (inr (x , north)) → refl - ; (inr (x , south)) → refl - ; (inr (x , merid a i₁)) → refl - ; (push a j) i → compPath-filler (cong₂ _+ₖ_ (snd (f a)) (snd (g a))) - (rUnitₖ n (0ₖ n)) (~ i) (~ j)}) - - funDirSect : ∀ {ℓ} {A : Pointed ℓ} → (x : (b : typ B) → Q b →∙ A) (b : typ B) (q : Q b .fst) - → funDir1 (funDir2 x) b .fst q ≡ x b .fst q - funDirSect f b north = refl - funDirSect f b south = refl - funDirSect f b (merid a i₁) = refl - - funDirRetr : ∀ {ℓ} {A : Pointed ℓ} (f : F̃ → typ A) (p : _) - → (x : F̃) → fst (funDir2 {A = A} (funDir1 (f , p))) x ≡ f x - funDirRetr f p (inl x) = sym p - funDirRetr f p (inr (x , north)) = refl - funDirRetr f p (inr (x , south)) = refl - funDirRetr f p (inr (x , merid a i₁)) = refl - funDirRetr f p (push a i₁) j = compPath-filler (sym (cong f (push a))) p (~ j) (~ i₁) - - Iso2 : ∀ {ℓ} {A : Pointed ℓ} - → Iso ((F̃ , inl tt) →∙ A) - ((b : typ B) → Q b →∙ A) - Iso.fun (Iso2 {A = A , a}) = funDir1 - Iso.inv (Iso2 {A = A , a}) = funDir2 - Iso.rightInv (Iso2 {A = A , a}) f = - funExt λ b → ΣPathP (funExt (funDirSect f b) - , sym (rUnit (snd (f b)))) - Iso.leftInv (Iso2 {A = A , a}) (f , p) = - ΣPathP ((funExt (funDirRetr f p)) - , λ i j → p (~ i ∨ j)) - - ι : (k : ℕ) → Iso ((b : typ B) → Q b →∙ coHomK-ptd k) - ((Ẽ , inl tt) →∙ coHomK-ptd k) - ι k = compIso (invIso Iso2) h - where - h : Iso ((F̃ , inl tt) →∙ coHomK-ptd k) - ((Ẽ , inl tt) →∙ coHomK-ptd k) - Iso.fun h G = (λ x → G .fst (Iso.inv IsoFE x)) - , (snd G) - Iso.inv h G = (λ x → G .fst (Iso.fun IsoFE x)) - , (snd G) - Iso.rightInv h G = - →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ x → cong (G .fst) (Iso.rightInv IsoFE x)) - Iso.leftInv h G = - →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ x → cong (G .fst) (Iso.leftInv IsoFE x)) - - ι-hom : (k : ℕ) → (f g : ((b : typ B) → Q b →∙ coHomK-ptd k)) - → Iso.fun (ι k) (λ b → f b ++ g b) - ≡ (Iso.fun (ι k) f ++ Iso.fun (ι k) g) - ι-hom k f g = - →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ x → funExt⁻ (cong fst (funDir2-hom _ f g)) (invFE x)) - -codomainIsoDep : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : A → Type ℓ'} {C : A → Type ℓ''} - → ((a : A) → Iso (B a) (C a)) - → Iso ((a : A) → B a) ((a : A) → C a) -Iso.fun (codomainIsoDep is) f a = Iso.fun (is a) (f a) -Iso.inv (codomainIsoDep is) f a = Iso.inv (is a) (f a) -Iso.rightInv (codomainIsoDep is) f = funExt λ a → Iso.rightInv (is a) (f a) -Iso.leftInv (codomainIsoDep is) f = funExt λ a → Iso.leftInv (is a) (f a) - -module Thom {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) - (conB : (x y : typ B) → ∥ x ≡ y ∥) - (n : ℕ) (Q-is : Iso (typ (Q B P (pt B))) (S₊ n)) - (Q-is-ptd : Iso.fun Q-is (pt (Q B P (pt B))) ≡ snd (S₊∙ n)) - (c : (b : typ B) → (Q B P b →∙ coHomK-ptd n)) - (c-pt : c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) where - - ϕ : (i : ℕ) → GroupEquiv (coHomGr i (typ B)) - (coHomRedGrDir (i +' n) (Ẽ B P , inl tt)) - fst (ϕ i) = - isoToEquiv - (setTruncIso - (compIso - (codomainIsoDep - λ b → equivToIso ((g B (Q B P) conB n Q-is Q-is-ptd c c-pt b i) - , g-equiv B (Q B P) conB n Q-is Q-is-ptd c c-pt b i)) - (ι B P (i +' n)))) - snd (ϕ i) = - makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) - λ F G → cong ∣_∣₂ (cong (Iso.fun (ι B P (i +' n))) - (funExt (λ a → g-hom B (Q B P) conB n Q-is Q-is-ptd c c-pt a i (F a) (G a))) - ∙ ι-hom B P (i +' n) _ _) - ∙ addAgree (i +' n) _ _) - -module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) - (conB : (x y : typ B) → ∥ x ≡ y ∥₂) - (n : ℕ) (Q-is : Iso (typ (Q B P (pt B))) (S₊ n)) - (Q-is-ptd : Iso.fun Q-is (pt (Q B P (pt B))) ≡ snd (S₊∙ n)) where - - 0-connB : (x y : typ B) → ∥ x ≡ y ∥ - 0-connB x y = sRec (isProp→isSet squash) (∥_∥.∣_∣) (conB x y) - - c = (c* B (Q B P) conB n Q-is Q-is-ptd .fst) - c-ptd = (c* B (Q B P) conB n Q-is Q-is-ptd .snd) - cEq = cEquiv B (Q B P) conB n Q-is Q-is-ptd - - module w = Thom B P 0-connB n Q-is Q-is-ptd c c-ptd - - ϕ = w.ϕ - - E' = E B P - - E'̃ = Ẽ B P - - p : E' → typ B - p = fst - - e : coHom n (typ B) - e = ∣ (λ b → c b .fst south) ∣₂ - - ⌣-hom : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr (i +' n) (typ B)) - fst (⌣-hom i) x = x ⌣ e - snd (⌣-hom i) = - makeIsGroupHom λ f g → rightDistr-⌣ _ _ f g e - - p-hom : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr i E') - p-hom i = coHomMorph _ p - - E-susp : (i : ℕ) → GroupHom (coHomGr i E') (coHomRedGrDir (suc i) (E'̃ , inl tt)) - fst (E-susp i) = sMap λ f → (λ { (inl x) → 0ₖ _ - ; (inr x) → 0ₖ _ - ; (push a j) → Kn→ΩKn+1 _ (f a) j}) , refl - snd (E-susp zero) = - makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) - λ f g → - cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn 1) - (funExt λ { (inl x) → refl - ; (inr x) → refl - ; (push a j) k → (Kn→ΩKn+1-hom zero (f a) (g a) - ∙ ∙≡+₁ (Kn→ΩKn+1 _ (f a)) (Kn→ΩKn+1 _ (g a))) k j}))) - snd (E-susp (suc i)) = - makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) - λ f g → - cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ { (inl x) → refl - ; (inr x) → refl - ; (push a j) k → (Kn→ΩKn+1-hom _ (f a) (g a) - ∙ ∙≡+₂ _ (Kn→ΩKn+1 _ (f a)) (Kn→ΩKn+1 _ (g a))) k j}))) - - module cofibSeq where - j* : (i : ℕ) → GroupHom (coHomRedGrDir i (E'̃ , (inl tt))) (coHomGr i (typ B)) - fst (j* i) = sMap λ f → λ x → fst f (inr x) - snd (j* zero) = - makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) - snd (j* (suc zero)) = - makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) - snd (j* (suc (suc i₁))) = - makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) - - Im-j⊂Ker-p : (i : ℕ) (x : _) → isInIm (j* i) x → isInKer (p-hom i) x - Im-j⊂Ker-p i = - sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) - λ f → pRec (squash₂ _ _) - (uncurry (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) - λ g P → subst (isInKer (p-hom i)) P - (cong ∣_∣₂ (funExt λ x → cong (g .fst) (sym (push x)) ∙ g .snd)))) - - Ker-p⊂Im-j : (i : ℕ) (x : _) → isInKer (p-hom i) x → isInIm (j* i) x - Ker-p⊂Im-j i = - sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) - λ f ker → pRec squash - (λ id → ∣ ∣ (λ { (inl x) → 0ₖ _ - ; (inr x) → f x - ; (push a i₁) → funExt⁻ id a (~ i₁)}) , refl ∣₂ , refl ∣) - (Iso.fun PathIdTrunc₀Iso ker) - - Im-p⊂Ker-Susp : (i : ℕ) (x : _) → isInIm (p-hom i) x → isInKer (E-susp i) x - Im-p⊂Ker-Susp i = - sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) - λ f → pRec (squash₂ _ _) - (uncurry (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) - λ g y → subst (isInKer (E-susp i)) y - (cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ { (inl x) → refl - ; (inr x) → sym (Kn→ΩKn+1 _ (g x)) - ; (push a j) k → Kn→ΩKn+1 i (g (fst a)) (~ k ∧ j)}))))) - open import Cubical.Foundations.Path - Ker-Susp⊂Im-p : (i : ℕ) (x : _) → isInKer (E-susp i) x → isInIm (p-hom i) x - Ker-Susp⊂Im-p i = - sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) - λ f ker → pRec squash - (λ id → ∣ ∣ (λ x → ΩKn+1→Kn i (sym (funExt⁻ (cong fst id) (inr x)))) ∣₂ - , cong ∣_∣₂ (funExt (λ { (a , b) → - cong (ΩKn+1→Kn i) (lUnit _ ∙ cong (_∙ sym (funExt⁻ (λ i₁ → fst (id i₁)) (inr a))) (sym (flipSquare (cong snd id)) - ∙∙ (PathP→compPathR (cong (funExt⁻ (cong fst id)) (push (a , b)))) - ∙∙ assoc _ _ _ - ∙ sym (rUnit _)) - ∙ (sym (assoc _ _ _) - ∙∙ cong (Kn→ΩKn+1 i (f (a , b)) ∙_) (rCancel _) - ∙∙ sym (rUnit _))) - ∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f (a , b))})) ∣) - (Iso.fun PathIdTrunc₀Iso ker) - - Im-Susp⊂Ker-j : (i : ℕ) (x : _) → isInIm (E-susp i) x → isInKer (cofibSeq.j* (suc i)) x - Im-Susp⊂Ker-j i = - sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) - λ g → pRec (squash₂ _ _) - (uncurry (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) - λ f id → pRec (squash₂ _ _) - (λ P → subst (isInKer (cofibSeq.j* (suc i))) (cong ∣_∣₂ P) - (cong ∣_∣₂ refl)) - (Iso.fun PathIdTrunc₀Iso id))) - - Ker-j⊂Im-Susp : (i : ℕ) (x : _) → isInKer (cofibSeq.j* (suc i)) x → isInIm (E-susp i) x - Ker-j⊂Im-Susp i = - sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) - λ f ker - → pRec squash - (λ p → ∣ ∣ (λ x → ΩKn+1→Kn _ (sym (snd f) ∙∙ cong (fst f) (push x) ∙∙ funExt⁻ p (fst x))) ∣₂ - , cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) - (funExt (λ { (inl x) → sym (snd f) - ; (inr x) → sym (funExt⁻ p x) - ; (push a j) k → h f p a k j}))) ∣) - (Iso.fun PathIdTrunc₀Iso ker) - where - h : (f : (E'̃ , inl tt) →∙ coHomK-ptd (suc i)) - → (p : (fst f ∘ inr) ≡ (λ _ → 0ₖ (suc i))) - → (a : E B P) - → PathP (λ i → snd f (~ i) ≡ p (~ i) (fst a)) - (Kn→ΩKn+1 i (ΩKn+1→Kn i (sym (snd f) ∙∙ cong (fst f) (push a) ∙∙ funExt⁻ p (fst a)))) - (cong (fst f) (push a)) - h f p a = Iso.rightInv (Iso-Kn-ΩKn+1 i) _ - ◁ λ i j → doubleCompPath-filler (sym (snd f)) (cong (fst f) (push a)) (funExt⁻ p (fst a)) (~ i) j - - ϕ∘j : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr (i +' n) (typ B)) - ϕ∘j i = compGroupHom (fst (fst (ϕ i)) , snd (ϕ i)) (cofibSeq.j* (i +' n)) - - +'-suc : (i n : ℕ) → (suc i +' n) ≡ suc (i +' n) - +'-suc zero zero = refl - +'-suc (suc i₁) zero = refl - +'-suc zero (suc n) = refl - +'-suc (suc i₁) (suc n) = refl - - private - h : ∀ {ℓ ℓ'} {n m : ℕ} (G : ℕ → Group ℓ) (H : Group ℓ') (p : n ≡ m) - → GroupEquiv (G n) H - → GroupEquiv (G m) H - h {n = n} G H = - J (λ m _ → GroupEquiv (G n) H → GroupEquiv (G m) H) - λ p → p - - h-ret : ∀ {ℓ ℓ'} {n m : ℕ} (G : ℕ → Group ℓ) (H : Group ℓ') - → (e : GroupEquiv (G n) H) - → (p : n ≡ m) - → (x : G m .fst) → invEq (fst e) (fst (fst (h G H p e)) x) ≡ subst (λ x → G x .fst) (sym p) x - h-ret G H e = - J (λ m p → ((x : G m .fst) → invEq (fst e) (fst (fst (h G H p e)) x) ≡ subst (λ x → G x .fst) (sym p) x)) - λ x → cong (invEq (fst e) ) - (λ i → transportRefl (transportRefl (fst (fst e) (transportRefl (transportRefl x i) i)) i) i) - ∙∙ retEq (fst e) x - ∙∙ sym (transportRefl _) - - isEquivϕ' : (i : ℕ) → GroupEquiv (coHomRedGrDir (suc (i +' n)) (E'̃ , inl tt)) - (coHomGr (suc i) (typ B)) - isEquivϕ' i = (h (λ n → coHomRedGrDir n (E'̃ , inl tt)) _ (+'-suc i n) - (invGroupEquiv (ϕ (suc i)))) - - ϕ' : (i : ℕ) → GroupHom (coHomRedGrDir (suc (i +' n)) (E'̃ , inl tt)) - (coHomGr (suc i) (typ B)) - ϕ' i = fst (fst (isEquivϕ' i)) , snd (isEquivϕ' i) - - susp∘ϕ : (i : ℕ) → GroupHom (coHomGr (i +' n) E') (coHomGr (suc i) (typ B)) - susp∘ϕ i = compGroupHom (E-susp (i +' n)) (ϕ' i) - - Im-ϕ∘j⊂Ker-p : (i : ℕ) (x : _) → isInIm (ϕ∘j i) x → isInKer (p-hom _) x - Im-ϕ∘j⊂Ker-p i x p = - cofibSeq.Im-j⊂Ker-p _ x - (pRec squash (uncurry (λ f p → ∣ fst (fst (ϕ _)) f , p ∣)) p) - - Ker-p⊂Im-ϕ∘j : (i : ℕ) (x : _) → isInKer (p-hom _) x → isInIm (ϕ∘j i) x - Ker-p⊂Im-ϕ∘j i x p = - pRec squash (uncurry (λ f p → ∣ (invEq (fst (ϕ _)) f) - , (cong (fst (cofibSeq.j* _)) (secEq (fst (ϕ _)) f) ∙ p) ∣)) - (cofibSeq.Ker-p⊂Im-j _ x p) - - - Im-p⊂KerSusp∘ϕ : (i : ℕ) (x : _) → isInIm (p-hom _) x → isInKer (susp∘ϕ i) x - Im-p⊂KerSusp∘ϕ i x p = cong (fst (ϕ' _)) (Im-p⊂Ker-Susp _ x p) ∙ IsGroupHom.pres1 (snd (ϕ' _)) - - KerSusp∘ϕ⊂Im-p : (i : ℕ) (x : _) → isInKer (susp∘ϕ i) x → isInIm (p-hom _) x - KerSusp∘ϕ⊂Im-p i x p = - Ker-Susp⊂Im-p _ x (sym (retEq (fst (isEquivϕ' _)) _) - ∙ (cong (invEq (fst (isEquivϕ' _))) p - ∙ IsGroupHom.pres1 (snd (invGroupEquiv (isEquivϕ' _))))) - - Im-Susp∘ϕ⊂Ker-ϕ∘j : (i : ℕ) → (x : _) → isInIm (susp∘ϕ i) x → isInKer (ϕ∘j (suc i)) x - Im-Susp∘ϕ⊂Ker-ϕ∘j i x = - pRec (squash₂ _ _) - (uncurry λ f → J (λ x p → isInKer (ϕ∘j (suc i)) x) - ((λ i → fst (cofibSeq.j* _) (fst (fst (ϕ _)) (fst (ϕ' _) (fst (E-susp _) f)))) - ∙∙ cong (fst (cofibSeq.j* _)) - ((h-ret (λ n → coHomRedGrDir n (E'̃ , inl tt)) _ - (invGroupEquiv (ϕ (suc i))) (+'-suc i n)) (fst (E-susp _) f)) - ∙∙ (natTranspLem _ (λ n → fst (cofibSeq.j* n)) (sym (+'-suc i n)) - ∙ cong (subst (λ z → coHomGr z (typ B) .fst) (sym (+'-suc i n))) - (Im-Susp⊂Ker-j _ (fst (E-susp (i +' n)) f) ∣ f , refl ∣) - ∙ tLem i n))) - where - tLem : (i n : ℕ) → subst (λ z → coHomGr z (typ B) .fst) (sym (+'-suc i n)) - (0ₕ _) ≡ 0ₕ _ - tLem zero zero = refl - tLem zero (suc n) = refl - tLem (suc i₁) zero = refl - tLem (suc i₁) (suc n) = refl - - - Ker-ϕ∘j⊂Im-Susp∘ϕ : (i : ℕ) (x : _) - → isInKer (ϕ∘j (suc i)) x → isInIm (susp∘ϕ i) x - Ker-ϕ∘j⊂Im-Susp∘ϕ i x p = - pRec squash - (uncurry (λ f p → ∣ f , cong (fst (fst (isEquivϕ' i))) p ∙ secEq (fst (isEquivϕ' _)) x ∣)) - (Ker-j⊂Im-Susp _ (invEq (fst (isEquivϕ' _)) x) - ((cong (cofibSeq.j* (suc (i +' n)) .fst ) lem2 - ∙ natTranspLem _ (λ n → cofibSeq.j* n .fst) (+'-suc i n)) - ∙∙ cong (transport (λ j → (coHomGr (+'-suc i n j) (typ B) .fst))) p - ∙∙ h2 i n)) - where - lem2 : invEq (fst (isEquivϕ' i)) x ≡ transport (λ j → coHomRed (+'-suc i n j) (E'̃ , inl tt)) (fst (fst (ϕ _)) x) - lem2 = cong (transport (λ j → coHomRed (+'-suc i n j) (E'̃ , inl tt))) - (transportRefl _ ∙ cong (fst (fst (ϕ _))) - λ i → transportRefl (transportRefl x i) i) - - h2 : (i n : ℕ) → transport (λ j → coHomGr (+'-suc i n j) (typ B) .fst) - (GroupStr.1g (coHomGr (suc i +' n) (typ B) .snd)) ≡ 0ₕ _ - h2 zero zero = refl - h2 zero (suc n) = refl - h2 (suc i₁) zero = refl - h2 (suc i₁) (suc n) = refl - - - ϕ∘j≡ : (i : ℕ) → ϕ∘j i ≡ ⌣-hom i - ϕ∘j≡ i = - Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (funExt (sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - λ _ → refl)) - -open import Cubical.Experiments.Brunerie -open import Cubical.HITs.Hopf - -open import Cubical.HITs.Join - -module fibS1 = hopfBase S1-AssocHSpace (sphereElim2 _ (λ _ _ → squash) ∣ refl ∣) - -S¹Hopf : S₊ 2 → Type -S¹Hopf = fibS1.Hopf - -TotalHopf' : Type _ -TotalHopf' = Σ (S₊ 2) S¹Hopf - -CP2 : Type _ -CP2 = fibS1.megaPush - -fibr : CP2 → Type _ -fibr = fibS1.P - -hopf : join S¹ S¹ → S₊ 2 -hopf x = fst (JoinS¹S¹→TotalHopf x) - -E* : Type _ -E* = fibS1.totalSpaceMegaPush - -IsoE' : Iso E* (join S¹ (join S¹ S¹)) -IsoE' = fibS1.IsoJoin₁ - -IsoE2 : (join S¹ (join S¹ S¹)) ≡ join S¹ (S₊ 3) -IsoE2 = cong (join S¹) (sym S³≡joinS¹S¹ ∙ isoToPath IsoS³S3) - -IsoTotalHopf' : Iso fibS1.TotalSpaceHopf' (join S¹ S¹) -IsoTotalHopf' = fibS1.joinIso₁ - -CP' : Type _ -CP' = Pushout (λ _ → tt) hopf - -conCP2 : (x y : CP2) → ∥ x ≡ y ∥₂ -conCP2 x y = sRec2 squash₂ (λ p q → ∣ p ∙ sym q ∣₂) (conCP2' x) (conCP2' y) - where - conCP2' : (x : CP2) → ∥ x ≡ inl tt ∥₂ - conCP2' (inl x) = ∣ refl ∣₂ - conCP2' (inr x) = sphereElim 1 {A = λ x → ∥ inr x ≡ inl tt ∥₂} (λ _ → squash₂) ∣ sym (push (inl base)) ∣₂ x - conCP2' (push a i₁) = ll a i₁ - where - h2 : ∀ {ℓ} {A : fibS1.TotalSpaceHopf' → Type ℓ} → ((a : _) → isProp (A a)) - → A (inl base) - → ((a : fibS1.TotalSpaceHopf') → A a) - h2 {A = A} p b = - PushoutToProp p (sphereElim 0 (λ _ → p _) b) - (sphereElim 0 (λ _ → p _) (subst A (push (base , base)) b)) - - ll : (a : fibS1.TotalSpaceHopf') → PathP (λ i → ∥ Path CP2 (push a i) (inl tt) ∥₂) (conCP2' (inl tt)) (conCP2' (inr (fibS1.induced a))) - ll = h2 (λ _ → isOfHLevelPathP' 1 squash₂ _ _) λ j → ∣ (λ i → push (inl base) (~ i ∧ j)) ∣₂ - -module GysinS1 = Gysin (CP2 , inl tt) fibr conCP2 2 idIso refl - -PushoutReplaceBase : - ∀ {ℓ ℓ' ℓ''} {A B : Type ℓ} {C : Type ℓ'} {D : Type ℓ''} {f : A → C} {g : A → D} - (e : B ≃ A) → Pushout (f ∘ fst e) (g ∘ fst e) ≡ Pushout f g -PushoutReplaceBase {f = f} {g = g} = - EquivJ (λ _ e → Pushout (f ∘ fst e) (g ∘ fst e) ≡ Pushout f g) - refl - -isContrH³E : isContr (coHom 3 (GysinS1.E')) -isContrH³E = - subst isContr - (sym (isoToPath (compIso (Iso1 0) (Iso2' 0))) - ∙ cong (coHom 3) (sym (isoToPath IsoE' ∙ IsoE2))) - (ll5 0 (snotz ∘ sym)) - -isContrH⁴E : isContr (coHom 4 (GysinS1.E')) -isContrH⁴E = - subst isContr - (sym (isoToPath (compIso (Iso1 1) (Iso2' 1))) - ∙ cong (coHom 4) (sym (isoToPath IsoE' ∙ IsoE2))) - (ll5 1 λ p → snotz (sym (cong predℕ p))) - -genH2 = GysinS1.e - -S³→Groupoid : ∀ {ℓ} (P : join S¹ S¹ → Type ℓ) - → ((x : _) → isGroupoid (P x)) - → P (inl base) - → (x : _) → P x -S³→Groupoid P grp b = - transport (λ i → (x : (sym (isoToPath S³IsojoinS¹S¹) ∙ isoToPath IsoS³S3) (~ i)) - → P (transp (λ j → (sym (isoToPath S³IsojoinS¹S¹) ∙ isoToPath IsoS³S3) (~ i ∧ ~ j)) i x)) - (sphereElim _ (λ _ → grp _) b) - -TotalHopf→Gpd : ∀ {ℓ} (P : fibS1.TotalSpaceHopf' → Type ℓ) - → ((x : _) → isGroupoid (P x)) - → P (inl base) - → (x : _) → P x -TotalHopf→Gpd P grp b = - transport (λ i → (x : sym (isoToPath IsoTotalHopf') i) - → P (transp (λ j → isoToPath IsoTotalHopf' (~ i ∧ ~ j)) i x)) - (S³→Groupoid _ (λ _ → grp _) b) - -TotalHopf→Gpd' : ∀ {ℓ} (P : Σ (S₊ 2) fibS1.Hopf → Type ℓ) - → ((x : _) → isGroupoid (P x)) - → P (north , base) - → (x : _) → P x -TotalHopf→Gpd' P grp b = - transport (λ i → (x : ua (_ , fibS1.isEquivTotalSpaceHopf'→TotalSpace) i) - → P (transp (λ j → ua (_ , fibS1.isEquivTotalSpaceHopf'→TotalSpace) (i ∨ j)) i x)) - (TotalHopf→Gpd _ (λ _ → grp _) b) - -CP2→Groupoid : ∀ {ℓ} {P : CP2 → Type ℓ} - → ((x : _) → is2Groupoid (P x)) - → (b : P (inl tt)) - → (s2 : ((x : _) → P (inr x))) - → PathP (λ i → P (push (inl base) i)) b (s2 north) - → (x : _) → P x -CP2→Groupoid {P = P} grp b s2 pp (inl x) = b -CP2→Groupoid {P = P} grp b s2 pp (inr x) = s2 x -CP2→Groupoid {P = P} grp b s2 pp (push a i₁) = h23 a i₁ - where - h23 : (a : fibS1.TotalSpaceHopf') → PathP (λ i → P (push a i)) b (s2 _) - h23 = TotalHopf→Gpd _ (λ _ → isOfHLevelPathP' 3 (grp _) _ _) pp - -mm : (S₊ 2 → coHomK 2) → (CP2 → coHomK 2) -mm f = λ { (inl x) → 0ₖ _ - ; (inr x) → f x -ₖ f north - ; (push a i) → TotalHopf→Gpd (λ x → 0ₖ 2 ≡ f (fibS1.TotalSpaceHopf'→TotalSpace x .fst) -ₖ f north) - (λ _ → isOfHLevelTrunc 4 _ _) - (sym (rCancelₖ 2 (f north))) - a i} - -e-inv : coHomGr 2 (S₊ 2) .fst → coHomGr 2 CP2 .fst -e-inv = sMap mm - -cancel : (f : CP2 → coHomK 2) → ∥ mm (f ∘ inr) ≡ f ∥ -cancel f = - pRec squash - (λ p → pRec squash - (λ f → ∣ funExt f ∣) - (zz2 p)) - (h1 (f (inl tt))) - where - h1 : (x : coHomK 2) → ∥ x ≡ 0ₖ 2 ∥ - h1 = coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash) ∣ refl ∣ - - zz2 : (p : f (inl tt) ≡ 0ₖ 2) → ∥ ((x : CP2) → mm (f ∘ inr) x ≡ f x) ∥ - zz2 p = trRec squash (λ pp → - ∣ CP2→Groupoid (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) - (sym p) - (λ x → (λ i → f (inr x) -ₖ f (push (inl base) (~ i))) - ∙∙ (λ i → f (inr x) -ₖ p i) - ∙∙ rUnitₖ 2 (f (inr x))) - pp ∣) - l - where - l : hLevelTrunc 1 (PathP ((λ i₁ → mm (f ∘ inr) (push (inl base) i₁) ≡ f (push (inl base) i₁))) - (sym p) - (((λ i₁ → f (inr north) -ₖ f (push (inl base) (~ i₁))) ∙∙ - (λ i₁ → f (inr north) -ₖ p i₁) ∙∙ rUnitₖ 2 (f (inr north))))) - l = isConnectedPathP 1 (isConnectedPath 2 (isConnectedKn 1) _ _) _ _ .fst - -e' : GroupIso (coHomGr 2 CP2) (coHomGr 2 (S₊ 2)) -Iso.fun (fst e') = sMap (_∘ inr) -Iso.inv (fst e') = e-inv -Iso.rightInv (fst e') = - sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - λ f → trRec {B = Iso.fun (fst e') (Iso.inv (fst e') ∣ f ∣₂) ≡ ∣ f ∣₂} - (isOfHLevelPath 2 squash₂ _ _) - (λ p → cong ∣_∣₂ (funExt λ x → cong (λ y → (f x) -ₖ y) p ∙ rUnitₖ 2 (f x))) - (Iso.fun (PathIdTruncIso _) (isContr→isProp (isConnectedKn 1) ∣ f north ∣ ∣ 0ₖ 2 ∣)) -Iso.leftInv (fst e') = - sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - λ f → pRec (squash₂ _ _) (cong ∣_∣₂) (cancel f) -snd e' = - makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ f g → refl) - -grrg : GroupIso (coHomGr 2 CP2) ℤGroup -grrg = compGroupIso e' (Hⁿ-Sⁿ≅ℤ 1) - -⌣hom : GroupEquiv (coHomGr 2 CP2) (coHomGr 4 CP2) -fst (fst ⌣hom) = GysinS1.⌣-hom 2 .fst -snd (fst ⌣hom) = subst isEquiv (cong fst (GysinS1.ϕ∘j≡ 2)) h23 - where - h23 : isEquiv (GysinS1.ϕ∘j 2 .fst) - h23 = - SES→Iso - (GroupPath _ _ .fst (invGroupEquiv - (isContr→≃Unit isContrH³E - , makeIsGroupHom λ _ _ → refl))) - (GroupPath _ _ .fst (invGroupEquiv - (isContr→≃Unit isContrH⁴E - , makeIsGroupHom λ _ _ → refl))) - (GysinS1.susp∘ϕ 1) - (GysinS1.ϕ∘j 2) - (GysinS1.p-hom 4) - (GysinS1.Ker-ϕ∘j⊂Im-Susp∘ϕ _) - (GysinS1.Ker-p⊂Im-ϕ∘j _) -snd ⌣hom = GysinS1.⌣-hom 2 .snd - -⌣hom2 : GroupEquiv (coHomGr 0 CP2) (coHomGr 2 CP2) -fst (fst ⌣hom2) = GysinS1.⌣-hom 0 .fst -snd (fst ⌣hom2) = subst isEquiv (cong fst (GysinS1.ϕ∘j≡ 0)) h23 - where - cr : GroupHom (coHomGr 1 (S₊ 2)) (coHomGr 0 CP2) - fst cr = - sMap λ f → λ { (inl x) → 0 - ; (inr x) → ΩKn+1→Kn 0 ({!!} ∙∙ cong f {!merid x!} ∙∙ {!!}) - ; (push a i₁) → {!!}} - snd cr = {!!} - h23 : isEquiv (GysinS1.ϕ∘j 0 .fst) - h23 = - SES→Iso - {!!} - {!!} - cr - (GysinS1.ϕ∘j 0) - (GysinS1.p-hom 2) - (sElim {!!} (λ f inker → pRec {!!} (λ g → {!funExt⁻ g (inl tt)!}) (Iso.fun PathIdTrunc₀Iso inker))) - (GysinS1.Ker-p⊂Im-ϕ∘j _) -snd ⌣hom2 = {!!} - -isGenerator≃ℤ : ∀ {ℓ} (G : Group ℓ) (g : fst G) - → Type _ -isGenerator≃ℤ G g = - Σ[ e ∈ GroupIso G ℤGroup ] abs (Iso.fun (fst e) g) ≡ 1 - -⌣-pres1 : ∀ {ℓ} (G : Type ℓ) (n : ℕ) → Type _ -⌣-pres1 G n = - (z : Σ[ x ∈ coHomGr n G .fst ] isGenerator≃ℤ (coHomGr n G) x) - → isGenerator≃ℤ (coHomGr (n +' n) G) (z .fst ⌣ z .fst) - - -genr : coHom 2 CP2 -genr = ∣ CP2→Groupoid (λ _ → isOfHLevelTrunc 4) - (0ₖ _) - ∣_∣ - refl ∣₂ - -≡CP2 : (f g : CP2 → coHomK 2) → ∥ f ∘ inr ≡ g ∘ inr ∥ → Path (coHom 2 CP2) ∣ f ∣₂ ∣ g ∣₂ -≡CP2 f g = pRec (squash₂ _ _) - (λ p → pRec (squash₂ _ _) (λ id → trRec (squash₂ _ _) - (λ pp → cong ∣_∣₂ - (funExt (CP2→Groupoid (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) - id - (funExt⁻ p) - (compPathR→PathP pp)))) - (pp2 (f (inl tt)) (g (inl tt)) id - (cong f (push (inl base)) ∙ (funExt⁻ p north) ∙ cong g (sym (push (inl base)))))) - (conn (f (inl tt)) (g (inl tt)))) - - where - conn : (x y : coHomK 2) → ∥ x ≡ y ∥ - conn = coHomK-elim _ (λ _ → isSetΠ λ _ → isOfHLevelSuc 1 squash) - (coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash) ∣ refl ∣) - - pp2 : (x y : coHomK 2) (p q : x ≡ y) → hLevelTrunc 1 (p ≡ q) - pp2 x y = λ p q → Iso.fun (PathIdTruncIso _) (isContr→isProp (isConnectedPath _ (isConnectedKn 1) x y) ∣ p ∣ ∣ q ∣) - -s2 : ℤ -s2 = Iso.fun (fst grrg) (genr) - --* : S¹ → S¹ --* x = {!!} - -rUnit* : (x : S¹) → x * base ≡ x -rUnit* base = refl -rUnit* (loop i₁) = refl - - -mmInv : (x : S¹) → (merid (invLooper x) ∙ sym (merid base)) ≡ sym (merid x ∙ sym (merid base)) -mmInv x = {!!} - -meridP : (a x : S¹) → Path (Path (coHomK 2) _ _) (cong ∣_∣ₕ (merid (a * x) ∙ sym (merid base))) - ((cong ∣_∣ₕ (merid a ∙ sym (merid base))) ∙ (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) -meridP = wedgeconFun _ _ (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) - (λ x → {!!}) - {!!} - {!!} - -commS1 : (a x : S¹) → a * x ≡ x * a -commS1 = wedgeconFun _ _ (λ _ _ → isGroupoidS¹ _ _) - (sym ∘ rUnit*) - rUnit* - refl - -s233 : (a x : S¹) → (invEq (fibS1.μ-eq a) x) * a ≡ (invLooper a * x) * a -s233 a x = secEq (fibS1.μ-eq a) x ∙∙ cong (_* x) (lemmie a) ∙∙ assocHSpace.μ-assoc S1-AssocHSpace a (invLooper a) x ∙ commS1 _ _ - where - lemmie : (x : S¹) → ptSn 1 ≡ x * (invLooper x) - lemmie base = refl - lemmie (loop i) j = - hcomp (λ r → λ {(i = i0) → base ; (i = i1) → base ; (j = i0) → base}) - base - -ss23 : (a x : S¹) → invEq (fibS1.μ-eq a) x ≡ invLooper a * x -ss23 a x = sym (retEq (fibS1.μ-eq a) (invEq (fibS1.μ-eq a) x)) - ∙∙ cong (invEq (fibS1.μ-eq a)) (s233 a x) - ∙∙ retEq (fibS1.μ-eq a) (invLooper a * x) - -ll : GysinS1.e ≡ genr -ll = ≡CP2 _ _ ∣ funExt (λ x → funExt⁻ (cong fst (ll32 x)) south) ∣ - where - merid* : (a b : S¹) → Path (Path (coHomK 2) _ _) {!!} {!!} - merid* = {!!} - - kr : (x : Σ (S₊ 2) fibS1.Hopf) → Path (hLevelTrunc 4 _) ∣ fst x ∣ ∣ north ∣ - kr = uncurry λ { north → λ y → cong ∣_∣ₕ (merid base ∙ sym (merid y)) - ; south → λ y → cong ∣_∣ₕ (sym (merid y)) - ; (merid a i) → lem a i} - where - lem : (a : S¹) → PathP (λ i → (y : fibS1.Hopf (merid a i)) → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a i ∣ ∣ north ∣) - (λ y → cong ∣_∣ₕ (merid base ∙ sym (merid y))) - λ y → cong ∣_∣ₕ (sym (merid y)) - lem a = toPathP (funExt λ x → cong (transport ((λ i₁ → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a i₁ ∣ ∣ north ∣))) - ((λ i → (λ z → cong ∣_∣ₕ (merid base - ∙ sym (merid (transport (λ j → uaInvEquiv (fibS1.μ-eq a) (~ i) j) x))) z)) - ∙ λ i → cong ∣_∣ₕ (merid base - ∙ sym (merid (transportRefl (invEq (fibS1.μ-eq a) x) i)))) - ∙∙ (λ i → transp (λ i₁ → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a (i₁ ∨ i) ∣ ∣ north ∣) i - (compPath-filler' (cong ∣_∣ₕ (sym (merid a))) - (cong ∣_∣ₕ (merid base ∙ sym (merid (ss23 a x i)))) i)) - ∙∙ {!!} - ∙∙ {!!} - ∙∙ {!!}) - - psst : (x : S₊ 2) → Q (CP2 , inl tt) fibr (inr x) →∙ coHomK-ptd 2 - fst (psst x) north = ∣ north ∣ - fst (psst x) south = ∣ x ∣ - fst (psst x) (merid a i₁) = kr (x , a) (~ i₁) - snd (psst x) = refl - - ll32-fst : GysinS1.c (inr north) .fst ≡ psst north .fst - ll32-fst = cong fst (GysinS1.cEq (inr north) ∣ push (inl base) ∣₂) - ∙ (funExt l) - where - l : (qb : _) → - ∣ (subst (fst ∘ Q (CP2 , inl tt) fibr) (sym (push (inl base))) qb) ∣ - ≡ psst north .fst qb - l north = refl - l south = cong ∣_∣ₕ (sym (merid base)) - l (merid a i) j = - hcomp (λ k → λ { (i = i0) → ∣ merid a (~ k ∧ j) ∣ - ; (i = i1) → ∣ merid base (~ j) ∣ - ; (j = i0) → ∣ transportRefl (merid a i) (~ k) ∣ - ; (j = i1) → ∣ compPath-filler (merid base) (sym (merid a)) k (~ i) ∣ₕ}) - (hcomp (λ k → λ { (i = i0) → ∣ merid a (j ∨ ~ k) ∣ - ; (i = i1) → ∣ merid base (~ j ∨ ~ k) ∣ - ; (j = i0) → ∣ merid a (~ k ∨ i) ∣ - ; (j = i1) → ∣ merid base (~ i ∨ ~ k) ∣ₕ}) - ∣ south ∣) - - is-setHepl : (x : S₊ 2) → isSet (Q (CP2 , inl tt) fibr (inr x) →∙ coHomK-ptd 2) - is-setHepl = sphereElim _ (λ _ → isProp→isOfHLevelSuc 1 (isPropIsOfHLevel 2)) - (isOfHLevel↑∙' 0 1) - - ll32 : (x : S₊ 2) → (GysinS1.c (inr x) ≡ psst x) - ll32 = sphereElim _ (λ x → isOfHLevelPath 2 (is-setHepl x) _ _) - (→∙Homogeneous≡ (isHomogeneousKn _) ll32-fst) -HopfA*A : ∀ {ℓ} {A : Type ℓ} - → (join A A → Type ℓ) - → {!(funExt⁻ (λ i₁ → fst (ll32 north i₁)) south)!} -HopfA*A = {!!} - -S₊2→S₊2 : CP' → Type ℓ-zero -S₊2→S₊2 (inl x) = {!hopfBase.joinIso₁!} -S₊2→S₊2 (inr x) = {!x!} -S₊2→S₊2 (push a i₁) = {!!} - - where - {- - fill (λ k → ua (μ-eq y) i) - (λ k → λ {(i = i0) → HSpace.μ e (pt A) x - ; (i = i1) → assocHSpace.μ-assoc e-ass (pt A) x y k}) - (inS ((ua-gluePt (μ-eq y) i (HSpace.μ e (pt A) x)))) - j -} - -{- -Goal: snd (v' (pt A) (push a i₁)) ≡ - ua-gluePt (μ-eq (snd a)) i₁ (fst a) -———— Boundary —————————————————————————————————————————————— -i₁ = i0 ⊢ HSpace.μₗ e (fst a) -i₁ = i1 ⊢ HSpace.μₗ e (HSpace.μ e (fst a) (snd a)) --} +open import Cubical.HITs.SetQuotients renaming (_/_ to _/'_) + + +S³+' : S₊ 3 → S₊ 3 → S₊ 3 +S³+' north y = y +S³+' south y = y +S³+' (merid a i) north = (merid a ∙ sym (merid north)) i +S³+' (merid a i) south = (sym (merid north) ∙ merid a) i +S³+' (merid north i) (merid b j) = hcomp (λ k → λ { (i = i0) → merid b j + ; (i = i1) → merid b j + ; (j = i0) → rCancel (merid north) (~ k) i + ; (j = i1) → lCancel (merid north) (~ k) i}) + (merid b j) +S³+' (merid south i) (merid b j) = hcomp (λ k → λ { (i = i0) → merid b j + ; (i = i1) → merid b j + ; (j = i0) → (merid (merid base k) ∙ sym (merid north)) i + ; (j = i1) → (sym (merid north) ∙ merid (merid base k)) i}) + (hcomp (λ k → λ { (i = i0) → merid b j + ; (i = i1) → merid b j + ; (j = i0) → rCancel (merid north) (~ k) i + ; (j = i1) → lCancel (merid north) (~ k) i}) + (merid b j)) +S³+' (merid (merid a j) i) (merid b ik) = {!!} + +S³+ : join S¹ S¹ → join S¹ S¹ → S₊ 3 +S³+ (inl x) y = north +S³+ (inr x) y = north +S³+ (push a b i) (inl x) = north +S³+ (push a b i) (inr x) = north +S³+ (push a b i) (push c d j) = + hcomp (λ k → λ { (i = i0) → north + ; (i = i1) → north + ; (j = i0) → rCancel (merid north) k i + ; (j = i1) → rCancel (merid north) k i}) + ((merid ((merid ((a * c) * (b * c)) ∙ sym (merid base)) j) ∙ sym (merid north)) i) + +rUnitS : (x : join S¹ S¹) → S³+ x (inl base) ≡ north +rUnitS (inl x) = refl +rUnitS (inr x) = refl +rUnitS (push a b i) = refl + + + +-- sss : ∥ typ ((Ω^ 3) (S₊∙ 2)) ∥₂ → ∥ typ ((Ω^ 4) (S₊∙ 3)) ∥₂ +-- sss = sMap (Ω→ (Ω→ (Ω→ ((λ x → merid x ∙ sym (merid north)) , (rCancel _)))) .fst) + +-- sss' : typ ((Ω^ 3) (S₊∙ 2)) → typ ((Ω^ 4) (S₊∙ 3)) +-- sss' x i j k l = +-- hcomp (λ r → λ { (i = i0) → rCancel (merid north) r l +-- ; (i = i1) → rCancel (merid north) r l +-- ; (j = i0) → rCancel (merid north) r l +-- ; (j = i1) → rCancel (merid north) r l +-- ; (k = i0) → rCancel (merid north) r l +-- ; (k = i1) → rCancel (merid north) r l +-- ; (l = i0) → north +-- ; (l = i1) → north}) +-- ((merid (x i j k) ∙ sym (merid north)) l) + + +-- sss≡ : (x : _) → sss' x ≡ sym (sss' x) +-- sss≡ x w i j k l = +-- hcomp ((λ r → λ { (i = i0) → rCancel (merid north) r l +-- ; (i = i1) → rCancel (merid north) r l +-- ; (j = i0) → rCancel (merid north) r l +-- ; (j = i1) → rCancel (merid north) r l +-- ; (k = i0) → rCancel (merid north) r l +-- ; (k = i1) → rCancel (merid north) r l +-- ; (l = i0) → north +-- ; (l = i1) → north})) +-- (hcomp ((λ r → λ { (i = i0) → symDistr (merid (x i j k)) (sym (merid north)) (~ r) (~ l) +-- ; (i = i1) → symDistr (merid (x i j k)) (sym (merid north)) (~ r) (~ l) +-- ; (j = i0) → symDistr (merid (x i j k)) (sym (merid north)) (~ r) (~ l) +-- ; (j = i1) → symDistr (merid (x i j k)) (sym (merid north)) (~ r) (~ l) +-- ; (k = i0) → symDistr (merid (x i j k)) (sym (merid north)) (~ r) (~ l) +-- ; (k = i1) → symDistr (merid (x i j k)) (sym (merid north)) (~ r) (~ l) +-- ; (l = i0) → north +-- ; (l = i1) → north +-- ; (w = i0) → symDistr (merid (x i j k)) (sym (merid north)) (~ r) (~ l) -- (merid (x i j k) ∙ sym (merid north)) l +-- ; (w = i1) → {!!}})) +-- {!!}) -- ((merid (x i j k) ∙ sym (merid north)) ((~ l ∧ w) ∨ (l ∧ ~ w)))) + +-- Π₄S³ : Type _ +-- Π₄S³ = ∥ typ ((Ω^ 3) (S₊∙ 2)) ∥₂ /' λ f g → sss f ≡ sss g +-- open import Cubical.Homotopy.Freudenthal + +-- ll123 : Type _ +-- ll123 = Pushout {A = {!!}} {B = {!!}} {!sss ∣ ? ∣₂!} (λ _ → tt) + +-- isEquiv→isEquivΩ→ : {!!} +-- isEquiv→isEquivΩ→ = {!!} + +-- module _ {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} {x y : A} (n : ℕ) +-- (f : A → B) (con : isConnectedFun (suc (suc n)) f) where +-- fib≡ : (b : B) (f1 f2 : fiber f b) → Iso (f1 ≡ f2) (fiber (cong f) (snd f1 ∙ sym (snd f2))) +-- fib≡ = {!!} + +-- ll : (p : f x ≡ f y) → isContr (hLevelTrunc (suc n) (fiber (cong f) p)) +-- ll p = subst (isContr ∘ hLevelTrunc (suc n)) +-- (isoToPath (fib≡ (f x) (x , refl) (y , sym p)) +-- ∙ cong (fiber (cong f)) (sym (lUnit p))) +-- (isOfHLevelRetractFromIso 0 +-- (invIso (PathIdTruncIso _)) +-- ((isContr→isProp (con (f x)) _ _) +-- , isProp→isSet (isContr→isProp (con (f x))) _ _ _)) + +-- conbase : isConnectedFun 4 {A = (S₊ 2)} (σ 1 {A = S₊∙ 2} (sphereConnected 2)) +-- conbase = isConnectedσ 1 (sphereConnected 2) + +-- conn-cong : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A} (n : ℕ) +-- → (f : A → B) +-- → isConnectedFun (suc (suc n)) f +-- → (p : _) → hLevelTrunc (suc n) (fiber (cong f) p) +-- conn-cong {x = x} {y = y} n f con p = +-- trRec (isOfHLevelTrunc (suc n)) +-- (λ pp → ∣ (cong fst pp) , (cong sym (rUnit _) +-- ∙∙ cong sym (cong (sym (λ i₁ → f (fst (pp i₁))) ∙_) (rUnit refl)) +-- ∙∙ cong sym (PathP→compPathL (cong snd pp))) ∣) l +-- where +-- l : hLevelTrunc (suc n) ((x , refl) ≡ (y , sym p)) +-- l = Iso.fun (PathIdTruncIso _) (isContr→isProp (con (f x)) ∣ x , refl ∣ ∣ y , sym p ∣) + +-- pr : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A} (n : ℕ) +-- → (f : A → B) +-- → isConnectedFun (suc (suc n)) f +-- → (p : _) → isProp (hLevelTrunc (suc n) (fiber {A = x ≡ y} (cong f) p)) +-- pr {x = x} {y = y} n f co p = +-- trElim2 (λ _ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) +-- (uncurry λ p2 P → uncurry λ q Q → trRec (isOfHLevelPath' n (isOfHLevelTrunc (suc n)) _ _) +-- (λ p3 → cong ∣_∣ₕ (ΣPathP ((sym (cong (cong fst) p3)) , {!cong (cong snd) p3!}))) +-- (l q p2 Q P)) +-- where +-- l : (q p2 : x ≡ y) (Q : cong f q ≡ p) (P : cong f p2 ≡ p) +-- → hLevelTrunc n (ΣPathP (q , λ j i → f (q (~ i ∧ j))) +-- ≡ ΣPathP (p2 , compPathL→PathP (cong (sym (cong f p2) ∙_) (sym (rUnit refl)) +-- ∙ sym (rUnit _) +-- ∙ cong sym (P ∙ sym Q)))) +-- l q p2 Q P = +-- Iso.fun (PathIdTruncIso _) (isContr→isProp +-- (isConnectedPath _ +-- (co (f x)) (x , refl) (y , (cong f (sym q)))) +-- ∣ ΣPathP (q , λ j i → f (q (~ i ∧ j))) ∣ +-- ∣ ΣPathP (p2 , compPathL→PathP (cong (sym (cong f p2) ∙_) (sym (rUnit refl)) +-- ∙ sym (rUnit _) +-- ∙ cong sym (P ∙ sym Q))) ∣) + +-- fib-sss : (f : ∥ typ ((Ω^ 4) (S₊∙ 3)) ∥₂) → fiber sss f +-- fib-sss = sElim {!!} λ f → trRec {!!} {!!} (conn-cong {!!} +-- (Ω→ (Ω→ ((λ x → merid x ∙ sym (merid north)) , (rCancel _))) .fst) +-- {!conn-cong _ ? ? ?!} +-- (Ω→ +-- (Ω→ +-- ((λ x → merid x ∙ (λ i → merid north (~ i))) , +-- rCancel (merid north))) +-- .snd ∙∙ f ∙∙ sym (Ω→ +-- (Ω→ +-- ((λ x → merid x ∙ (λ i → merid north (~ i))) , +-- rCancel (merid north))) +-- .snd))) + + +-- -- SES→Iso : ∀ {ℓ ℓ'} {L R : Group ℓ-zero} +-- -- → {G : Group ℓ} {H : Group ℓ'} +-- -- → UnitGroup ≡ L +-- -- → UnitGroup ≡ R +-- -- → (lhom : GroupHom L G) (midhom : GroupHom G H) (rhom : GroupHom H R) +-- -- → ((x : _) → isInKer midhom x → isInIm lhom x) +-- -- → ((x : _) → isInKer rhom x → isInIm midhom x) +-- -- → isEquiv (fst midhom) +-- -- SES→Iso {R = R} {G = G} {H = H} = +-- -- J (λ L _ → UnitGroup ≡ R → +-- -- (lhom : GroupHom L G) (midhom : GroupHom G H) +-- -- (rhom : GroupHom H R) → +-- -- ((x : fst G) → isInKer midhom x → isInIm lhom x) → +-- -- ((x : fst H) → isInKer rhom x → isInIm midhom x) → +-- -- isEquiv (fst midhom)) +-- -- ((J (λ R _ → (lhom : GroupHom UnitGroup G) (midhom : GroupHom G H) +-- -- (rhom : GroupHom H R) → +-- -- ((x : fst G) → isInKer midhom x → isInIm lhom x) → +-- -- ((x : _) → isInKer rhom x → isInIm midhom x) → +-- -- isEquiv (fst midhom)) +-- -- main)) +-- -- where +-- -- main : (lhom : GroupHom UnitGroup G) (midhom : GroupHom G H) +-- -- (rhom : GroupHom H UnitGroup) → +-- -- ((x : fst G) → isInKer midhom x → isInIm lhom x) → +-- -- ((x : fst H) → isInKer rhom x → isInIm midhom x) → +-- -- isEquiv (fst midhom) +-- -- main lhom midhom rhom lexact rexact = +-- -- BijectionIsoToGroupEquiv {G = G} {H = H} +-- -- bijIso' .fst .snd +-- -- where +-- -- bijIso' : BijectionIso G H +-- -- BijectionIso.fun bijIso' = midhom +-- -- BijectionIso.inj bijIso' x inker = +-- -- pRec (GroupStr.is-set (snd G) _ _) +-- -- (λ s → sym (snd s) ∙ IsGroupHom.pres1 (snd lhom)) (lexact _ inker) +-- -- BijectionIso.surj bijIso' x = rexact x refl + +-- -- HopfInvariantPush : (n : ℕ) +-- -- → (f : S₊ (3 +ℕ n +ℕ n) → S₊ (2 +ℕ n)) +-- -- → Type _ +-- -- HopfInvariantPush n f = Pushout (λ _ → tt) f + + +-- -- Hopfα : (n : ℕ) +-- -- → (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) +-- -- → coHom (2 +ℕ n) (HopfInvariantPush n (fst f)) +-- -- Hopfα n f = ∣ (λ { (inl x) → 0ₖ _ +-- -- ; (inr x) → ∣ x ∣ +-- -- ; (push a i) → help a (~ i)}) ∣₂ +-- -- where +-- -- help : (a : S₊ (3 +ℕ n +ℕ n)) → ∣ fst f a ∣ ≡ 0ₖ (2 +ℕ n) +-- -- help = sphereElim _ (λ _ → isOfHLevelPlus' {n = n} (3 +ℕ n) +-- -- (isOfHLevelPath' (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n)) _ _)) (cong ∣_∣ₕ (snd f)) + +-- -- Hopfβ : (n : ℕ) +-- -- → (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) +-- -- → coHom (4 +ℕ (n +ℕ n)) (HopfInvariantPush n (fst f)) +-- -- Hopfβ n f = fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) ∣ ∣_∣ ∣₂ + +-- -- module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) where +-- -- module M = MV _ _ _ (λ _ → tt) (fst f) +-- -- ¬lem : (n m : ℕ) → ¬ suc (suc (m +ℕ n)) ≡ suc n +-- -- ¬lem zero zero p = snotz (cong predℕ p) +-- -- ¬lem (suc n) zero p = ¬lem n zero (cong predℕ p) +-- -- ¬lem zero (suc m) p = snotz (cong predℕ p) +-- -- ¬lem (suc n) (suc m) p = +-- -- ¬lem n (suc m) (cong (suc ∘ suc ) +-- -- (sym (+-suc m n)) ∙ (cong predℕ p)) + +-- -- GroupEquiv1 : +-- -- GroupEquiv +-- -- (coHomGr (3 +ℕ n +ℕ n) (fst (S₊∙ (3 +ℕ n +ℕ n)))) +-- -- ((coHomGr (suc (3 +ℕ n +ℕ n)) (Pushout (λ _ → tt) (fst f)))) +-- -- fst (fst GroupEquiv1) = (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) .fst +-- -- snd (fst GroupEquiv1) = +-- -- SES→Iso +-- -- (GroupPath _ _ .fst +-- -- (invGroupEquiv (isContr→≃Unit +-- -- (isOfHLevelΣ 0 +-- -- (isOfHLevelRetractFromIso 0 (fst (Hⁿ-Unit≅0 _)) isContrUnit) +-- -- λ _ → isOfHLevelRetractFromIso 0 +-- -- (fst (Hⁿ-Sᵐ≅0 _ _ (¬lem n n))) isContrUnit) +-- -- , makeIsGroupHom λ _ _ → refl))) +-- -- (GroupPath _ _ .fst +-- -- (invGroupEquiv +-- -- (isContr→≃Unit +-- -- (isContrΣ +-- -- (isOfHLevelRetractFromIso 0 (fst (Hⁿ-Unit≅0 _)) +-- -- isContrUnit) +-- -- λ _ → isOfHLevelRetractFromIso 0 +-- -- (fst (Hⁿ-Sᵐ≅0 _ _ (¬lem n (suc n)))) isContrUnit) +-- -- , makeIsGroupHom λ _ _ → refl))) +-- -- (M.Δ (3 +ℕ n +ℕ n)) +-- -- (M.d (3 +ℕ n +ℕ n)) +-- -- (M.i (4 +ℕ n +ℕ n)) +-- -- (M.Ker-d⊂Im-Δ _) +-- -- (M.Ker-i⊂Im-d _) +-- -- snd GroupEquiv1 = (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) .snd + +-- -- Hopfβ-Iso : GroupIso ((coHomGr (suc (3 +ℕ n +ℕ n)) (Pushout (λ _ → tt) (fst f)))) ℤGroup +-- -- Hopfβ-Iso = compGroupIso (GroupEquiv→GroupIso (invGroupEquiv GroupEquiv1)) (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) + +-- -- h33 : (n : ℕ) → Iso.inv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) 1 ≡ ∣ ∣_∣ ∣₂ +-- -- h33 zero = Iso.leftInv (fst (Hⁿ-Sⁿ≅ℤ (suc zero))) _ +-- -- h33 (suc n) = (λ i → Iso.inv +-- -- (fst +-- -- (suspensionAx-Sn (suc n) (suc n))) (h33 n i)) +-- -- ∙ cong ∣_∣₂ (funExt λ { north → refl +-- -- ; south → cong ∣_∣ₕ (merid north) +-- -- ; (merid a i) j → ∣ compPath-filler (merid a) (sym (merid north)) (~ j) i ∣}) + +-- -- d-β : (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) +-- -- → Iso.fun (fst (Hopfβ-Iso n f)) (Hopfβ n f) ≡ 1 +-- -- d-β n f = sym (cong (Iso.fun (fst (Hopfβ-Iso n f))) h) ∙ Iso.rightInv (fst (Hopfβ-Iso n f)) (pos 1) +-- -- where +-- -- h : Iso.inv (fst (Hopfβ-Iso n f)) (pos 1) ≡ Hopfβ n f +-- -- h = (λ i → fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) +-- -- (h33 _ i)) +-- -- ∙ cong ∣_∣₂ (funExt (λ { (inl x) → refl +-- -- ; (inr x) → refl +-- -- ; (push a i) → refl})) + +-- -- module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) where +-- -- 2+n = 2 +ℕ n +-- -- H = HopfInvariantPush n (fst f) + +-- -- ff : coHom 2+n H → coHom 2+n (S₊ (suc (suc n))) +-- -- ff = sMap (_∘ inr) + +-- -- grHom : IsGroupHom (snd (coHomGr 2+n H)) ff (snd (coHomGr 2+n (S₊ (suc (suc n))))) +-- -- grHom = +-- -- makeIsGroupHom +-- -- (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) +-- -- λ f g → refl) + +-- -- h' : (g : (S₊ (suc (suc n)) → coHomK 2+n)) → H → coHomK (2 +ℕ n) +-- -- h' g (inl x) = 0ₖ _ +-- -- h' g (inr x) = g x -ₖ g north +-- -- h' g (push a i) = h23 a i +-- -- where +-- -- h23 : (a : S₊ (suc (suc (suc (n +ℕ n))))) → 0ₖ (suc (suc n)) ≡ g (fst f a) -ₖ g north +-- -- h23 = sphereElim _ (λ x → isOfHLevelPlus' {n = n} (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n) _ _)) +-- -- (sym (rCancelₖ _ (g north)) ∙ cong (λ x → g x -ₖ g north) (sym (snd f))) + +-- -- ff⁻ : coHom 2+n (S₊ (suc (suc n))) → coHom 2+n H +-- -- ff⁻ = sMap h' + +-- -- h23 : (x : _) → ∥ x ≡ 0ₖ (suc (suc n)) ∥ +-- -- h23 = coHomK-elim _ (λ _ → isProp→isOfHLevelSuc (suc n) squash) ∣ refl ∣ + +-- -- h : Iso (coHom (2 +ℕ n) (HopfInvariantPush n (fst f))) (coHom (2 +ℕ n) ((S₊ (suc (suc n))))) +-- -- Iso.fun h = ff +-- -- Iso.inv h = ff⁻ +-- -- Iso.rightInv h = +-- -- sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) +-- -- λ g → pRec (squash₂ _ _) +-- -- (λ p → cong ∣_∣₂ (funExt λ x → cong (g x +ₖ_) (cong (-ₖ_) p) ∙ rUnitₖ _ (g x))) +-- -- (h23 (g north)) +-- -- Iso.leftInv h = +-- -- sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) +-- -- λ g → pRec (squash₂ _ _) +-- -- (pRec (isPropΠ (λ _ → squash₂ _ _)) +-- -- (λ gn gtt → +-- -- trRec (isProp→isOfHLevelSuc n (squash₂ _ _)) +-- -- (λ r → cong ∣_∣₂ (funExt λ { +-- -- (inl x) → sym gtt +-- -- ; (inr x) → (λ i → g (inr x) -ₖ gn i) ∙ rUnitₖ _ (g (inr x)) +-- -- ; (push a i) +-- -- → sphereElim _ +-- -- {A = λ a → PathP (λ i → h' (λ x → g (inr x)) (push a i) ≡ g (push a i)) +-- -- (sym gtt) +-- -- ((λ i → g (inr (fst f a)) -ₖ gn i) ∙ rUnitₖ _ (g (inr (fst f a))))} +-- -- (λ _ → isOfHLevelPathP' (suc (suc (suc (n +ℕ n)))) +-- -- (isOfHLevelPath (suc (suc (suc (suc (n +ℕ n))))) +-- -- (isOfHLevelPlus' {n = n} (suc (suc (suc (suc n)))) +-- -- (isOfHLevelTrunc (suc (suc (suc (suc n)))))) _ _) _ _) +-- -- r a i})) +-- -- (l g gtt gn)) +-- -- (h23 (g (inr north)))) +-- -- (h23 (g (inl tt))) +-- -- where +-- -- l : (g : HopfInvariantPush n (fst f) → coHomK (suc (suc n))) +-- -- → (gtt : g (inl tt) ≡ 0ₖ (suc (suc n))) +-- -- → (gn : g (inr north) ≡ 0ₖ (suc (suc n))) +-- -- → hLevelTrunc (suc n) (PathP (λ i → h' (λ x → g (inr x)) (push north i) ≡ g (push north i)) +-- -- (sym gtt) ((λ i → g (inr (fst f north)) -ₖ gn i) ∙ rUnitₖ _ (g (inr (fst f north))))) +-- -- l g gtt gn = isConnectedPathP _ (isConnectedPath _ (isConnectedKn _) _ _) _ _ .fst + + +-- -- Hopfα-Iso : GroupIso (coHomGr (2 +ℕ n) (HopfInvariantPush n (fst f))) ℤGroup +-- -- Hopfα-Iso = +-- -- compGroupIso +-- -- (h , grHom) +-- -- (Hⁿ-Sⁿ≅ℤ (suc n)) + +-- -- Hopfα-Iso-α : (n : ℕ) (f : _) +-- -- → Iso.fun (fst (Hopfα-Iso n f)) (Hopfα n f) +-- -- ≡ 1 +-- -- Hopfα-Iso-α n f = +-- -- sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (h33 n)) +-- -- ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) +-- -- where +-- -- hz : Iso.fun (h n f) (Hopfα n f) ≡ ∣ ∣_∣ ∣₂ +-- -- hz = refl + +-- -- _·₀ₕ_ : ∀ {ℓ} {n : ℕ} {A : Type ℓ} → ℤ → coHom n A → coHom n A +-- -- _·₀ₕ_ {n = n} a = sMap λ f x → a ·₀ f x + +-- -- HopfInvariant : (n : ℕ) +-- -- → (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → Type _ +-- -- HopfInvariant n f = +-- -- Σ[ x ∈ ℤ ] +-- -- Hopfα n f ⌣ Hopfα n f +-- -- ≡ (x ·₀ₕ subst (λ x → coHom x (HopfInvariantPush n (fst f))) +-- -- (λ i → suc (suc (suc (+-suc n n (~ i))))) +-- -- (Hopfβ n f)) + +-- -- pres- : (e : GroupHom ℤGroup ℤGroup) → (a : ℤ) → fst e (- a) ≡ - fst e a +-- -- pres- e a = +Comm (fst e (- a)) (pos 0) +-- -- ∙ cong (_+ (fst e (- a))) (sym (l (fst e a)) ∙ +Comm (fst e a) (- fst e a)) +-- -- ∙∙ sym (+Assoc _ _ _) +-- -- ∙∙ cong (- (fst e a) +_) (pp ∙ l (fst e a)) +-- -- where +-- -- l : (a : ℤ) → a + (- a) ≡ 0 +-- -- l (pos zero) = refl +-- -- l (pos (suc zero)) = refl +-- -- l (pos (suc (suc n))) = predℤ+negsuc n (pos (suc (suc n))) ∙ l (pos (suc n)) +-- -- l (negsuc zero) = refl +-- -- l (negsuc (suc n)) = (cong sucℤ (sucℤ+pos n (negsuc (suc n)))) ∙ l (negsuc n) + +-- -- pp : fst e a + fst e (- a) ≡ fst e a + (- fst e a) +-- -- pp = sym (IsGroupHom.pres· (snd e) a (- a)) +-- -- ∙∙ cong (fst e) (l a) +-- -- ∙∙ (IsGroupHom.pres1 (snd e) +-- -- ∙ sym (l _)) + + + +-- -- GroupHomPres· : (e : GroupHom ℤGroup ℤGroup) → (a b : ℤ) → fst e (a · b) ≡ a · fst e b +-- -- GroupHomPres· e (pos zero) b = IsGroupHom.pres1 (snd e) +-- -- GroupHomPres· e (pos (suc n)) b = +-- -- IsGroupHom.pres· (snd e) b (pos n · b) ∙ cong (fst e b +_) (GroupHomPres· e (pos n) b) +-- -- GroupHomPres· e (negsuc zero) b = pres- e b +-- -- GroupHomPres· e (negsuc (suc n)) b = +-- -- IsGroupHom.pres· (snd e) (- b) (negsuc n · b) +-- -- ∙ cong₂ _+_ (pres- e b) (GroupHomPres· e (negsuc n) b) + +-- -- lUnit·₀ₕ : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → (x : coHom n A) → (pos zero) ·₀ₕ x ≡ 0ₕ _ +-- -- lUnit·₀ₕ n = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl + +-- -- minus≡0- : (x : ℤ) → - x ≡ (0 - x) +-- -- minus≡0- x = sym (GroupStr.lid (snd ℤGroup) (- x)) + +-- -- GroupHomPres·₀ : ∀ {ℓ} {A : Type ℓ} (n : ℕ) +-- -- (e : GroupHom ℤGroup (coHomGr n A)) +-- -- → (a b : ℤ) +-- -- → fst e (a · b) +-- -- ≡ (a ·₀ₕ fst e b) +-- -- GroupHomPres·₀ n e (pos zero) b = IsGroupHom.pres1 (snd e) ∙ sym (lUnit·₀ₕ n (fst e b)) +-- -- GroupHomPres·₀ {A = A} n e (pos (suc a)) b = +-- -- IsGroupHom.pres· (snd e) b (pos a · b) +-- -- ∙ (λ i → fst e b +ₕ GroupHomPres·₀ n e (pos a) b i) +-- -- ∙ s (fst e b) +-- -- where +-- -- s : (x : coHom n A) → x +ₕ (pos a ·₀ₕ x) ≡ pos (suc a) ·₀ₕ x +-- -- s = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl +-- -- GroupHomPres·₀ n e (negsuc zero) b = +-- -- (cong (fst e) (sym (GroupStr.lid (snd ℤGroup) (- b))) +-- -- ∙ IsGroupHom.presinv (snd e) b) +-- -- GroupHomPres·₀ {A = A} n e (negsuc (suc a)) b = +-- -- IsGroupHom.pres· (snd e) (- b) (negsuc a · b) +-- -- ∙∙ cong₂ _+ₕ_ ((cong (fst e) (sym (GroupStr.lid (snd ℤGroup) (- b))) +-- -- ∙ IsGroupHom.presinv (snd e) b)) +-- -- (GroupHomPres·₀ n e (negsuc a) b) +-- -- ∙∙ helper (fst e b) +-- -- where +-- -- helper : (x : coHom n A) → (-ₕ x) +ₕ (negsuc a ·₀ₕ x) ≡ (negsuc (suc a) ·₀ₕ x) +-- -- helper = +-- -- sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) +-- -- λ f → cong ∣_∣₂ (funExt λ x → commₖ n (-ₖ (f x)) (negsuc a ·₀ f x)) + +-- -- genH : ∀ {ℓ} {A : Type ℓ} (n : ℕ) +-- -- → (e : GroupIso (coHomGr n A) +-- -- ℤGroup) +-- -- → Σ[ x ∈ coHom n A ] +-- -- ((y : coHom n A) +-- -- → Σ[ a ∈ ℤ ] a ·₀ₕ x ≡ y) +-- -- genH {A = A} n e = (Iso.inv (fst e) 1) +-- -- , λ y → (Iso.fun (fst e) y) +-- -- , (sym (GroupHomPres·₀ n (_ , snd (invGroupIso e)) (Iso.fun (fst e) y) (pos 1)) +-- -- ∙∙ cong (Iso.inv (fst e)) (·Comm (Iso.fun (fst e) y) (pos 1)) +-- -- ∙∙ Iso.leftInv (fst e) y) + +-- -- -- open import Cubical.Data.Sum +-- -- -- open import Cubical.Relation.Nullary +-- -- -- l2 : (n m : ℕ) → Σ[ a ∈ ℕ ] (negsuc n · pos (suc m)) ≡ negsuc a +-- -- -- l2 n zero = n , ·Comm (negsuc n) (pos 1) +-- -- -- l2 n (suc m) = h _ _ .fst , +-- -- -- (·Comm (negsuc n) (pos (suc (suc m))) +-- -- -- ∙∙ cong (negsuc n +_) (·Comm (pos (suc m)) (negsuc n) ∙ (l2 n m .snd)) +-- -- -- ∙∙ h _ _ .snd) +-- -- -- where +-- -- -- h : (x y : ℕ) → Σ[ a ∈ ℕ ] negsuc x + negsuc y ≡ negsuc a +-- -- -- h zero zero = 1 , refl +-- -- -- h zero (suc y) = (suc (suc y)) , +Comm (negsuc zero) (negsuc (suc y)) +-- -- -- h (suc x) zero = (suc (suc x)) , refl +-- -- -- h (suc x) (suc y) = (h (suc (suc x)) y .fst) , (predℤ+negsuc y (negsuc (suc x)) ∙ snd ((h (suc (suc x))) y)) + +-- -- -- zz : (n x : ℕ) → Σ[ a ∈ ℕ ] ((pos (suc x)) · pos (suc (suc n)) ≡ pos (suc (suc a))) +-- -- -- zz n zero = n , refl +-- -- -- zz n (suc x) = h _ _ (zz n x) +-- -- -- where +-- -- -- h : (x : ℤ) (n : ℕ) → Σ[ a ∈ ℕ ] (x ≡ pos (suc (suc a))) +-- -- -- → Σ[ a ∈ ℕ ] pos n + x ≡ pos (suc (suc a)) +-- -- -- h x n (a , p) = n +ℕ a +-- -- -- , cong (pos n +_) p ∙ cong sucℤ (sucℤ+pos a (pos n)) +-- -- -- ∙ sucℤ+pos a (pos (suc n)) ∙ (sym (pos+ (suc (suc n)) a)) + +-- -- -- lem22 : (n : ℕ) (x : ℤ) → ¬ pos 1 ≡ x · pos (suc (suc n)) +-- -- -- lem22 n (pos zero) p = snotz (injPos p) +-- -- -- lem22 n (pos (suc n₁)) p = snotz (injPos (sym (cong predℤ (snd (zz n n₁))) ∙ sym (cong predℤ p))) +-- -- -- lem22 n (negsuc n₁) p = posNotnegsuc _ _ (p ∙ l2 _ _ .snd) + + +-- -- -- grr : (e : GroupEquiv ℤGroup ℤGroup) (x : ℤ) → (fst (fst e) 1) ≡ x → abs (fst (fst e) 1) ≡ 1 +-- -- -- grr e (pos zero) p = +-- -- -- ⊥-rec (snotz (injPos (sym (retEq (fst e) 1) +-- -- -- ∙∙ (cong (fst (fst (invGroupEquiv e))) p) +-- -- -- ∙∙ IsGroupHom.pres1 (snd (invGroupEquiv e))))) +-- -- -- grr e (pos (suc zero)) p = cong abs p +-- -- -- grr e (pos (suc (suc n))) p = ⊥-rec (lem22 _ _ (h ∙ ·Comm (pos (suc (suc n))) (invEq (fst e) 1))) +-- -- -- where + +-- -- -- h : pos 1 ≡ _ +-- -- -- h = sym (retEq (fst e) 1) +-- -- -- ∙∙ (cong (fst (fst (invGroupEquiv e))) (p ∙ ·Comm 1 (pos (suc (suc n))))) +-- -- -- ∙∙ GroupHomPres· (_ , snd (invGroupEquiv e)) (pos (suc (suc n))) 1 +-- -- -- grr e (negsuc zero) p = cong abs p +-- -- -- grr e (negsuc (suc n)) p = ⊥-rec (lem22 _ _ l33) +-- -- -- where +-- -- -- h : fst (fst e) (negsuc zero) ≡ pos (suc (suc n)) +-- -- -- h = pres- (_ , snd e) (pos 1) ∙ cong -_ p + +-- -- -- compGroup : GroupEquiv ℤGroup ℤGroup +-- -- -- fst compGroup = isoToEquiv (iso -_ -_ -Involutive -Involutive) +-- -- -- snd compGroup = makeIsGroupHom -Dist+ + +-- -- -- compHom : GroupEquiv ℤGroup ℤGroup +-- -- -- compHom = compGroupEquiv compGroup e + +-- -- -- l32 : fst (fst compHom) (pos 1) ≡ pos (suc (suc n)) +-- -- -- l32 = h + +-- -- -- l33 : pos 1 ≡ invEq (fst compHom) (pos 1) · pos (suc (suc n)) +-- -- -- l33 = sym (retEq (fst compHom) (pos 1)) +-- -- -- ∙∙ cong (invEq (fst compHom)) l32 +-- -- -- ∙∙ (cong (invEq (fst compHom)) (·Comm (pos 1) (pos (suc (suc n)))) +-- -- -- ∙ GroupHomPres· (_ , (snd (invGroupEquiv compHom))) (pos (suc (suc n))) (pos 1) +-- -- -- ∙ ·Comm (pos (suc (suc n))) (invEq (fst compHom) (pos 1))) + +-- -- -- abs→⊎ : (x : ℤ) (n : ℕ) → abs x ≡ n → (x ≡ pos n) ⊎ (x ≡ - pos n) +-- -- -- abs→⊎ x n = J (λ n _ → (x ≡ pos n) ⊎ (x ≡ - pos n)) (help x) +-- -- -- where +-- -- -- help : (x : ℤ) → (x ≡ pos (abs x)) ⊎ (x ≡ - pos (abs x)) +-- -- -- help (pos n) = inl refl +-- -- -- help (negsuc n) = inr refl + +-- -- -- JGroupEquiv : ∀ {ℓ ℓ'} {G : Group ℓ} (P : (H : Group ℓ) → GroupEquiv G H → Type ℓ') +-- -- -- → P G idGroupEquiv +-- -- -- → ∀ {H} e → P H e +-- -- -- JGroupEquiv {G = G} P p {H} e = +-- -- -- transport (λ i → P (GroupPath G H .fst e i) +-- -- -- (transp (λ j → GroupEquiv G (GroupPath G H .fst e (i ∨ ~ j))) i e)) +-- -- -- (subst (P G) (sym l) p) +-- -- -- where +-- -- -- l : (transp (λ j → GroupEquiv G (GroupPath G H .fst e (~ j))) i0 e) ≡ idGroupEquiv +-- -- -- l = Σ≡Prop (λ _ → isPropIsGroupHom _ _) +-- -- -- (Σ≡Prop (λ _ → isPropIsEquiv _) +-- -- -- (funExt λ x → (λ i → fst (fst (fst e .snd .equiv-proof +-- -- -- (transportRefl (fst (fst e) (transportRefl x i)) i)))) +-- -- -- ∙ retEq (fst e) x)) + +-- -- -- sesIsoPresGen : +-- -- -- ∀ (G : Group ℓ-zero) +-- -- -- → (iso : GroupEquiv ℤGroup G) +-- -- -- → (H : Group ℓ-zero) +-- -- -- → (iso2 : GroupEquiv ℤGroup H) +-- -- -- → (e : fst G) +-- -- -- → invEq (fst iso) e ≡ 1 +-- -- -- → (hom : GroupEquiv G H) +-- -- -- → abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1 +-- -- -- sesIsoPresGen G = JGroupEquiv (λ G iso → +-- -- -- (H : Group ℓ-zero) +-- -- -- (iso2 : GroupEquiv ℤGroup H) +-- -- -- → (e : fst G) +-- -- -- → invEq (fst iso) e ≡ 1 +-- -- -- → (hom : GroupEquiv G H) +-- -- -- → abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1) +-- -- -- λ H → JGroupEquiv (λ H iso2 → (e : ℤ) → e ≡ pos 1 → +-- -- -- (hom : GroupEquiv ℤGroup H) → +-- -- -- abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1) +-- -- -- λ e p hom → cong (abs ∘ (fst (fst hom))) p ∙ autoPres1 hom +-- -- -- where +-- -- -- autoPres1 : (e : GroupEquiv ℤGroup ℤGroup) +-- -- -- → abs (fst (fst e) 1) ≡ 1 +-- -- -- autoPres1 e = grr e _ refl + +-- -- -- characFunSpaceS¹ : ∀ {ℓ} {A : Type ℓ} → +-- -- -- Iso (S₊ 1 → A) (Σ[ x ∈ A ] x ≡ x) +-- -- -- Iso.fun characFunSpaceS¹ f = f base , cong f loop +-- -- -- Iso.inv characFunSpaceS¹ (x , p) base = x +-- -- -- Iso.inv characFunSpaceS¹ (x , p) (loop i) = p i +-- -- -- Iso.rightInv characFunSpaceS¹ _ = refl +-- -- -- Iso.leftInv characFunSpaceS¹ f = funExt λ { base → refl ; (loop i) → refl} + +-- -- -- characFunSpace : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} +-- -- -- → Iso (join A B → C) +-- -- -- (Σ[ f ∈ (A → C) ] Σ[ g ∈ (B → C) ] +-- -- -- ((a : A) (b : B) → f a ≡ g b)) +-- -- -- Iso.fun characFunSpace f = (f ∘ inl) , ((f ∘ inr) , (λ a b → cong f (push a b))) +-- -- -- Iso.inv characFunSpace (f , g , p) (inl x) = f x +-- -- -- Iso.inv characFunSpace (f , g , p) (inr x) = g x +-- -- -- Iso.inv characFunSpace (f , g , p) (push a b i) = p a b i +-- -- -- Iso.rightInv characFunSpace (f , g , p) = refl +-- -- -- Iso.leftInv characFunSpace f = +-- -- -- funExt λ { (inl x) → refl ; (inr x) → refl ; (push a b i) → refl} + +-- -- -- coHomS¹-ish : (n : ℕ) → Type _ +-- -- -- coHomS¹-ish n = hLevelTrunc 3 (S₊ 1 → coHomK (3 +ℕ n)) + +-- -- -- fib : (n : ℕ) → coHomS¹-ish n → Type _ +-- -- -- fib n x = +-- -- -- trRec {B = TypeOfHLevel ℓ-zero 2} (isOfHLevelTypeOfHLevel 2) +-- -- -- (λ f → ∥ (Σ[ g ∈ (S₊ 3 → coHomK (3 +ℕ n)) ] +-- -- -- ((a : S₊ 1) (b : S₊ 3) → f a ≡ g b)) ∥₂ , squash₂) x .fst + +-- -- -- contrFstΣ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} +-- -- -- → (e : isContr A) +-- -- -- → Iso (Σ A B) (B (fst e)) +-- -- -- Iso.fun (contrFstΣ {B = B} e) (a , b) = subst B (sym (snd e a)) b +-- -- -- Iso.inv (contrFstΣ {B = B} e) b = (fst e) , b +-- -- -- Iso.rightInv (contrFstΣ {B = B} e) b = cong (λ x → subst B x b) h ∙ transportRefl b +-- -- -- where +-- -- -- h : sym (snd e (fst e)) ≡ refl +-- -- -- h = isProp→isSet (isContr→isProp e) _ _ _ _ +-- -- -- Iso.leftInv (contrFstΣ {B = B} e) b = +-- -- -- ΣPathP ((snd e (fst b)) +-- -- -- , (λ j → transp (λ i → B (snd e (fst b) (~ i ∨ j))) j (snd b))) + +-- -- -- Iso1 : (n : ℕ) → Iso (coHom (3 +ℕ n) (join S¹ (S₊ 3))) ∥ Σ (coHomS¹-ish n) (fib n) ∥₂ +-- -- -- Iso1 n = compIso (setTruncIso characFunSpace) Iso2 +-- -- -- where +-- -- -- Iso2 : Iso _ ∥ Σ (coHomS¹-ish n) (fib n) ∥₂ +-- -- -- Iso.fun Iso2 = sMap λ F → ∣ fst F ∣ , ∣ (fst (snd F)) , (snd (snd F)) ∣₂ +-- -- -- Iso.inv Iso2 = +-- -- -- sRec squash₂ +-- -- -- (uncurry (trElim (λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelSuc 2 squash₂) +-- -- -- λ f → sRec squash₂ λ p → ∣ f , p ∣₂)) +-- -- -- Iso.rightInv Iso2 = +-- -- -- sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) +-- -- -- (uncurry (trElim (λ _ → isOfHLevelΠ 3 λ _ → isProp→isOfHLevelSuc 2 (squash₂ _ _)) +-- -- -- λ f → sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl )) +-- -- -- Iso.leftInv Iso2 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl + +-- -- -- isContrcoHomS¹-ish : (n : ℕ) → isContr (coHomS¹-ish n) +-- -- -- isContrcoHomS¹-ish n = isOfHLevelRetractFromIso 0 (mapCompIso characFunSpaceS¹) isContrEnd +-- -- -- where +-- -- -- isContrEnd : isContr (hLevelTrunc 3 (Σ[ x ∈ coHomK (3 +ℕ n) ] x ≡ x)) +-- -- -- fst isContrEnd = ∣ 0ₖ _ , refl ∣ +-- -- -- snd isContrEnd = +-- -- -- trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) +-- -- -- (uncurry (coHomK-elim _ +-- -- -- (λ _ → isOfHLevelΠ (suc (suc (suc n))) +-- -- -- λ _ → isOfHLevelPlus' {n = n} 3 (isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _)) +-- -- -- (λ p → (trRec (isOfHLevelPlus' {n = n} 3 (isOfHLevelTrunc 3) _ _) +-- -- -- (λ p i → ∣ (0ₖ (3 +ℕ n) , p i) ∣) +-- -- -- (Iso.fun (PathIdTruncIso _) +-- -- -- (isContr→isProp +-- -- -- (isConnectedPath _ (isConnectedKn (2 +ℕ n)) (0ₖ _) (0ₖ _)) ∣ refl ∣ ∣ p ∣)))))) + +-- -- -- Iso2' : (n : ℕ) → Iso (∥ Σ (coHomS¹-ish n) (fib n) ∥₂) (fib n ∣ (λ _ → 0ₖ _) ∣) +-- -- -- Iso2' n = compIso (setTruncIso (contrFstΣ centerChange)) +-- -- -- (setTruncIdempotentIso squash₂) +-- -- -- where +-- -- -- centerChange : isContr (coHomS¹-ish n) +-- -- -- fst centerChange = ∣ (λ _ → 0ₖ _) ∣ +-- -- -- snd centerChange y = isContr→isProp (isContrcoHomS¹-ish n) _ _ + +-- -- -- open import Cubical.Foundations.Equiv.HalfAdjoint +-- -- -- open import Cubical.Relation.Nullary +-- -- -- ll5 : (n : ℕ) → ¬ (n ≡ 2) → isContr (fib n ∣ (λ _ → 0ₖ _) ∣) +-- -- -- ll5 n p = +-- -- -- isOfHLevelRetractFromIso 0 +-- -- -- (compIso +-- -- -- (setTruncIso +-- -- -- (compIso (Σ-cong-iso-snd λ f → +-- -- -- (compIso characFunSpaceS¹ (invIso (Σ-cong-iso-fst (iso funExt⁻ funExt (λ _ → refl) λ _ → refl))))) +-- -- -- (compIso ΣΣ (contrFstΣ (isContrSingl _))))) +-- -- -- (compIso (setTruncIso (iso funExt⁻ funExt (λ _ → refl) λ _ → refl)) +-- -- -- (compIso (setTruncIso (codomainIso (congIso (invIso (Iso-Kn-ΩKn+1 _))))) +-- -- -- ((compIso (invIso (fst (coHom≅coHomΩ _ (S₊ _)))) +-- -- -- (fst (Hⁿ-Sᵐ≅0 n 2 p))))))) +-- -- -- isContrUnit + +-- -- -- record ExactSeqℤ {ℓ : Level} (G : ℤ → Group ℓ) : Type ℓ where +-- -- -- field +-- -- -- maps : ∀ n → GroupHom (G n) (G (sucℤ n)) +-- -- -- im⊂ker : ∀ n → ∀ g → isInIm (maps n) g → isInKer (maps (sucℤ n)) g +-- -- -- ker⊂im : ∀ n → ∀ g → isInKer (maps (sucℤ n)) g → isInIm (maps n) g + +-- -- -- record ExactSeqℕ {ℓ : Level} (G : ℕ → Group ℓ) : Type ℓ where +-- -- -- field +-- -- -- maps : ∀ n → GroupHom (G n) (G (suc n)) +-- -- -- im⊂ker : ∀ n → ∀ g → isInIm (maps n) g → isInKer (maps (suc n)) g +-- -- -- ker⊂im : ∀ n → ∀ g → isInKer (maps (suc n)) g → isInIm (maps n) g + + +-- -- -- module _ {ℓ : Level} {A : Pointed ℓ} {n : ℕ} +-- -- -- where +-- -- -- funTyp : Type _ +-- -- -- funTyp = A →∙ coHomK-ptd n + +-- -- -- _++_ : funTyp → funTyp → funTyp +-- -- -- fst (f ++ g) x = fst f x +ₖ fst g x +-- -- -- snd (f ++ g) = cong₂ _+ₖ_ (snd f) (snd g) ∙ rUnitₖ _ (0ₖ _) + +-- -- -- addAgree : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (x y : funTyp {A = A} {n = n}) +-- -- -- → Path (fst (coHomRedGrDir n A)) +-- -- -- ∣ x ++ y ∣₂ +-- -- -- (∣ x ∣₂ +ₕ∙ ∣ y ∣₂) +-- -- -- addAgree {A = A} zero f g = +-- -- -- cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) refl) +-- -- -- addAgree {A = A} (suc zero) f g = +-- -- -- cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) refl) +-- -- -- addAgree {A = A} (suc (suc n)) f g = +-- -- -- cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) refl) + +-- -- -- ss : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Type ℓ''} +-- -- -- → isHomogeneous B +-- -- -- → (x y : (A →∙ B)) (f : C → x ≡ y) +-- -- -- → isEquiv (cong fst ∘ f) +-- -- -- → isEquiv f +-- -- -- ss homb x y f e = +-- -- -- isoToIsEquiv +-- -- -- (iso _ +-- -- -- (λ p → invEq (_ , e) (cong fst p)) +-- -- -- (λ p → →∙Homogeneous≡Path homb _ _ (secEq (_ , e) (cong fst p))) +-- -- -- (retEq (_ , e))) + +-- -- -- Whitehead : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} +-- -- -- → (f : A → B) +-- -- -- → isEmbedding f +-- -- -- → (∀ (b : B) → ∃[ a ∈ A ] f a ≡ b) +-- -- -- → isEquiv f +-- -- -- equiv-proof (Whitehead f emb p) b = +-- -- -- pRec isPropIsContr +-- -- -- (λ fib → fib , isEmbedding→hasPropFibers emb b fib) +-- -- -- (p b) + +-- -- -- Int-ind : ∀ {ℓ} (P : ℤ → Type ℓ) +-- -- -- → P (pos zero) → P (pos 1) +-- -- -- → P (negsuc zero) +-- -- -- → ((x y : ℤ) → P x → P y → P (x + y)) → (x : ℤ) → P x +-- -- -- Int-ind P z one min ind (pos zero) = z +-- -- -- Int-ind P z one min ind (pos (suc zero)) = one +-- -- -- Int-ind P z one min ind (pos (suc (suc n))) = +-- -- -- ind (pos (suc n)) 1 (Int-ind P z one min ind (pos (suc n))) one +-- -- -- Int-ind P z one min ind (negsuc zero) = min +-- -- -- Int-ind P z one min ind (negsuc (suc n)) = +-- -- -- ind (negsuc n) (negsuc zero) (Int-ind P z one min ind (negsuc n)) min + +-- -- -- genFunSpace : (n : ℕ) → S₊∙ n →∙ coHomK-ptd n +-- -- -- fst (genFunSpace zero) false = 1 +-- -- -- fst (genFunSpace zero) true = 0 +-- -- -- snd (genFunSpace zero) = refl +-- -- -- fst (genFunSpace (suc n)) = ∣_∣ +-- -- -- snd (genFunSpace (suc zero)) = refl +-- -- -- snd (genFunSpace (suc (suc n))) = refl + +-- -- -- module _ where +-- -- -- open import Cubical.Algebra.Monoid +-- -- -- open import Cubical.Algebra.Semigroup +-- -- -- open GroupStr +-- -- -- open IsGroup +-- -- -- open IsMonoid +-- -- -- open IsSemigroup +-- -- -- πS : (n : ℕ) → Group ℓ-zero +-- -- -- fst (πS n) = S₊∙ n →∙ coHomK-ptd n +-- -- -- 1g (snd (πS n)) = (λ _ → 0ₖ n) , refl +-- -- -- GroupStr._·_ (snd (πS n)) = +-- -- -- λ f g → (λ x → fst f x +ₖ fst g x) +-- -- -- , cong₂ _+ₖ_ (snd f) (snd g) ∙ rUnitₖ n (0ₖ n) +-- -- -- inv (snd (πS n)) f = (λ x → -ₖ fst f x) , cong -ₖ_ (snd f) ∙ -0ₖ {n = n} +-- -- -- is-set (isSemigroup (isMonoid (isGroup (snd (πS zero))))) = +-- -- -- isOfHLevelΣ 2 (isSetΠ (λ _ → isSetℤ)) +-- -- -- λ _ → isOfHLevelPath 2 isSetℤ _ _ +-- -- -- is-set (isSemigroup (isMonoid (isGroup (snd (πS (suc n)))))) = isOfHLevel↑∙' 0 n +-- -- -- IsSemigroup.assoc (isSemigroup (isMonoid (isGroup (snd (πS n))))) x y z = +-- -- -- →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ w → assocₖ n (fst x w) (fst y w) (fst z w)) +-- -- -- fst (identity (isMonoid (isGroup (snd (πS n)))) (f , p)) = +-- -- -- →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → rUnitₖ n (f x)) +-- -- -- snd (identity (isMonoid (isGroup (snd (πS n)))) (f , p)) = +-- -- -- →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → lUnitₖ n (f x)) +-- -- -- fst (inverse (isGroup (snd (πS n))) (f , p)) = +-- -- -- →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → rCancelₖ n (f x)) +-- -- -- snd (inverse (isGroup (snd (πS n))) (f , p)) = +-- -- -- →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → lCancelₖ n (f x)) + +-- -- -- isSetπS : (n : ℕ) → isSet (S₊∙ n →∙ coHomK-ptd n) +-- -- -- isSetπS = λ n → is-set (snd (πS n)) + +-- -- -- K : (n : ℕ) → GroupIso (coHomRedGrDir n (S₊∙ n)) (πS n) +-- -- -- fst (K n) = setTruncIdempotentIso (isSetπS n) +-- -- -- snd (K zero) = +-- -- -- makeIsGroupHom +-- -- -- (sElim2 (λ _ _ → isOfHLevelPath 2 (isSetπS 0) _ _) +-- -- -- λ f g → →∙Homogeneous≡ (isHomogeneousKn 0) refl) +-- -- -- snd (K (suc zero)) = +-- -- -- makeIsGroupHom +-- -- -- (sElim2 (λ _ _ → isOfHLevelPath 2 (isSetπS 1) _ _) +-- -- -- λ f g → →∙Homogeneous≡ (isHomogeneousKn 1) refl) +-- -- -- snd (K (suc (suc n))) = +-- -- -- makeIsGroupHom +-- -- -- (sElim2 (λ _ _ → isOfHLevelPath 2 (isSetπS _) _ _) +-- -- -- λ f g → →∙Homogeneous≡ (isHomogeneousKn _) refl) + +-- -- -- t : ∀ {ℓ} {A : Pointed ℓ} → Iso ((Bool , true) →∙ A) (typ A) +-- -- -- Iso.fun t f = fst f false +-- -- -- fst (Iso.inv t a) false = a +-- -- -- fst (Iso.inv (t {A = A}) a) true = pt A +-- -- -- snd (Iso.inv t a) = refl +-- -- -- Iso.rightInv t a = refl +-- -- -- Iso.leftInv t (f , p) = +-- -- -- ΣPathP ((funExt (λ { false → refl ; true → sym p})) , λ i j → p (~ i ∨ j)) + +-- -- -- S1' : (n : ℕ) → GroupIso (πS n) ℤGroup +-- -- -- fst (S1' zero) = t +-- -- -- snd (S1' zero) = makeIsGroupHom λ _ _ → refl +-- -- -- S1' (suc n) = +-- -- -- compGroupIso +-- -- -- (invGroupIso (K (suc n))) +-- -- -- (compGroupIso +-- -- -- (GroupEquiv→GroupIso (coHomGr≅coHomRedGr n (S₊∙ (suc n)))) +-- -- -- (Hⁿ-Sⁿ≅ℤ n)) + +-- -- -- S1 : (n : ℕ) → Iso (S₊∙ (suc n) →∙ coHomK-ptd (suc n)) ℤ +-- -- -- S1 n = compIso (invIso (setTruncIdempotentIso (isOfHLevel↑∙' 0 n))) +-- -- -- (compIso (equivToIso (fst (coHomGr≅coHomRedGr n (S₊∙ (suc n))))) +-- -- -- (fst (Hⁿ-Sⁿ≅ℤ n))) + +-- -- -- connFunSpace : (n i : ℕ) → (x y : S₊∙ n →∙ coHomK-ptd (suc i +' n)) → ∥ x ≡ y ∥ +-- -- -- connFunSpace n i f g = Iso.fun PathIdTrunc₀Iso (isContr→isProp (lem n) ∣ f ∣₂ ∣ g ∣₂) +-- -- -- where +-- -- -- open import Cubical.Relation.Nullary +-- -- -- natLem : (n i : ℕ) → ¬ (suc (i +ℕ n) ≡ n) +-- -- -- natLem zero i = snotz +-- -- -- natLem (suc n) i p = natLem n i (sym (+-suc i n) ∙ (cong predℕ p)) + +-- -- -- lem : (n : ℕ) → isContr ∥ (S₊∙ n →∙ coHomK-ptd (suc i +' n)) ∥₂ +-- -- -- fst (lem zero) = 0ₕ∙ (suc i) +-- -- -- snd (lem zero) = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) +-- -- -- λ f → pRec (squash₂ _ _) +-- -- -- (λ id → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn (suc i)) +-- -- -- (funExt λ { false → sym id ; true → sym (snd f)}))) +-- -- -- (help (f .fst false)) +-- -- -- where +-- -- -- help : (x : coHomK (suc i)) → ∥ x ≡ 0ₖ _ ∥ +-- -- -- help = coHomK-elim _ (λ _ → isProp→isOfHLevelSuc i squash) ∣ refl ∣ +-- -- -- lem (suc n) = +-- -- -- isOfHLevelRetractFromIso 0 +-- -- -- (compIso (equivToIso (fst (coHomGr≅coHomRedGr (suc (i +ℕ n)) (S₊∙ (suc n))))) +-- -- -- (fst (Hⁿ-Sᵐ≅0 (suc (i +ℕ n)) n (natLem n i)))) +-- -- -- isContrUnit + +-- -- -- S1Pres1 : (n : ℕ) → Iso.fun (fst (S1' (suc n))) (∣_∣ , refl) ≡ pos 1 +-- -- -- S1Pres1 zero = refl +-- -- -- S1Pres1 (suc n) = cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ n))) (lem n) ∙ S1Pres1 n +-- -- -- where +-- -- -- lem : (n : ℕ) → Iso.fun (fst (suspensionAx-Sn n n)) ∣ ∣_∣ ∣₂ ≡ ∣ ∣_∣ ∣₂ +-- -- -- lem zero = cong ∣_∣₂ (funExt λ x → transportRefl (∣ x ∣ +ₖ (0ₖ 1)) ∙ rUnitₖ 1 ∣ x ∣) +-- -- -- lem (suc n) = cong ∣_∣₂ +-- -- -- (funExt λ x → (λ i → transportRefl ((ΩKn+1→Kn (suc (suc n)) +-- -- -- (transp (λ j → 0ₖ (suc (suc (suc n))) ≡ ∣ merid north (~ j ∧ ~ i) ∣) i +-- -- -- (λ z → ∣ compPath-filler (merid (transportRefl (transportRefl x i) i)) (sym (merid north)) i z +-- -- -- ∣)))) i) +-- -- -- ∙ Iso.leftInv (Iso-Kn-ΩKn+1 (suc (suc n))) ∣ x ∣) + +-- -- -- S1Pres1← : (n : ℕ) → Iso.inv (fst (S1' (suc n))) (pos 1) ≡ (∣_∣ , refl) +-- -- -- S1Pres1← n = sym (cong (Iso.inv (fst (S1' (suc n)))) (S1Pres1 n)) ∙ Iso.leftInv (fst (S1' (suc n))) (∣_∣ , refl) + +-- -- -- IsoFunSpace : (n : ℕ) → Iso (S₊∙ n →∙ coHomK-ptd n) ℤ +-- -- -- IsoFunSpace n = fst (S1' n) + +-- -- -- module g-base where +-- -- -- g : (n : ℕ) (i : ℕ) → coHomK i → (S₊∙ n →∙ coHomK-ptd (i +' n)) +-- -- -- fst (g n i x) y = x ⌣ₖ (genFunSpace n) .fst y +-- -- -- snd (g n i x) = cong (x ⌣ₖ_) ((genFunSpace n) .snd) ∙ ⌣ₖ-0ₖ i n x + +-- -- -- G : (n : ℕ) (i : ℕ) → coHomK i → (S₊∙ n →∙ coHomK-ptd (n +' i)) +-- -- -- fst (G n i x) y = (genFunSpace n) .fst y ⌣ₖ x +-- -- -- snd (G n i x) = cong (_⌣ₖ x) ((genFunSpace n) .snd) ∙ 0ₖ-⌣ₖ n i x + +-- -- -- eq1 : (n : ℕ) (i : ℕ) → (S₊∙ n →∙ coHomK-ptd (i +' n)) ≃ (S₊∙ n →∙ coHomK-ptd (i +' n)) +-- -- -- eq1 n i = isoToEquiv (iso F F FF FF) +-- -- -- where +-- -- -- lem : (i n : ℕ) → (-ₖ^ i · n) (snd (coHomK-ptd (i +' n))) ≡ 0ₖ _ +-- -- -- lem zero zero = refl +-- -- -- lem zero (suc zero) = refl +-- -- -- lem zero (suc (suc n)) = refl +-- -- -- lem (suc zero) zero = refl +-- -- -- lem (suc (suc i)) zero = refl +-- -- -- lem (suc i) (suc n) = refl + +-- -- -- F : S₊∙ n →∙ coHomK-ptd (i +' n) → S₊∙ n →∙ coHomK-ptd (i +' n) +-- -- -- fst (F f) x = (-ₖ^ i · n) (fst f x) +-- -- -- snd (F f) = cong (-ₖ^ i · n) (snd f) ∙ lem i n + +-- -- -- FF : (x : _) → F (F x) ≡ x +-- -- -- FF x = +-- -- -- →∙Homogeneous≡ (isHomogeneousKn _) +-- -- -- (funExt λ y → -ₖ-gen² i n _ _ (fst x y)) + +-- -- -- rCancel'' : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) → sym p ∙∙ refl ∙∙ p ≡ refl +-- -- -- rCancel'' p = (λ j → (λ i → p (~ i ∨ j)) ∙∙ refl ∙∙ λ i → p (i ∨ j)) ∙ sym (rUnit refl) + +-- -- -- transpPres0 : ∀ {k m : ℕ} (p : k ≡ m) → subst coHomK p (0ₖ k) ≡ 0ₖ m +-- -- -- transpPres0 {k = k} = J (λ m p → subst coHomK p (0ₖ k) ≡ 0ₖ m) (transportRefl _) + +-- -- -- eq2 : (n : ℕ) (i : ℕ) → (S₊∙ n →∙ coHomK-ptd (n +' i)) ≃ (S₊∙ n →∙ coHomK-ptd (i +' n)) +-- -- -- eq2 n i = +-- -- -- isoToEquiv (iso (λ f → (λ x → subst coHomK (+'-comm n i) (fst f x)) , +-- -- -- cong (subst coHomK (+'-comm n i)) (snd f) ∙ transpPres0 (+'-comm n i)) +-- -- -- (λ f → (λ x → subst coHomK (sym (+'-comm n i)) (fst f x)) +-- -- -- , (cong (subst coHomK (sym (+'-comm n i))) (snd f) ∙ transpPres0 (sym (+'-comm n i)))) +-- -- -- (λ f → →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ x → transportTransport⁻ _ (f .fst x))) +-- -- -- λ f → →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ x → transportTransport⁻ _ (f .fst x))) + +-- -- -- g≡ : (n : ℕ) (i : ℕ) → g n i ≡ λ x → fst (compEquiv (eq2 n i) (eq1 n i)) ((G n i) x) +-- -- -- g≡ n i = funExt (λ f → →∙Homogeneous≡ (isHomogeneousKn _) +-- -- -- (funExt λ y → gradedComm-⌣ₖ _ _ f (genFunSpace n .fst y))) + +-- -- -- glIso3-fun : (n m : ℕ) → +-- -- -- (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) +-- -- -- → (S₊ n → coHomK m) +-- -- -- glIso3-fun zero m (f , p) false = ΩKn+1→Kn _ (sym p ∙∙ cong f loop ∙∙ p) +-- -- -- glIso3-fun zero m (f , p) true = 0ₖ _ +-- -- -- glIso3-fun (suc n) m (f , p) x = +-- -- -- ΩKn+1→Kn _ (sym p ∙∙ cong f (merid x ∙ sym (merid (ptSn _))) ∙∙ p) + +-- -- -- glIso3-fun∙ : (n m : ℕ) → (f : _) → glIso3-fun n m f (ptSn _) ≡ 0ₖ m +-- -- -- glIso3-fun∙ zero m = λ _ → refl +-- -- -- glIso3-fun∙ (suc n) m (f , p) = +-- -- -- cong (ΩKn+1→Kn m) (cong (sym p ∙∙_∙∙ p) (cong (cong f) (rCancel (merid (ptSn _))))) +-- -- -- ∙∙ cong (ΩKn+1→Kn m) (rCancel'' p) +-- -- -- ∙∙ ΩKn+1→Kn-refl m + + +-- -- -- glIso3-inv : (n m : ℕ) → (S₊∙ n →∙ coHomK-ptd m) → (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) +-- -- -- fst (glIso3-inv zero m (f , p)) base = 0ₖ _ +-- -- -- fst (glIso3-inv zero m (f , p)) (loop i) = Kn→ΩKn+1 m (f false) i +-- -- -- snd (glIso3-inv zero m (f , p)) = refl +-- -- -- fst (glIso3-inv (suc n) m (f , p)) north = 0ₖ _ +-- -- -- fst (glIso3-inv (suc n) m (f , p)) south = 0ₖ _ +-- -- -- fst (glIso3-inv (suc n) m (f , p)) (merid a i) = Kn→ΩKn+1 m (f a) i +-- -- -- snd (glIso3-inv (suc n) m (f , p)) = refl + +-- -- -- glIso3 : (n m : ℕ) → +-- -- -- Iso (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) +-- -- -- (S₊∙ n →∙ coHomK-ptd m) +-- -- -- Iso.fun (glIso3 n m) f = (glIso3-fun n m f) , (glIso3-fun∙ n m f) +-- -- -- Iso.inv (glIso3 n m) = glIso3-inv n m +-- -- -- Iso.rightInv (glIso3 zero m) (f , p) = +-- -- -- →∙Homogeneous≡ (isHomogeneousKn _) +-- -- -- (funExt λ { false → cong (ΩKn+1→Kn m) (sym (rUnit _)) ∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f false) +-- -- -- ; true → sym p}) +-- -- -- Iso.rightInv (glIso3 (suc n) m) (f , p) = +-- -- -- →∙Homogeneous≡ (isHomogeneousKn _) +-- -- -- (funExt λ x → (λ i → ΩKn+1→Kn m (sym (rUnit (cong-∙ (glIso3-inv (suc n) m (f , p) .fst) (merid x) (sym (merid (ptSn _))) i)) i)) +-- -- -- ∙∙ cong (ΩKn+1→Kn m) (cong (Kn→ΩKn+1 m (f x) ∙_) (cong sym (cong (Kn→ΩKn+1 m) p ∙ Kn→ΩKn+10ₖ m)) ∙ sym (rUnit _)) +-- -- -- ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f x)) +-- -- -- Iso.leftInv (glIso3 zero m) (f , p) = →∙Homogeneous≡ (isHomogeneousKn _) +-- -- -- (funExt λ { base → sym p +-- -- -- ; (loop i) j → h j i}) +-- -- -- where +-- -- -- h : PathP (λ i → p (~ i) ≡ p (~ i)) (Kn→ΩKn+1 m (ΩKn+1→Kn m (sym p ∙∙ cong f loop ∙∙ p))) (cong f loop) +-- -- -- h = Iso.rightInv (Iso-Kn-ΩKn+1 _) _ +-- -- -- ◁ λ i → doubleCompPath-filler (sym p) (cong f loop) p (~ i) +-- -- -- Iso.leftInv (glIso3 (suc n) m) (f , p) = +-- -- -- →∙Homogeneous≡ (isHomogeneousKn _) +-- -- -- (funExt λ { north → sym p +-- -- -- ; south → sym p ∙ cong f (merid (ptSn _)) +-- -- -- ; (merid a i) j → h a j i}) +-- -- -- where +-- -- -- h : (a : S₊ (suc n)) → PathP (λ i → p (~ i) ≡ (sym p ∙ cong f (merid (ptSn (suc n)))) i) +-- -- -- (Kn→ΩKn+1 m (ΩKn+1→Kn m (sym p ∙∙ cong f (merid a ∙ sym (merid (ptSn _))) ∙∙ p))) +-- -- -- (cong f (merid a)) +-- -- -- h a = Iso.rightInv (Iso-Kn-ΩKn+1 _) _ +-- -- -- ◁ λ i j → hcomp (λ k → λ { (i = i1) → (f (merid a j)) +-- -- -- ; (j = i0) → p (k ∧ ~ i) +-- -- -- ; (j = i1) → compPath-filler' (sym p) (cong f (merid (ptSn _))) k i }) +-- -- -- (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) + +-- -- -- glIsoInvHom : (n m : ℕ) (x y : coHomK n) (z : S₊ (suc m)) +-- -- -- → Iso.inv (glIso3 _ _) (G m n (x +ₖ y)) .fst z +-- -- -- ≡ Iso.inv (glIso3 _ _) (G m n x) .fst z +-- -- -- +ₖ Iso.inv (glIso3 _ _) (G m n y) .fst z +-- -- -- glIsoInvHom zero zero x y base = refl +-- -- -- glIsoInvHom (suc n) zero x y base = refl +-- -- -- glIsoInvHom zero zero x y (loop i) j = h j i +-- -- -- where +-- -- -- h : (cong (Iso.inv (glIso3 _ _) (G zero zero (x + y)) .fst) loop) +-- -- -- ≡ cong₂ _+ₖ_ (cong (Iso.inv (glIso3 _ _) (G zero zero x) .fst) loop) +-- -- -- (cong (Iso.inv (glIso3 _ _) (G zero zero y) .fst) loop) +-- -- -- h = Kn→ΩKn+1-hom 0 x y +-- -- -- ∙ ∙≡+₁ (cong (Iso.inv (glIso3 _ _) (G zero zero x) .fst) loop) +-- -- -- (cong (Iso.inv (glIso3 _ _) (G zero zero y) .fst) loop) +-- -- -- glIsoInvHom (suc n) zero x y (loop i) j = h j i +-- -- -- where +-- -- -- h : Kn→ΩKn+1 (suc n) ((pos (suc zero)) ·₀ (x +ₖ y)) +-- -- -- ≡ cong₂ _+ₖ_ (cong (Iso.inv (glIso3 zero (zero +' suc n)) (G zero (suc n) x) .fst) loop) +-- -- -- (cong (Iso.inv (glIso3 zero (zero +' suc n)) (G zero (suc n) y) .fst) loop) +-- -- -- h = cong (Kn→ΩKn+1 (suc n)) (lUnit⌣ₖ _ (x +ₖ y)) +-- -- -- ∙∙ Kn→ΩKn+1-hom (suc n) x y +-- -- -- ∙∙ (λ i → ∙≡+₂ n (Kn→ΩKn+1 (suc n) (lUnit⌣ₖ _ x (~ i))) +-- -- -- (Kn→ΩKn+1 (suc n) (lUnit⌣ₖ _ y (~ i))) i) +-- -- -- glIsoInvHom zero (suc m) x y north = refl +-- -- -- glIsoInvHom zero (suc m) x y south = refl +-- -- -- glIsoInvHom zero (suc m) x y (merid a i) j = h j i +-- -- -- where +-- -- -- h : Kn→ΩKn+1 (suc m) (_⌣ₖ_ {n = suc m} {m = zero} ∣ a ∣ (x + y)) +-- -- -- ≡ cong₂ _+ₖ_ (Kn→ΩKn+1 (suc m) (_⌣ₖ_ {n = suc m} {m = zero} ∣ a ∣ x)) +-- -- -- (Kn→ΩKn+1 (suc m) (_⌣ₖ_ {n = suc m} {m = zero} ∣ a ∣ y)) +-- -- -- h = cong (Kn→ΩKn+1 (suc m)) (leftDistr-⌣ₖ (suc m) 0 ∣ a ∣ x y) +-- -- -- ∙∙ Kn→ΩKn+1-hom (suc m) _ _ +-- -- -- ∙∙ ∙≡+₂ _ _ _ +-- -- -- glIsoInvHom (suc n) (suc m) x y north = refl +-- -- -- glIsoInvHom (suc n) (suc m) x y south = refl +-- -- -- glIsoInvHom (suc n) (suc m) x y (merid a i) j = h j i +-- -- -- where +-- -- -- h : Kn→ΩKn+1 (suc (suc (m +ℕ n))) (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ (x +ₖ y)) +-- -- -- ≡ cong₂ _+ₖ_ (Kn→ΩKn+1 (suc (suc (m +ℕ n))) (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ x)) +-- -- -- (Kn→ΩKn+1 (suc (suc (m +ℕ n))) (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ y)) +-- -- -- h = cong (Kn→ΩKn+1 (suc (suc (m +ℕ n)))) +-- -- -- (leftDistr-⌣ₖ (suc m) (suc n) ∣ a ∣ x y) +-- -- -- ∙∙ Kn→ΩKn+1-hom _ _ _ +-- -- -- ∙∙ ∙≡+₂ _ _ _ + +-- -- -- +'-suc : (n m : ℕ) → suc n +' m ≡ suc (n +' m) +-- -- -- +'-suc zero zero = refl +-- -- -- +'-suc (suc n) zero = refl +-- -- -- +'-suc zero (suc m) = refl +-- -- -- +'-suc (suc n) (suc m) = refl + +-- -- -- LEM : (i n : ℕ) (x : coHomK i) +-- -- -- → Path _ (G (suc n) i x) +-- -- -- (subst (λ x → S₊∙ (suc n) →∙ coHomK-ptd x) +-- -- -- (sym (+'-suc n i)) +-- -- -- ((Iso.inv (glIso3 n _)) (G n i x))) +-- -- -- LEM zero zero x = +-- -- -- →∙Homogeneous≡ (isHomogeneousKn _) +-- -- -- (funExt λ z → (λ i → x ·₀ ∣ z ∣) ∙ h x z ∙ sym (transportRefl _)) +-- -- -- where +-- -- -- h : (x : ℤ) (z : S¹) → _·₀_ {n = 1} x ∣ z ∣ ≡ fst (Iso.inv (glIso3 0 zero) (G zero zero x)) z +-- -- -- h = Int-ind _ (λ { base → refl ; (loop i) j → Kn→ΩKn+10ₖ zero (~ j) i}) +-- -- -- (λ { base → refl ; (loop i) j → rUnit (cong ∣_∣ₕ (lUnit loop j)) j i}) +-- -- -- (λ { base → refl ; (loop i) j → rUnit (cong ∣_∣ₕ (sym loop)) j i}) +-- -- -- λ x y inx iny z +-- -- -- → rightDistr-⌣ₖ 0 1 x y ∣ z ∣ +-- -- -- ∙∙ cong₂ _+ₖ_ (inx z) (iny z) +-- -- -- ∙∙ sym (glIsoInvHom zero zero x y z) +-- -- -- LEM (suc i) zero x = +-- -- -- →∙Homogeneous≡ (isHomogeneousKn _) +-- -- -- (funExt λ z → h z +-- -- -- ∙ sym (transportRefl +-- -- -- ((Iso.inv (glIso3 zero (suc i)) (G zero (suc i) x)) .fst z))) +-- -- -- where +-- -- -- h : (z : S₊ 1) → _ ≡ Iso.inv (glIso3 zero (suc i)) (G zero (suc i) x) .fst z +-- -- -- h base = refl +-- -- -- h (loop k) j = Kn→ΩKn+1 (suc i) (lUnit⌣ₖ _ x (~ j)) k +-- -- -- LEM zero (suc n) x = +-- -- -- →∙Homogeneous≡ (isHomogeneousKn _) +-- -- -- (funExt λ z → h x z ∙ λ i → transportRefl (Iso.inv (glIso3 (suc n) (suc n)) (G (suc n) zero x)) (~ i) .fst z) +-- -- -- where +-- -- -- +merid : rUnitₖ (suc (suc n)) ∣ south ∣ ≡ cong ∣_∣ₕ (merid (ptSn _)) +-- -- -- +merid = sym (lUnit _) +-- -- -- ∙ cong (cong ∣_∣ₕ) +-- -- -- λ j i → transp (λ _ → S₊ (suc (suc n))) (i ∨ j) (merid (ptSn (suc n)) i) +-- -- -- open import Cubical.Foundations.Path + +-- -- -- pp : (a : S₊ (suc n)) → PathP (λ i → ∣ merid a i ∣ₕ ≡ Kn→ΩKn+1 (suc n) (∣ a ∣ +ₖ (0ₖ _)) i) +-- -- -- refl (sym (rUnitₖ (suc (suc n)) ∣ south ∣)) +-- -- -- pp a = flipSquare ((λ i j → ∣ compPath-filler (merid a) (sym (merid (ptSn _))) i j ∣ₕ ) +-- -- -- ▷ cong (Kn→ΩKn+1 (suc n)) (sym (rUnitₖ (suc n) ∣ a ∣ₕ))) +-- -- -- ▷ sym (cong sym +merid) + +-- -- -- pp2 : (a : S₊ (suc n)) → (λ i → -ₖ ∣ merid a i ∣) +-- -- -- ≡ Kn→ΩKn+1 (suc n) (-ₖ ∣ a ∣) +-- -- -- pp2 a = cong (cong ∣_∣ₕ) (sym (symDistr (merid a) (sym (merid (ptSn (suc n)))))) +-- -- -- ∙ sym (Kn→ΩKn+1-ₖ (suc n) ∣ a ∣) + +-- -- -- h : (x : ℤ) (z : S₊ (suc (suc n))) +-- -- -- → _⌣ₖ_ {n = suc (suc n)} {m = 0} ∣ z ∣ x +-- -- -- ≡ Iso.inv (glIso3 (suc n) (suc n)) (G (suc n) zero x) .fst z +-- -- -- h = Int-ind _ +-- -- -- (λ { north → refl ; south → refl ; (merid a i) j → Kn→ΩKn+10ₖ (suc n) (~ j) i}) +-- -- -- (λ { north → refl ; south → refl +-- -- -- ; (merid a i) j → hcomp (λ k → λ { (i = i0) → ∣ north ∣ +-- -- -- ; (i = i1) → rUnitₖ (suc (suc n)) ∣ south ∣ (~ k) +-- -- -- ; (j = i0) → rUnitₖ (suc (suc n)) ∣ merid a i ∣ (~ k) +-- -- -- ; (j = i1) → pp a i k}) +-- -- -- ∣ merid a i ∣ₕ}) +-- -- -- (λ { north → refl +-- -- -- ; south → refl +-- -- -- ; (merid a i) j → pp2 a j i}) +-- -- -- λ x y indx indy z → leftDistr-⌣ₖ _ _ ∣ z ∣ x y +-- -- -- ∙ cong₂ _+ₖ_ (indx z) (indy z) +-- -- -- ∙ sym (glIsoInvHom _ _ _ _ _) +-- -- -- LEM (suc i) (suc n) x = +-- -- -- →∙Homogeneous≡ (isHomogeneousKn _) +-- -- -- (funExt λ z → h z +-- -- -- ∙ λ j → transportRefl ((Iso.inv (glIso3 (suc n) (suc (suc (n +ℕ i)))) +-- -- -- (G (suc n) (suc i) x))) (~ j) .fst z) +-- -- -- where +-- -- -- h : (z : S₊ (suc (suc n))) → _ +-- -- -- h north = refl +-- -- -- h south = refl +-- -- -- h (merid a i) = refl + +-- -- -- isEquivGzero : (i : ℕ) → isEquiv (G zero i) +-- -- -- isEquivGzero i = +-- -- -- isoToIsEquiv +-- -- -- (iso _ (λ f → fst f false) +-- -- -- (λ {(f , p) → →∙Homogeneous≡ (isHomogeneousKn _) +-- -- -- (funExt λ { false → rUnitₖ _ (f false) ; true → sym p})}) +-- -- -- (lUnit⌣ₖ _)) + +-- -- -- isEquivG : (n i : ℕ) → isEquiv (G n i) +-- -- -- isEquivG zero i = +-- -- -- isoToIsEquiv +-- -- -- (iso _ (λ f → fst f false) +-- -- -- (λ {(f , p) → →∙Homogeneous≡ (isHomogeneousKn _) +-- -- -- (funExt λ { false → rUnitₖ _ (f false) ; true → sym p})}) +-- -- -- (lUnit⌣ₖ _)) +-- -- -- isEquivG (suc n) i = +-- -- -- subst isEquiv (sym (funExt (LEM i n))) +-- -- -- (compEquiv (compEquiv (G n i , isEquivG n i) +-- -- -- (isoToEquiv (invIso (glIso3 n (n +' i))))) +-- -- -- (pathToEquiv (λ j → S₊∙ (suc n) →∙ coHomK-ptd (+'-suc n i (~ j)))) .snd) + + +-- -- -- isEquiv-g : (n i : ℕ) → isEquiv (g n i) +-- -- -- isEquiv-g n i = +-- -- -- subst isEquiv (sym (g≡ n i)) +-- -- -- (compEquiv (G n i , isEquivG n i) (compEquiv (eq2 n i) (eq1 n i)) .snd) + +-- -- -- module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) +-- -- -- (conB : (x y : typ B) → ∥ x ≡ y ∥) +-- -- -- (n : ℕ) (Q-is : Iso (typ (Q (pt B))) (S₊ n)) +-- -- -- (Q-is-ptd : Iso.fun Q-is (pt (Q (pt B))) ≡ snd (S₊∙ n)) +-- -- -- (c : (b : typ B) → (Q b →∙ coHomK-ptd n)) +-- -- -- (c-pt : c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) where + +-- -- -- g : (b : typ B) (i : ℕ) → coHomK i → (Q b →∙ coHomK-ptd (i +' n)) +-- -- -- fst (g b i x) y = x ⌣ₖ c b .fst y +-- -- -- snd (g b i x) = cong (x ⌣ₖ_) (c b .snd) ∙ ⌣ₖ-0ₖ i n x + +-- -- -- g' : (b : typ B) (i : ℕ) → coHomK i → coHomK i → (Q b →∙ coHomK-ptd (i +' n)) +-- -- -- fst (g' b i x y) z = x ⌣ₖ c b .fst z +ₖ y ⌣ₖ c b .fst z +-- -- -- snd (g' b i x y) = cong₂ _+ₖ_ (cong (x ⌣ₖ_) (c b .snd) ∙ ⌣ₖ-0ₖ i n x) +-- -- -- (cong (y ⌣ₖ_) (c b .snd) ∙ ⌣ₖ-0ₖ i n y) ∙ rUnitₖ _ (0ₖ _) + +-- -- -- g-hom : (b : typ B) (i : ℕ) → (x y : coHomK i) → g b i (x +ₖ y) ≡ ((g b i x) ++ (g b i y)) +-- -- -- g-hom b i x y = →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ z → rightDistr-⌣ₖ i n x y (c b .fst z)) + +-- -- -- gPathP' : (i : ℕ) → PathP (λ j → coHomK i → (isoToPath Q-is j , ua-gluePath (isoToEquiv Q-is) (Q-is-ptd) j) +-- -- -- →∙ coHomK-ptd (i +' n)) (g (pt B) i) (g-base.g n i) +-- -- -- gPathP' i = +-- -- -- toPathP +-- -- -- (funExt +-- -- -- λ x → →∙Homogeneous≡ (isHomogeneousKn _) +-- -- -- (funExt λ y → (λ i → transportRefl (transportRefl x i ⌣ₖ c (pt B) .fst (Iso.inv Q-is (transportRefl y i))) i) +-- -- -- ∙ cong (x ⌣ₖ_) (funExt⁻ c-pt (Iso.inv Q-is y) ∙ cong (genFunSpace n .fst) (Iso.rightInv Q-is y)))) + +-- -- -- g-base : (i : ℕ) → isEquiv (g (pt B) i) +-- -- -- g-base i = transport (λ j → isEquiv (gPathP' i (~ j))) (g-base.isEquiv-g n i) + +-- -- -- g-equiv : (b : typ B) (i : ℕ) → isEquiv (g b i) +-- -- -- g-equiv b i = +-- -- -- pRec (isPropIsEquiv _) +-- -- -- (J (λ b _ → isEquiv (g b i)) +-- -- -- (g-base i)) +-- -- -- (conB (pt B) b) + +-- -- -- module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) +-- -- -- (conB : (x y : typ B) → ∥ x ≡ y ∥₂) +-- -- -- (n : ℕ) (Q-is : Iso (typ (Q (pt B))) (S₊ n)) +-- -- -- (Q-is-ptd : Iso.fun Q-is (pt (Q (pt B))) ≡ snd (S₊∙ n)) where + +-- -- -- is-setQ→K : (b : typ B) → isSet (Q b →∙ coHomK-ptd n) +-- -- -- is-setQ→K b = +-- -- -- sRec (isProp→isOfHLevelSuc 1 isPropIsSet) +-- -- -- (J (λ b _ → isSet (Q b →∙ coHomK-ptd n)) +-- -- -- (subst isSet (cong (_→∙ coHomK-ptd n) +-- -- -- (ua∙ (isoToEquiv (invIso Q-is)) (cong (Iso.inv Q-is) (sym Q-is-ptd) ∙ Iso.leftInv Q-is _))) +-- -- -- (isOfHLevelRetractFromIso 2 (fst (S1' n)) isSetℤ))) +-- -- -- (conB (pt B) b) + + +-- -- -- isConnB : isConnected 3 (typ B) +-- -- -- fst isConnB = ∣ pt B ∣ +-- -- -- snd isConnB = +-- -- -- trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) +-- -- -- λ a → sRec (isOfHLevelTrunc 3 _ _) (cong ∣_∣ₕ) (conB (pt B) a) + +-- -- -- isPropPath : isProp (∥ pt B ≡ pt B ∥₂) +-- -- -- isPropPath = +-- -- -- isOfHLevelRetractFromIso 1 setTruncTrunc2Iso +-- -- -- (isContr→isProp (isConnectedPath _ isConnB (pt B) (pt B))) + +-- -- -- c* : Σ[ c ∈ ((b : typ B) → (Q b →∙ coHomK-ptd n)) ] +-- -- -- (c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) +-- -- -- fst c* b = +-- -- -- sRec (is-setQ→K b) +-- -- -- (J (λ b _ → Q b →∙ coHomK-ptd n) +-- -- -- ((λ x → genFunSpace n .fst (Iso.fun Q-is x)) +-- -- -- , cong (genFunSpace n .fst) Q-is-ptd ∙ genFunSpace n .snd)) (conB (pt B) b) +-- -- -- snd c* = +-- -- -- funExt λ x → (λ i → sRec (is-setQ→K (pt B)) +-- -- -- (J (λ b _ → Q b →∙ coHomK-ptd n) +-- -- -- ((λ x₁ → genFunSpace n .fst (Iso.fun Q-is x₁)) , +-- -- -- (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) +-- -- -- (isPropPath (conB (pt B) (pt B)) ∣ refl ∣₂ i) .fst x) +-- -- -- ∙ (λ i → transportRefl (genFunSpace n .fst (Iso.fun Q-is (transportRefl x i))) i) + +-- -- -- p-help : {b : fst B} (p : pt B ≡ b) → (subst (fst ∘ Q) (sym p) (snd (Q b))) ≡ (snd (Q (pt B))) +-- -- -- p-help {b = b} = +-- -- -- J (λ b p → subst (fst ∘ Q) (sym p) (snd (Q b)) ≡ snd (Q (pt B))) (transportRefl _) + +-- -- -- cEquiv : (b : fst B) (p : ∥ pt B ≡ b ∥₂) +-- -- -- → (c* .fst b) +-- -- -- ≡ sRec (is-setQ→K b) +-- -- -- (λ pp → (λ qb → genFunSpace n .fst (Iso.fun Q-is (subst (fst ∘ Q) (sym pp) qb))) +-- -- -- , cong (genFunSpace n .fst ∘ Iso.fun Q-is) (p-help pp) +-- -- -- ∙ ((λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) p +-- -- -- cEquiv b = +-- -- -- sElim (λ _ → isOfHLevelPath 2 (is-setQ→K b) _ _) +-- -- -- (J (λ b a → c* .fst b ≡ +-- -- -- sRec (is-setQ→K b) (λ pp → +-- -- -- (λ qb → +-- -- -- genFunSpace n .fst (Iso.fun Q-is (subst (fst ∘ Q) (sym pp) qb))) +-- -- -- , +-- -- -- cong (genFunSpace n .fst ∘ Iso.fun Q-is) (p-help pp) ∙ +-- -- -- (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd) +-- -- -- ∣ a ∣₂) +-- -- -- ((λ i → sRec (is-setQ→K (pt B)) +-- -- -- (J (λ b₁ _ → Q b₁ →∙ coHomK-ptd n) +-- -- -- ((λ x → genFunSpace n .fst (Iso.fun Q-is x)) , +-- -- -- (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) +-- -- -- (isPropPath (conB (pt B) (pt B)) ∣ refl ∣₂ i)) +-- -- -- ∙ →∙Homogeneous≡ (isHomogeneousKn n) +-- -- -- (transportRefl ((λ x → genFunSpace n .fst (Iso.fun Q-is x))) +-- -- -- ∙ funExt λ x → cong (genFunSpace n .fst ∘ Iso.fun Q-is) (sym (transportRefl x))))) + +-- -- -- module _ {ℓ ℓ'} (B : Pointed ℓ) (P : typ B → Type ℓ') where +-- -- -- E : Type _ +-- -- -- E = Σ (typ B) P + +-- -- -- Ẽ : Type _ +-- -- -- Ẽ = Pushout {A = E} (λ _ → tt) fst + +-- -- -- i : (n : ℕ) → (P-base : Iso (P (pt B)) (S₊ n)) → S₊ (suc n) → Ẽ +-- -- -- i zero P-base base = inr (pt B) +-- -- -- i zero P-base (loop j) = (sym (push (pt B , Iso.inv P-base false)) +-- -- -- ∙ push ((pt B) , Iso.inv P-base true)) j +-- -- -- i (suc n) P-base north = inl tt +-- -- -- i (suc n) P-base south = inr (pt B) +-- -- -- i (suc n) P-base (merid a i₁) = push (pt B , Iso.inv P-base a) i₁ + +-- -- -- Q : typ B → Pointed ℓ' +-- -- -- Q x = Susp (P x) , north + +-- -- -- F : Type _ +-- -- -- F = Σ (typ B) λ x → Q x .fst + +-- -- -- F̃ : Type _ +-- -- -- F̃ = Pushout {A = typ B} {C = F} (λ _ → tt) λ b → b , north + +-- -- -- invFE : Ẽ → F̃ +-- -- -- invFE (inl x) = inl tt +-- -- -- invFE (inr x) = inr (x , south) +-- -- -- invFE (push (x , a) i₁) = ((push x) ∙ λ i → inr (x , merid a i)) i₁ + +-- -- -- funFE : F̃ → Ẽ +-- -- -- funFE (inl x) = inl tt +-- -- -- funFE (inr (x , north)) = inl tt +-- -- -- funFE (inr (x , south)) = inr x +-- -- -- funFE (inr (x , merid a i₁)) = push (x , a) i₁ +-- -- -- funFE (push a i₁) = inl tt + +-- -- -- IsoFE : Iso F̃ Ẽ +-- -- -- Iso.fun IsoFE = funFE +-- -- -- Iso.inv IsoFE = invFE +-- -- -- Iso.rightInv IsoFE (inl x) = refl +-- -- -- Iso.rightInv IsoFE (inr x) = refl +-- -- -- Iso.rightInv IsoFE (push (x , a) i₁) k = h k i₁ +-- -- -- where +-- -- -- h : cong funFE (((push x) ∙ λ i → inr (x , merid a i))) +-- -- -- ≡ push (x , a) +-- -- -- h = congFunct funFE (push x) (λ i → inr (x , merid a i)) +-- -- -- ∙ sym (lUnit (push (x , a))) +-- -- -- Iso.leftInv IsoFE (inl x) = refl +-- -- -- Iso.leftInv IsoFE (inr (x , north)) = push x +-- -- -- Iso.leftInv IsoFE (inr (x , south)) = refl +-- -- -- Iso.leftInv IsoFE (inr (x , merid a i)) j = +-- -- -- compPath-filler' (push x) (λ i₁ → inr (x , merid a i₁)) (~ j) i +-- -- -- Iso.leftInv IsoFE (push a i₁) k = push a (i₁ ∧ k) + + +-- -- -- funDir1 : ∀ {ℓ} {A : Pointed ℓ} → (F̃ , inl tt) →∙ A → (b : typ B) → Q b →∙ A +-- -- -- fst (funDir1 {A = A , a} (f , p) b) north = f (inr (b , north)) +-- -- -- fst (funDir1 {A = A , a} (f , p) b) south = f (inr (b , south)) +-- -- -- fst (funDir1 {A = A , a} (f , p) b) (merid a₁ i₁) = f (inr (b , merid a₁ i₁)) +-- -- -- snd (funDir1 {A = A , a} (f , p) b) = sym (cong f (push b)) ∙ p + +-- -- -- funDir2 : ∀ {ℓ} {A : Pointed ℓ} → ((b : typ B) → Q b →∙ A) → (F̃ , inl tt) →∙ A +-- -- -- fst (funDir2 {A = A , a} f) (inl x) = a +-- -- -- fst (funDir2 {A = A , a} f) (inr (x , north)) = f x .fst north +-- -- -- fst (funDir2 {A = A , a} f) (inr (x , south)) = f x .fst south +-- -- -- fst (funDir2 {A = A , a} f) (inr (x , merid a₁ i₁)) = f x .fst (merid a₁ i₁) +-- -- -- fst (funDir2 {A = A , a} f) (push a₁ i₁) = snd (f a₁) (~ i₁) +-- -- -- snd (funDir2 {A = A , a} f) = refl + +-- -- -- funDir2-hom : (n : ℕ) → (f g : ((b : typ B) → Q b →∙ coHomK-ptd n)) +-- -- -- → funDir2 (λ b → f b ++ g b) ≡ (funDir2 f ++ funDir2 g) +-- -- -- funDir2-hom n f g = +-- -- -- →∙Homogeneous≡ (isHomogeneousKn _) +-- -- -- (funExt λ { (inl x) → sym (rUnitₖ _ (0ₖ _)) +-- -- -- ; (inr (x , north)) → refl +-- -- -- ; (inr (x , south)) → refl +-- -- -- ; (inr (x , merid a i₁)) → refl +-- -- -- ; (push a j) i → compPath-filler (cong₂ _+ₖ_ (snd (f a)) (snd (g a))) +-- -- -- (rUnitₖ n (0ₖ n)) (~ i) (~ j)}) + +-- -- -- funDirSect : ∀ {ℓ} {A : Pointed ℓ} → (x : (b : typ B) → Q b →∙ A) (b : typ B) (q : Q b .fst) +-- -- -- → funDir1 (funDir2 x) b .fst q ≡ x b .fst q +-- -- -- funDirSect f b north = refl +-- -- -- funDirSect f b south = refl +-- -- -- funDirSect f b (merid a i₁) = refl + +-- -- -- funDirRetr : ∀ {ℓ} {A : Pointed ℓ} (f : F̃ → typ A) (p : _) +-- -- -- → (x : F̃) → fst (funDir2 {A = A} (funDir1 (f , p))) x ≡ f x +-- -- -- funDirRetr f p (inl x) = sym p +-- -- -- funDirRetr f p (inr (x , north)) = refl +-- -- -- funDirRetr f p (inr (x , south)) = refl +-- -- -- funDirRetr f p (inr (x , merid a i₁)) = refl +-- -- -- funDirRetr f p (push a i₁) j = compPath-filler (sym (cong f (push a))) p (~ j) (~ i₁) + +-- -- -- Iso2 : ∀ {ℓ} {A : Pointed ℓ} +-- -- -- → Iso ((F̃ , inl tt) →∙ A) +-- -- -- ((b : typ B) → Q b →∙ A) +-- -- -- Iso.fun (Iso2 {A = A , a}) = funDir1 +-- -- -- Iso.inv (Iso2 {A = A , a}) = funDir2 +-- -- -- Iso.rightInv (Iso2 {A = A , a}) f = +-- -- -- funExt λ b → ΣPathP (funExt (funDirSect f b) +-- -- -- , sym (rUnit (snd (f b)))) +-- -- -- Iso.leftInv (Iso2 {A = A , a}) (f , p) = +-- -- -- ΣPathP ((funExt (funDirRetr f p)) +-- -- -- , λ i j → p (~ i ∨ j)) + +-- -- -- ι : (k : ℕ) → Iso ((b : typ B) → Q b →∙ coHomK-ptd k) +-- -- -- ((Ẽ , inl tt) →∙ coHomK-ptd k) +-- -- -- ι k = compIso (invIso Iso2) h +-- -- -- where +-- -- -- h : Iso ((F̃ , inl tt) →∙ coHomK-ptd k) +-- -- -- ((Ẽ , inl tt) →∙ coHomK-ptd k) +-- -- -- Iso.fun h G = (λ x → G .fst (Iso.inv IsoFE x)) +-- -- -- , (snd G) +-- -- -- Iso.inv h G = (λ x → G .fst (Iso.fun IsoFE x)) +-- -- -- , (snd G) +-- -- -- Iso.rightInv h G = +-- -- -- →∙Homogeneous≡ (isHomogeneousKn _) +-- -- -- (funExt λ x → cong (G .fst) (Iso.rightInv IsoFE x)) +-- -- -- Iso.leftInv h G = +-- -- -- →∙Homogeneous≡ (isHomogeneousKn _) +-- -- -- (funExt λ x → cong (G .fst) (Iso.leftInv IsoFE x)) + +-- -- -- ι-hom : (k : ℕ) → (f g : ((b : typ B) → Q b →∙ coHomK-ptd k)) +-- -- -- → Iso.fun (ι k) (λ b → f b ++ g b) +-- -- -- ≡ (Iso.fun (ι k) f ++ Iso.fun (ι k) g) +-- -- -- ι-hom k f g = +-- -- -- →∙Homogeneous≡ (isHomogeneousKn _) +-- -- -- (funExt λ x → funExt⁻ (cong fst (funDir2-hom _ f g)) (invFE x)) + +-- -- -- codomainIsoDep : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : A → Type ℓ'} {C : A → Type ℓ''} +-- -- -- → ((a : A) → Iso (B a) (C a)) +-- -- -- → Iso ((a : A) → B a) ((a : A) → C a) +-- -- -- Iso.fun (codomainIsoDep is) f a = Iso.fun (is a) (f a) +-- -- -- Iso.inv (codomainIsoDep is) f a = Iso.inv (is a) (f a) +-- -- -- Iso.rightInv (codomainIsoDep is) f = funExt λ a → Iso.rightInv (is a) (f a) +-- -- -- Iso.leftInv (codomainIsoDep is) f = funExt λ a → Iso.leftInv (is a) (f a) + +-- -- -- module Thom {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) +-- -- -- (conB : (x y : typ B) → ∥ x ≡ y ∥) +-- -- -- (n : ℕ) (Q-is : Iso (typ (Q B P (pt B))) (S₊ n)) +-- -- -- (Q-is-ptd : Iso.fun Q-is (pt (Q B P (pt B))) ≡ snd (S₊∙ n)) +-- -- -- (c : (b : typ B) → (Q B P b →∙ coHomK-ptd n)) +-- -- -- (c-pt : c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) where + +-- -- -- ϕ : (i : ℕ) → GroupEquiv (coHomGr i (typ B)) +-- -- -- (coHomRedGrDir (i +' n) (Ẽ B P , inl tt)) +-- -- -- fst (ϕ i) = +-- -- -- isoToEquiv +-- -- -- (setTruncIso +-- -- -- (compIso +-- -- -- (codomainIsoDep +-- -- -- λ b → equivToIso ((g B (Q B P) conB n Q-is Q-is-ptd c c-pt b i) +-- -- -- , g-equiv B (Q B P) conB n Q-is Q-is-ptd c c-pt b i)) +-- -- -- (ι B P (i +' n)))) +-- -- -- snd (ϕ i) = +-- -- -- makeIsGroupHom +-- -- -- (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) +-- -- -- λ F G → cong ∣_∣₂ (cong (Iso.fun (ι B P (i +' n))) +-- -- -- (funExt (λ a → g-hom B (Q B P) conB n Q-is Q-is-ptd c c-pt a i (F a) (G a))) +-- -- -- ∙ ι-hom B P (i +' n) _ _) +-- -- -- ∙ addAgree (i +' n) _ _) + +-- -- -- module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) +-- -- -- (conB : (x y : typ B) → ∥ x ≡ y ∥₂) +-- -- -- (n : ℕ) (Q-is : Iso (typ (Q B P (pt B))) (S₊ n)) +-- -- -- (Q-is-ptd : Iso.fun Q-is (pt (Q B P (pt B))) ≡ snd (S₊∙ n)) where + +-- -- -- 0-connB : (x y : typ B) → ∥ x ≡ y ∥ +-- -- -- 0-connB x y = sRec (isProp→isSet squash) (∥_∥.∣_∣) (conB x y) + +-- -- -- c = (c* B (Q B P) conB n Q-is Q-is-ptd .fst) +-- -- -- c-ptd = (c* B (Q B P) conB n Q-is Q-is-ptd .snd) +-- -- -- cEq = cEquiv B (Q B P) conB n Q-is Q-is-ptd + +-- -- -- module w = Thom B P 0-connB n Q-is Q-is-ptd c c-ptd + +-- -- -- ϕ = w.ϕ + +-- -- -- E' = E B P + +-- -- -- E'̃ = Ẽ B P + +-- -- -- p : E' → typ B +-- -- -- p = fst + +-- -- -- e : coHom n (typ B) +-- -- -- e = ∣ (λ b → c b .fst south) ∣₂ + +-- -- -- ⌣-hom : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr (i +' n) (typ B)) +-- -- -- fst (⌣-hom i) x = x ⌣ e +-- -- -- snd (⌣-hom i) = +-- -- -- makeIsGroupHom λ f g → rightDistr-⌣ _ _ f g e + +-- -- -- p-hom : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr i E') +-- -- -- p-hom i = coHomMorph _ p + +-- -- -- E-susp : (i : ℕ) → GroupHom (coHomGr i E') (coHomRedGrDir (suc i) (E'̃ , inl tt)) +-- -- -- fst (E-susp i) = sMap λ f → (λ { (inl x) → 0ₖ _ +-- -- -- ; (inr x) → 0ₖ _ +-- -- -- ; (push a j) → Kn→ΩKn+1 _ (f a) j}) , refl +-- -- -- snd (E-susp zero) = +-- -- -- makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) +-- -- -- λ f g → +-- -- -- cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn 1) +-- -- -- (funExt λ { (inl x) → refl +-- -- -- ; (inr x) → refl +-- -- -- ; (push a j) k → (Kn→ΩKn+1-hom zero (f a) (g a) +-- -- -- ∙ ∙≡+₁ (Kn→ΩKn+1 _ (f a)) (Kn→ΩKn+1 _ (g a))) k j}))) +-- -- -- snd (E-susp (suc i)) = +-- -- -- makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) +-- -- -- λ f g → +-- -- -- cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) +-- -- -- (funExt λ { (inl x) → refl +-- -- -- ; (inr x) → refl +-- -- -- ; (push a j) k → (Kn→ΩKn+1-hom _ (f a) (g a) +-- -- -- ∙ ∙≡+₂ _ (Kn→ΩKn+1 _ (f a)) (Kn→ΩKn+1 _ (g a))) k j}))) + +-- -- -- module cofibSeq where +-- -- -- j* : (i : ℕ) → GroupHom (coHomRedGrDir i (E'̃ , (inl tt))) (coHomGr i (typ B)) +-- -- -- fst (j* i) = sMap λ f → λ x → fst f (inr x) +-- -- -- snd (j* zero) = +-- -- -- makeIsGroupHom +-- -- -- (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) +-- -- -- snd (j* (suc zero)) = +-- -- -- makeIsGroupHom +-- -- -- (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) +-- -- -- snd (j* (suc (suc i₁))) = +-- -- -- makeIsGroupHom +-- -- -- (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) + +-- -- -- Im-j⊂Ker-p : (i : ℕ) (x : _) → isInIm (j* i) x → isInKer (p-hom i) x +-- -- -- Im-j⊂Ker-p i = +-- -- -- sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) +-- -- -- λ f → pRec (squash₂ _ _) +-- -- -- (uncurry (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) +-- -- -- λ g P → subst (isInKer (p-hom i)) P +-- -- -- (cong ∣_∣₂ (funExt λ x → cong (g .fst) (sym (push x)) ∙ g .snd)))) + +-- -- -- Ker-p⊂Im-j : (i : ℕ) (x : _) → isInKer (p-hom i) x → isInIm (j* i) x +-- -- -- Ker-p⊂Im-j i = +-- -- -- sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) +-- -- -- λ f ker → pRec squash +-- -- -- (λ id → ∣ ∣ (λ { (inl x) → 0ₖ _ +-- -- -- ; (inr x) → f x +-- -- -- ; (push a i₁) → funExt⁻ id a (~ i₁)}) , refl ∣₂ , refl ∣) +-- -- -- (Iso.fun PathIdTrunc₀Iso ker) + +-- -- -- Im-p⊂Ker-Susp : (i : ℕ) (x : _) → isInIm (p-hom i) x → isInKer (E-susp i) x +-- -- -- Im-p⊂Ker-Susp i = +-- -- -- sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) +-- -- -- λ f → pRec (squash₂ _ _) +-- -- -- (uncurry (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) +-- -- -- λ g y → subst (isInKer (E-susp i)) y +-- -- -- (cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) +-- -- -- (funExt λ { (inl x) → refl +-- -- -- ; (inr x) → sym (Kn→ΩKn+1 _ (g x)) +-- -- -- ; (push a j) k → Kn→ΩKn+1 i (g (fst a)) (~ k ∧ j)}))))) +-- -- -- open import Cubical.Foundations.Path +-- -- -- Ker-Susp⊂Im-p : (i : ℕ) (x : _) → isInKer (E-susp i) x → isInIm (p-hom i) x +-- -- -- Ker-Susp⊂Im-p i = +-- -- -- sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) +-- -- -- λ f ker → pRec squash +-- -- -- (λ id → ∣ ∣ (λ x → ΩKn+1→Kn i (sym (funExt⁻ (cong fst id) (inr x)))) ∣₂ +-- -- -- , cong ∣_∣₂ (funExt (λ { (a , b) → +-- -- -- cong (ΩKn+1→Kn i) (lUnit _ ∙ cong (_∙ sym (funExt⁻ (λ i₁ → fst (id i₁)) (inr a))) (sym (flipSquare (cong snd id)) +-- -- -- ∙∙ (PathP→compPathR (cong (funExt⁻ (cong fst id)) (push (a , b)))) +-- -- -- ∙∙ assoc _ _ _ +-- -- -- ∙ sym (rUnit _)) +-- -- -- ∙ (sym (assoc _ _ _) +-- -- -- ∙∙ cong (Kn→ΩKn+1 i (f (a , b)) ∙_) (rCancel _) +-- -- -- ∙∙ sym (rUnit _))) +-- -- -- ∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f (a , b))})) ∣) +-- -- -- (Iso.fun PathIdTrunc₀Iso ker) + +-- -- -- Im-Susp⊂Ker-j : (i : ℕ) (x : _) → isInIm (E-susp i) x → isInKer (cofibSeq.j* (suc i)) x +-- -- -- Im-Susp⊂Ker-j i = +-- -- -- sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) +-- -- -- λ g → pRec (squash₂ _ _) +-- -- -- (uncurry (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) +-- -- -- λ f id → pRec (squash₂ _ _) +-- -- -- (λ P → subst (isInKer (cofibSeq.j* (suc i))) (cong ∣_∣₂ P) +-- -- -- (cong ∣_∣₂ refl)) +-- -- -- (Iso.fun PathIdTrunc₀Iso id))) + +-- -- -- Ker-j⊂Im-Susp : (i : ℕ) (x : _) → isInKer (cofibSeq.j* (suc i)) x → isInIm (E-susp i) x +-- -- -- Ker-j⊂Im-Susp i = +-- -- -- sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) +-- -- -- λ f ker +-- -- -- → pRec squash +-- -- -- (λ p → ∣ ∣ (λ x → ΩKn+1→Kn _ (sym (snd f) ∙∙ cong (fst f) (push x) ∙∙ funExt⁻ p (fst x))) ∣₂ +-- -- -- , cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) +-- -- -- (funExt (λ { (inl x) → sym (snd f) +-- -- -- ; (inr x) → sym (funExt⁻ p x) +-- -- -- ; (push a j) k → h f p a k j}))) ∣) +-- -- -- (Iso.fun PathIdTrunc₀Iso ker) +-- -- -- where +-- -- -- h : (f : (E'̃ , inl tt) →∙ coHomK-ptd (suc i)) +-- -- -- → (p : (fst f ∘ inr) ≡ (λ _ → 0ₖ (suc i))) +-- -- -- → (a : E B P) +-- -- -- → PathP (λ i → snd f (~ i) ≡ p (~ i) (fst a)) +-- -- -- (Kn→ΩKn+1 i (ΩKn+1→Kn i (sym (snd f) ∙∙ cong (fst f) (push a) ∙∙ funExt⁻ p (fst a)))) +-- -- -- (cong (fst f) (push a)) +-- -- -- h f p a = Iso.rightInv (Iso-Kn-ΩKn+1 i) _ +-- -- -- ◁ λ i j → doubleCompPath-filler (sym (snd f)) (cong (fst f) (push a)) (funExt⁻ p (fst a)) (~ i) j + +-- -- -- ϕ∘j : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr (i +' n) (typ B)) +-- -- -- ϕ∘j i = compGroupHom (fst (fst (ϕ i)) , snd (ϕ i)) (cofibSeq.j* (i +' n)) + +-- -- -- +'-suc : (i n : ℕ) → (suc i +' n) ≡ suc (i +' n) +-- -- -- +'-suc zero zero = refl +-- -- -- +'-suc (suc i₁) zero = refl +-- -- -- +'-suc zero (suc n) = refl +-- -- -- +'-suc (suc i₁) (suc n) = refl + +-- -- -- private +-- -- -- h : ∀ {ℓ ℓ'} {n m : ℕ} (G : ℕ → Group ℓ) (H : Group ℓ') (p : n ≡ m) +-- -- -- → GroupEquiv (G n) H +-- -- -- → GroupEquiv (G m) H +-- -- -- h {n = n} G H = +-- -- -- J (λ m _ → GroupEquiv (G n) H → GroupEquiv (G m) H) +-- -- -- λ p → p + +-- -- -- h-ret : ∀ {ℓ ℓ'} {n m : ℕ} (G : ℕ → Group ℓ) (H : Group ℓ') +-- -- -- → (e : GroupEquiv (G n) H) +-- -- -- → (p : n ≡ m) +-- -- -- → (x : G m .fst) → invEq (fst e) (fst (fst (h G H p e)) x) ≡ subst (λ x → G x .fst) (sym p) x +-- -- -- h-ret G H e = +-- -- -- J (λ m p → ((x : G m .fst) → invEq (fst e) (fst (fst (h G H p e)) x) ≡ subst (λ x → G x .fst) (sym p) x)) +-- -- -- λ x → cong (invEq (fst e) ) +-- -- -- (λ i → transportRefl (transportRefl (fst (fst e) (transportRefl (transportRefl x i) i)) i) i) +-- -- -- ∙∙ retEq (fst e) x +-- -- -- ∙∙ sym (transportRefl _) + +-- -- -- isEquivϕ' : (i : ℕ) → GroupEquiv (coHomRedGrDir (suc (i +' n)) (E'̃ , inl tt)) +-- -- -- (coHomGr (suc i) (typ B)) +-- -- -- isEquivϕ' i = (h (λ n → coHomRedGrDir n (E'̃ , inl tt)) _ (+'-suc i n) +-- -- -- (invGroupEquiv (ϕ (suc i)))) + +-- -- -- ϕ' : (i : ℕ) → GroupHom (coHomRedGrDir (suc (i +' n)) (E'̃ , inl tt)) +-- -- -- (coHomGr (suc i) (typ B)) +-- -- -- ϕ' i = fst (fst (isEquivϕ' i)) , snd (isEquivϕ' i) + +-- -- -- susp∘ϕ : (i : ℕ) → GroupHom (coHomGr (i +' n) E') (coHomGr (suc i) (typ B)) +-- -- -- susp∘ϕ i = compGroupHom (E-susp (i +' n)) (ϕ' i) + +-- -- -- Im-ϕ∘j⊂Ker-p : (i : ℕ) (x : _) → isInIm (ϕ∘j i) x → isInKer (p-hom _) x +-- -- -- Im-ϕ∘j⊂Ker-p i x p = +-- -- -- cofibSeq.Im-j⊂Ker-p _ x +-- -- -- (pRec squash (uncurry (λ f p → ∣ fst (fst (ϕ _)) f , p ∣)) p) + +-- -- -- Ker-p⊂Im-ϕ∘j : (i : ℕ) (x : _) → isInKer (p-hom _) x → isInIm (ϕ∘j i) x +-- -- -- Ker-p⊂Im-ϕ∘j i x p = +-- -- -- pRec squash (uncurry (λ f p → ∣ (invEq (fst (ϕ _)) f) +-- -- -- , (cong (fst (cofibSeq.j* _)) (secEq (fst (ϕ _)) f) ∙ p) ∣)) +-- -- -- (cofibSeq.Ker-p⊂Im-j _ x p) + + +-- -- -- Im-p⊂KerSusp∘ϕ : (i : ℕ) (x : _) → isInIm (p-hom _) x → isInKer (susp∘ϕ i) x +-- -- -- Im-p⊂KerSusp∘ϕ i x p = cong (fst (ϕ' _)) (Im-p⊂Ker-Susp _ x p) ∙ IsGroupHom.pres1 (snd (ϕ' _)) + +-- -- -- KerSusp∘ϕ⊂Im-p : (i : ℕ) (x : _) → isInKer (susp∘ϕ i) x → isInIm (p-hom _) x +-- -- -- KerSusp∘ϕ⊂Im-p i x p = +-- -- -- Ker-Susp⊂Im-p _ x (sym (retEq (fst (isEquivϕ' _)) _) +-- -- -- ∙ (cong (invEq (fst (isEquivϕ' _))) p +-- -- -- ∙ IsGroupHom.pres1 (snd (invGroupEquiv (isEquivϕ' _))))) + +-- -- -- Im-Susp∘ϕ⊂Ker-ϕ∘j : (i : ℕ) → (x : _) → isInIm (susp∘ϕ i) x → isInKer (ϕ∘j (suc i)) x +-- -- -- Im-Susp∘ϕ⊂Ker-ϕ∘j i x = +-- -- -- pRec (squash₂ _ _) +-- -- -- (uncurry λ f → J (λ x p → isInKer (ϕ∘j (suc i)) x) +-- -- -- ((λ i → fst (cofibSeq.j* _) (fst (fst (ϕ _)) (fst (ϕ' _) (fst (E-susp _) f)))) +-- -- -- ∙∙ cong (fst (cofibSeq.j* _)) +-- -- -- ((h-ret (λ n → coHomRedGrDir n (E'̃ , inl tt)) _ +-- -- -- (invGroupEquiv (ϕ (suc i))) (+'-suc i n)) (fst (E-susp _) f)) +-- -- -- ∙∙ (natTranspLem _ (λ n → fst (cofibSeq.j* n)) (sym (+'-suc i n)) +-- -- -- ∙ cong (subst (λ z → coHomGr z (typ B) .fst) (sym (+'-suc i n))) +-- -- -- (Im-Susp⊂Ker-j _ (fst (E-susp (i +' n)) f) ∣ f , refl ∣) +-- -- -- ∙ tLem i n))) +-- -- -- where +-- -- -- tLem : (i n : ℕ) → subst (λ z → coHomGr z (typ B) .fst) (sym (+'-suc i n)) +-- -- -- (0ₕ _) ≡ 0ₕ _ +-- -- -- tLem zero zero = refl +-- -- -- tLem zero (suc n) = refl +-- -- -- tLem (suc i₁) zero = refl +-- -- -- tLem (suc i₁) (suc n) = refl + + +-- -- -- Ker-ϕ∘j⊂Im-Susp∘ϕ : (i : ℕ) (x : _) +-- -- -- → isInKer (ϕ∘j (suc i)) x → isInIm (susp∘ϕ i) x +-- -- -- Ker-ϕ∘j⊂Im-Susp∘ϕ i x p = +-- -- -- pRec squash +-- -- -- (uncurry (λ f p → ∣ f , cong (fst (fst (isEquivϕ' i))) p ∙ secEq (fst (isEquivϕ' _)) x ∣)) +-- -- -- (Ker-j⊂Im-Susp _ (invEq (fst (isEquivϕ' _)) x) +-- -- -- ((cong (cofibSeq.j* (suc (i +' n)) .fst ) lem2 +-- -- -- ∙ natTranspLem _ (λ n → cofibSeq.j* n .fst) (+'-suc i n)) +-- -- -- ∙∙ cong (transport (λ j → (coHomGr (+'-suc i n j) (typ B) .fst))) p +-- -- -- ∙∙ h2 i n)) +-- -- -- where +-- -- -- lem2 : invEq (fst (isEquivϕ' i)) x ≡ transport (λ j → coHomRed (+'-suc i n j) (E'̃ , inl tt)) (fst (fst (ϕ _)) x) +-- -- -- lem2 = cong (transport (λ j → coHomRed (+'-suc i n j) (E'̃ , inl tt))) +-- -- -- (transportRefl _ ∙ cong (fst (fst (ϕ _))) +-- -- -- λ i → transportRefl (transportRefl x i) i) + +-- -- -- h2 : (i n : ℕ) → transport (λ j → coHomGr (+'-suc i n j) (typ B) .fst) +-- -- -- (GroupStr.1g (coHomGr (suc i +' n) (typ B) .snd)) ≡ 0ₕ _ +-- -- -- h2 zero zero = refl +-- -- -- h2 zero (suc n) = refl +-- -- -- h2 (suc i₁) zero = refl +-- -- -- h2 (suc i₁) (suc n) = refl + + +-- -- -- ϕ∘j≡ : (i : ℕ) → ϕ∘j i ≡ ⌣-hom i +-- -- -- ϕ∘j≡ i = +-- -- -- Σ≡Prop (λ _ → isPropIsGroupHom _ _) +-- -- -- (funExt (sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) +-- -- -- λ _ → refl)) + +-- -- -- open import Cubical.Experiments.Brunerie +-- -- -- open import Cubical.HITs.Hopf + +-- -- -- open import Cubical.HITs.Join + +-- -- -- module fibS1 = hopfBase S1-AssocHSpace (sphereElim2 _ (λ _ _ → squash) ∣ refl ∣) + +-- -- -- S¹Hopf : S₊ 2 → Type +-- -- -- S¹Hopf = fibS1.Hopf + +-- -- -- TotalHopf' : Type _ +-- -- -- TotalHopf' = Σ (S₊ 2) S¹Hopf + +-- -- -- CP2 : Type _ +-- -- -- CP2 = fibS1.megaPush + +-- -- -- fibr : CP2 → Type _ +-- -- -- fibr = fibS1.P + +-- -- -- hopf : join S¹ S¹ → S₊ 2 +-- -- -- hopf x = fst (JoinS¹S¹→TotalHopf x) + +-- -- -- E* : Type _ +-- -- -- E* = fibS1.totalSpaceMegaPush + +-- -- -- IsoE' : Iso E* (join S¹ (join S¹ S¹)) +-- -- -- IsoE' = fibS1.IsoJoin₁ + +-- -- -- IsoE2 : (join S¹ (join S¹ S¹)) ≡ join S¹ (S₊ 3) +-- -- -- IsoE2 = cong (join S¹) (sym S³≡joinS¹S¹ ∙ isoToPath IsoS³S3) + +-- -- -- IsoTotalHopf' : Iso fibS1.TotalSpaceHopf' (join S¹ S¹) +-- -- -- IsoTotalHopf' = fibS1.joinIso₁ + +-- -- -- CP' : Type _ +-- -- -- CP' = Pushout (λ _ → tt) hopf + +-- -- -- conCP2 : (x y : CP2) → ∥ x ≡ y ∥₂ +-- -- -- conCP2 x y = sRec2 squash₂ (λ p q → ∣ p ∙ sym q ∣₂) (conCP2' x) (conCP2' y) +-- -- -- where +-- -- -- conCP2' : (x : CP2) → ∥ x ≡ inl tt ∥₂ +-- -- -- conCP2' (inl x) = ∣ refl ∣₂ +-- -- -- conCP2' (inr x) = sphereElim 1 {A = λ x → ∥ inr x ≡ inl tt ∥₂} (λ _ → squash₂) ∣ sym (push (inl base)) ∣₂ x +-- -- -- conCP2' (push a i₁) = ll a i₁ +-- -- -- where +-- -- -- h2 : ∀ {ℓ} {A : fibS1.TotalSpaceHopf' → Type ℓ} → ((a : _) → isProp (A a)) +-- -- -- → A (inl base) +-- -- -- → ((a : fibS1.TotalSpaceHopf') → A a) +-- -- -- h2 {A = A} p b = +-- -- -- PushoutToProp p (sphereElim 0 (λ _ → p _) b) +-- -- -- (sphereElim 0 (λ _ → p _) (subst A (push (base , base)) b)) + +-- -- -- ll : (a : fibS1.TotalSpaceHopf') → PathP (λ i → ∥ Path CP2 (push a i) (inl tt) ∥₂) (conCP2' (inl tt)) (conCP2' (inr (fibS1.induced a))) +-- -- -- ll = h2 (λ _ → isOfHLevelPathP' 1 squash₂ _ _) λ j → ∣ (λ i → push (inl base) (~ i ∧ j)) ∣₂ + +-- -- -- module GysinS1 = Gysin (CP2 , inl tt) fibr conCP2 2 idIso refl + +-- -- -- PushoutReplaceBase : +-- -- -- ∀ {ℓ ℓ' ℓ''} {A B : Type ℓ} {C : Type ℓ'} {D : Type ℓ''} {f : A → C} {g : A → D} +-- -- -- (e : B ≃ A) → Pushout (f ∘ fst e) (g ∘ fst e) ≡ Pushout f g +-- -- -- PushoutReplaceBase {f = f} {g = g} = +-- -- -- EquivJ (λ _ e → Pushout (f ∘ fst e) (g ∘ fst e) ≡ Pushout f g) +-- -- -- refl + +-- -- -- isContrH³E : isContr (coHom 3 (GysinS1.E')) +-- -- -- isContrH³E = +-- -- -- subst isContr +-- -- -- (sym (isoToPath (compIso (Iso1 0) (Iso2' 0))) +-- -- -- ∙ cong (coHom 3) (sym (isoToPath IsoE' ∙ IsoE2))) +-- -- -- (ll5 0 (snotz ∘ sym)) + +-- -- -- isContrH⁴E : isContr (coHom 4 (GysinS1.E')) +-- -- -- isContrH⁴E = +-- -- -- subst isContr +-- -- -- (sym (isoToPath (compIso (Iso1 1) (Iso2' 1))) +-- -- -- ∙ cong (coHom 4) (sym (isoToPath IsoE' ∙ IsoE2))) +-- -- -- (ll5 1 λ p → snotz (sym (cong predℕ p))) + +-- -- -- genH2 = GysinS1.e + +-- -- -- S³→Groupoid : ∀ {ℓ} (P : join S¹ S¹ → Type ℓ) +-- -- -- → ((x : _) → isGroupoid (P x)) +-- -- -- → P (inl base) +-- -- -- → (x : _) → P x +-- -- -- S³→Groupoid P grp b = +-- -- -- transport (λ i → (x : (sym (isoToPath S³IsojoinS¹S¹) ∙ isoToPath IsoS³S3) (~ i)) +-- -- -- → P (transp (λ j → (sym (isoToPath S³IsojoinS¹S¹) ∙ isoToPath IsoS³S3) (~ i ∧ ~ j)) i x)) +-- -- -- (sphereElim _ (λ _ → grp _) b) + +-- -- -- TotalHopf→Gpd : ∀ {ℓ} (P : fibS1.TotalSpaceHopf' → Type ℓ) +-- -- -- → ((x : _) → isGroupoid (P x)) +-- -- -- → P (inl base) +-- -- -- → (x : _) → P x +-- -- -- TotalHopf→Gpd P grp b = +-- -- -- transport (λ i → (x : sym (isoToPath IsoTotalHopf') i) +-- -- -- → P (transp (λ j → isoToPath IsoTotalHopf' (~ i ∧ ~ j)) i x)) +-- -- -- (S³→Groupoid _ (λ _ → grp _) b) + +-- -- -- TotalHopf→Gpd' : ∀ {ℓ} (P : Σ (S₊ 2) fibS1.Hopf → Type ℓ) +-- -- -- → ((x : _) → isGroupoid (P x)) +-- -- -- → P (north , base) +-- -- -- → (x : _) → P x +-- -- -- TotalHopf→Gpd' P grp b = +-- -- -- transport (λ i → (x : ua (_ , fibS1.isEquivTotalSpaceHopf'→TotalSpace) i) +-- -- -- → P (transp (λ j → ua (_ , fibS1.isEquivTotalSpaceHopf'→TotalSpace) (i ∨ j)) i x)) +-- -- -- (TotalHopf→Gpd _ (λ _ → grp _) b) + +-- -- -- CP2→Groupoid : ∀ {ℓ} {P : CP2 → Type ℓ} +-- -- -- → ((x : _) → is2Groupoid (P x)) +-- -- -- → (b : P (inl tt)) +-- -- -- → (s2 : ((x : _) → P (inr x))) +-- -- -- → PathP (λ i → P (push (inl base) i)) b (s2 north) +-- -- -- → (x : _) → P x +-- -- -- CP2→Groupoid {P = P} grp b s2 pp (inl x) = b +-- -- -- CP2→Groupoid {P = P} grp b s2 pp (inr x) = s2 x +-- -- -- CP2→Groupoid {P = P} grp b s2 pp (push a i₁) = h23 a i₁ +-- -- -- where +-- -- -- h23 : (a : fibS1.TotalSpaceHopf') → PathP (λ i → P (push a i)) b (s2 _) +-- -- -- h23 = TotalHopf→Gpd _ (λ _ → isOfHLevelPathP' 3 (grp _) _ _) pp + +-- -- -- mm : (S₊ 2 → coHomK 2) → (CP2 → coHomK 2) +-- -- -- mm f = λ { (inl x) → 0ₖ _ +-- -- -- ; (inr x) → f x -ₖ f north +-- -- -- ; (push a i) → TotalHopf→Gpd (λ x → 0ₖ 2 ≡ f (fibS1.TotalSpaceHopf'→TotalSpace x .fst) -ₖ f north) +-- -- -- (λ _ → isOfHLevelTrunc 4 _ _) +-- -- -- (sym (rCancelₖ 2 (f north))) +-- -- -- a i} + +-- -- -- e-inv : coHomGr 2 (S₊ 2) .fst → coHomGr 2 CP2 .fst +-- -- -- e-inv = sMap mm + +-- -- -- cancel : (f : CP2 → coHomK 2) → ∥ mm (f ∘ inr) ≡ f ∥ +-- -- -- cancel f = +-- -- -- pRec squash +-- -- -- (λ p → pRec squash +-- -- -- (λ f → ∣ funExt f ∣) +-- -- -- (zz2 p)) +-- -- -- (h1 (f (inl tt))) +-- -- -- where +-- -- -- h1 : (x : coHomK 2) → ∥ x ≡ 0ₖ 2 ∥ +-- -- -- h1 = coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash) ∣ refl ∣ + +-- -- -- zz2 : (p : f (inl tt) ≡ 0ₖ 2) → ∥ ((x : CP2) → mm (f ∘ inr) x ≡ f x) ∥ +-- -- -- zz2 p = trRec squash (λ pp → +-- -- -- ∣ CP2→Groupoid (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) +-- -- -- (sym p) +-- -- -- (λ x → (λ i → f (inr x) -ₖ f (push (inl base) (~ i))) +-- -- -- ∙∙ (λ i → f (inr x) -ₖ p i) +-- -- -- ∙∙ rUnitₖ 2 (f (inr x))) +-- -- -- pp ∣) +-- -- -- l +-- -- -- where +-- -- -- l : hLevelTrunc 1 (PathP ((λ i₁ → mm (f ∘ inr) (push (inl base) i₁) ≡ f (push (inl base) i₁))) +-- -- -- (sym p) +-- -- -- (((λ i₁ → f (inr north) -ₖ f (push (inl base) (~ i₁))) ∙∙ +-- -- -- (λ i₁ → f (inr north) -ₖ p i₁) ∙∙ rUnitₖ 2 (f (inr north))))) +-- -- -- l = isConnectedPathP 1 (isConnectedPath 2 (isConnectedKn 1) _ _) _ _ .fst + +-- -- -- e' : GroupIso (coHomGr 2 CP2) (coHomGr 2 (S₊ 2)) +-- -- -- Iso.fun (fst e') = sMap (_∘ inr) +-- -- -- Iso.inv (fst e') = e-inv +-- -- -- Iso.rightInv (fst e') = +-- -- -- sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) +-- -- -- λ f → trRec {B = Iso.fun (fst e') (Iso.inv (fst e') ∣ f ∣₂) ≡ ∣ f ∣₂} +-- -- -- (isOfHLevelPath 2 squash₂ _ _) +-- -- -- (λ p → cong ∣_∣₂ (funExt λ x → cong (λ y → (f x) -ₖ y) p ∙ rUnitₖ 2 (f x))) +-- -- -- (Iso.fun (PathIdTruncIso _) (isContr→isProp (isConnectedKn 1) ∣ f north ∣ ∣ 0ₖ 2 ∣)) +-- -- -- Iso.leftInv (fst e') = +-- -- -- sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) +-- -- -- λ f → pRec (squash₂ _ _) (cong ∣_∣₂) (cancel f) +-- -- -- snd e' = +-- -- -- makeIsGroupHom +-- -- -- (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ f g → refl) + +-- -- -- grrg : GroupIso (coHomGr 2 CP2) ℤGroup +-- -- -- grrg = compGroupIso e' (Hⁿ-Sⁿ≅ℤ 1) + +-- -- -- ⌣hom : GroupEquiv (coHomGr 2 CP2) (coHomGr 4 CP2) +-- -- -- fst (fst ⌣hom) = GysinS1.⌣-hom 2 .fst +-- -- -- snd (fst ⌣hom) = subst isEquiv (cong fst (GysinS1.ϕ∘j≡ 2)) h23 +-- -- -- where +-- -- -- h23 : isEquiv (GysinS1.ϕ∘j 2 .fst) +-- -- -- h23 = +-- -- -- SES→Iso +-- -- -- (GroupPath _ _ .fst (invGroupEquiv +-- -- -- (isContr→≃Unit isContrH³E +-- -- -- , makeIsGroupHom λ _ _ → refl))) +-- -- -- (GroupPath _ _ .fst (invGroupEquiv +-- -- -- (isContr→≃Unit isContrH⁴E +-- -- -- , makeIsGroupHom λ _ _ → refl))) +-- -- -- (GysinS1.susp∘ϕ 1) +-- -- -- (GysinS1.ϕ∘j 2) +-- -- -- (GysinS1.p-hom 4) +-- -- -- (GysinS1.Ker-ϕ∘j⊂Im-Susp∘ϕ _) +-- -- -- (GysinS1.Ker-p⊂Im-ϕ∘j _) +-- -- -- snd ⌣hom = GysinS1.⌣-hom 2 .snd + +-- -- -- isGenerator≃ℤ : ∀ {ℓ} (G : Group ℓ) (g : fst G) +-- -- -- → Type _ +-- -- -- isGenerator≃ℤ G g = +-- -- -- ∃[ e ∈ GroupIso G ℤGroup ] abs (Iso.fun (fst e) g) ≡ 1 + + + +-- -- -- -- ⌣-pres1 : ∀ {ℓ} (G : Type ℓ) (n : ℕ) → Type _ +-- -- -- -- ⌣-pres1 G n = +-- -- -- -- Σ[ x ∈ coHomGr n G .fst ] +-- -- -- -- isGenerator≃ℤ (coHomGr n G) x × isGenerator≃ℤ (coHomGr (n +' n) G) (x ⌣ x) + +-- -- -- -- genr : coHom 2 CP2 +-- -- -- -- genr = ∣ CP2→Groupoid (λ _ → isOfHLevelTrunc 4) +-- -- -- -- (0ₖ _) +-- -- -- -- ∣_∣ +-- -- -- -- refl ∣₂ + +-- -- -- -- ≡CP2 : (f g : CP2 → coHomK 2) → ∥ f ∘ inr ≡ g ∘ inr ∥ → Path (coHom 2 CP2) ∣ f ∣₂ ∣ g ∣₂ +-- -- -- -- ≡CP2 f g = pRec (squash₂ _ _) +-- -- -- -- (λ p → pRec (squash₂ _ _) (λ id → trRec (squash₂ _ _) +-- -- -- -- (λ pp → cong ∣_∣₂ +-- -- -- -- (funExt (CP2→Groupoid (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) +-- -- -- -- id +-- -- -- -- (funExt⁻ p) +-- -- -- -- (compPathR→PathP pp)))) +-- -- -- -- (pp2 (f (inl tt)) (g (inl tt)) id +-- -- -- -- (cong f (push (inl base)) ∙ (funExt⁻ p north) ∙ cong g (sym (push (inl base)))))) +-- -- -- -- (conn (f (inl tt)) (g (inl tt)))) + +-- -- -- -- where +-- -- -- -- conn : (x y : coHomK 2) → ∥ x ≡ y ∥ +-- -- -- -- conn = coHomK-elim _ (λ _ → isSetΠ λ _ → isOfHLevelSuc 1 squash) +-- -- -- -- (coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash) ∣ refl ∣) + +-- -- -- -- pp2 : (x y : coHomK 2) (p q : x ≡ y) → hLevelTrunc 1 (p ≡ q) +-- -- -- -- pp2 x y = λ p q → Iso.fun (PathIdTruncIso _) (isContr→isProp (isConnectedPath _ (isConnectedKn 1) x y) ∣ p ∣ ∣ q ∣) + +-- -- -- -- rUnit* : (x : S¹) → x * base ≡ x +-- -- -- -- rUnit* base = refl +-- -- -- -- rUnit* (loop i₁) = refl + +-- -- -- -- meridP : (a x : S¹) → Path (Path (coHomK 2) _ _) (cong ∣_∣ₕ (merid (a * x) ∙ sym (merid base))) +-- -- -- -- ((cong ∣_∣ₕ (merid a ∙ sym (merid base))) ∙ (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) +-- -- -- -- meridP = wedgeconFun _ _ (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) +-- -- -- -- (λ x → lUnit _ ∙ cong (_∙ cong ∣_∣ₕ (merid x ∙ sym (merid base))) (cong (cong ∣_∣ₕ) (sym (rCancel (merid base))))) +-- -- -- -- (λ x → (λ i → cong ∣_∣ₕ (merid (rUnit* x i) ∙ sym (merid base))) +-- -- -- -- ∙∙ rUnit _ +-- -- -- -- ∙∙ cong (cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙_) +-- -- -- -- (cong (cong ∣_∣ₕ) (sym (rCancel (merid base))))) +-- -- -- -- (sym (l (cong ∣_∣ₕ (merid base ∙ sym (merid base))) +-- -- -- -- (cong (cong ∣_∣ₕ) (sym (rCancel (merid base)))))) +-- -- -- -- where +-- -- -- -- l : ∀ {ℓ} {A : Type ℓ} {x : A} (p : x ≡ x) (P : refl ≡ p) +-- -- -- -- → lUnit p ∙ cong (_∙ p) P ≡ rUnit p ∙ cong (p ∙_) P +-- -- -- -- l p = J (λ p P → lUnit p ∙ cong (_∙ p) P ≡ rUnit p ∙ cong (p ∙_) P) refl + +-- -- -- -- lemmie : (x : S¹) → ptSn 1 ≡ x * (invLooper x) +-- -- -- -- lemmie base = refl +-- -- -- -- lemmie (loop i) j = +-- -- -- -- hcomp (λ r → λ {(i = i0) → base ; (i = i1) → base ; (j = i0) → base}) +-- -- -- -- base + +-- -- -- -- mmInv : (x : S¹) → Path (Path (coHomK 2) _ _) +-- -- -- -- (cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base))) +-- -- -- -- (cong ∣_∣ₕ (sym (merid x ∙ sym (merid base)))) +-- -- -- -- mmInv x = (lUnit _ +-- -- -- -- ∙∙ cong (_∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base))) (sym (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base))))) ∙∙ sym (assoc _ _ _)) +-- -- -- -- ∙∙ cong (sym (cong ∣_∣ₕ (merid x ∙ sym (merid base))) ∙_) hh +-- -- -- -- ∙∙ (assoc _ _ _ +-- -- -- -- ∙∙ cong (_∙ (cong ∣_∣ₕ (sym (merid x ∙ sym (merid base))))) (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) +-- -- -- -- ∙∙ sym (lUnit _)) +-- -- -- -- where +-- -- -- -- hh : cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base)) +-- -- -- -- ≡ cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙ cong ∣_∣ₕ (sym (merid x ∙ sym (merid base))) +-- -- -- -- hh = sym (meridP x (invLooper x)) ∙ ((λ i → cong ∣_∣ₕ (merid (lemmie x (~ i)) ∙ sym (merid base))) ∙ cong (cong ∣_∣ₕ) (rCancel (merid base))) ∙ sym (rCancel _) + +-- -- -- -- commS1 : (a x : S¹) → a * x ≡ x * a +-- -- -- -- commS1 = wedgeconFun _ _ (λ _ _ → isGroupoidS¹ _ _) +-- -- -- -- (sym ∘ rUnit*) +-- -- -- -- rUnit* +-- -- -- -- refl + +-- -- -- -- s233 : (a x : S¹) → (invEq (fibS1.μ-eq a) x) * a ≡ (invLooper a * x) * a +-- -- -- -- s233 a x = secEq (fibS1.μ-eq a) x ∙∙ cong (_* x) (lemmie a) ∙∙ assocHSpace.μ-assoc S1-AssocHSpace a (invLooper a) x ∙ commS1 _ _ + +-- -- -- -- ss23 : (a x : S¹) → invEq (fibS1.μ-eq a) x ≡ invLooper a * x +-- -- -- -- ss23 a x = sym (retEq (fibS1.μ-eq a) (invEq (fibS1.μ-eq a) x)) +-- -- -- -- ∙∙ cong (invEq (fibS1.μ-eq a)) (s233 a x) +-- -- -- -- ∙∙ retEq (fibS1.μ-eq a) (invLooper a * x) + +-- -- -- -- ll : GysinS1.e ≡ genr +-- -- -- -- ll = ≡CP2 _ _ ∣ funExt (λ x → funExt⁻ (cong fst (ll32 x)) south) ∣ +-- -- -- -- where +-- -- -- -- kr : (x : Σ (S₊ 2) fibS1.Hopf) → Path (hLevelTrunc 4 _) ∣ fst x ∣ ∣ north ∣ +-- -- -- -- kr = uncurry λ { north → λ y → cong ∣_∣ₕ (merid base ∙ sym (merid y)) +-- -- -- -- ; south → λ y → cong ∣_∣ₕ (sym (merid y)) +-- -- -- -- ; (merid a i) → lem a i} +-- -- -- -- where +-- -- -- -- lem : (a : S¹) → PathP (λ i → (y : fibS1.Hopf (merid a i)) → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a i ∣ ∣ north ∣) +-- -- -- -- (λ y → cong ∣_∣ₕ (merid base ∙ sym (merid y))) +-- -- -- -- λ y → cong ∣_∣ₕ (sym (merid y)) +-- -- -- -- lem a = toPathP (funExt λ x → cong (transport ((λ i₁ → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a i₁ ∣ ∣ north ∣))) +-- -- -- -- ((λ i → (λ z → cong ∣_∣ₕ (merid base +-- -- -- -- ∙ sym (merid (transport (λ j → uaInvEquiv (fibS1.μ-eq a) (~ i) j) x))) z)) +-- -- -- -- ∙ λ i → cong ∣_∣ₕ (merid base +-- -- -- -- ∙ sym (merid (transportRefl (invEq (fibS1.μ-eq a) x) i)))) +-- -- -- -- ∙∙ (λ i → transp (λ i₁ → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a (i₁ ∨ i) ∣ ∣ north ∣) i +-- -- -- -- (compPath-filler' (cong ∣_∣ₕ (sym (merid a))) +-- -- -- -- (cong ∣_∣ₕ (merid base ∙ sym (merid (ss23 a x i)))) i)) +-- -- -- -- ∙∙ cong ((cong ∣_∣ₕ) (sym (merid a)) ∙_) (cong (cong ∣_∣ₕ) (cong sym (symDistr (merid base) (sym (merid (invLooper a * x))))) +-- -- -- -- ∙ cong sym (meridP (invLooper a) x) +-- -- -- -- ∙ symDistr ((cong ∣_∣ₕ) (merid (invLooper a) ∙ sym (merid base))) +-- -- -- -- ((cong ∣_∣ₕ) (merid x ∙ sym (merid base))) +-- -- -- -- ∙ isCommΩK 2 (sym (λ i₁ → ∣ (merid x ∙ (λ i₂ → merid base (~ i₂))) i₁ ∣)) +-- -- -- -- (sym (λ i₁ → ∣ (merid (invLooper a) ∙ (λ i₂ → merid base (~ i₂))) i₁ ∣)) +-- -- -- -- ∙ cong₂ _∙_ (cong sym (mmInv a) ∙ cong-∙ ∣_∣ₕ (merid a) (sym (merid base))) +-- -- -- -- (cong (cong ∣_∣ₕ) (symDistr (merid x) (sym (merid base))) ∙ cong-∙ ∣_∣ₕ (merid base) (sym (merid x)))) +-- -- -- -- ∙∙ (λ j → (λ i₁ → ∣ merid a (~ i₁ ∨ j) ∣) ∙ +-- -- -- -- ((λ i₁ → ∣ merid a (i₁ ∨ j) ∣) ∙ +-- -- -- -- (λ i₁ → ∣ merid base (~ i₁ ∨ j) ∣)) +-- -- -- -- ∙ +-- -- -- -- (λ i₁ → ∣ merid base (i₁ ∨ j) ∣) ∙ +-- -- -- -- (λ i₁ → ∣ merid x (~ i₁) ∣ₕ)) +-- -- -- -- ∙∙ sym (lUnit _) +-- -- -- -- ∙∙ sym (assoc _ _ _) +-- -- -- -- ∙∙ (sym (lUnit _) ∙ sym (lUnit _) ∙ sym (lUnit _))) + +-- -- -- -- psst : (x : S₊ 2) → Q (CP2 , inl tt) fibr (inr x) →∙ coHomK-ptd 2 +-- -- -- -- fst (psst x) north = ∣ north ∣ +-- -- -- -- fst (psst x) south = ∣ x ∣ +-- -- -- -- fst (psst x) (merid a i₁) = kr (x , a) (~ i₁) +-- -- -- -- snd (psst x) = refl + +-- -- -- -- ll32-fst : GysinS1.c (inr north) .fst ≡ psst north .fst +-- -- -- -- ll32-fst = cong fst (GysinS1.cEq (inr north) ∣ push (inl base) ∣₂) +-- -- -- -- ∙ (funExt l) +-- -- -- -- where +-- -- -- -- l : (qb : _) → +-- -- -- -- ∣ (subst (fst ∘ Q (CP2 , inl tt) fibr) (sym (push (inl base))) qb) ∣ +-- -- -- -- ≡ psst north .fst qb +-- -- -- -- l north = refl +-- -- -- -- l south = cong ∣_∣ₕ (sym (merid base)) +-- -- -- -- l (merid a i) j = +-- -- -- -- hcomp (λ k → λ { (i = i0) → ∣ merid a (~ k ∧ j) ∣ +-- -- -- -- ; (i = i1) → ∣ merid base (~ j) ∣ +-- -- -- -- ; (j = i0) → ∣ transportRefl (merid a i) (~ k) ∣ +-- -- -- -- ; (j = i1) → ∣ compPath-filler (merid base) (sym (merid a)) k (~ i) ∣ₕ}) +-- -- -- -- (hcomp (λ k → λ { (i = i0) → ∣ merid a (j ∨ ~ k) ∣ +-- -- -- -- ; (i = i1) → ∣ merid base (~ j ∨ ~ k) ∣ +-- -- -- -- ; (j = i0) → ∣ merid a (~ k ∨ i) ∣ +-- -- -- -- ; (j = i1) → ∣ merid base (~ i ∨ ~ k) ∣ₕ}) +-- -- -- -- ∣ south ∣) + +-- -- -- -- is-setHepl : (x : S₊ 2) → isSet (Q (CP2 , inl tt) fibr (inr x) →∙ coHomK-ptd 2) +-- -- -- -- is-setHepl = sphereElim _ (λ _ → isProp→isOfHLevelSuc 1 (isPropIsOfHLevel 2)) +-- -- -- -- (isOfHLevel↑∙' 0 1) + +-- -- -- -- ll32 : (x : S₊ 2) → (GysinS1.c (inr x) ≡ psst x) +-- -- -- -- ll32 = sphereElim _ (λ x → isOfHLevelPath 2 (is-setHepl x) _ _) +-- -- -- -- (→∙Homogeneous≡ (isHomogeneousKn _) ll32-fst) + +-- -- -- -- isGenerator≃ℤ-e : isGenerator≃ℤ (coHomGr 2 CP2) GysinS1.e +-- -- -- -- isGenerator≃ℤ-e = +-- -- -- -- subst (isGenerator≃ℤ (coHomGr 2 CP2)) (sym ll) +-- -- -- -- ∣ grrg , refl ∣ + +-- -- -- -- ⌣-pres1-CP2 : ⌣-pres1 CP2 2 +-- -- -- -- ⌣-pres1-CP2 = +-- -- -- -- GysinS1.e +-- -- -- -- , (isGenerator≃ℤ-e +-- -- -- -- , ∣ compGroupIso (GroupEquiv→GroupIso (invGroupEquiv ⌣hom)) grrg +-- -- -- -- , {!!} ∣) + + +-- -- -- -- {- +-- -- -- -- Goal: snd (v' (pt A) (push a i₁)) ≡ +-- -- -- -- ua-gluePt (μ-eq (snd a)) i₁ (fst a) +-- -- -- -- ———— Boundary —————————————————————————————————————————————— +-- -- -- -- i₁ = i0 ⊢ HSpace.μₗ e (fst a) +-- -- -- -- i₁ = i1 ⊢ HSpace.μₗ e (HSpace.μ e (fst a) (snd a)) +-- -- -- -- -} - -- help : (x : _) → (v' (pt A)) x ≡ TotalSpaceHopf'→TotalSpace x - -- help (inl x) = ΣPathP (refl , HSpace.μₗ e x) - -- help (inr x) = ΣPathP (refl , (HSpace.μₗ e x)) - -- help (push (x , y) i) j = - -- comp (λ _ → Σ (Susp (typ A)) Hopf) - -- (λ k → λ {(i = i0) → merid y i , HSpace.μₗ e x j - -- ; (i = i1) → merid y i , assocHSpace.μ-assoc-filler e-ass x y j k - -- ; (j = i0) → merid y i , hfill - -- (λ j → λ { (i = i0) → HSpace.μ e (pt A) x - -- ; (i = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j - -- }) - -- (inS (ua-gluePt (μ-eq y) i (HSpace.μ e (pt A) x))) - -- k - -- ; (j = i1) → merid y i , ua-gluePt (μ-eq y) i x}) - -- (merid y i , ua-gluePt (μ-eq y) i (HSpace.μₗ e x j)) - -- where - -- open import Cubical.Foundations.Path - - -- PPΣ : ∀ {ℓ} {A : Type ℓ} {f : A ≃ A} (p : f ≡ f) → {!!} - -- PPΣ = {!!} - - -- V : PathP (λ i₁ → hcomp - -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x - -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j - -- }) - -- (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x)) ≡ - -- ua-gluePt (μ-eq y) i₁ x) - -- (HSpace.μₗ e x) - -- (HSpace.μₗ e (HSpace.μ e x y)) -- (HSpace.μₗ e (HSpace.μ e (fst a) (snd a))) - -- V = transport (λ z → {!PathP (λ i₁ → hfill - -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x - -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j - -- }) - -- (inS (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x))) z ≡ ua-gluePt (μ-eq y) i₁ x) - -- ? ?!}) - -- {!hfill - -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x - -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j - -- }) - -- (inS (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x))) ?!} -- toPathP ({!!} ∙∙ {!!} ∙∙ {!!}) -- toPathP (flipSquare {!!}) -- hcomp {!!} {!!} - - -- P : Pushout {A = TotalSpaceHopf'} (λ _ → tt) induced → Type _ - -- P (inl x) = typ A - -- P (inr x) = Hopf x - -- P (push a i₁) = ua (v a) i₁ +-- -- -- -- -- help : (x : _) → (v' (pt A)) x ≡ TotalSpaceHopf'→TotalSpace x +-- -- -- -- -- help (inl x) = ΣPathP (refl , HSpace.μₗ e x) +-- -- -- -- -- help (inr x) = ΣPathP (refl , (HSpace.μₗ e x)) +-- -- -- -- -- help (push (x , y) i) j = +-- -- -- -- -- comp (λ _ → Σ (Susp (typ A)) Hopf) +-- -- -- -- -- (λ k → λ {(i = i0) → merid y i , HSpace.μₗ e x j +-- -- -- -- -- ; (i = i1) → merid y i , assocHSpace.μ-assoc-filler e-ass x y j k +-- -- -- -- -- ; (j = i0) → merid y i , hfill +-- -- -- -- -- (λ j → λ { (i = i0) → HSpace.μ e (pt A) x +-- -- -- -- -- ; (i = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j +-- -- -- -- -- }) +-- -- -- -- -- (inS (ua-gluePt (μ-eq y) i (HSpace.μ e (pt A) x))) +-- -- -- -- -- k +-- -- -- -- -- ; (j = i1) → merid y i , ua-gluePt (μ-eq y) i x}) +-- -- -- -- -- (merid y i , ua-gluePt (μ-eq y) i (HSpace.μₗ e x j)) +-- -- -- -- -- where +-- -- -- -- -- open import Cubical.Foundations.Path + +-- -- -- -- -- PPΣ : ∀ {ℓ} {A : Type ℓ} {f : A ≃ A} (p : f ≡ f) → {!!} +-- -- -- -- -- PPΣ = {!!} + +-- -- -- -- -- V : PathP (λ i₁ → hcomp +-- -- -- -- -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x +-- -- -- -- -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j +-- -- -- -- -- }) +-- -- -- -- -- (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x)) ≡ +-- -- -- -- -- ua-gluePt (μ-eq y) i₁ x) +-- -- -- -- -- (HSpace.μₗ e x) +-- -- -- -- -- (HSpace.μₗ e (HSpace.μ e x y)) -- (HSpace.μₗ e (HSpace.μ e (fst a) (snd a))) +-- -- -- -- -- V = transport (λ z → {!PathP (λ i₁ → hfill +-- -- -- -- -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x +-- -- -- -- -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j +-- -- -- -- -- }) +-- -- -- -- -- (inS (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x))) z ≡ ua-gluePt (μ-eq y) i₁ x) +-- -- -- -- -- ? ?!}) +-- -- -- -- -- {!hfill +-- -- -- -- -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x +-- -- -- -- -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j +-- -- -- -- -- }) +-- -- -- -- -- (inS (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x))) ?!} -- toPathP ({!!} ∙∙ {!!} ∙∙ {!!}) -- toPathP (flipSquare {!!}) -- hcomp {!!} {!!} + +-- -- -- -- -- P : Pushout {A = TotalSpaceHopf'} (λ _ → tt) induced → Type _ +-- -- -- -- -- P (inl x) = typ A +-- -- -- -- -- P (inr x) = Hopf x +-- -- -- -- -- P (push a i₁) = ua (v a) i₁ diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index e2392008f..97ef86156 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -52,603 +52,893 @@ private variable ℓ ℓ' : Level -------------------- Connectedness --------------------- -is2ConnectedKn : (n : ℕ) → isConnected 2 (coHomK (suc n)) -is2ConnectedKn zero = ∣ ∣ base ∣ ∣ - , trElim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) - (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 (isOfHLevelTrunc 2)) _ _) - (toPropElim (λ _ → isOfHLevelTrunc 2 _ _) refl)) -is2ConnectedKn (suc n) = ∣ ∣ north ∣ ∣ - , trElim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) - (trElim (λ _ → isProp→isOfHLevelSuc (3 + n) (isOfHLevelTrunc 2 _ _)) - (suspToPropElim (ptSn (suc n)) (λ _ → isOfHLevelTrunc 2 _ _) refl)) - -isConnectedKn : (n : ℕ) → isConnected (2 + n) (coHomK (suc n)) -isConnectedKn n = isOfHLevelRetractFromIso 0 (invIso (truncOfTruncIso (2 + n) 1)) (sphereConnected (suc n)) - --- direct proof of connectedness of ΩKₙ₊₁ not relying on the equivalence ∥ a ≡ b ∥ₙ ≃ (∣ a ∣ₙ₊₁ ≡ ∣ b ∣ₙ₊₁) -isConnectedPathKn : (n : ℕ) (x y : (coHomK (suc n))) → isConnected (suc n) (x ≡ y) -isConnectedPathKn n = - trElim (λ _ → isProp→isOfHLevelSuc (2 + n) (isPropΠ λ _ → isPropIsContr)) - (sphereElim _ (λ _ → isProp→isOfHLevelSuc n (isPropΠ λ _ → isPropIsContr)) - λ y → isContrRetractOfConstFun - {B = (hLevelTrunc (suc n) (ptSn (suc n) ≡ ptSn (suc n)))} ∣ refl ∣ - (fun⁻ n y - , trElim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) - (J (λ y p → fun⁻ n y _ ≡ _) (funExt⁻ (fun⁻Id n) ∣ refl ∣)))) - where - fun⁻ : (n : ℕ) → (y : coHomK (suc n)) → - hLevelTrunc (suc n) (ptSn (suc n) ≡ ptSn (suc n)) - → hLevelTrunc (suc n) (∣ ptSn (suc n) ∣ ≡ y) - fun⁻ n = - trElim (λ _ → isOfHLevelΠ (3 + n) λ _ → isOfHLevelSuc (2 + n) (isOfHLevelSuc (suc n) (isOfHLevelTrunc (suc n)))) - (sphereElim n (λ _ → isOfHLevelΠ (suc n) λ _ → isOfHLevelTrunc (suc n)) λ _ → ∣ refl ∣) - - fun⁻Id : (n : ℕ) → fun⁻ n ∣ ptSn (suc n) ∣ ≡ λ _ → ∣ refl ∣ - fun⁻Id zero = refl - fun⁻Id (suc n) = refl - -------------------- --- Induction principles for cohomology groups (n ≥ 1) --- If we want to show a proposition about some x : Hⁿ(A), it suffices to show it under the --- assumption that x = ∣ f ∣₂ for some f : A → Kₙ and that f is pointed - -coHomPointedElim : {A : Type ℓ} (n : ℕ) (a : A) {B : coHom (suc n) A → Type ℓ'} - → ((x : coHom (suc n) A) → isProp (B x)) - → ((f : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → B ∣ f ∣₂) - → (x : coHom (suc n) A) → B x -coHomPointedElim {ℓ' = ℓ'} {A = A} n a isprop indp = - sElim (λ _ → isOfHLevelSuc 1 (isprop _)) - λ f → helper n isprop indp f (f a) refl - where - helper : (n : ℕ) {B : coHom (suc n) A → Type ℓ'} - → ((x : coHom (suc n) A) → isProp (B x)) - → ((f : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → B ∣ f ∣₂) - → (f : A → coHomK (suc n)) - → (x : coHomK (suc n)) - → f a ≡ x → B ∣ f ∣₂ - -- pattern matching a bit extra to avoid isOfHLevelPlus' - helper zero isprop ind f = - trElim (λ _ → isOfHLevelPlus {n = 1} 2 (isPropΠ λ _ → isprop _)) - (toPropElim (λ _ → isPropΠ λ _ → isprop _) (ind f)) - helper (suc zero) isprop ind f = - trElim (λ _ → isOfHLevelPlus {n = 1} 3 (isPropΠ λ _ → isprop _)) - (suspToPropElim base (λ _ → isPropΠ λ _ → isprop _) (ind f)) - helper (suc (suc zero)) isprop ind f = - trElim (λ _ → isOfHLevelPlus {n = 1} 4 (isPropΠ λ _ → isprop _)) - (suspToPropElim north (λ _ → isPropΠ λ _ → isprop _) (ind f)) - helper (suc (suc (suc n))) isprop ind f = - trElim (λ _ → isOfHLevelPlus' {n = 5 + n} 1 (isPropΠ λ _ → isprop _)) - (suspToPropElim north (λ _ → isPropΠ λ _ → isprop _) (ind f)) - - -coHomPointedElim2 : {A : Type ℓ} (n : ℕ) (a : A) {B : coHom (suc n) A → coHom (suc n) A → Type ℓ'} - → ((x y : coHom (suc n) A) → isProp (B x y)) - → ((f g : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → g a ≡ coHom-pt (suc n) → B ∣ f ∣₂ ∣ g ∣₂) - → (x y : coHom (suc n) A) → B x y -coHomPointedElim2 {ℓ' = ℓ'} {A = A} n a isprop indp = sElim2 (λ _ _ → isOfHLevelSuc 1 (isprop _ _)) - λ f g → helper n a isprop indp f g (f a) (g a) refl refl - where - helper : (n : ℕ) (a : A) {B : coHom (suc n) A → coHom (suc n) A → Type ℓ'} - → ((x y : coHom (suc n) A) → isProp (B x y)) - → ((f g : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → g a ≡ coHom-pt (suc n) → B ∣ f ∣₂ ∣ g ∣₂) - → (f g : A → coHomK (suc n)) - → (x y : coHomK (suc n)) - → f a ≡ x → g a ≡ y - → B ∣ f ∣₂ ∣ g ∣₂ - helper zero a isprop indp f g = - elim2 (λ _ _ → isOfHLevelPlus {n = 1} 2 (isPropΠ2 λ _ _ → isprop _ _)) - (toPropElim2 (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) - helper (suc zero) a isprop indp f g = - elim2 (λ _ _ → isOfHLevelPlus {n = 1} 3 (isPropΠ2 λ _ _ → isprop _ _)) - (suspToPropElim2 base (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) - helper (suc (suc zero)) a isprop indp f g = - elim2 (λ _ _ → isOfHLevelPlus {n = 1} 4 (isPropΠ2 λ _ _ → isprop _ _)) - (suspToPropElim2 north (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) - helper (suc (suc (suc n))) a isprop indp f g = - elim2 (λ _ _ → isOfHLevelPlus' {n = 5 + n} 1 (isPropΠ2 λ _ _ → isprop _ _)) - (suspToPropElim2 north (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) - -coHomK-elim : ∀ {ℓ} (n : ℕ) {B : coHomK (suc n) → Type ℓ} - → ((x : _) → isOfHLevel (suc n) (B x)) - → B (0ₖ (suc n)) - → (x : _) → B x -coHomK-elim n {B = B } hlev b = - trElim (λ _ → isOfHLevelPlus {n = (suc n)} 2 (hlev _)) - (sphereElim _ (hlev ∘ ∣_∣) b) - -{- Equivalence between cohomology of A and reduced cohomology of (A + 1) -} -coHomRed+1Equiv : (n : ℕ) → - (A : Type ℓ) → - (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) -coHomRed+1Equiv zero A i = ∥ helpLemma {C = (ℤ , pos 0)} i ∥₂ - module coHomRed+1 where - helpLemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →∙ C))) - helpLemma {C = C} = isoToPath (iso map1 - map2 - (λ b → linvPf b) - (λ _ → refl)) - where - map1 : (A → typ C) → ((((A ⊎ Unit) , inr (tt)) →∙ C)) - map1 f = map1' , refl - module helpmap where - map1' : A ⊎ Unit → fst C - map1' (inl x) = f x - map1' (inr x) = pt C - - map2 : ((((A ⊎ Unit) , inr (tt)) →∙ C)) → (A → typ C) - map2 (g , pf) x = g (inl x) - - linvPf : (b :((((A ⊎ Unit) , inr (tt)) →∙ C))) → map1 (map2 b) ≡ b - linvPf (f , snd) i = (λ x → helper x i) , λ j → snd ((~ i) ∨ j) - where - helper : (x : A ⊎ Unit) → ((helpmap.map1') (map2 (f , snd)) x) ≡ f x - helper (inl x) = refl - helper (inr tt) = sym snd -coHomRed+1Equiv (suc zero) A i = ∥ coHomRed+1.helpLemma A i {C = (coHomK 1 , ∣ base ∣)} i ∥₂ -coHomRed+1Equiv (suc (suc n)) A i = ∥ coHomRed+1.helpLemma A i {C = (coHomK (2 + n) , ∣ north ∣)} i ∥₂ - -Iso-coHom-coHomRed : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → Iso (coHomRed (suc n) A) (coHom (suc n) (typ A)) -fun (Iso-coHom-coHomRed {A = A , a} n) = map fst -inv' (Iso-coHom-coHomRed {A = A , a} n) = map λ f → (λ x → f x -ₖ f a) , rCancelₖ _ _ -rightInv (Iso-coHom-coHomRed {A = A , a} n) = - sElim (λ _ → isOfHLevelPath 2 § _ _) - λ f → trRec (isProp→isOfHLevelSuc _ (§ _ _)) - (λ p → cong ∣_∣₂ (funExt λ x → cong (λ y → f x +ₖ y) (cong -ₖ_ p ∙ -0ₖ) ∙ rUnitₖ _ (f x))) - (Iso.fun (PathIdTruncIso (suc n)) (isContr→isProp (isConnectedKn n) ∣ f a ∣ ∣ 0ₖ _ ∣)) -leftInv (Iso-coHom-coHomRed {A = A , a} n) = - sElim (λ _ → isOfHLevelPath 2 § _ _) - λ {(f , p) → cong ∣_∣₂ (ΣPathP (((funExt λ x → (cong (λ y → f x -ₖ y) p - ∙∙ cong (λ y → f x +ₖ y) -0ₖ - ∙∙ rUnitₖ _ (f x)) ∙ refl)) - , helper n (f a) (sym p)))} - where - path : (n : ℕ) (x : coHomK (suc n)) (p : 0ₖ _ ≡ x) → _ - path n x p = (cong (λ y → x -ₖ y) (sym p) ∙∙ cong (λ y → x +ₖ y) -0ₖ ∙∙ rUnitₖ _ x) ∙ refl - - helper : (n : ℕ) (x : coHomK (suc n)) (p : 0ₖ _ ≡ x) - → PathP (λ i → path n x p i ≡ 0ₖ _) (rCancelₖ _ x) (sym p) - helper zero x = - J (λ x p → PathP (λ i → path 0 x p i ≡ 0ₖ _) - (rCancelₖ _ x) (sym p)) - λ i j → rUnit (rUnit (λ _ → 0ₖ 1) (~ j)) (~ j) i - helper (suc n) x = - J (λ x p → PathP (λ i → path (suc n) x p i ≡ 0ₖ _) (rCancelₖ _ x) (sym p)) - λ i j → rCancelₖ (suc (suc n)) (0ₖ (suc (suc n))) (~ i ∧ ~ j) - -+∙≡+ : (n : ℕ) {A : Pointed ℓ} (x y : coHomRed (suc n) A) - → Iso.fun (Iso-coHom-coHomRed n) (x +ₕ∙ y) - ≡ Iso.fun (Iso-coHom-coHomRed n) x +ₕ Iso.fun (Iso-coHom-coHomRed n) y - -+∙≡+ zero = sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl -+∙≡+ (suc n) = sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl - -private - homhelp : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) (x y : coHom (suc n) (typ A)) - → Iso.inv (Iso-coHom-coHomRed {A = A} n) (x +ₕ y) - ≡ Iso.inv (Iso-coHom-coHomRed n) x +ₕ∙ Iso.inv (Iso-coHom-coHomRed n) y - homhelp n A = morphLemmas.isMorphInv _+ₕ∙_ _+ₕ_ - (Iso.fun (Iso-coHom-coHomRed n)) (+∙≡+ n) _ - (Iso.rightInv (Iso-coHom-coHomRed n)) (Iso.leftInv (Iso-coHom-coHomRed n)) - -coHomGr≅coHomRedGr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) - → GroupEquiv (coHomRedGrDir (suc n) A) (coHomGr (suc n) (typ A)) -fst (coHomGr≅coHomRedGr n A) = isoToEquiv (Iso-coHom-coHomRed n) -snd (coHomGr≅coHomRedGr n A) = makeIsGroupHom (+∙≡+ n) - -coHomRedGroup : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → AbGroup ℓ -coHomRedGroup zero A = coHomRedGroupDir zero A -coHomRedGroup (suc n) A = - InducedAbGroup (coHomGroup (suc n) (typ A)) - _+ₕ∙_ - (isoToEquiv (invIso (Iso-coHom-coHomRed n))) - (homhelp n A) - -abstract - coHomGroup≡coHomRedGroup : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) - → coHomGroup (suc n) (typ A) ≡ coHomRedGroup (suc n) A - coHomGroup≡coHomRedGroup n A = - InducedAbGroupPath (coHomGroup (suc n) (typ A)) - _+ₕ∙_ - (isoToEquiv (invIso (Iso-coHom-coHomRed n))) - (homhelp n A) - -------------------- Kₙ ≃ ΩKₙ₊₁ --------------------- --- This proof uses the encode-decode method rather than Freudenthal - --- We define the map σ : Kₙ → ΩKₙ₊₁ and prove that it is a morphism -private - module _ (n : ℕ) where - σ : {n : ℕ} → coHomK (suc n) → Path (coHomK (2 + n)) ∣ north ∣ ∣ north ∣ - σ {n = n} = trRec (isOfHLevelTrunc (4 + n) _ _) - λ a → cong ∣_∣ (merid a ∙ sym (merid (ptSn (suc n)))) - - σ-hom-helper : ∀ {ℓ} {A : Type ℓ} {a : A} (p : a ≡ a) (r : refl ≡ p) - → lUnit p ∙ cong (_∙ p) r ≡ rUnit p ∙ cong (p ∙_) r - σ-hom-helper p = J (λ p r → lUnit p ∙ cong (_∙ p) r ≡ rUnit p ∙ cong (p ∙_) r) refl - - σ-hom : {n : ℕ} (x y : coHomK (suc n)) → σ (x +ₖ y) ≡ σ x ∙ σ y - σ-hom {n = zero} = - elim2 (λ _ _ → isOfHLevelPath 3 (isOfHLevelTrunc 4 _ _) _ _) - (wedgeconFun _ _ - (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) - (λ x → lUnit _ - ∙ cong (_∙ σ ∣ x ∣) (cong (cong ∣_∣) (sym (rCancel (merid base))))) - (λ y → cong σ (rUnitₖ 1 ∣ y ∣) - ∙∙ rUnit _ - ∙∙ cong (σ ∣ y ∣ ∙_) (cong (cong ∣_∣) (sym (rCancel (merid base))))) - (sym (σ-hom-helper (σ ∣ base ∣) (cong (cong ∣_∣) (sym (rCancel (merid base))))))) - σ-hom {n = suc n} = - elim2 (λ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (5 + n) _ _) _ _) - (wedgeconFun _ _ (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (wedgeConHLev' n) _ _) - (λ x → lUnit _ - ∙ cong (_∙ σ ∣ x ∣) (cong (cong ∣_∣) (sym (rCancel (merid north))))) - (λ y → cong σ (rUnitₖ (2 + n) ∣ y ∣) - ∙∙ rUnit _ - ∙∙ cong (σ ∣ y ∣ ∙_) (cong (cong ∣_∣) (sym (rCancel (merid north))))) - (sym (σ-hom-helper (σ ∣ north ∣) (cong (cong ∣_∣) (sym (rCancel (merid north))))))) - - -- We will need to following lemma - σ-minusDistr : {n : ℕ} (x y : coHomK (suc n)) → σ (x -ₖ y) ≡ σ x ∙ sym (σ y) - σ-minusDistr {n = n} = - morphLemmas.distrMinus' - _+ₖ_ _∙_ - σ σ-hom ∣ (ptSn (suc n)) ∣ refl - -ₖ_ sym - (λ x → sym (lUnit x)) (λ x → sym (rUnit x)) - (rUnitₖ (suc n)) - (lCancelₖ (suc n)) rCancel - (assocₖ (suc n)) assoc∙ - (cong (cong ∣_∣) (rCancel (merid (ptSn (suc n))))) - - -- we define the code using addIso - Code : (n : ℕ) → coHomK (2 + n) → Type₀ - Code n x = (trRec {B = TypeOfHLevel ℓ-zero (3 + n)} (isOfHLevelTypeOfHLevel (3 + n)) - λ a → Code' a , hLevCode' a) x .fst - where - Code' : (S₊ (2 + n)) → Type₀ - Code' north = coHomK (suc n) - Code' south = coHomK (suc n) - Code' (merid a i) = isoToPath (addIso (suc n) ∣ a ∣) i - - hLevCode' : (x : S₊ (2 + n)) → isOfHLevel (3 + n) (Code' x) - hLevCode' = suspToPropElim (ptSn (suc n)) (λ _ → isPropIsOfHLevel (3 + n)) (isOfHLevelTrunc (3 + n)) - - symMeridLem : (n : ℕ) → (x : S₊ (suc n)) (y : coHomK (suc n)) - → subst (Code n) (cong ∣_∣ (sym (merid x))) y ≡ y -ₖ ∣ x ∣ - symMeridLem n x = trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) - (λ y → cong (_-ₖ ∣ x ∣) (transportRefl ∣ y ∣)) - - decode : {n : ℕ} (x : coHomK (2 + n)) → Code n x → ∣ north ∣ ≡ x - decode {n = n} = trElim (λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) - decode-elim - where - north≡merid : (a : S₊ (suc n)) - → Path (coHomK (2 + n)) ∣ north ∣ ∣ north ∣ - ≡ (Path (coHomK (2 + n)) ∣ north ∣ ∣ south ∣) - north≡merid a i = Path (coHomK (2 + n)) ∣ north ∣ ∣ merid a i ∣ - - decode-elim : (a : S₊ (2 + n)) → Code n ∣ a ∣ → Path (coHomK (2 + n)) ∣ north ∣ ∣ a ∣ - decode-elim north = σ - decode-elim south = trRec (isOfHLevelTrunc (4 + n) _ _) - λ a → cong ∣_∣ (merid a) - decode-elim (merid a i) = - hcomp (λ k → λ { (i = i0) → σ - ; (i = i1) → mainPath a k}) - (funTypeTransp (Code n) (λ x → ∣ north ∣ ≡ x) (cong ∣_∣ (merid a)) σ i) - where - mainPath : (a : (S₊ (suc n))) → - transport (north≡merid a) ∘ σ ∘ transport (λ i → Code n ∣ merid a (~ i) ∣) - ≡ trRec (isOfHLevelTrunc (4 + n) _ _) λ a → cong ∣_∣ (merid a) - mainPath a = funExt (trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (4 + n) _ _) _ _) - (λ x → (λ i → transport (north≡merid a) (σ (symMeridLem n a ∣ x ∣ i))) - ∙∙ cong (transport (north≡merid a)) (-distrHelp x) - ∙∙ (substAbove x))) - where - -distrHelp : (x : S₊ (suc n)) → σ (∣ x ∣ -ₖ ∣ a ∣) ≡ cong ∣_∣ (merid x) ∙ cong ∣_∣ (sym (merid a)) - -distrHelp x = - σ-minusDistr ∣ x ∣ ∣ a ∣ - ∙ (λ i → (cong ∣_∣ (compPath-filler (merid x) (λ j → merid (ptSn (suc n)) (~ j ∨ i)) (~ i))) - ∙ (cong ∣_∣ (sym (compPath-filler (merid a) (λ j → merid (ptSn (suc n)) (~ j ∨ i)) (~ i))))) - - substAbove : (x : S₊ (suc n)) → transport (north≡merid a) (cong ∣_∣ (merid x) ∙ cong ∣_∣ (sym (merid a))) - ≡ cong ∣_∣ (merid x) - substAbove x i = transp (λ j → north≡merid a (i ∨ j)) i - (compPath-filler (cong ∣_∣ (merid x)) (λ j → ∣ merid a (~ j ∨ i) ∣) (~ i)) - - encode : {n : ℕ} {x : coHomK (2 + n)} → Path (coHomK (2 + n)) ∣ north ∣ x → Code n x - encode {n = n} p = transport (cong (Code n) p) ∣ (ptSn (suc n)) ∣ - - decode-encode : {n : ℕ} {x : coHomK (2 + n)} (p : Path (coHomK (2 + n)) ∣ north ∣ x) - → decode _ (encode p) ≡ p - decode-encode {n = n} = - J (λ y p → decode _ (encode p) ≡ p) - (cong (decode ∣ north ∣) (transportRefl ∣ ptSn (suc n) ∣) - ∙ cong (cong ∣_∣) (rCancel (merid (ptSn (suc n))))) - --- We define an addition operation on Code which we can use in order to show that encode is a --- morphism (in a very loose sense) - hLevCode : {n : ℕ} (x : coHomK (2 + n)) → isOfHLevel (3 + n) (Code n x) - hLevCode {n = n} = - trElim (λ _ → isProp→isOfHLevelSuc (3 + n) (isPropIsOfHLevel (3 + n))) - (sphereToPropElim _ - (λ _ → (isPropIsOfHLevel (3 + n))) (isOfHLevelTrunc (3 + n))) - - Code-add' : {n : ℕ} (x : _) → Code n ∣ north ∣ → Code n ∣ x ∣ → Code n ∣ x ∣ - Code-add' {n = n} north = _+ₖ_ - Code-add' {n = n} south = _+ₖ_ - Code-add' {n = n} (merid a i) = helper n a i - where - help : (n : ℕ) → (x y a : S₊ (suc n)) - → transport (λ i → Code n ∣ north ∣ → Code n ∣ merid a i ∣ → Code n ∣ merid a i ∣) - (_+ₖ_) ∣ x ∣ ∣ y ∣ - ≡ ∣ x ∣ +ₖ ∣ y ∣ - help n x y a = - (λ i → transportRefl ((∣ transportRefl x i ∣ +ₖ (∣ transportRefl y i ∣ -ₖ ∣ a ∣)) +ₖ ∣ a ∣) i) - ∙∙ cong (_+ₖ ∣ a ∣) (assocₖ _ ∣ x ∣ ∣ y ∣ (-ₖ ∣ a ∣)) - ∙∙ sym (assocₖ _ (∣ x ∣ +ₖ ∣ y ∣) (-ₖ ∣ a ∣) ∣ a ∣) - ∙∙ cong ((∣ x ∣ +ₖ ∣ y ∣) +ₖ_) (lCancelₖ _ ∣ a ∣) - ∙∙ rUnitₖ _ _ - - helper : (n : ℕ) (a : S₊ (suc n)) - → PathP (λ i → Code n ∣ north ∣ → Code n ∣ merid a i ∣ → Code n ∣ merid a i ∣) _+ₖ_ _+ₖ_ - helper n a = - toPathP (funExt - (trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelΠ (3 + n) (λ _ → isOfHLevelTrunc (3 + n))) _ _) - λ x → funExt - (trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) - λ y → help n x y a))) - - Code-add : {n : ℕ} (x : _) → Code n ∣ north ∣ → Code n x → Code n x - Code-add {n = n} = - trElim (λ x → isOfHLevelΠ (4 + n) λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelSuc (3 + n) (hLevCode {n = n} x)) - Code-add' - - encode-hom : {n : ℕ} {x : _} (q : 0ₖ _ ≡ 0ₖ _) (p : 0ₖ _ ≡ x) - → encode (q ∙ p) ≡ Code-add {n = n} x (encode q) (encode p) - encode-hom {n = n} q = J (λ x p → encode (q ∙ p) ≡ Code-add {n = n} x (encode q) (encode p)) - (cong encode (sym (rUnit q)) - ∙∙ sym (rUnitₖ _ (encode q)) - ∙∙ cong (encode q +ₖ_) (cong ∣_∣ (sym (transportRefl _)))) - -stabSpheres : (n : ℕ) → Iso (coHomK (suc n)) (typ (Ω (coHomK-ptd (2 + n)))) -fun (stabSpheres n) = decode _ -inv' (stabSpheres n) = encode -rightInv (stabSpheres n) p = decode-encode p -leftInv (stabSpheres n) = - trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) - λ a → cong encode (congFunct ∣_∣ (merid a) (sym (merid (ptSn (suc n))))) - ∙∙ (λ i → transport (congFunct (Code n) (cong ∣_∣ (merid a)) - (cong ∣_∣ (sym (merid (ptSn (suc n))))) i) ∣ ptSn (suc n) ∣) - ∙∙ (substComposite (λ x → x) - (cong (Code n) (cong ∣_∣ (merid a))) - (cong (Code n) (cong ∣_∣ (sym (merid (ptSn (suc n)))))) ∣ ptSn (suc n) ∣ - ∙∙ cong (transport (λ i → Code n ∣ merid (ptSn (suc n)) (~ i) ∣)) - (transportRefl (∣ (ptSn (suc n)) ∣ +ₖ ∣ a ∣) ∙ lUnitₖ (suc n) ∣ a ∣) - ∙∙ symMeridLem n (ptSn (suc n)) ∣ a ∣ - ∙∙ cong (∣ a ∣ +ₖ_) -0ₖ - ∙∙ rUnitₖ (suc n) ∣ a ∣) - -Iso-Kn-ΩKn+1 : (n : HLevel) → Iso (coHomK n) (typ (Ω (coHomK-ptd (suc n)))) -Iso-Kn-ΩKn+1 zero = invIso (compIso (congIso (truncIdempotentIso _ isGroupoidS¹)) ΩS¹Isoℤ) -Iso-Kn-ΩKn+1 (suc n) = stabSpheres n - -Kn≃ΩKn+1 : {n : ℕ} → coHomK n ≃ typ (Ω (coHomK-ptd (suc n))) -Kn≃ΩKn+1 {n = n} = isoToEquiv (Iso-Kn-ΩKn+1 n) - --- Some properties of the Iso -Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) -Kn→ΩKn+1 n = Iso.fun (Iso-Kn-ΩKn+1 n) - -ΩKn+1→Kn : (n : ℕ) → typ (Ω (coHomK-ptd (suc n))) → coHomK n -ΩKn+1→Kn n = Iso.inv (Iso-Kn-ΩKn+1 n) - -Kn→ΩKn+10ₖ : (n : ℕ) → Kn→ΩKn+1 n (0ₖ n) ≡ refl -Kn→ΩKn+10ₖ zero = sym (rUnit refl) -Kn→ΩKn+10ₖ (suc n) i j = ∣ (rCancel (merid (ptSn (suc n))) i j) ∣ - -ΩKn+1→Kn-refl : (n : ℕ) → ΩKn+1→Kn n refl ≡ 0ₖ n -ΩKn+1→Kn-refl zero = refl -ΩKn+1→Kn-refl (suc zero) = refl -ΩKn+1→Kn-refl (suc (suc n)) = refl - -Kn≃ΩKn+1∙ : {n : ℕ} → coHomK-ptd n ≡ (Ω (coHomK-ptd (suc n))) -Kn≃ΩKn+1∙ {n = n} = ua∙ Kn≃ΩKn+1 (Kn→ΩKn+10ₖ n) - -Kn→ΩKn+1-hom : (n : ℕ) (x y : coHomK n) → Kn→ΩKn+1 n (x +[ n ]ₖ y) ≡ Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y -Kn→ΩKn+1-hom zero x y = (λ j i → hfill (doubleComp-faces (λ i₁ → ∣ base ∣) (λ _ → ∣ base ∣) i) - (inS (∣ intLoop (x ℤ+ y) i ∣)) (~ j)) - ∙∙ (λ j i → ∣ intLoop-hom x y (~ j) i ∣) - ∙∙ (congFunct ∣_∣ (intLoop x) (intLoop y) - ∙ cong₂ _∙_ (λ j i → hfill (doubleComp-faces (λ i₁ → ∣ base ∣) (λ _ → ∣ base ∣) i) - (inS (∣ intLoop x i ∣)) j) - λ j i → hfill (doubleComp-faces (λ i₁ → ∣ base ∣) (λ _ → ∣ base ∣) i) - (inS (∣ intLoop y i ∣)) j) -Kn→ΩKn+1-hom (suc n) = σ-hom - -ΩKn+1→Kn-hom : (n : ℕ) (x y : Path (coHomK (suc n)) (0ₖ _) (0ₖ _)) - → ΩKn+1→Kn n (x ∙ y) ≡ ΩKn+1→Kn n x +[ n ]ₖ ΩKn+1→Kn n y -ΩKn+1→Kn-hom zero p q = - cong winding (congFunct (trRec isGroupoidS¹ (λ x → x)) p q) - ∙ winding-hom (cong (trRec isGroupoidS¹ (λ x → x)) p) (cong (trRec isGroupoidS¹ (λ x → x)) q) -ΩKn+1→Kn-hom (suc n) = encode-hom - -Kn→ΩKn+1-ₖ : (n : ℕ) (x : coHomK n) → Kn→ΩKn+1 n (-ₖ x) ≡ sym (Kn→ΩKn+1 n x) -Kn→ΩKn+1-ₖ n x = - lUnit _ - ∙∙ cong (_∙ Kn→ΩKn+1 n (-ₖ x)) (sym (lCancel _)) - ∙∙ sym (assoc∙ _ _ _) - ∙∙ cong (sym (Kn→ΩKn+1 n x) ∙_) help - ∙∙ sym (rUnit _) - where - help : Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n (-ₖ x) ≡ refl - help = sym (Kn→ΩKn+1-hom n x (-ₖ x)) ∙∙ cong (Kn→ΩKn+1 n) (rCancelₖ n x) ∙∙ Kn→ΩKn+10ₖ n - -isHomogeneousKn : (n : HLevel) → isHomogeneous (coHomK-ptd n) -isHomogeneousKn n = - subst isHomogeneous - (sym (ΣPathP (ua Kn≃ΩKn+1 , ua-gluePath _ (Kn→ΩKn+10ₖ n)))) - (isHomogeneousPath _ _) - --- With the equivalence Kn≃ΩKn+1, we get that the two definitions of cohomology groups agree -open IsGroupHom - -coHom≅coHomΩ : ∀ {ℓ} (n : ℕ) (A : Type ℓ) → GroupIso (coHomGr n A) (coHomGrΩ n A) -fun (fst (coHom≅coHomΩ n A)) = map λ f a → Kn→ΩKn+1 n (f a) -inv' (fst (coHom≅coHomΩ n A)) = map λ f a → ΩKn+1→Kn n (f a) -rightInv (fst (coHom≅coHomΩ n A)) = - sElim (λ _ → isOfHLevelPath 2 § _ _) - λ f → cong ∣_∣₂ (funExt λ x → rightInv (Iso-Kn-ΩKn+1 n) (f x)) -leftInv (fst (coHom≅coHomΩ n A)) = - sElim (λ _ → isOfHLevelPath 2 § _ _) - λ f → cong ∣_∣₂ (funExt λ x → leftInv (Iso-Kn-ΩKn+1 n) (f x)) -snd (coHom≅coHomΩ n A) = - makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) - λ f g → cong ∣_∣₂ (funExt λ x → Kn→ΩKn+1-hom n (f x) (g x))) - -module lockedKnIso (key : Unit') where - Kn→ΩKn+1' : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) - Kn→ΩKn+1' n = lock key (Iso.fun (Iso-Kn-ΩKn+1 n)) - - ΩKn+1→Kn' : (n : ℕ) → typ (Ω (coHomK-ptd (suc n))) → coHomK n - ΩKn+1→Kn' n = lock key (Iso.inv (Iso-Kn-ΩKn+1 n)) - - ΩKn+1→Kn→ΩKn+1 : (n : ℕ) → (x : typ (Ω (coHomK-ptd (suc n)))) → Kn→ΩKn+1' n (ΩKn+1→Kn' n x) ≡ x - ΩKn+1→Kn→ΩKn+1 n x = pm key - where - pm : (key : Unit') → lock key (Iso.fun (Iso-Kn-ΩKn+1 n)) (lock key (Iso.inv (Iso-Kn-ΩKn+1 n)) x) ≡ x - pm unlock = Iso.rightInv (Iso-Kn-ΩKn+1 n) x - - Kn→ΩKn+1→Kn : (n : ℕ) → (x : coHomK n) → ΩKn+1→Kn' n (Kn→ΩKn+1' n x) ≡ x - Kn→ΩKn+1→Kn n x = pm key - where - pm : (key : Unit') → lock key (Iso.inv (Iso-Kn-ΩKn+1 n)) (lock key (Iso.fun (Iso-Kn-ΩKn+1 n)) x) ≡ x - pm unlock = Iso.leftInv (Iso-Kn-ΩKn+1 n) x - --distrLemma : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n m : ℕ) (f : GroupHom (coHomGr n A) (coHomGr m B)) - (x y : coHom n A) - → fst f (x -[ n ]ₕ y) ≡ fst f x -[ m ]ₕ fst f y --distrLemma n m f' x y = sym (-cancelRₕ m (f y) (f (x -[ n ]ₕ y))) - ∙∙ cong (λ x → x -[ m ]ₕ f y) (sym (f' .snd .pres· (x -[ n ]ₕ y) y)) - ∙∙ cong (λ x → x -[ m ]ₕ f y) ( cong f (-+cancelₕ n _ _)) - where - f = fst f' - --- HLevel stuff for cup product -isContr-↓∙ : (n : ℕ) → isContr (coHomK-ptd (suc n) →∙ coHomK-ptd n) -fst (isContr-↓∙ zero) = (λ _ → 0) , refl -snd (isContr-↓∙ zero) (f , p) = - Σ≡Prop (λ f → isSetℤ _ _) - (funExt (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 isSetℤ) _ _) - (toPropElim (λ _ → isSetℤ _ _) (sym p)))) -fst (isContr-↓∙ (suc n)) = (λ _ → 0ₖ _) , refl -fst (snd (isContr-↓∙ (suc n)) f i) x = - trElim {B = λ x → 0ₖ (suc n) ≡ fst f x} - (λ _ → isOfHLevelPath (4 + n) (isOfHLevelSuc (3 + n) (isOfHLevelTrunc (3 + n))) _ _) - (sphereElim _ (λ _ → isOfHLevelTrunc (3 + n) _ _) - (sym (snd f))) x i -snd (snd (isContr-↓∙ (suc n)) f i) j = snd f (~ i ∨ j) - -isContr-↓∙' : (n : ℕ) → isContr (S₊∙ (suc n) →∙ coHomK-ptd n) -fst (isContr-↓∙' zero) = (λ _ → 0) , refl -snd (isContr-↓∙' zero) (f , p) = - Σ≡Prop (λ f → isSetℤ _ _) - (funExt (toPropElim (λ _ → isSetℤ _ _) (sym p))) -fst (isContr-↓∙' (suc n)) = (λ _ → 0ₖ _) , refl -fst (snd (isContr-↓∙' (suc n)) f i) x = - sphereElim _ {A = λ x → 0ₖ (suc n) ≡ fst f x} - (λ _ → isOfHLevelTrunc (3 + n) _ _) - (sym (snd f)) x i -snd (snd (isContr-↓∙' (suc n)) f i) j = snd f (~ i ∨ j) - -isOfHLevel→∙Kn : {A : Pointed₀} (n m : ℕ) - → isOfHLevel (suc m) (A →∙ coHomK-ptd n) → isOfHLevel (suc (suc m)) (A →∙ coHomK-ptd (suc n)) -isOfHLevel→∙Kn {A = A} n m hlev = step₃ - where - step₁ : isOfHLevel (suc m) (A →∙ Ω (coHomK-ptd (suc n))) - step₁ = - subst (isOfHLevel (suc m)) - (λ i → A →∙ ua∙ {A = Ω (coHomK-ptd (suc n))} {B = coHomK-ptd n} - (invEquiv Kn≃ΩKn+1) - (ΩKn+1→Kn-refl n) (~ i)) hlev - - step₂ : isOfHLevel (suc m) (typ (Ω (A →∙ coHomK-ptd (suc n) ∙))) - step₂ = isOfHLevelRetractFromIso (suc m) (invIso (invIso (ΩfunExtIso _ _))) step₁ - - step₃ : isOfHLevel (suc (suc m)) (A →∙ coHomK-ptd (suc n)) - step₃ = - isOfHLevelΩ→isOfHLevel m - λ f → subst (λ x → isOfHLevel (suc m) (typ (Ω x))) - (isHomogeneous→∙ (isHomogeneousKn (suc n)) f) - step₂ - -isOfHLevel↑∙ : ∀ n m → isOfHLevel (2 + n) (coHomK-ptd (suc m) →∙ coHomK-ptd (suc (n + m))) -isOfHLevel↑∙ zero m = isOfHLevel→∙Kn m 0 (isContr→isProp (isContr-↓∙ m)) -isOfHLevel↑∙ (suc n) m = isOfHLevel→∙Kn (suc (n + m)) (suc n) (isOfHLevel↑∙ n m) - -isOfHLevel↑∙' : ∀ n m → isOfHLevel (2 + n) (S₊∙ (suc m) →∙ coHomK-ptd (suc (n + m))) -isOfHLevel↑∙' zero m = isOfHLevel→∙Kn m 0 (isContr→isProp (isContr-↓∙' m)) -isOfHLevel↑∙' (suc n) m = isOfHLevel→∙Kn (suc (n + m)) (suc n) (isOfHLevel↑∙' n m) - -→∙KnPath : (A : Pointed₀) (n : ℕ) → Ω (A →∙ coHomK-ptd (suc n) ∙) ≡ (A →∙ coHomK-ptd n ∙) -→∙KnPath A n = - ua∙ (isoToEquiv (ΩfunExtIso A (coHomK-ptd (suc n)))) refl - ∙ (λ i → (A →∙ Kn≃ΩKn+1∙ {n = n} (~ i) ∙)) - -contr∙-lem : (n m : ℕ) - → isContr (coHomK-ptd (suc n) - →∙ (coHomK-ptd (suc m) - →∙ coHomK-ptd (suc (n + m)) ∙)) -fst (contr∙-lem n m) = (λ _ → (λ _ → 0ₖ _) , refl) , refl -snd (contr∙-lem n m) (f , p) = →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousKn _)) (sym (funExt help)) - where - help : (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) - help = - trElim (λ _ → isOfHLevelPath (3 + n) - (subst (λ x → isOfHLevel x (coHomK-ptd (suc m) →∙ coHomK-ptd (suc (n + m)))) (λ i → suc (suc (+-comm n 1 i))) - (isOfHLevelPlus' {n = 1} (2 + n) (isOfHLevel↑∙ n m))) _ _) - (sphereElim _ (λ _ → isOfHLevel↑∙ n m _ _) p) - -isOfHLevel↑∙∙ : ∀ n m l - → isOfHLevel (2 + l) (coHomK-ptd (suc n) - →∙ (coHomK-ptd (suc m) - →∙ coHomK-ptd (suc (suc (l + n + m))) ∙)) -isOfHLevel↑∙∙ n m zero = - isOfHLevelΩ→isOfHLevel 0 λ f - → subst - isProp (cong (λ x → typ (Ω x)) - (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousKn _)) f)) - (isOfHLevelRetractFromIso 1 (ΩfunExtIso _ _) h) - where - h : isProp (coHomK-ptd (suc n) - →∙ (Ω (coHomK-ptd (suc m) - →∙ coHomK-ptd (suc (suc (n + m))) ∙))) - h = - subst isProp (λ i → coHomK-ptd (suc n) →∙ (→∙KnPath (coHomK-ptd (suc m)) (suc (n + m)) (~ i))) - (isContr→isProp (contr∙-lem n m)) -isOfHLevel↑∙∙ n m (suc l) = - isOfHLevelΩ→isOfHLevel (suc l) - λ f → - subst - (isOfHLevel (2 + l)) (cong (λ x → typ (Ω x)) - (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousKn _)) f)) - (isOfHLevelRetractFromIso (2 + l) (ΩfunExtIso _ _) h) - where - h : isOfHLevel (2 + l) - (coHomK-ptd (suc n) - →∙ (Ω (coHomK-ptd (suc m) - →∙ coHomK-ptd (suc (suc (suc (l + n + m)))) ∙))) - h = - subst (isOfHLevel (2 + l)) - (λ i → coHomK-ptd (suc n) →∙ →∙KnPath (coHomK-ptd (suc m)) (suc (suc (l + n + m))) (~ i)) - (isOfHLevel↑∙∙ n m l) - --- Misc. -isoType→isoCohom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) - → Iso A B - → GroupIso (coHomGr n A) (coHomGr n B) -fst (isoType→isoCohom n is) = setTruncIso (domIso is) -snd (isoType→isoCohom n is) = - makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) - (λ _ _ → refl)) +cup-pure : {!!} +cup-pure = {!!} +open import Cubical.HITs.S1 renaming (_·_ to _·'_) + +S₊3 : S₊ 3 → S₊ 3 → S₊ 3 +S₊3 north y = north +S₊3 south y = north +S₊3 (merid north i) north = north +S₊3 (merid south i) north = north +S₊3 (merid (merid base i) j) north = north +S₊3 (merid (merid (loop i₁) i) j) north = north +S₊3 (merid north i) south = north +S₊3 (merid south i) south = north +S₊3 (merid (merid base j) i) south = north +S₊3 (merid (merid (loop k) i) j) south = {!!} +S₊3 (merid north i) (merid north j) = north +S₊3 (merid north i) (merid south j) = north +S₊3 (merid north i) (merid (merid a i₁) j) = {!!} +S₊3 (merid south i) (merid north j) = north +S₊3 (merid south i) (merid south j) = {!north!} +S₊3 (merid south i) (merid (merid a i₁) j) = {!!} +S₊3 (merid (merid a i₁) i) (merid north j) = {!!} +S₊3 (merid (merid a i₁) i) (merid south j) = {!!} +S₊3 (merid (merid a i) j) (merid (merid a₁ k) l) = {!!} + +-- open import Cubical.HITs.Join +-- rofl : join S¹ S¹ → Type _ +-- rofl (inl x) = {!!} +-- rofl (inr x) = {!!} +-- rofl (push a b i) = {!!} +-- where +-- zz : (x : S₊ 2) → hLevelTrunc 5 (S₊ 3) → hLevelTrunc 5 (S₊ 3) +-- zz x = {!trElim ? ?!} + +-- r4 : hLevelTrunc 7 {!!} → {!!} → {!!} +-- r4 = {!!} + +-- r3 : hLevelTrunc 6 (S₊ 2) → S₊ 4 → hLevelTrunc 6 (S₊ 2) +-- r3 = trRec (isOfHLevelΠ 6 (λ _ → isOfHLevelTrunc 6)) +-- (wedgeconFun 1 3 (λ _ _ → isOfHLevelTrunc 6) +-- {!!} ∣_∣ +-- {!!}) + +-- t : hLevelTrunc 6 (S₊ 3) → hLevelTrunc 5 (S₊ 2) → hLevelTrunc 6 (S₊ 3) +-- t = trRec {!!} {!!} +-- where +-- pp : S₊ 3 → hLevelTrunc 5 (S₊ 2) → hLevelTrunc 6 (S₊ 3) +-- pp north = λ _ → ∣ north ∣ +-- pp south = λ _ → ∣ north ∣ +-- pp (merid a i) x = {!!} +-- where +-- h : (x : hLevelTrunc 5 (S₊ 2)) → Path (hLevelTrunc 6 (S₊ 3)) ∣ north ∣ ∣ north ∣ +-- h = trElim (λ _ → isOfHLevelTrunc 6 _ _) +-- λ b → {!!} + +-- S³∥ : Type _ +-- S³∥ = hLevelTrunc 6 (S₊ 3) + +-- _+3_ : S³∥ → S³∥ → S³∥ +-- _+3_ = elim2 (λ _ _ → isOfHLevelTrunc 6) (wedgeconFun _ _ (λ _ _ → isOfHLevelTrunc 6) +-- ∣_∣ +-- ∣_∣ +-- refl) + +-- 0₃ : hLevelTrunc 6 (S₊ 3) +-- 0₃ = ∣ north ∣ + +-- rUnit3 : ∀ x → x +3 0₃ ≡ x +-- rUnit3 = trElim (λ _ → isOfHLevelPath 6 (isOfHLevelTrunc 6) _ _) +-- (wedgeconRight 2 2 {λ _ _ → S³∥} +-- (λ _ _ → isOfHLevelTrunc 6) ∣_∣ₕ ∣_∣ₕ refl) + +-- lUnit3 : ∀ x → 0₃ +3 x ≡ x +-- lUnit3 = trElim (λ _ → isOfHLevelPath 6 (isOfHLevelTrunc 6) _ _) +-- λ _ → refl + +-- min3 : hLevelTrunc 6 (S₊ 3) → hLevelTrunc 6 (S₊ 3) +-- min3 = trMap m +-- where +-- m : S₊ 3 → S₊ 3 +-- m north = north +-- m south = north +-- m (merid a i) = (merid a ∙ sym (merid north)) i + +-- comm3 : ∀ x y → x +3 y ≡ y +3 x +-- comm3 = elim2 (λ _ _ → isOfHLevelPath 6 (isOfHLevelTrunc 6) _ _) +-- (wedgeconFun 2 2 +-- (λ _ _ → isOfHLevelPath 6 (isOfHLevelTrunc 6) _ _) +-- (sym ∘ rUnit3 ∘ ∣_∣ₕ) +-- (rUnit3 ∘ ∣_∣ₕ) +-- refl) + +-- rCancel3 : (x : hLevelTrunc 6 (S₊ 3)) → x +3 min3 x ≡ 0₃ +-- rCancel3 = +-- trElim {!!} +-- λ { north → refl +-- ; south → refl +-- ; (merid a i) → {!!}} +-- where +-- p+ : ∀ {x : S³∥} (p : 0₃ ≡ x) → cong₂ _+3_ p (sym p) ≡ lUnit3 x ∙ sym (rUnit3 x) +-- p+ = J (λ x p → cong₂ _+3_ p (sym p) ≡ lUnit3 x ∙ sym (rUnit3 x)) (rUnit refl) + +-- p+merid : ∀ (x : S₊ 2) → cong₂ _+3_ (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid x))) +-- ≡ {!!} +-- p+merid = {!!} + +-- π-iso : Iso (∥ (S₊∙ 4 →∙ (S³∥ , ∣ north ∣)) ∥₂) (∥ (S₊ 4 → S³∥) ∥₂) +-- fun π-iso = map fst +-- inv' π-iso = map λ f → (λ x → (f x) +3 min3 (f north)) , rCancel3 (f north) +-- rightInv π-iso = sElim {!!} {!!} +-- leftInv π-iso = {!!} + +-- p2 : Iso ∥ (S₊ 3 → typ (Ω (S³∥ , ∣ north ∣))) ∥₂ +-- ∥ (S₊ 2 → typ ((Ω^ 2) (S³∥ , ∣ north ∣))) ∥₂ +-- p2 = {!!} + +-- data RP2 : Type ℓ-zero where +-- b : RP2 +-- l1 l2 : b ≡ b +-- sq : l1 ≡ sym l2 + +-- rar : S₊ 2 → RP2 +-- rar north = b +-- rar south = b +-- rar (merid base i) = (l1 ∙ l2) i +-- rar (merid (loop j) i) = +-- hcomp (λ k → λ { (i = i0) → b +-- ; (i = i1) → b +-- ; (j = i0) → (sym (lCancel l2) ∙ cong (_∙ l2) (sym sq)) k i +-- ; (j = i1) → (sym (lCancel l2) ∙ cong (_∙ l2) (sym sq)) k i}) +-- b + + + +-- jmap : join S¹ S¹ → S₊ 2 +-- jmap (inl x) = north +-- jmap (inr x) = south +-- jmap (push a b₁ i) = merid (a ·' b₁) i + +-- gl : Susp (join S¹ S¹) → S₊ 3 +-- gl north = north +-- gl south = south +-- gl (merid a i) = merid (jmap a) i + +-- _+j_ : (Susp (join S¹ S¹) , north) →∙ S₊∙ 3 +-- → (Susp (join S¹ S¹) , north) →∙ S₊∙ 3 +-- → (Susp (join S¹ S¹) , north) →∙ S₊∙ 3 +-- fst (f +j g) north = north +-- fst (f +j g) south = north +-- fst (f +j g) (merid a i) = +-- ((sym (snd f) ∙∙ cong (fst f) (merid a ∙ sym (merid (inl base))) ∙∙ snd f) +-- ∙ ((sym (snd g) ∙∙ cong (fst g) (merid a ∙ sym (merid (inl base))) ∙∙ snd g))) i +-- snd (f +j g) = refl + +-- ss123 : {!!} +-- ss123 = {!!} + +-- thegen : {!(Susp (join S¹ S¹) , north) →∙ S₊∙ 3!} +-- thegen = {!!} + +-- thegen2 : (Susp (join S¹ S¹) , north) →∙ (S³∥ , ∣ north ∣) +-- fst thegen2 north = ∣ north ∣ +-- fst thegen2 south = ∣ north ∣ +-- fst thegen2 (merid a i) = +-- ∣ (cong gl (merid a ∙ sym (merid (inl base))) +-- ∙ cong gl (merid a ∙ sym (merid (inl base)))) i ∣ +-- snd thegen2 = refl + +-- s : join S¹ S¹ → {!!} +-- s = {!!} + +-- rs : (x : _) → fst thegen2 x ≡ ∣ north ∣ +-- rs north = refl +-- rs south = refl +-- rs (merid a i) j = pp a j i +-- where +-- open import Cubical.HITs.Wedge +-- wedge→gpd : ∀ {ℓ} (P : (S₊∙ 1) ⋁ S₊∙ 1 → Type ℓ) +-- → ((s : (S₊∙ 1) ⋁ S₊∙ 1) → isGroupoid (P s)) +-- → {!!} +-- → {!!} +-- wedge→gpd = {!!} + +-- pp3 : (a : join S¹ S¹) → +-- cong (∣_∣ₕ {n = 6}) (cong gl (merid a ∙ sym (merid (inl base)))) ≡ +-- sym (cong (∣_∣ₕ {n = 6}) (cong gl (merid a ∙ sym (merid (inl base))))) -- 4 +-- pp3 (inl x) = {!!} +-- pp3 (inr x) = {!!} +-- pp3 (push a b₁ i) = {!!} -- 3 + +-- pp4 : (a : join S¹ S¹) → cong (∣_∣ₕ {n = 6}) (merid (jmap a) ∙ sym (merid north)) +-- ≡ sym (cong (∣_∣ₕ {n = 6}) (merid (jmap a) ∙ sym (merid north))) +-- -- 4-type +-- pp4 (inl x) = cong (cong (∣_∣ₕ {n = 6})) (rCancel (merid north)) +-- ∙ cong sym (cong (cong (∣_∣ₕ {n = 6})) (sym (rCancel (merid north)))) +-- pp4 (inr x) = (cong (cong (∣_∣ₕ {n = 6})) λ i → merid (merid base (~ i)) ∙ sym (merid north)) +-- ∙∙ (cong (cong (∣_∣ₕ {n = 6})) (rCancel (merid north)) +-- ∙ cong sym (cong (cong (∣_∣ₕ {n = 6})) (sym (rCancel (merid north))))) +-- ∙∙ cong sym (cong (cong (∣_∣ₕ {n = 6})) +-- λ i → merid (merid base i) ∙ sym (merid north)) +-- -- 3-type +-- pp4 (push a c i) j k = {!!} +-- where +-- c' : ∥ cong merid (merid a) ≡ cong merid (merid base) ∥ 1 +-- c' = isConnectedPath _ (isConnectedPath 2 (isConnectedPath 3 (sphereConnected 3) _ _) _ _) _ _ .fst + +-- pr : (a c : S¹) → cong merid (merid (a ·' c)) ≡ cong merid (merid a ∙ sym (merid c)) +-- ∙ (cong merid (merid base)) +-- pr base base = {!!} +-- pr base (loop i) = {!!} +-- pr (loop i) base = {!cong ∣_∣ₕ (merid (cong jmap (push a c)) ∙ sym (merid north))!} +-- pr (loop i) (loop j) = {!!} + +-- hej : Cube (λ j → cong ∣_∣ₕ (merid (merid (a ·' c) j) ∙ sym (merid north))) -- i j k +-- (λ j → cong ∣_∣ₕ (merid (merid a j) ∙ sym (merid (merid c j)))) +-- {!!} +-- {!!} +-- {!!} +-- {!!} +-- hej = {!!} +-- pp : (a : join S¹ S¹) → +-- cong (∣_∣ₕ {n = 6}) ((cong gl (merid a ∙ sym (merid (inl base))) +-- ∙ cong gl (merid a ∙ sym (merid (inl base))))) ≡ refl +-- pp a = {!merid a!} ∙ {!!} +-- +gen : ℤ → (Susp (join S¹ S¹) , north) →∙ S₊∙ 3 +-- +gen (pos zero) = (λ x → north) , refl +-- fst (+gen (pos (suc n))) north = north +-- fst (+gen (pos (suc n))) south = north +-- fst (+gen (pos (suc n))) (merid a i) = +-- {!cong (fst (+gen (pos n))) (merid a ∙ sym (merid (inl base)))!} +-- snd (+gen (pos (suc n))) = {!!} +-- +gen (negsuc zero) = {!!} +-- +gen (negsuc (suc n)) = {!!} + +-- _+∙_ : {!!} -- {n m : ℕ} → (S₊∙ n → S₊∙ m) → (S₊∙ n → S₊∙ m) → (S₊∙ n → S₊∙ m) +-- _+∙_ = {!!} + +-- sn : Iso ∥ (typ ((Ω^ 4) (S₊∙ 3))) ∥₂ +-- (∥ (Σ[ x ∈ typ ((Ω^ 3) (S₊∙ 3)) ] +-- x ≡ sym x) ∥₂) +-- fun sn = map λ p → (λ i j k → p (i ∨ j) (j ∨ k) (k ∨ i) (~ k ∨ ~ i ∨ ~ j)) +-- , {!!} +-- inv' sn = {!!} +-- rightInv sn = {!!} +-- leftInv sn = {!!} + +-- gaza : ∥ {!!} ≃ {!!} ∥₂ +-- gaza = {!!} + +-- Σ- : ∀ {ℓ} {A : Type ℓ} → Susp A → Susp A +-- Σ- north = south +-- Σ- south = north +-- Σ- (merid a i) = merid a (~ i) + +-- rr : Iso (∥ (Σ[ f ∈ S₊∙ 3 →∙ (S³∥ , ∣ north ∣) ] ((x : S₊ 3) → fst f x ≡ fst f x)) ∥₂) +-- (∥ (Σ[ f ∈ S₊∙ 3 →∙ coHomK-ptd 3 ] ((x : S₊ 3) → fst f x +ₖ fst f x ≡ 0ₖ 3)) ∥₂) +-- fun rr = {!sMap ?!} +-- inv' rr = {!!} +-- rightInv rr = {!!} +-- leftInv rr = {!!} + +-- open import Cubical.HITs.Pushout +-- RP3 : Type ℓ-zero +-- RP3 = Pushout rar λ _ → tt + +-- RP4 : S₊ 3 → RP3 +-- RP4 north = inl b +-- RP4 south = inr tt +-- RP4 (merid north i) = push north i +-- RP4 (merid south i) = push south i +-- RP4 (merid (merid base j) i) = push {!rar!} i +-- RP4 (merid (merid (loop k) j) i) = {!!} + +-- ll : hLevelTrunc 5 (S₊ 2) → hLevelTrunc 5 (S₊ 2) → {!!} +-- ll = {!!} + +-- ss : Iso ∥ (S₊ 1 → (S₊∙ 1 →∙ ((Ω^ 2) (S³∥ , ∣ north ∣)))) ∥₂ +-- ∥ ((S₊ 1 → (S₊ 1 → typ ((Ω^ 2) (S³∥ , ∣ north ∣))))) ∥₂ +-- fun ss = map λ f → λ x y → f x .fst y +-- inv' ss = map λ f x → {!f!} +-- rightInv ss = {!!} +-- leftInv ss = {!!} + +-- p3 : Iso ∥ (S₊ 2 → coHomK-ptd 1 →∙ (coHomK-ptd 1 →∙ S³∥ , ∣ north ∣ ∙)) ∥₂ +-- ∥ {!!} ∥₂ +-- p3 = {!S₊∙ 3 !} + +-- cong2 : {!!} +-- cong2 = {!!} + +-- -- ------------------- Connectedness --------------------- +-- -- is2ConnectedKn : (n : ℕ) → isConnected 2 (coHomK (suc n)) +-- -- is2ConnectedKn zero = ∣ ∣ base ∣ ∣ +-- -- , trElim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) +-- -- (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 (isOfHLevelTrunc 2)) _ _) +-- -- (toPropElim (λ _ → isOfHLevelTrunc 2 _ _) refl)) +-- -- is2ConnectedKn (suc n) = ∣ ∣ north ∣ ∣ +-- -- , trElim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) +-- -- (trElim (λ _ → isProp→isOfHLevelSuc (3 + n) (isOfHLevelTrunc 2 _ _)) +-- -- (suspToPropElim (ptSn (suc n)) (λ _ → isOfHLevelTrunc 2 _ _) refl)) + +-- -- isConnectedKn : (n : ℕ) → isConnected (2 + n) (coHomK (suc n)) +-- -- isConnectedKn n = isOfHLevelRetractFromIso 0 (invIso (truncOfTruncIso (2 + n) 1)) (sphereConnected (suc n)) + +-- -- -- direct proof of connectedness of ΩKₙ₊₁ not relying on the equivalence ∥ a ≡ b ∥ₙ ≃ (∣ a ∣ₙ₊₁ ≡ ∣ b ∣ₙ₊₁) +-- -- isConnectedPathKn : (n : ℕ) (x y : (coHomK (suc n))) → isConnected (suc n) (x ≡ y) +-- -- isConnectedPathKn n = +-- -- trElim (λ _ → isProp→isOfHLevelSuc (2 + n) (isPropΠ λ _ → isPropIsContr)) +-- -- (sphereElim _ (λ _ → isProp→isOfHLevelSuc n (isPropΠ λ _ → isPropIsContr)) +-- -- λ y → isContrRetractOfConstFun +-- -- {B = (hLevelTrunc (suc n) (ptSn (suc n) ≡ ptSn (suc n)))} ∣ refl ∣ +-- -- (fun⁻ n y +-- -- , trElim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) +-- -- (J (λ y p → fun⁻ n y _ ≡ _) (funExt⁻ (fun⁻Id n) ∣ refl ∣)))) +-- -- where +-- -- fun⁻ : (n : ℕ) → (y : coHomK (suc n)) → +-- -- hLevelTrunc (suc n) (ptSn (suc n) ≡ ptSn (suc n)) +-- -- → hLevelTrunc (suc n) (∣ ptSn (suc n) ∣ ≡ y) +-- -- fun⁻ n = +-- -- trElim (λ _ → isOfHLevelΠ (3 + n) λ _ → isOfHLevelSuc (2 + n) (isOfHLevelSuc (suc n) (isOfHLevelTrunc (suc n)))) +-- -- (sphereElim n (λ _ → isOfHLevelΠ (suc n) λ _ → isOfHLevelTrunc (suc n)) λ _ → ∣ refl ∣) + +-- -- fun⁻Id : (n : ℕ) → fun⁻ n ∣ ptSn (suc n) ∣ ≡ λ _ → ∣ refl ∣ +-- -- fun⁻Id zero = refl +-- -- fun⁻Id (suc n) = refl + +-- -- ------------------- +-- -- -- Induction principles for cohomology groups (n ≥ 1) +-- -- -- If we want to show a proposition about some x : Hⁿ(A), it suffices to show it under the +-- -- -- assumption that x = ∣ f ∣₂ for some f : A → Kₙ and that f is pointed + +-- -- coHomPointedElim : {A : Type ℓ} (n : ℕ) (a : A) {B : coHom (suc n) A → Type ℓ'} +-- -- → ((x : coHom (suc n) A) → isProp (B x)) +-- -- → ((f : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → B ∣ f ∣₂) +-- -- → (x : coHom (suc n) A) → B x +-- -- coHomPointedElim {ℓ' = ℓ'} {A = A} n a isprop indp = +-- -- sElim (λ _ → isOfHLevelSuc 1 (isprop _)) +-- -- λ f → helper n isprop indp f (f a) refl +-- -- where +-- -- helper : (n : ℕ) {B : coHom (suc n) A → Type ℓ'} +-- -- → ((x : coHom (suc n) A) → isProp (B x)) +-- -- → ((f : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → B ∣ f ∣₂) +-- -- → (f : A → coHomK (suc n)) +-- -- → (x : coHomK (suc n)) +-- -- → f a ≡ x → B ∣ f ∣₂ +-- -- -- pattern matching a bit extra to avoid isOfHLevelPlus' +-- -- helper zero isprop ind f = +-- -- trElim (λ _ → isOfHLevelPlus {n = 1} 2 (isPropΠ λ _ → isprop _)) +-- -- (toPropElim (λ _ → isPropΠ λ _ → isprop _) (ind f)) +-- -- helper (suc zero) isprop ind f = +-- -- trElim (λ _ → isOfHLevelPlus {n = 1} 3 (isPropΠ λ _ → isprop _)) +-- -- (suspToPropElim base (λ _ → isPropΠ λ _ → isprop _) (ind f)) +-- -- helper (suc (suc zero)) isprop ind f = +-- -- trElim (λ _ → isOfHLevelPlus {n = 1} 4 (isPropΠ λ _ → isprop _)) +-- -- (suspToPropElim north (λ _ → isPropΠ λ _ → isprop _) (ind f)) +-- -- helper (suc (suc (suc n))) isprop ind f = +-- -- trElim (λ _ → isOfHLevelPlus' {n = 5 + n} 1 (isPropΠ λ _ → isprop _)) +-- -- (suspToPropElim north (λ _ → isPropΠ λ _ → isprop _) (ind f)) + + +-- -- coHomPointedElim2 : {A : Type ℓ} (n : ℕ) (a : A) {B : coHom (suc n) A → coHom (suc n) A → Type ℓ'} +-- -- → ((x y : coHom (suc n) A) → isProp (B x y)) +-- -- → ((f g : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → g a ≡ coHom-pt (suc n) → B ∣ f ∣₂ ∣ g ∣₂) +-- -- → (x y : coHom (suc n) A) → B x y +-- -- coHomPointedElim2 {ℓ' = ℓ'} {A = A} n a isprop indp = sElim2 (λ _ _ → isOfHLevelSuc 1 (isprop _ _)) +-- -- λ f g → helper n a isprop indp f g (f a) (g a) refl refl +-- -- where +-- -- helper : (n : ℕ) (a : A) {B : coHom (suc n) A → coHom (suc n) A → Type ℓ'} +-- -- → ((x y : coHom (suc n) A) → isProp (B x y)) +-- -- → ((f g : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → g a ≡ coHom-pt (suc n) → B ∣ f ∣₂ ∣ g ∣₂) +-- -- → (f g : A → coHomK (suc n)) +-- -- → (x y : coHomK (suc n)) +-- -- → f a ≡ x → g a ≡ y +-- -- → B ∣ f ∣₂ ∣ g ∣₂ +-- -- helper zero a isprop indp f g = +-- -- elim2 (λ _ _ → isOfHLevelPlus {n = 1} 2 (isPropΠ2 λ _ _ → isprop _ _)) +-- -- (toPropElim2 (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) +-- -- helper (suc zero) a isprop indp f g = +-- -- elim2 (λ _ _ → isOfHLevelPlus {n = 1} 3 (isPropΠ2 λ _ _ → isprop _ _)) +-- -- (suspToPropElim2 base (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) +-- -- helper (suc (suc zero)) a isprop indp f g = +-- -- elim2 (λ _ _ → isOfHLevelPlus {n = 1} 4 (isPropΠ2 λ _ _ → isprop _ _)) +-- -- (suspToPropElim2 north (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) +-- -- helper (suc (suc (suc n))) a isprop indp f g = +-- -- elim2 (λ _ _ → isOfHLevelPlus' {n = 5 + n} 1 (isPropΠ2 λ _ _ → isprop _ _)) +-- -- (suspToPropElim2 north (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) + +-- -- coHomK-elim : ∀ {ℓ} (n : ℕ) {B : coHomK (suc n) → Type ℓ} +-- -- → ((x : _) → isOfHLevel (suc n) (B x)) +-- -- → B (0ₖ (suc n)) +-- -- → (x : _) → B x +-- -- coHomK-elim n {B = B } hlev b = +-- -- trElim (λ _ → isOfHLevelPlus {n = (suc n)} 2 (hlev _)) +-- -- (sphereElim _ (hlev ∘ ∣_∣) b) + +-- -- {- Equivalence between cohomology of A and reduced cohomology of (A + 1) -} +-- -- coHomRed+1Equiv : (n : ℕ) → +-- -- (A : Type ℓ) → +-- -- (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) +-- -- coHomRed+1Equiv zero A i = ∥ helpLemma {C = (ℤ , pos 0)} i ∥₂ +-- -- module coHomRed+1 where +-- -- helpLemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →∙ C))) +-- -- helpLemma {C = C} = isoToPath (iso map1 +-- -- map2 +-- -- (λ b → linvPf b) +-- -- (λ _ → refl)) +-- -- where +-- -- map1 : (A → typ C) → ((((A ⊎ Unit) , inr (tt)) →∙ C)) +-- -- map1 f = map1' , refl +-- -- module helpmap where +-- -- map1' : A ⊎ Unit → fst C +-- -- map1' (inl x) = f x +-- -- map1' (inr x) = pt C + +-- -- map2 : ((((A ⊎ Unit) , inr (tt)) →∙ C)) → (A → typ C) +-- -- map2 (g , pf) x = g (inl x) + +-- -- linvPf : (b :((((A ⊎ Unit) , inr (tt)) →∙ C))) → map1 (map2 b) ≡ b +-- -- linvPf (f , snd) i = (λ x → helper x i) , λ j → snd ((~ i) ∨ j) +-- -- where +-- -- helper : (x : A ⊎ Unit) → ((helpmap.map1') (map2 (f , snd)) x) ≡ f x +-- -- helper (inl x) = refl +-- -- helper (inr tt) = sym snd +-- -- coHomRed+1Equiv (suc zero) A i = ∥ coHomRed+1.helpLemma A i {C = (coHomK 1 , ∣ base ∣)} i ∥₂ +-- -- coHomRed+1Equiv (suc (suc n)) A i = ∥ coHomRed+1.helpLemma A i {C = (coHomK (2 + n) , ∣ north ∣)} i ∥₂ + +-- -- Iso-coHom-coHomRed : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → Iso (coHomRed (suc n) A) (coHom (suc n) (typ A)) +-- -- fun (Iso-coHom-coHomRed {A = A , a} n) = map fst +-- -- inv' (Iso-coHom-coHomRed {A = A , a} n) = map λ f → (λ x → f x -ₖ f a) , rCancelₖ _ _ +-- -- rightInv (Iso-coHom-coHomRed {A = A , a} n) = +-- -- sElim (λ _ → isOfHLevelPath 2 § _ _) +-- -- λ f → trRec (isProp→isOfHLevelSuc _ (§ _ _)) +-- -- (λ p → cong ∣_∣₂ (funExt λ x → cong (λ y → f x +ₖ y) (cong -ₖ_ p ∙ -0ₖ) ∙ rUnitₖ _ (f x))) +-- -- (Iso.fun (PathIdTruncIso (suc n)) (isContr→isProp (isConnectedKn n) ∣ f a ∣ ∣ 0ₖ _ ∣)) +-- -- leftInv (Iso-coHom-coHomRed {A = A , a} n) = +-- -- sElim (λ _ → isOfHLevelPath 2 § _ _) +-- -- λ {(f , p) → cong ∣_∣₂ (ΣPathP (((funExt λ x → (cong (λ y → f x -ₖ y) p +-- -- ∙∙ cong (λ y → f x +ₖ y) -0ₖ +-- -- ∙∙ rUnitₖ _ (f x)) ∙ refl)) +-- -- , helper n (f a) (sym p)))} +-- -- where +-- -- path : (n : ℕ) (x : coHomK (suc n)) (p : 0ₖ _ ≡ x) → _ +-- -- path n x p = (cong (λ y → x -ₖ y) (sym p) ∙∙ cong (λ y → x +ₖ y) -0ₖ ∙∙ rUnitₖ _ x) ∙ refl + +-- -- helper : (n : ℕ) (x : coHomK (suc n)) (p : 0ₖ _ ≡ x) +-- -- → PathP (λ i → path n x p i ≡ 0ₖ _) (rCancelₖ _ x) (sym p) +-- -- helper zero x = +-- -- J (λ x p → PathP (λ i → path 0 x p i ≡ 0ₖ _) +-- -- (rCancelₖ _ x) (sym p)) +-- -- λ i j → rUnit (rUnit (λ _ → 0ₖ 1) (~ j)) (~ j) i +-- -- helper (suc n) x = +-- -- J (λ x p → PathP (λ i → path (suc n) x p i ≡ 0ₖ _) (rCancelₖ _ x) (sym p)) +-- -- λ i j → rCancelₖ (suc (suc n)) (0ₖ (suc (suc n))) (~ i ∧ ~ j) + +-- -- +∙≡+ : (n : ℕ) {A : Pointed ℓ} (x y : coHomRed (suc n) A) +-- -- → Iso.fun (Iso-coHom-coHomRed n) (x +ₕ∙ y) +-- -- ≡ Iso.fun (Iso-coHom-coHomRed n) x +ₕ Iso.fun (Iso-coHom-coHomRed n) y + +-- -- +∙≡+ zero = sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl +-- -- +∙≡+ (suc n) = sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl + +-- -- private +-- -- homhelp : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) (x y : coHom (suc n) (typ A)) +-- -- → Iso.inv (Iso-coHom-coHomRed {A = A} n) (x +ₕ y) +-- -- ≡ Iso.inv (Iso-coHom-coHomRed n) x +ₕ∙ Iso.inv (Iso-coHom-coHomRed n) y +-- -- homhelp n A = morphLemmas.isMorphInv _+ₕ∙_ _+ₕ_ +-- -- (Iso.fun (Iso-coHom-coHomRed n)) (+∙≡+ n) _ +-- -- (Iso.rightInv (Iso-coHom-coHomRed n)) (Iso.leftInv (Iso-coHom-coHomRed n)) + +-- -- coHomGr≅coHomRedGr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) +-- -- → GroupEquiv (coHomRedGrDir (suc n) A) (coHomGr (suc n) (typ A)) +-- -- fst (coHomGr≅coHomRedGr n A) = isoToEquiv (Iso-coHom-coHomRed n) +-- -- snd (coHomGr≅coHomRedGr n A) = makeIsGroupHom (+∙≡+ n) + +-- -- coHomRedGroup : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → AbGroup ℓ +-- -- coHomRedGroup zero A = coHomRedGroupDir zero A +-- -- coHomRedGroup (suc n) A = +-- -- InducedAbGroup (coHomGroup (suc n) (typ A)) +-- -- _+ₕ∙_ +-- -- (isoToEquiv (invIso (Iso-coHom-coHomRed n))) +-- -- (homhelp n A) + +-- -- abstract +-- -- coHomGroup≡coHomRedGroup : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) +-- -- → coHomGroup (suc n) (typ A) ≡ coHomRedGroup (suc n) A +-- -- coHomGroup≡coHomRedGroup n A = +-- -- InducedAbGroupPath (coHomGroup (suc n) (typ A)) +-- -- _+ₕ∙_ +-- -- (isoToEquiv (invIso (Iso-coHom-coHomRed n))) +-- -- (homhelp n A) + +-- -- ------------------- Kₙ ≃ ΩKₙ₊₁ --------------------- +-- -- -- This proof uses the encode-decode method rather than Freudenthal + +-- -- -- We define the map σ : Kₙ → ΩKₙ₊₁ and prove that it is a morphism +-- -- private +-- -- module _ (n : ℕ) where +-- -- σ : {n : ℕ} → coHomK (suc n) → Path (coHomK (2 + n)) ∣ north ∣ ∣ north ∣ +-- -- σ {n = n} = trRec (isOfHLevelTrunc (4 + n) _ _) +-- -- λ a → cong ∣_∣ (merid a ∙ sym (merid (ptSn (suc n)))) + +-- -- σ-hom-helper : ∀ {ℓ} {A : Type ℓ} {a : A} (p : a ≡ a) (r : refl ≡ p) +-- -- → lUnit p ∙ cong (_∙ p) r ≡ rUnit p ∙ cong (p ∙_) r +-- -- σ-hom-helper p = J (λ p r → lUnit p ∙ cong (_∙ p) r ≡ rUnit p ∙ cong (p ∙_) r) refl + +-- -- σ-hom : {n : ℕ} (x y : coHomK (suc n)) → σ (x +ₖ y) ≡ σ x ∙ σ y +-- -- σ-hom {n = zero} = +-- -- elim2 (λ _ _ → isOfHLevelPath 3 (isOfHLevelTrunc 4 _ _) _ _) +-- -- (wedgeconFun _ _ +-- -- (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) +-- -- (λ x → lUnit _ +-- -- ∙ cong (_∙ σ ∣ x ∣) (cong (cong ∣_∣) (sym (rCancel (merid base))))) +-- -- (λ y → cong σ (rUnitₖ 1 ∣ y ∣) +-- -- ∙∙ rUnit _ +-- -- ∙∙ cong (σ ∣ y ∣ ∙_) (cong (cong ∣_∣) (sym (rCancel (merid base))))) +-- -- (sym (σ-hom-helper (σ ∣ base ∣) (cong (cong ∣_∣) (sym (rCancel (merid base))))))) +-- -- σ-hom {n = suc n} = +-- -- elim2 (λ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (5 + n) _ _) _ _) +-- -- (wedgeconFun _ _ (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (wedgeConHLev' n) _ _) +-- -- (λ x → lUnit _ +-- -- ∙ cong (_∙ σ ∣ x ∣) (cong (cong ∣_∣) (sym (rCancel (merid north))))) +-- -- (λ y → cong σ (rUnitₖ (2 + n) ∣ y ∣) +-- -- ∙∙ rUnit _ +-- -- ∙∙ cong (σ ∣ y ∣ ∙_) (cong (cong ∣_∣) (sym (rCancel (merid north))))) +-- -- (sym (σ-hom-helper (σ ∣ north ∣) (cong (cong ∣_∣) (sym (rCancel (merid north))))))) + +-- -- -- We will need to following lemma +-- -- σ-minusDistr : {n : ℕ} (x y : coHomK (suc n)) → σ (x -ₖ y) ≡ σ x ∙ sym (σ y) +-- -- σ-minusDistr {n = n} = +-- -- morphLemmas.distrMinus' +-- -- _+ₖ_ _∙_ +-- -- σ σ-hom ∣ (ptSn (suc n)) ∣ refl +-- -- -ₖ_ sym +-- -- (λ x → sym (lUnit x)) (λ x → sym (rUnit x)) +-- -- (rUnitₖ (suc n)) +-- -- (lCancelₖ (suc n)) rCancel +-- -- (assocₖ (suc n)) assoc∙ +-- -- (cong (cong ∣_∣) (rCancel (merid (ptSn (suc n))))) + +-- -- -- we define the code using addIso +-- -- Code : (n : ℕ) → coHomK (2 + n) → Type₀ +-- -- Code n x = (trRec {B = TypeOfHLevel ℓ-zero (3 + n)} (isOfHLevelTypeOfHLevel (3 + n)) +-- -- λ a → Code' a , hLevCode' a) x .fst +-- -- where +-- -- Code' : (S₊ (2 + n)) → Type₀ +-- -- Code' north = coHomK (suc n) +-- -- Code' south = coHomK (suc n) +-- -- Code' (merid a i) = isoToPath (addIso (suc n) ∣ a ∣) i + +-- -- hLevCode' : (x : S₊ (2 + n)) → isOfHLevel (3 + n) (Code' x) +-- -- hLevCode' = suspToPropElim (ptSn (suc n)) (λ _ → isPropIsOfHLevel (3 + n)) (isOfHLevelTrunc (3 + n)) + +-- -- symMeridLem : (n : ℕ) → (x : S₊ (suc n)) (y : coHomK (suc n)) +-- -- → subst (Code n) (cong ∣_∣ (sym (merid x))) y ≡ y -ₖ ∣ x ∣ +-- -- symMeridLem n x = trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) +-- -- (λ y → cong (_-ₖ ∣ x ∣) (transportRefl ∣ y ∣)) + +-- -- decode : {n : ℕ} (x : coHomK (2 + n)) → Code n x → ∣ north ∣ ≡ x +-- -- decode {n = n} = trElim (λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) +-- -- decode-elim +-- -- where +-- -- north≡merid : (a : S₊ (suc n)) +-- -- → Path (coHomK (2 + n)) ∣ north ∣ ∣ north ∣ +-- -- ≡ (Path (coHomK (2 + n)) ∣ north ∣ ∣ south ∣) +-- -- north≡merid a i = Path (coHomK (2 + n)) ∣ north ∣ ∣ merid a i ∣ + +-- -- decode-elim : (a : S₊ (2 + n)) → Code n ∣ a ∣ → Path (coHomK (2 + n)) ∣ north ∣ ∣ a ∣ +-- -- decode-elim north = σ +-- -- decode-elim south = trRec (isOfHLevelTrunc (4 + n) _ _) +-- -- λ a → cong ∣_∣ (merid a) +-- -- decode-elim (merid a i) = +-- -- hcomp (λ k → λ { (i = i0) → σ +-- -- ; (i = i1) → mainPath a k}) +-- -- (funTypeTransp (Code n) (λ x → ∣ north ∣ ≡ x) (cong ∣_∣ (merid a)) σ i) +-- -- where +-- -- mainPath : (a : (S₊ (suc n))) → +-- -- transport (north≡merid a) ∘ σ ∘ transport (λ i → Code n ∣ merid a (~ i) ∣) +-- -- ≡ trRec (isOfHLevelTrunc (4 + n) _ _) λ a → cong ∣_∣ (merid a) +-- -- mainPath a = funExt (trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (4 + n) _ _) _ _) +-- -- (λ x → (λ i → transport (north≡merid a) (σ (symMeridLem n a ∣ x ∣ i))) +-- -- ∙∙ cong (transport (north≡merid a)) (-distrHelp x) +-- -- ∙∙ (substAbove x))) +-- -- where +-- -- -distrHelp : (x : S₊ (suc n)) → σ (∣ x ∣ -ₖ ∣ a ∣) ≡ cong ∣_∣ (merid x) ∙ cong ∣_∣ (sym (merid a)) +-- -- -distrHelp x = +-- -- σ-minusDistr ∣ x ∣ ∣ a ∣ +-- -- ∙ (λ i → (cong ∣_∣ (compPath-filler (merid x) (λ j → merid (ptSn (suc n)) (~ j ∨ i)) (~ i))) +-- -- ∙ (cong ∣_∣ (sym (compPath-filler (merid a) (λ j → merid (ptSn (suc n)) (~ j ∨ i)) (~ i))))) + +-- -- substAbove : (x : S₊ (suc n)) → transport (north≡merid a) (cong ∣_∣ (merid x) ∙ cong ∣_∣ (sym (merid a))) +-- -- ≡ cong ∣_∣ (merid x) +-- -- substAbove x i = transp (λ j → north≡merid a (i ∨ j)) i +-- -- (compPath-filler (cong ∣_∣ (merid x)) (λ j → ∣ merid a (~ j ∨ i) ∣) (~ i)) + +-- -- encode : {n : ℕ} {x : coHomK (2 + n)} → Path (coHomK (2 + n)) ∣ north ∣ x → Code n x +-- -- encode {n = n} p = transport (cong (Code n) p) ∣ (ptSn (suc n)) ∣ + +-- -- decode-encode : {n : ℕ} {x : coHomK (2 + n)} (p : Path (coHomK (2 + n)) ∣ north ∣ x) +-- -- → decode _ (encode p) ≡ p +-- -- decode-encode {n = n} = +-- -- J (λ y p → decode _ (encode p) ≡ p) +-- -- (cong (decode ∣ north ∣) (transportRefl ∣ ptSn (suc n) ∣) +-- -- ∙ cong (cong ∣_∣) (rCancel (merid (ptSn (suc n))))) + +-- -- -- We define an addition operation on Code which we can use in order to show that encode is a +-- -- -- morphism (in a very loose sense) +-- -- hLevCode : {n : ℕ} (x : coHomK (2 + n)) → isOfHLevel (3 + n) (Code n x) +-- -- hLevCode {n = n} = +-- -- trElim (λ _ → isProp→isOfHLevelSuc (3 + n) (isPropIsOfHLevel (3 + n))) +-- -- (sphereToPropElim _ +-- -- (λ _ → (isPropIsOfHLevel (3 + n))) (isOfHLevelTrunc (3 + n))) + +-- -- Code-add' : {n : ℕ} (x : _) → Code n ∣ north ∣ → Code n ∣ x ∣ → Code n ∣ x ∣ +-- -- Code-add' {n = n} north = _+ₖ_ +-- -- Code-add' {n = n} south = _+ₖ_ +-- -- Code-add' {n = n} (merid a i) = helper n a i +-- -- where +-- -- help : (n : ℕ) → (x y a : S₊ (suc n)) +-- -- → transport (λ i → Code n ∣ north ∣ → Code n ∣ merid a i ∣ → Code n ∣ merid a i ∣) +-- -- (_+ₖ_) ∣ x ∣ ∣ y ∣ +-- -- ≡ ∣ x ∣ +ₖ ∣ y ∣ +-- -- help n x y a = +-- -- (λ i → transportRefl ((∣ transportRefl x i ∣ +ₖ (∣ transportRefl y i ∣ -ₖ ∣ a ∣)) +ₖ ∣ a ∣) i) +-- -- ∙∙ cong (_+ₖ ∣ a ∣) (assocₖ _ ∣ x ∣ ∣ y ∣ (-ₖ ∣ a ∣)) +-- -- ∙∙ sym (assocₖ _ (∣ x ∣ +ₖ ∣ y ∣) (-ₖ ∣ a ∣) ∣ a ∣) +-- -- ∙∙ cong ((∣ x ∣ +ₖ ∣ y ∣) +ₖ_) (lCancelₖ _ ∣ a ∣) +-- -- ∙∙ rUnitₖ _ _ + +-- -- helper : (n : ℕ) (a : S₊ (suc n)) +-- -- → PathP (λ i → Code n ∣ north ∣ → Code n ∣ merid a i ∣ → Code n ∣ merid a i ∣) _+ₖ_ _+ₖ_ +-- -- helper n a = +-- -- toPathP (funExt +-- -- (trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelΠ (3 + n) (λ _ → isOfHLevelTrunc (3 + n))) _ _) +-- -- λ x → funExt +-- -- (trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) +-- -- λ y → help n x y a))) + +-- -- Code-add : {n : ℕ} (x : _) → Code n ∣ north ∣ → Code n x → Code n x +-- -- Code-add {n = n} = +-- -- trElim (λ x → isOfHLevelΠ (4 + n) λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelSuc (3 + n) (hLevCode {n = n} x)) +-- -- Code-add' + +-- -- encode-hom : {n : ℕ} {x : _} (q : 0ₖ _ ≡ 0ₖ _) (p : 0ₖ _ ≡ x) +-- -- → encode (q ∙ p) ≡ Code-add {n = n} x (encode q) (encode p) +-- -- encode-hom {n = n} q = J (λ x p → encode (q ∙ p) ≡ Code-add {n = n} x (encode q) (encode p)) +-- -- (cong encode (sym (rUnit q)) +-- -- ∙∙ sym (rUnitₖ _ (encode q)) +-- -- ∙∙ cong (encode q +ₖ_) (cong ∣_∣ (sym (transportRefl _)))) + +-- -- stabSpheres : (n : ℕ) → Iso (coHomK (suc n)) (typ (Ω (coHomK-ptd (2 + n)))) +-- -- fun (stabSpheres n) = decode _ +-- -- inv' (stabSpheres n) = encode +-- -- rightInv (stabSpheres n) p = decode-encode p +-- -- leftInv (stabSpheres n) = +-- -- trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) +-- -- λ a → cong encode (congFunct ∣_∣ (merid a) (sym (merid (ptSn (suc n))))) +-- -- ∙∙ (λ i → transport (congFunct (Code n) (cong ∣_∣ (merid a)) +-- -- (cong ∣_∣ (sym (merid (ptSn (suc n))))) i) ∣ ptSn (suc n) ∣) +-- -- ∙∙ (substComposite (λ x → x) +-- -- (cong (Code n) (cong ∣_∣ (merid a))) +-- -- (cong (Code n) (cong ∣_∣ (sym (merid (ptSn (suc n)))))) ∣ ptSn (suc n) ∣ +-- -- ∙∙ cong (transport (λ i → Code n ∣ merid (ptSn (suc n)) (~ i) ∣)) +-- -- (transportRefl (∣ (ptSn (suc n)) ∣ +ₖ ∣ a ∣) ∙ lUnitₖ (suc n) ∣ a ∣) +-- -- ∙∙ symMeridLem n (ptSn (suc n)) ∣ a ∣ +-- -- ∙∙ cong (∣ a ∣ +ₖ_) -0ₖ +-- -- ∙∙ rUnitₖ (suc n) ∣ a ∣) + +-- -- Iso-Kn-ΩKn+1 : (n : HLevel) → Iso (coHomK n) (typ (Ω (coHomK-ptd (suc n)))) +-- -- Iso-Kn-ΩKn+1 zero = invIso (compIso (congIso (truncIdempotentIso _ isGroupoidS¹)) ΩS¹Isoℤ) +-- -- Iso-Kn-ΩKn+1 (suc n) = stabSpheres n + +-- -- Kn≃ΩKn+1 : {n : ℕ} → coHomK n ≃ typ (Ω (coHomK-ptd (suc n))) +-- -- Kn≃ΩKn+1 {n = n} = isoToEquiv (Iso-Kn-ΩKn+1 n) + +-- -- -- Some properties of the Iso +-- -- Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) +-- -- Kn→ΩKn+1 n = Iso.fun (Iso-Kn-ΩKn+1 n) + +-- -- ΩKn+1→Kn : (n : ℕ) → typ (Ω (coHomK-ptd (suc n))) → coHomK n +-- -- ΩKn+1→Kn n = Iso.inv (Iso-Kn-ΩKn+1 n) + +-- -- Kn→ΩKn+10ₖ : (n : ℕ) → Kn→ΩKn+1 n (0ₖ n) ≡ refl +-- -- Kn→ΩKn+10ₖ zero = sym (rUnit refl) +-- -- Kn→ΩKn+10ₖ (suc n) i j = ∣ (rCancel (merid (ptSn (suc n))) i j) ∣ + +-- -- ΩKn+1→Kn-refl : (n : ℕ) → ΩKn+1→Kn n refl ≡ 0ₖ n +-- -- ΩKn+1→Kn-refl zero = refl +-- -- ΩKn+1→Kn-refl (suc zero) = refl +-- -- ΩKn+1→Kn-refl (suc (suc n)) = refl + +-- -- Kn≃ΩKn+1∙ : {n : ℕ} → coHomK-ptd n ≡ (Ω (coHomK-ptd (suc n))) +-- -- Kn≃ΩKn+1∙ {n = n} = ua∙ Kn≃ΩKn+1 (Kn→ΩKn+10ₖ n) + +-- -- Kn→ΩKn+1-hom : (n : ℕ) (x y : coHomK n) → Kn→ΩKn+1 n (x +[ n ]ₖ y) ≡ Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y +-- -- Kn→ΩKn+1-hom zero x y = (λ j i → hfill (doubleComp-faces (λ i₁ → ∣ base ∣) (λ _ → ∣ base ∣) i) +-- -- (inS (∣ intLoop (x ℤ+ y) i ∣)) (~ j)) +-- -- ∙∙ (λ j i → ∣ intLoop-hom x y (~ j) i ∣) +-- -- ∙∙ (congFunct ∣_∣ (intLoop x) (intLoop y) +-- -- ∙ cong₂ _∙_ (λ j i → hfill (doubleComp-faces (λ i₁ → ∣ base ∣) (λ _ → ∣ base ∣) i) +-- -- (inS (∣ intLoop x i ∣)) j) +-- -- λ j i → hfill (doubleComp-faces (λ i₁ → ∣ base ∣) (λ _ → ∣ base ∣) i) +-- -- (inS (∣ intLoop y i ∣)) j) +-- -- Kn→ΩKn+1-hom (suc n) = σ-hom + +-- -- ΩKn+1→Kn-hom : (n : ℕ) (x y : Path (coHomK (suc n)) (0ₖ _) (0ₖ _)) +-- -- → ΩKn+1→Kn n (x ∙ y) ≡ ΩKn+1→Kn n x +[ n ]ₖ ΩKn+1→Kn n y +-- -- ΩKn+1→Kn-hom zero p q = +-- -- cong winding (congFunct (trRec isGroupoidS¹ (λ x → x)) p q) +-- -- ∙ winding-hom (cong (trRec isGroupoidS¹ (λ x → x)) p) (cong (trRec isGroupoidS¹ (λ x → x)) q) +-- -- ΩKn+1→Kn-hom (suc n) = encode-hom + +-- -- Kn→ΩKn+1-ₖ : (n : ℕ) (x : coHomK n) → Kn→ΩKn+1 n (-ₖ x) ≡ sym (Kn→ΩKn+1 n x) +-- -- Kn→ΩKn+1-ₖ n x = +-- -- lUnit _ +-- -- ∙∙ cong (_∙ Kn→ΩKn+1 n (-ₖ x)) (sym (lCancel _)) +-- -- ∙∙ sym (assoc∙ _ _ _) +-- -- ∙∙ cong (sym (Kn→ΩKn+1 n x) ∙_) help +-- -- ∙∙ sym (rUnit _) +-- -- where +-- -- help : Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n (-ₖ x) ≡ refl +-- -- help = sym (Kn→ΩKn+1-hom n x (-ₖ x)) ∙∙ cong (Kn→ΩKn+1 n) (rCancelₖ n x) ∙∙ Kn→ΩKn+10ₖ n + +-- -- isHomogeneousKn : (n : HLevel) → isHomogeneous (coHomK-ptd n) +-- -- isHomogeneousKn n = +-- -- subst isHomogeneous +-- -- (sym (ΣPathP (ua Kn≃ΩKn+1 , ua-gluePath _ (Kn→ΩKn+10ₖ n)))) +-- -- (isHomogeneousPath _ _) + +-- -- -- With the equivalence Kn≃ΩKn+1, we get that the two definitions of cohomology groups agree +-- -- open IsGroupHom + +-- -- coHom≅coHomΩ : ∀ {ℓ} (n : ℕ) (A : Type ℓ) → GroupIso (coHomGr n A) (coHomGrΩ n A) +-- -- fun (fst (coHom≅coHomΩ n A)) = map λ f a → Kn→ΩKn+1 n (f a) +-- -- inv' (fst (coHom≅coHomΩ n A)) = map λ f a → ΩKn+1→Kn n (f a) +-- -- rightInv (fst (coHom≅coHomΩ n A)) = +-- -- sElim (λ _ → isOfHLevelPath 2 § _ _) +-- -- λ f → cong ∣_∣₂ (funExt λ x → rightInv (Iso-Kn-ΩKn+1 n) (f x)) +-- -- leftInv (fst (coHom≅coHomΩ n A)) = +-- -- sElim (λ _ → isOfHLevelPath 2 § _ _) +-- -- λ f → cong ∣_∣₂ (funExt λ x → leftInv (Iso-Kn-ΩKn+1 n) (f x)) +-- -- snd (coHom≅coHomΩ n A) = +-- -- makeIsGroupHom +-- -- (sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) +-- -- λ f g → cong ∣_∣₂ (funExt λ x → Kn→ΩKn+1-hom n (f x) (g x))) + +-- -- module lockedKnIso (key : Unit') where +-- -- Kn→ΩKn+1' : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) +-- -- Kn→ΩKn+1' n = lock key (Iso.fun (Iso-Kn-ΩKn+1 n)) + +-- -- ΩKn+1→Kn' : (n : ℕ) → typ (Ω (coHomK-ptd (suc n))) → coHomK n +-- -- ΩKn+1→Kn' n = lock key (Iso.inv (Iso-Kn-ΩKn+1 n)) + +-- -- ΩKn+1→Kn→ΩKn+1 : (n : ℕ) → (x : typ (Ω (coHomK-ptd (suc n)))) → Kn→ΩKn+1' n (ΩKn+1→Kn' n x) ≡ x +-- -- ΩKn+1→Kn→ΩKn+1 n x = pm key +-- -- where +-- -- pm : (key : Unit') → lock key (Iso.fun (Iso-Kn-ΩKn+1 n)) (lock key (Iso.inv (Iso-Kn-ΩKn+1 n)) x) ≡ x +-- -- pm unlock = Iso.rightInv (Iso-Kn-ΩKn+1 n) x + +-- -- Kn→ΩKn+1→Kn : (n : ℕ) → (x : coHomK n) → ΩKn+1→Kn' n (Kn→ΩKn+1' n x) ≡ x +-- -- Kn→ΩKn+1→Kn n x = pm key +-- -- where +-- -- pm : (key : Unit') → lock key (Iso.inv (Iso-Kn-ΩKn+1 n)) (lock key (Iso.fun (Iso-Kn-ΩKn+1 n)) x) ≡ x +-- -- pm unlock = Iso.leftInv (Iso-Kn-ΩKn+1 n) x + +-- -- -distrLemma : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n m : ℕ) (f : GroupHom (coHomGr n A) (coHomGr m B)) +-- -- (x y : coHom n A) +-- -- → fst f (x -[ n ]ₕ y) ≡ fst f x -[ m ]ₕ fst f y +-- -- -distrLemma n m f' x y = sym (-cancelRₕ m (f y) (f (x -[ n ]ₕ y))) +-- -- ∙∙ cong (λ x → x -[ m ]ₕ f y) (sym (f' .snd .pres· (x -[ n ]ₕ y) y)) +-- -- ∙∙ cong (λ x → x -[ m ]ₕ f y) ( cong f (-+cancelₕ n _ _)) +-- -- where +-- -- f = fst f' + +-- -- -- HLevel stuff for cup product +-- -- isContr-↓∙ : (n : ℕ) → isContr (coHomK-ptd (suc n) →∙ coHomK-ptd n) +-- -- fst (isContr-↓∙ zero) = (λ _ → 0) , refl +-- -- snd (isContr-↓∙ zero) (f , p) = +-- -- Σ≡Prop (λ f → isSetℤ _ _) +-- -- (funExt (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 isSetℤ) _ _) +-- -- (toPropElim (λ _ → isSetℤ _ _) (sym p)))) +-- -- fst (isContr-↓∙ (suc n)) = (λ _ → 0ₖ _) , refl +-- -- fst (snd (isContr-↓∙ (suc n)) f i) x = +-- -- trElim {B = λ x → 0ₖ (suc n) ≡ fst f x} +-- -- (λ _ → isOfHLevelPath (4 + n) (isOfHLevelSuc (3 + n) (isOfHLevelTrunc (3 + n))) _ _) +-- -- (sphereElim _ (λ _ → isOfHLevelTrunc (3 + n) _ _) +-- -- (sym (snd f))) x i +-- -- snd (snd (isContr-↓∙ (suc n)) f i) j = snd f (~ i ∨ j) + +-- -- isContr-↓∙' : (n : ℕ) → isContr (S₊∙ (suc n) →∙ coHomK-ptd n) +-- -- fst (isContr-↓∙' zero) = (λ _ → 0) , refl +-- -- snd (isContr-↓∙' zero) (f , p) = +-- -- Σ≡Prop (λ f → isSetℤ _ _) +-- -- (funExt (toPropElim (λ _ → isSetℤ _ _) (sym p))) +-- -- fst (isContr-↓∙' (suc n)) = (λ _ → 0ₖ _) , refl +-- -- fst (snd (isContr-↓∙' (suc n)) f i) x = +-- -- sphereElim _ {A = λ x → 0ₖ (suc n) ≡ fst f x} +-- -- (λ _ → isOfHLevelTrunc (3 + n) _ _) +-- -- (sym (snd f)) x i +-- -- snd (snd (isContr-↓∙' (suc n)) f i) j = snd f (~ i ∨ j) + +-- -- isOfHLevel→∙Kn : {A : Pointed₀} (n m : ℕ) +-- -- → isOfHLevel (suc m) (A →∙ coHomK-ptd n) → isOfHLevel (suc (suc m)) (A →∙ coHomK-ptd (suc n)) +-- -- isOfHLevel→∙Kn {A = A} n m hlev = step₃ +-- -- where +-- -- step₁ : isOfHLevel (suc m) (A →∙ Ω (coHomK-ptd (suc n))) +-- -- step₁ = +-- -- subst (isOfHLevel (suc m)) +-- -- (λ i → A →∙ ua∙ {A = Ω (coHomK-ptd (suc n))} {B = coHomK-ptd n} +-- -- (invEquiv Kn≃ΩKn+1) +-- -- (ΩKn+1→Kn-refl n) (~ i)) hlev + +-- -- step₂ : isOfHLevel (suc m) (typ (Ω (A →∙ coHomK-ptd (suc n) ∙))) +-- -- step₂ = isOfHLevelRetractFromIso (suc m) (invIso (invIso (ΩfunExtIso _ _))) step₁ + +-- -- step₃ : isOfHLevel (suc (suc m)) (A →∙ coHomK-ptd (suc n)) +-- -- step₃ = +-- -- isOfHLevelΩ→isOfHLevel m +-- -- λ f → subst (λ x → isOfHLevel (suc m) (typ (Ω x))) +-- -- (isHomogeneous→∙ (isHomogeneousKn (suc n)) f) +-- -- step₂ + +-- -- isOfHLevel↑∙ : ∀ n m → isOfHLevel (2 + n) (coHomK-ptd (suc m) →∙ coHomK-ptd (suc (n + m))) +-- -- isOfHLevel↑∙ zero m = isOfHLevel→∙Kn m 0 (isContr→isProp (isContr-↓∙ m)) +-- -- isOfHLevel↑∙ (suc n) m = isOfHLevel→∙Kn (suc (n + m)) (suc n) (isOfHLevel↑∙ n m) + +-- -- isOfHLevel↑∙' : ∀ n m → isOfHLevel (2 + n) (S₊∙ (suc m) →∙ coHomK-ptd (suc (n + m))) +-- -- isOfHLevel↑∙' zero m = isOfHLevel→∙Kn m 0 (isContr→isProp (isContr-↓∙' m)) +-- -- isOfHLevel↑∙' (suc n) m = isOfHLevel→∙Kn (suc (n + m)) (suc n) (isOfHLevel↑∙' n m) + +-- -- →∙KnPath : (A : Pointed₀) (n : ℕ) → Ω (A →∙ coHomK-ptd (suc n) ∙) ≡ (A →∙ coHomK-ptd n ∙) +-- -- →∙KnPath A n = +-- -- ua∙ (isoToEquiv (ΩfunExtIso A (coHomK-ptd (suc n)))) refl +-- -- ∙ (λ i → (A →∙ Kn≃ΩKn+1∙ {n = n} (~ i) ∙)) + +-- -- contr∙-lem : (n m : ℕ) +-- -- → isContr (coHomK-ptd (suc n) +-- -- →∙ (coHomK-ptd (suc m) +-- -- →∙ coHomK-ptd (suc (n + m)) ∙)) +-- -- fst (contr∙-lem n m) = (λ _ → (λ _ → 0ₖ _) , refl) , refl +-- -- snd (contr∙-lem n m) (f , p) = →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousKn _)) (sym (funExt help)) +-- -- where +-- -- help : (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) +-- -- help = +-- -- trElim (λ _ → isOfHLevelPath (3 + n) +-- -- (subst (λ x → isOfHLevel x (coHomK-ptd (suc m) →∙ coHomK-ptd (suc (n + m)))) (λ i → suc (suc (+-comm n 1 i))) +-- -- (isOfHLevelPlus' {n = 1} (2 + n) (isOfHLevel↑∙ n m))) _ _) +-- -- (sphereElim _ (λ _ → isOfHLevel↑∙ n m _ _) p) + +-- -- isOfHLevel↑∙∙ : ∀ n m l +-- -- → isOfHLevel (2 + l) (coHomK-ptd (suc n) +-- -- →∙ (coHomK-ptd (suc m) +-- -- →∙ coHomK-ptd (suc (suc (l + n + m))) ∙)) +-- -- isOfHLevel↑∙∙ n m zero = +-- -- isOfHLevelΩ→isOfHLevel 0 λ f +-- -- → subst +-- -- isProp (cong (λ x → typ (Ω x)) +-- -- (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousKn _)) f)) +-- -- (isOfHLevelRetractFromIso 1 (ΩfunExtIso _ _) h) +-- -- where +-- -- h : isProp (coHomK-ptd (suc n) +-- -- →∙ (Ω (coHomK-ptd (suc m) +-- -- →∙ coHomK-ptd (suc (suc (n + m))) ∙))) +-- -- h = +-- -- subst isProp (λ i → coHomK-ptd (suc n) →∙ (→∙KnPath (coHomK-ptd (suc m)) (suc (n + m)) (~ i))) +-- -- (isContr→isProp (contr∙-lem n m)) +-- -- isOfHLevel↑∙∙ n m (suc l) = +-- -- isOfHLevelΩ→isOfHLevel (suc l) +-- -- λ f → +-- -- subst +-- -- (isOfHLevel (2 + l)) (cong (λ x → typ (Ω x)) +-- -- (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousKn _)) f)) +-- -- (isOfHLevelRetractFromIso (2 + l) (ΩfunExtIso _ _) h) +-- -- where +-- -- h : isOfHLevel (2 + l) +-- -- (coHomK-ptd (suc n) +-- -- →∙ (Ω (coHomK-ptd (suc m) +-- -- →∙ coHomK-ptd (suc (suc (suc (l + n + m)))) ∙))) +-- -- h = +-- -- subst (isOfHLevel (2 + l)) +-- -- (λ i → coHomK-ptd (suc n) →∙ →∙KnPath (coHomK-ptd (suc m)) (suc (suc (l + n + m))) (~ i)) +-- -- (isOfHLevel↑∙∙ n m l) + +-- -- -- Misc. +-- -- isoType→isoCohom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) +-- -- → Iso A B +-- -- → GroupIso (coHomGr n A) (coHomGr n B) +-- -- fst (isoType→isoCohom n is) = setTruncIso (domIso is) +-- -- snd (isoType→isoCohom n is) = +-- -- makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) +-- -- (λ _ _ → refl)) From 6750ccdaeaa6cff759f0404ec22cf11e66c8e025 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Fri, 24 Sep 2021 00:55:42 +0200 Subject: [PATCH 66/89] stuff --- Cubical/Algebra/Group/GroupPath.agda | 166 + Cubical/Algebra/Group/MorphismProperties.agda | 37 +- Cubical/ZCohomology/Gysin.agda | 4147 +++++++++-------- Cubical/ZCohomology/Gysin2.agda | 1221 +++++ Cubical/ZCohomology/Properties.agda | 1490 +++--- .../ZCohomology/RingStructure/RingLaws.agda | 5 + 6 files changed, 4094 insertions(+), 2972 deletions(-) create mode 100644 Cubical/ZCohomology/Gysin2.agda diff --git a/Cubical/Algebra/Group/GroupPath.agda b/Cubical/Algebra/Group/GroupPath.agda index 43de6de64..60d5900cc 100644 --- a/Cubical/Algebra/Group/GroupPath.agda +++ b/Cubical/Algebra/Group/GroupPath.agda @@ -127,3 +127,169 @@ uaCompGroupEquiv f g = caracGroup≡ _ _ ( cong ⟨_⟩ (uaGroup f) ∙ cong ⟨_⟩ (uaGroup g) ≡⟨ sym (cong-∙ ⟨_⟩ (uaGroup f) (uaGroup g)) ⟩ cong ⟨_⟩ (uaGroup f ∙ uaGroup g) ∎) + +open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) +open import Cubical.Data.Int renaming (_·_ to _*_ ; _+_ to _+ℤ_ ; _-_ to _-ℤ_) +open import Cubical.Data.Nat +open import Cubical.Data.List + +_ℤ[_]·_ : ∀ {ℓ} → ℤ → (G : Group ℓ) → fst G → fst G +pos zero ℤ[ G' ]· g = GroupStr.1g (snd G') +pos (suc n) ℤ[ G' ]· g = GroupStr._·_ (snd G') g (pos n ℤ[ G' ]· g) +negsuc zero ℤ[ G' ]· g = GroupStr.inv (snd G') g +negsuc (suc n) ℤ[ G' ]· g = + GroupStr._·_ (snd G') (GroupStr.inv (snd G') g) (negsuc n ℤ[ G' ]· g) + +gen₁-by : ∀ {ℓ} (G : Group ℓ) + → (g : fst G) + → Type _ +gen₁-by G g = (h : fst G) + → Σ[ a ∈ ℤ ] h ≡ (a ℤ[ G ]· g) + +gen₂-by : ∀ {ℓ} (G : Group ℓ) + → (g₁ g₂ : fst G) + → Type _ +gen₂-by G g₁ g₂ = + (h : fst G) + → Σ[ a ∈ ℤ × ℤ ] h + ≡ GroupStr._·_ (snd G) + ((fst a) ℤ[ G ]· g₁) + ((snd a) ℤ[ G ]· g₂) + +hompres· : ∀ {ℓ ℓ'} {G : Group ℓ} {H : Group ℓ'} + → (ϕ : GroupHom G H) + → (x : fst G) (z : ℤ) + → fst ϕ (z ℤ[ G ]· x) ≡ (z ℤ[ H ]· fst ϕ x) +hompres· (ϕ , ϕhom) x (pos zero) = IsGroupHom.pres1 ϕhom +hompres· {H = H} (ϕ , ϕhom) x (pos (suc n)) = + IsGroupHom.pres· ϕhom _ _ + ∙ cong (GroupStr._·_ (snd H) (ϕ x)) (hompres· (ϕ , ϕhom) x (pos n)) +hompres· (ϕ , ϕhom) x (negsuc zero) = IsGroupHom.presinv ϕhom _ +hompres· {H = H} (ϕ , ϕhom) x (negsuc (suc n)) = + IsGroupHom.pres· ϕhom _ _ + ∙ cong₂ (GroupStr._·_ (snd H)) (IsGroupHom.presinv ϕhom x) + ((hompres· (ϕ , ϕhom) x (negsuc n))) + +Iso-pres-gen₁ : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') + → (g : fst G) + → gen₁-by G g + → (e : GroupIso G H) + → gen₁-by H (Iso.fun (fst e) g) +Iso-pres-gen₁ G H g genG is h = + (fst (genG (Iso.inv (fst is) h))) + , (sym (Iso.rightInv (fst is) h) + ∙∙ cong (Iso.fun (fst is)) (snd (genG (Iso.inv (fst is) h))) + ∙∙ (hompres· (_ , snd is) g (fst (genG (Iso.inv (fst is) h))))) + +Iso-pres-gen₂ : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') + → (g₁ g₂ : fst G) + → gen₂-by G g₁ g₂ + → (e : GroupIso G H) + → gen₂-by H (Iso.fun (fst e) g₁) (Iso.fun (fst e) g₂) +fst (Iso-pres-gen₂ G H g₁ g₂ genG is h) = genG (Iso.inv (fst is) h) .fst +snd (Iso-pres-gen₂ G H g₁ g₂ genG is h) = + sym (Iso.rightInv (fst is) h) + ∙∙ cong (fun (fst is)) (snd (genG (Iso.inv (fst is) h))) + ∙∙ (IsGroupHom.pres· (snd is) _ _ + ∙ cong₂ (GroupStr._·_ (snd H)) + (hompres· (_ , snd is) g₁ (fst (fst (genG (inv (fst is) h))))) + (hompres· (_ , snd is) g₂ (snd (fst (genG (inv (fst is) h)))))) + +-- module _ {ℓ : Level} {G' : Group ℓ} where +-- private +-- G = fst G' + +-- _+G_ = GroupStr._·_ (snd G') +-- -G = GroupStr.inv (snd G') +-- 1G = GroupStr.1g (snd G') + +-- _ℤ·_ : ℤ → G → G +-- pos zero ℤ· g = 1G +-- pos (suc n) ℤ· g = g +G (pos n ℤ· g) +-- negsuc zero ℤ· g = -G g +-- negsuc (suc n) ℤ· g = -G g +G (negsuc n ℤ· g) + +-- rUnitℤ : (g : G) → 1 ℤ· g ≡ g +-- rUnitℤ = {!!} + +-- copyList : {!length!} +-- copyList = {!!} +-- open import Cubical.Data.Empty renaming (rec to ⊥-rec) +-- module _ {ℓ : Level} (G' : Group ℓ) where +-- private +-- G = fst G' + +-- _+G_ = GroupStr._·_ (snd G') +-- -G = GroupStr.inv (snd G') +-- 1G = GroupStr.1g (snd G') +-- sum : List (ℤ × G) → G +-- sum [] = 1G +-- sum (x ∷ x₁) = (_ℤ·_ {G' = G'} (fst x) (snd x)) +G sum x₁ + +-- sum'help : (l1 : List ℤ) (l2 : List G) (n m : ℕ) +-- → (n ≡ m) +-- → length l1 ≡ n +-- → length l2 ≡ m +-- → G +-- sum'help [] [] n m p q r = 1G +-- sum'help [] (x ∷ l2) n m p q r = +-- ⊥-rec (snotz (r ∙∙ sym p ∙∙ sym q)) +-- sum'help (x ∷ l1) [] n m p q r = +-- ⊥-rec (snotz (q ∙∙ p ∙∙ sym r)) +-- sum'help (x ∷ l1) (x₁ ∷ l2) n m p q r = +-- _ℤ·_ {G' = G'} x x₁ +G sum'help l1 l2 _ _ (cong predℕ p) (cong predℕ q) (cong predℕ r) + +-- sum' : (l1 : List ℤ) (l2 : List G) → length l1 ≡ length l2 +-- → G +-- sum' l1 l2 p = sum'help l1 l2 _ _ p refl refl + +-- isFinGen : Type _ +-- isFinGen = +-- Σ[ l ∈ List G ] +-- ((g : G) → Σ[ l2 ∈ List ℤ ] +-- Σ[ p ∈ (length l2 ≡ length l) ] +-- sum' l2 l p ≡ g) + +-- isFinGenℤ : isFinGen ℤGroup +-- isFinGenℤ = (1 ∷ []) +-- , λ g → (g ∷ []) , refl , help g +-- where +-- help : (g : ℤ) → sum' ℤGroup (g ∷ []) (pos 1 ∷ []) (λ _ → 1) ≡ g +-- help (pos zero) = refl +-- help (pos (suc n)) = cong (pos 1 +ℤ_) (help (pos n)) ∙ +Comm (pos 1) (pos n) +-- help (negsuc zero) = refl +-- help (negsuc (suc n)) = cong (negsuc 0 +ℤ_) (help (negsuc n)) ∙ +Comm (negsuc 0) (negsuc n) + +-- open import Cubical.Algebra.Group.DirProd +-- isFinGenℤ×ℤ : isFinGen (DirProd ℤGroup ℤGroup) +-- fst isFinGenℤ×ℤ = (1 , 0) ∷ (0 , 1) ∷ [] +-- fst (snd isFinGenℤ×ℤ (x , y)) = x ∷ y ∷ [] +-- fst (snd (snd isFinGenℤ×ℤ (x , y))) = refl +-- snd (snd (snd isFinGenℤ×ℤ (x , y))) = +-- ΣPathP ((λ i → fst ((distrLem 1 0 x) i) +ℤ fst (distrLem 0 1 y i)) +-- ∙ (λ i → (·Comm x 1 i) +ℤ (·Comm y 0 i)) +-- , (λ i → snd ((distrLem 1 0 x) i) +ℤ snd (distrLem 0 1 y i)) +-- ∙ (λ i → (·Comm x 0 i) +ℤ (·Comm y 1 i)) +-- ∙ sym (+Comm y 0)) +-- where +-- ll : (x : ℤ) → - x ≡ 0 -ℤ x +-- ll (pos zero) = refl +-- ll (pos (suc zero)) = refl +-- ll (pos (suc (suc n))) = +-- cong predℤ (ll (pos (suc n))) +-- ll (negsuc zero) = refl +-- ll (negsuc (suc n)) = cong sucℤ (ll (negsuc n)) + +-- ℤ×ℤ = DirProd ℤGroup ℤGroup +-- _+''_ = GroupStr._·_ (snd ℤ×ℤ) +-- distrLem : (x y : ℤ) (z : ℤ) +-- → Path (ℤ × ℤ) (_ℤ·_ {G' = DirProd ℤGroup ℤGroup} z (x , y)) (z * x , z * y) +-- distrLem x y (pos zero) = refl +-- distrLem x y (pos (suc n)) = cong ((x , y) +''_) (distrLem x y (pos n)) +-- distrLem x y (negsuc zero) = ΣPathP (sym (ll x) , sym (ll y)) +-- distrLem x y (negsuc (suc n)) = +-- cong₂ _+''_ (ΣPathP (sym (ll x) , sym (ll y))) (distrLem x y (negsuc n)) + +-- maini : ∀ {ℓ} {G : Type ℓ} +-- → {!GroupIso!} +-- maini = {!!} diff --git a/Cubical/Algebra/Group/MorphismProperties.agda b/Cubical/Algebra/Group/MorphismProperties.agda index 373da7b87..c00e05652 100644 --- a/Cubical/Algebra/Group/MorphismProperties.agda +++ b/Cubical/Algebra/Group/MorphismProperties.agda @@ -184,25 +184,26 @@ compGroupEquiv : GroupEquiv F G → GroupEquiv G H → GroupEquiv F H fst (compGroupEquiv f g) = compEquiv (fst f) (fst g) snd (compGroupEquiv f g) = isGroupHomComp (_ , f .snd) (_ , g .snd) +isGroupHomInv : (f : GroupEquiv G H) → IsGroupHom (H .snd) (invEq (fst f)) (G .snd) +isGroupHomInv {G = G} {H = H} f = makeIsGroupHom λ h h' → + isInj-f _ _ + (f' (g (h ⋆² h')) ≡⟨ secEq (fst f) _ ⟩ + (h ⋆² h') ≡⟨ sym (cong₂ _⋆²_ (secEq (fst f) h) (secEq (fst f) h')) ⟩ + (f' (g h) ⋆² f' (g h')) ≡⟨ sym (pres· (snd f) _ _) ⟩ + f' (g h ⋆¹ g h') ∎) + where + f' = fst (fst f) + _⋆¹_ = _·_ (snd G) + _⋆²_ = _·_ (snd H) + g = invEq (fst f) + + isInj-f : (x y : ⟨ G ⟩) → f' x ≡ f' y → x ≡ y + isInj-f x y = invEq (_ , isEquiv→isEmbedding (snd (fst f)) x y) + invGroupEquiv : GroupEquiv G H → GroupEquiv H G fst (invGroupEquiv f) = invEquiv (fst f) snd (invGroupEquiv f) = isGroupHomInv f - where - isGroupHomInv : (f : GroupEquiv G H) → IsGroupHom (H .snd) (invEq (fst f)) (G .snd) - isGroupHomInv {G = G} {H = H} f = makeIsGroupHom λ h h' → - isInj-f _ _ - (f' (g (h ⋆² h')) ≡⟨ secEq (fst f) _ ⟩ - (h ⋆² h') ≡⟨ sym (cong₂ _⋆²_ (secEq (fst f) h) (secEq (fst f) h')) ⟩ - (f' (g h) ⋆² f' (g h')) ≡⟨ sym (pres· (snd f) _ _) ⟩ - f' (g h ⋆¹ g h') ∎) - where - f' = fst (fst f) - _⋆¹_ = _·_ (snd G) - _⋆²_ = _·_ (snd H) - g = invEq (fst f) - isInj-f : (x y : ⟨ G ⟩) → f' x ≡ f' y → x ≡ y - isInj-f x y = invEq (_ , isEquiv→isEmbedding (snd (fst f)) x y) GroupEquivDirProd : {A : Group ℓ} {B : Group ℓ'} {C : Group ℓ''} {D : Group ℓ'''} → GroupEquiv A C → GroupEquiv B D @@ -229,10 +230,10 @@ snd (compGroupIso iso1 iso2) = isGroupHomComp (_ , snd iso1) (_ , snd iso2) invGroupIso : GroupIso G H → GroupIso H G fst (invGroupIso iso1) = invIso (fst iso1) -snd (invGroupIso iso1) = isGroupHomInv iso1 +snd (invGroupIso iso1) = isGroupHomInv' iso1 where - isGroupHomInv : (f : GroupIso G H) → IsGroupHom (H .snd) (inv (fst f)) (G .snd) - isGroupHomInv {G = G} {H = H} f = makeIsGroupHom λ h h' → + isGroupHomInv' : (f : GroupIso G H) → IsGroupHom (H .snd) (inv (fst f)) (G .snd) + isGroupHomInv' {G = G} {H = H} f = makeIsGroupHom λ h h' → isInj-f _ _ (f' (g (h ⋆² h')) ≡⟨ (rightInv (fst f)) _ ⟩ (h ⋆² h') ≡⟨ sym (cong₂ _⋆²_ (rightInv (fst f) h) (rightInv (fst f) h')) ⟩ diff --git a/Cubical/ZCohomology/Gysin.agda b/Cubical/ZCohomology/Gysin.agda index e7bb4babe..1687d4e33 100644 --- a/Cubical/ZCohomology/Gysin.agda +++ b/Cubical/ZCohomology/Gysin.agda @@ -12,6 +12,9 @@ open import Cubical.ZCohomology.RingStructure.CupProduct open import Cubical.ZCohomology.RingStructure.RingLaws open import Cubical.ZCohomology.RingStructure.GradedCommutativity +open import Cubical.Data.Sum +open import Cubical.Relation.Nullary + open import Cubical.Functions.Embedding open import Cubical.Data.Fin @@ -65,2069 +68,2085 @@ open import Cubical.Homotopy.Hopf open import Cubical.HITs.SetQuotients renaming (_/_ to _/'_) +SES→Iso : ∀ {ℓ ℓ'} {L R : Group ℓ-zero} + → {G : Group ℓ} {H : Group ℓ'} + → UnitGroup ≡ L + → UnitGroup ≡ R + → (lhom : GroupHom L G) (midhom : GroupHom G H) (rhom : GroupHom H R) + → ((x : _) → isInKer midhom x → isInIm lhom x) + → ((x : _) → isInKer rhom x → isInIm midhom x) + → isEquiv (fst midhom) +SES→Iso {R = R} {G = G} {H = H} = + J (λ L _ → UnitGroup ≡ R → + (lhom : GroupHom L G) (midhom : GroupHom G H) + (rhom : GroupHom H R) → + ((x : fst G) → isInKer midhom x → isInIm lhom x) → + ((x : fst H) → isInKer rhom x → isInIm midhom x) → + isEquiv (fst midhom)) + ((J (λ R _ → (lhom : GroupHom UnitGroup G) (midhom : GroupHom G H) + (rhom : GroupHom H R) → + ((x : fst G) → isInKer midhom x → isInIm lhom x) → + ((x : _) → isInKer rhom x → isInIm midhom x) → + isEquiv (fst midhom)) + main)) + where + main : (lhom : GroupHom UnitGroup G) (midhom : GroupHom G H) + (rhom : GroupHom H UnitGroup) → + ((x : fst G) → isInKer midhom x → isInIm lhom x) → + ((x : fst H) → isInKer rhom x → isInIm midhom x) → + isEquiv (fst midhom) + main lhom midhom rhom lexact rexact = + BijectionIsoToGroupEquiv {G = G} {H = H} + bijIso' .fst .snd + where + bijIso' : BijectionIso G H + BijectionIso.fun bijIso' = midhom + BijectionIso.inj bijIso' x inker = + pRec (GroupStr.is-set (snd G) _ _) + (λ s → sym (snd s) ∙ IsGroupHom.pres1 (snd lhom)) (lexact _ inker) + BijectionIso.surj bijIso' x = rexact x refl + +HopfInvariantPush : (n : ℕ) + → (f : S₊ (3 +ℕ n +ℕ n) → S₊ (2 +ℕ n)) + → Type _ +HopfInvariantPush n f = Pushout (λ _ → tt) f + + +Hopfα : (n : ℕ) + → (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) + → coHom (2 +ℕ n) (HopfInvariantPush n (fst f)) +Hopfα n f = ∣ (λ { (inl x) → 0ₖ _ + ; (inr x) → ∣ x ∣ + ; (push a i) → help a (~ i)}) ∣₂ + where + help : (a : S₊ (3 +ℕ n +ℕ n)) → ∣ fst f a ∣ ≡ 0ₖ (2 +ℕ n) + help = sphereElim _ (λ _ → isOfHLevelPlus' {n = n} (3 +ℕ n) + (isOfHLevelPath' (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n)) _ _)) (cong ∣_∣ₕ (snd f)) + +Hopfβ : (n : ℕ) + → (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) + → coHom (4 +ℕ (n +ℕ n)) (HopfInvariantPush n (fst f)) +Hopfβ n f = fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) ∣ ∣_∣ ∣₂ + +module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) where + module M = MV _ _ _ (λ _ → tt) (fst f) + ¬lem : (n m : ℕ) → ¬ suc (suc (m +ℕ n)) ≡ suc n + ¬lem zero zero p = snotz (cong predℕ p) + ¬lem (suc n) zero p = ¬lem n zero (cong predℕ p) + ¬lem zero (suc m) p = snotz (cong predℕ p) + ¬lem (suc n) (suc m) p = + ¬lem n (suc m) (cong (suc ∘ suc ) + (sym (+-suc m n)) ∙ (cong predℕ p)) + + GroupEquiv1 : + GroupEquiv + (coHomGr (3 +ℕ n +ℕ n) (fst (S₊∙ (3 +ℕ n +ℕ n)))) + ((coHomGr (suc (3 +ℕ n +ℕ n)) (Pushout (λ _ → tt) (fst f)))) + fst (fst GroupEquiv1) = fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) + snd (fst GroupEquiv1) = help + where + abstract + help : isEquiv (fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n))) + help = + SES→Iso + (fst (GroupPath _ _) + (invGroupEquiv (isContr→≃Unit + (isOfHLevelΣ 0 + (isOfHLevelRetractFromIso 0 (fst (Hⁿ-Unit≅0 _)) isContrUnit) + λ _ → isOfHLevelRetractFromIso 0 + (fst (Hⁿ-Sᵐ≅0 _ _ (¬lem n n))) isContrUnit) + , makeIsGroupHom λ _ _ → refl))) + (fst (GroupPath _ _) + (invGroupEquiv + (isContr→≃Unit + (isContrΣ + (isOfHLevelRetractFromIso 0 (fst (Hⁿ-Unit≅0 _)) + isContrUnit) + λ _ → isOfHLevelRetractFromIso 0 + (fst (Hⁿ-Sᵐ≅0 _ _ (¬lem n (suc n)))) isContrUnit) + , makeIsGroupHom λ _ _ → refl))) + (M.Δ (3 +ℕ n +ℕ n)) + (M.d (3 +ℕ n +ℕ n)) + (M.i (4 +ℕ n +ℕ n)) + (M.Ker-d⊂Im-Δ _) + (M.Ker-i⊂Im-d _) + snd GroupEquiv1 = s + where + abstract + s : IsGroupHom (snd (coHomGr (3 +ℕ n +ℕ n) (fst (S₊∙ (3 +ℕ n +ℕ n))))) + (fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n))) + (snd (coHomGr (suc (3 +ℕ n +ℕ n)) (Pushout (λ _ → tt) (fst f)))) + s = snd (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) + + Hopfβ-Iso : GroupIso ((coHomGr (suc (3 +ℕ n +ℕ n)) (Pushout (λ _ → tt) (fst f)))) ℤGroup + Iso.fun (fst Hopfβ-Iso) x = Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (invEq (fst GroupEquiv1) x) + Iso.inv (fst Hopfβ-Iso) f = (fst (fst GroupEquiv1)) (Iso.inv (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) f) + Iso.rightInv (fst Hopfβ-Iso) f = + cong (Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) ) + (retEq (fst GroupEquiv1) (Iso.inv (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) f)) + ∙ Iso.rightInv (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) f + Iso.leftInv (fst Hopfβ-Iso) x = + cong (fst (fst GroupEquiv1)) + (Iso.leftInv (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) + (invEq (fst GroupEquiv1) x)) + ∙ secEq (fst GroupEquiv1) x + snd Hopfβ-Iso = grHom + where + abstract + grHom : IsGroupHom (coHomGr (suc (3 +ℕ n +ℕ n)) (Pushout (λ _ → tt) (fst f)) .snd) + (λ x → Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (invEq (fst GroupEquiv1) x)) + (ℤGroup .snd) + grHom = makeIsGroupHom λ x y → + cong (Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) + (IsGroupHom.pres· (isGroupHomInv GroupEquiv1) x y) + ∙ IsGroupHom.pres· (snd (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) _ _ + + -- compGroupIso (GroupEquiv→GroupIso (invGroupEquiv GroupEquiv1)) (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) + +h33 : (n : ℕ) → Iso.inv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) 1 ≡ ∣ ∣_∣ ∣₂ +h33 zero = Iso.leftInv (fst (Hⁿ-Sⁿ≅ℤ (suc zero))) _ +h33 (suc n) = (λ i → Iso.inv + (fst + (suspensionAx-Sn (suc n) (suc n))) (h33 n i)) + ∙ cong ∣_∣₂ (funExt λ { north → refl + ; south → cong ∣_∣ₕ (merid north) + ; (merid a i) j → ∣ compPath-filler (merid a) (sym (merid north)) (~ j) i ∣}) + +d-β : (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) + → Iso.fun (fst (Hopfβ-Iso n f)) (Hopfβ n f) ≡ 1 +d-β n f = sym (cong (Iso.fun (fst (Hopfβ-Iso n f))) h) ∙ Iso.rightInv (fst (Hopfβ-Iso n f)) (pos 1) + where + h : Iso.inv (fst (Hopfβ-Iso n f)) (pos 1) ≡ Hopfβ n f + h = (λ i → fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) + (h33 _ i)) + ∙ cong ∣_∣₂ (funExt (λ { (inl x) → refl + ; (inr x) → refl + ; (push a i) → refl})) + +module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) where + 2+n = 2 +ℕ n + H = HopfInvariantPush n (fst f) + + ff : coHom 2+n H → coHom 2+n (S₊ (suc (suc n))) + ff = sMap (_∘ inr) + + grHom : IsGroupHom (snd (coHomGr 2+n H)) ff (snd (coHomGr 2+n (S₊ (suc (suc n))))) + grHom = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f g → refl) + + h' : (g : (S₊ (suc (suc n)) → coHomK 2+n)) → H → coHomK (2 +ℕ n) + h' g (inl x) = 0ₖ _ + h' g (inr x) = g x -ₖ g north + h' g (push a i) = h23 a i + where + h23 : (a : S₊ (suc (suc (suc (n +ℕ n))))) → 0ₖ (suc (suc n)) ≡ g (fst f a) -ₖ g north + h23 = sphereElim _ (λ x → isOfHLevelPlus' {n = n} (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n) _ _)) + (sym (rCancelₖ _ (g north)) ∙ cong (λ x → g x -ₖ g north) (sym (snd f))) + + ff⁻ : coHom 2+n (S₊ (suc (suc n))) → coHom 2+n H + ff⁻ = sMap h' + + h23 : (x : _) → ∥ x ≡ 0ₖ (suc (suc n)) ∥ + h23 = coHomK-elim _ (λ _ → isProp→isOfHLevelSuc (suc n) squash) ∣ refl ∣ + + h : Iso (coHom (2 +ℕ n) (HopfInvariantPush n (fst f))) (coHom (2 +ℕ n) ((S₊ (suc (suc n))))) + Iso.fun h = ff + Iso.inv h = ff⁻ + Iso.rightInv h = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ g → pRec (squash₂ _ _) + (λ p → cong ∣_∣₂ (funExt λ x → cong (g x +ₖ_) (cong (-ₖ_) p) ∙ rUnitₖ _ (g x))) + (h23 (g north)) + Iso.leftInv h = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ g → pRec (squash₂ _ _) + (pRec (isPropΠ (λ _ → squash₂ _ _)) + (λ gn gtt → + trRec (isProp→isOfHLevelSuc n (squash₂ _ _)) + (λ r → cong ∣_∣₂ (funExt λ { + (inl x) → sym gtt + ; (inr x) → (λ i → g (inr x) -ₖ gn i) ∙ rUnitₖ _ (g (inr x)) + ; (push a i) + → sphereElim _ + {A = λ a → PathP (λ i → h' (λ x → g (inr x)) (push a i) ≡ g (push a i)) + (sym gtt) + ((λ i → g (inr (fst f a)) -ₖ gn i) ∙ rUnitₖ _ (g (inr (fst f a))))} + (λ _ → isOfHLevelPathP' (suc (suc (suc (n +ℕ n)))) + (isOfHLevelPath (suc (suc (suc (suc (n +ℕ n))))) + (isOfHLevelPlus' {n = n} (suc (suc (suc (suc n)))) + (isOfHLevelTrunc (suc (suc (suc (suc n)))))) _ _) _ _) + r a i})) + (l g gtt gn)) + (h23 (g (inr north)))) + (h23 (g (inl tt))) + where + l : (g : HopfInvariantPush n (fst f) → coHomK (suc (suc n))) + → (gtt : g (inl tt) ≡ 0ₖ (suc (suc n))) + → (gn : g (inr north) ≡ 0ₖ (suc (suc n))) + → hLevelTrunc (suc n) (PathP (λ i → h' (λ x → g (inr x)) (push north i) ≡ g (push north i)) + (sym gtt) ((λ i → g (inr (fst f north)) -ₖ gn i) ∙ rUnitₖ _ (g (inr (fst f north))))) + l g gtt gn = isConnectedPathP _ (isConnectedPath _ (isConnectedKn _) _ _) _ _ .fst + + + Hopfα-Iso : GroupIso (coHomGr (2 +ℕ n) (HopfInvariantPush n (fst f))) ℤGroup + Hopfα-Iso = + compGroupIso + (h , grHom) + (Hⁿ-Sⁿ≅ℤ (suc n)) + +Hopfα-Iso-α : (n : ℕ) (f : _) + → Iso.fun (fst (Hopfα-Iso n f)) (Hopfα n f) + ≡ 1 +Hopfα-Iso-α n f = + sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (h33 n)) + ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) + where + hz : Iso.fun (h n f) (Hopfα n f) ≡ ∣ ∣_∣ ∣₂ + hz = refl + +_·₀ₕ_ : ∀ {ℓ} {n : ℕ} {A : Type ℓ} → ℤ → coHom n A → coHom n A +_·₀ₕ_ {n = n} a = sMap λ f x → a ·₀ f x + +HopfInvariant : (n : ℕ) + → (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → Type _ +HopfInvariant n f = + Σ[ x ∈ ℤ ] + Hopfα n f ⌣ Hopfα n f + ≡ (x ·₀ₕ subst (λ x → coHom x (HopfInvariantPush n (fst f))) + (λ i → suc (suc (suc (+-suc n n (~ i))))) + (Hopfβ n f)) + + +⌣-α : (n : ℕ) + → (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → _ +⌣-α n f = Hopfα n f ⌣ Hopfα n f + +HopfFunction : (n : ℕ) + → (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → ℤ +HopfFunction n f = Iso.fun (fst (Hopfβ-Iso n f)) + ((subst (λ x → coHom x (HopfInvariantPush n (fst f))) + (λ i → suc (suc (suc (+-suc n n i))))) (⌣-α n f)) -- (f) + +pres- : (e : GroupHom ℤGroup ℤGroup) → (a : ℤ) → fst e (- a) ≡ - fst e a +pres- e a = +Comm (fst e (- a)) (pos 0) + ∙ cong (_+ (fst e (- a))) (sym (l (fst e a)) ∙ +Comm (fst e a) (- fst e a)) + ∙∙ sym (+Assoc _ _ _) + ∙∙ cong (- (fst e a) +_) (pp ∙ l (fst e a)) + where + l : (a : ℤ) → a + (- a) ≡ 0 + l (pos zero) = refl + l (pos (suc zero)) = refl + l (pos (suc (suc n))) = predℤ+negsuc n (pos (suc (suc n))) ∙ l (pos (suc n)) + l (negsuc zero) = refl + l (negsuc (suc n)) = (cong sucℤ (sucℤ+pos n (negsuc (suc n)))) ∙ l (negsuc n) + + pp : fst e a + fst e (- a) ≡ fst e a + (- fst e a) + pp = sym (IsGroupHom.pres· (snd e) a (- a)) + ∙∙ cong (fst e) (l a) + ∙∙ (IsGroupHom.pres1 (snd e) + ∙ sym (l _)) + + + +GroupHomPres· : (e : GroupHom ℤGroup ℤGroup) → (a b : ℤ) → fst e (a · b) ≡ a · fst e b +GroupHomPres· e (pos zero) b = IsGroupHom.pres1 (snd e) +GroupHomPres· e (pos (suc n)) b = + IsGroupHom.pres· (snd e) b (pos n · b) ∙ cong (fst e b +_) (GroupHomPres· e (pos n) b) +GroupHomPres· e (negsuc zero) b = pres- e b +GroupHomPres· e (negsuc (suc n)) b = + IsGroupHom.pres· (snd e) (- b) (negsuc n · b) + ∙ cong₂ _+_ (pres- e b) (GroupHomPres· e (negsuc n) b) + +lUnit·₀ₕ : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → (x : coHom n A) → (pos zero) ·₀ₕ x ≡ 0ₕ _ +lUnit·₀ₕ n = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl + +minus≡0- : (x : ℤ) → - x ≡ (0 - x) +minus≡0- x = sym (GroupStr.lid (snd ℤGroup) (- x)) + +GroupHomPres·₀ : ∀ {ℓ} {A : Type ℓ} (n : ℕ) + (e : GroupHom ℤGroup (coHomGr n A)) + → (a b : ℤ) + → fst e (a · b) + ≡ (a ·₀ₕ fst e b) +GroupHomPres·₀ n e (pos zero) b = IsGroupHom.pres1 (snd e) ∙ sym (lUnit·₀ₕ n (fst e b)) +GroupHomPres·₀ {A = A} n e (pos (suc a)) b = + IsGroupHom.pres· (snd e) b (pos a · b) + ∙ (λ i → fst e b +ₕ GroupHomPres·₀ n e (pos a) b i) + ∙ s (fst e b) + where + s : (x : coHom n A) → x +ₕ (pos a ·₀ₕ x) ≡ pos (suc a) ·₀ₕ x + s = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl +GroupHomPres·₀ n e (negsuc zero) b = + (cong (fst e) (sym (GroupStr.lid (snd ℤGroup) (- b))) + ∙ IsGroupHom.presinv (snd e) b) +GroupHomPres·₀ {A = A} n e (negsuc (suc a)) b = + IsGroupHom.pres· (snd e) (- b) (negsuc a · b) + ∙∙ cong₂ _+ₕ_ ((cong (fst e) (sym (GroupStr.lid (snd ℤGroup) (- b))) + ∙ IsGroupHom.presinv (snd e) b)) + (GroupHomPres·₀ n e (negsuc a) b) + ∙∙ helper (fst e b) + where + helper : (x : coHom n A) → (-ₕ x) +ₕ (negsuc a ·₀ₕ x) ≡ (negsuc (suc a) ·₀ₕ x) + helper = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt λ x → commₖ n (-ₖ (f x)) (negsuc a ·₀ f x)) + +genH : ∀ {ℓ} {A : Type ℓ} (n : ℕ) + → (e : GroupIso (coHomGr n A) + ℤGroup) + → Σ[ x ∈ coHom n A ] + ((y : coHom n A) + → Σ[ a ∈ ℤ ] a ·₀ₕ x ≡ y) +genH {A = A} n e = (Iso.inv (fst e) 1) + , λ y → (Iso.fun (fst e) y) + , (sym (GroupHomPres·₀ n (_ , snd (invGroupIso e)) (Iso.fun (fst e) y) (pos 1)) + ∙∙ cong (Iso.inv (fst e)) (·Comm (Iso.fun (fst e) y) (pos 1)) + ∙∙ Iso.leftInv (fst e) y) + +l2 : (n m : ℕ) → Σ[ a ∈ ℕ ] (negsuc n · pos (suc m)) ≡ negsuc a +l2 n zero = n , ·Comm (negsuc n) (pos 1) +l2 n (suc m) = h3 _ _ .fst , + (·Comm (negsuc n) (pos (suc (suc m))) + ∙∙ cong (negsuc n +_) (·Comm (pos (suc m)) (negsuc n) ∙ (l2 n m .snd)) + ∙∙ h3 _ _ .snd) + where + h3 : (x y : ℕ) → Σ[ a ∈ ℕ ] negsuc x + negsuc y ≡ negsuc a + h3 zero zero = 1 , refl + h3 zero (suc y) = (suc (suc y)) , +Comm (negsuc zero) (negsuc (suc y)) + h3 (suc x) zero = (suc (suc x)) , refl + h3 (suc x) (suc y) = (h3 (suc (suc x)) y .fst) , (predℤ+negsuc y (negsuc (suc x)) ∙ snd ((h3 (suc (suc x))) y)) + +zz : (n x : ℕ) → Σ[ a ∈ ℕ ] ((pos (suc x)) · pos (suc (suc n)) ≡ pos (suc (suc a))) +zz n zero = n , refl +zz n (suc x) = h3 _ _ (zz n x) + where + h3 : (x : ℤ) (n : ℕ) → Σ[ a ∈ ℕ ] (x ≡ pos (suc (suc a))) + → Σ[ a ∈ ℕ ] pos n + x ≡ pos (suc (suc a)) + h3 x n (a , p) = n +ℕ a + , cong (pos n +_) p ∙ cong sucℤ (sucℤ+pos a (pos n)) + ∙ sucℤ+pos a (pos (suc n)) ∙ (sym (pos+ (suc (suc n)) a)) + +lem22 : (n : ℕ) (x : ℤ) → ¬ pos 1 ≡ x · pos (suc (suc n)) +lem22 n (pos zero) p = snotz (injPos p) +lem22 n (pos (suc n₁)) p = snotz (injPos (sym (cong predℤ (snd (zz n n₁))) ∙ sym (cong predℤ p))) +lem22 n (negsuc n₁) p = posNotnegsuc _ _ (p ∙ l2 _ _ .snd) + + +grr : (e : GroupEquiv ℤGroup ℤGroup) (x : ℤ) → (fst (fst e) 1) ≡ x → abs (fst (fst e) 1) ≡ 1 +grr e (pos zero) p = + ⊥-rec (snotz (injPos (sym (retEq (fst e) 1) + ∙∙ (cong (fst (fst (invGroupEquiv e))) p) + ∙∙ IsGroupHom.pres1 (snd (invGroupEquiv e))))) +grr e (pos (suc zero)) p = cong abs p +grr e (pos (suc (suc n))) p = ⊥-rec (lem22 _ _ (h3 ∙ ·Comm (pos (suc (suc n))) (invEq (fst e) 1))) + where + + h3 : pos 1 ≡ _ + h3 = sym (retEq (fst e) 1) + ∙∙ (cong (fst (fst (invGroupEquiv e))) (p ∙ ·Comm 1 (pos (suc (suc n))))) + ∙∙ GroupHomPres· (_ , snd (invGroupEquiv e)) (pos (suc (suc n))) 1 +grr e (negsuc zero) p = cong abs p +grr e (negsuc (suc n)) p = ⊥-rec (lem22 _ _ l33) + where + h3 : fst (fst e) (negsuc zero) ≡ pos (suc (suc n)) + h3 = pres- (_ , snd e) (pos 1) ∙ cong -_ p + + compGroup : GroupEquiv ℤGroup ℤGroup + fst compGroup = isoToEquiv (iso -_ -_ -Involutive -Involutive) + snd compGroup = makeIsGroupHom -Dist+ + + compHom : GroupEquiv ℤGroup ℤGroup + compHom = compGroupEquiv compGroup e + + l32 : fst (fst compHom) (pos 1) ≡ pos (suc (suc n)) + l32 = h3 + + l33 : pos 1 ≡ invEq (fst compHom) (pos 1) · pos (suc (suc n)) + l33 = sym (retEq (fst compHom) (pos 1)) + ∙∙ cong (invEq (fst compHom)) l32 + ∙∙ (cong (invEq (fst compHom)) (·Comm (pos 1) (pos (suc (suc n)))) + ∙ GroupHomPres· (_ , (snd (invGroupEquiv compHom))) (pos (suc (suc n))) (pos 1) + ∙ ·Comm (pos (suc (suc n))) (invEq (fst compHom) (pos 1))) + +abs→⊎ : (x : ℤ) (n : ℕ) → abs x ≡ n → (x ≡ pos n) ⊎ (x ≡ - pos n) +abs→⊎ x n = J (λ n _ → (x ≡ pos n) ⊎ (x ≡ - pos n)) (help x) + where + help : (x : ℤ) → (x ≡ pos (abs x)) ⊎ (x ≡ - pos (abs x)) + help (pos n) = inl refl + help (negsuc n) = inr refl + +⊎→abs : (x : ℤ) (n : ℕ) → (x ≡ pos n) ⊎ (x ≡ - pos n) → abs x ≡ n +⊎→abs (pos n₁) n (inl x₁) = cong abs x₁ +⊎→abs (negsuc n₁) n (inl x₁) = cong abs x₁ +⊎→abs x zero (inr x₁) = cong abs x₁ +⊎→abs x (suc n) (inr x₁) = cong abs x₁ + +JGroupEquiv : ∀ {ℓ ℓ'} {G : Group ℓ} (P : (H : Group ℓ) → GroupEquiv G H → Type ℓ') + → P G idGroupEquiv + → ∀ {H} e → P H e +JGroupEquiv {G = G} P p {H} e = + transport (λ i → P (GroupPath G H .fst e i) + (transp (λ j → GroupEquiv G (GroupPath G H .fst e (i ∨ ~ j))) i e)) + (subst (P G) (sym l) p) + where + l : (transp (λ j → GroupEquiv G (GroupPath G H .fst e (~ j))) i0 e) ≡ idGroupEquiv + l = Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (Σ≡Prop (λ _ → isPropIsEquiv _) + (funExt λ x → (λ i → fst (fst (fst e .snd .equiv-proof + (transportRefl (fst (fst e) (transportRefl x i)) i)))) + ∙ retEq (fst e) x)) + +sesIsoPresGen : + ∀ (G : Group ℓ-zero) + → (iso : GroupEquiv ℤGroup G) + → (H : Group ℓ-zero) + → (iso2 : GroupEquiv ℤGroup H) + → (e : fst G) + → invEq (fst iso) e ≡ 1 + → (hom : GroupEquiv G H) + → abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1 +sesIsoPresGen G = JGroupEquiv (λ G iso → + (H : Group ℓ-zero) + (iso2 : GroupEquiv ℤGroup H) + → (e : fst G) + → invEq (fst iso) e ≡ 1 + → (hom : GroupEquiv G H) + → abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1) + λ H → JGroupEquiv (λ H iso2 → (e : ℤ) → e ≡ pos 1 → + (hom : GroupEquiv ℤGroup H) → + abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1) + λ e p hom → cong (abs ∘ (fst (fst hom))) p ∙ autoPres1 hom + where + autoPres1 : (e : GroupEquiv ℤGroup ℤGroup) + → abs (fst (fst e) 1) ≡ 1 + autoPres1 e = grr e _ refl + + +allisoPresGen : ∀ {ℓ} (G : Group ℓ) (ϕ : GroupEquiv G ℤGroup) (x : fst G) + → (fst (fst ϕ) x ≡ 1) ⊎ (fst (fst ϕ) x ≡ - 1) + → (ψ : GroupEquiv G ℤGroup) → (fst (fst ψ) x ≡ 1) ⊎ (fst (fst ψ) x ≡ - 1) +allisoPresGen G (ϕeq , ϕhom) x (inl r) (ψeq , ψhom) = + abs→⊎ _ _ (cong abs (cong (fst ψeq) (sym (retEq ϕeq x) ∙ cong (invEq ϕeq) r)) + ∙ grr (compGroupEquiv (invGroupEquiv (ϕeq , ϕhom)) (ψeq , ψhom)) _ refl) +allisoPresGen G (ϕeq , ϕhom) x (inr r) (ψeq , ψhom) = + abs→⊎ _ _ + (cong abs (cong (fst ψeq) (sym (retEq ϕeq x) ∙ cong (invEq ϕeq) r)) + ∙ cong abs (IsGroupHom.presinv (snd (compGroupEquiv (invGroupEquiv (ϕeq , ϕhom)) (ψeq , ψhom))) 1) + ∙ absLem _ (grr (compGroupEquiv (invGroupEquiv (ϕeq , ϕhom)) (ψeq , ψhom)) _ refl)) + where + absLem : ∀ x → abs x ≡ 1 → abs (pos 0 - x) ≡ 1 + absLem (pos zero) p = ⊥-rec (snotz (sym p)) + absLem (pos (suc zero)) p = refl + absLem (pos (suc (suc n))) p = ⊥-rec (snotz (cong predℕ p)) + absLem (negsuc zero) p = refl + absLem (negsuc (suc n)) p = ⊥-rec (snotz (cong predℕ p)) + + +l3 : ∀ {ℓ} {G H : Type ℓ} → (G ≡ H) → (g₁ : coHom 2 G) (h₁ : coHom 2 H) + → (fstEq : (Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] abs (fst (fst ϕ) g₁) ≡ 1)) + → (sndEq : ((Σ[ ϕ ∈ GroupEquiv (coHomGr 2 H) ℤGroup ] abs (fst (fst ϕ) h₁) ≡ 1))) + → isEquiv {A = coHom 2 G} {B = coHom 4 G} (_⌣ g₁) + → (3rdEq : GroupEquiv (coHomGr 4 H) ℤGroup) + → abs (fst (fst 3rdEq) (h₁ ⌣ h₁)) ≡ 1 +l3 {G = G} = J (λ H _ → (g₁ : coHom 2 G) (h₁ : coHom 2 H) + → (fstEq : (Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] abs (fst (fst ϕ) g₁) ≡ 1)) + → (sndEq : ((Σ[ ϕ ∈ GroupEquiv (coHomGr 2 H) ℤGroup ] abs (fst (fst ϕ) h₁) ≡ 1))) + → isEquiv {A = coHom 2 G} {B = coHom 4 G} (_⌣ g₁) + → (3rdEq : GroupEquiv (coHomGr 4 H) ℤGroup) + → abs (fst (fst 3rdEq) (h₁ ⌣ h₁)) ≡ 1) + help + where + help : (g₁ h₁ : coHom 2 G) + → (fstEq : (Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] abs (fst (fst ϕ) g₁) ≡ 1)) + → (sndEq : ((Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] abs (fst (fst ϕ) h₁) ≡ 1))) + → isEquiv {A = coHom 2 G} {B = coHom 4 G} (_⌣ g₁) + → (3rdEq : GroupEquiv (coHomGr 4 G) ℤGroup) + → abs (fst (fst 3rdEq) (h₁ ⌣ h₁)) ≡ 1 + help g h (ϕ , idg) (ψ , idh) ⌣eq ξ = + ⊎→abs _ _ + (allisoPresGen _ + (compGroupEquiv main4 (compGroupEquiv (invGroupEquiv main4) ψ)) + h (abs→⊎ _ _ (cong abs (cong (fst (fst ψ)) (retEq (fst main4) h)) ∙ idh)) (compGroupEquiv main4 ξ)) + where + k1 : ((fst (fst ψ) h) ≡ 1) ⊎ (fst (fst ψ) h ≡ -1) + → abs (fst (fst ϕ) h) ≡ 1 + k1 p = ⊎→abs _ _ (allisoPresGen _ ψ h p ϕ) + + help2 : ((fst (fst ϕ) h) ≡ 1) ⊎ (fst (fst ϕ) h ≡ -1) + → ((fst (fst ϕ) g) ≡ 1) ⊎ (fst (fst ϕ) g ≡ -1) + → (h ≡ g) ⊎ (h ≡ (-ₕ g)) + help2 (inl x) (inl x₁) = + inl (sym (retEq (fst ϕ) h) + ∙∙ cong (invEq (fst ϕ)) (x ∙ sym x₁) + ∙∙ retEq (fst ϕ) g) + help2 (inl x) (inr x₁) = + inr (sym (retEq (fst ϕ) h) + ∙∙ cong (invEq (fst ϕ)) x + ∙ IsGroupHom.presinv (snd (invGroupEquiv ϕ)) (negsuc zero) + ∙∙ cong (-ₕ_) (cong (invEq (fst ϕ)) (sym x₁) ∙ (retEq (fst ϕ) g))) + help2 (inr x) (inl x₁) = + inr (sym (retEq (fst ϕ) h) + ∙∙ cong (invEq (fst ϕ)) x + ∙∙ (IsGroupHom.presinv (snd (invGroupEquiv ϕ)) 1 + ∙ cong (-ₕ_) (cong (invEq (fst ϕ)) (sym x₁) ∙ (retEq (fst ϕ) g)))) + help2 (inr x) (inr x₁) = + inl (sym (retEq (fst ϕ) h) + ∙∙ cong (invEq (fst ϕ)) (x ∙ sym x₁) + ∙∙ retEq (fst ϕ) g) + + -ₕeq : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → Iso (coHom n A) (coHom n A) + Iso.fun (-ₕeq n) = -ₕ_ + Iso.inv (-ₕeq n) = -ₕ_ + Iso.rightInv (-ₕeq n) = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt λ x → -ₖ^2 (f x)) + Iso.leftInv (-ₕeq n) = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt λ x → -ₖ^2 (f x)) + + theEq : coHom 2 G ≃ coHom 4 G -- (λ x → -ₕ (x ⌣ g)) + theEq = compEquiv (_ , ⌣eq) (isoToEquiv (-ₕeq 4)) + + main3 : (h ≡ g) ⊎ (h ≡ (-ₕ g)) → isEquiv {A = coHom 2 G} (_⌣ h) + main3 (inl x) = subst isEquiv (λ i → _⌣ (x (~ i))) ⌣eq + main3 (inr x) = + subst isEquiv (funExt (λ x → -ₕDistᵣ 2 2 x g) ∙ (λ i → _⌣ (x (~ i)))) + (theEq .snd) + + main4 : GroupEquiv (coHomGr 2 G) (coHomGr 4 G) + fst main4 = _ , (main3 (help2 (abs→⊎ _ _ (k1 (abs→⊎ _ _ idh))) (abs→⊎ _ _ idg))) + snd main4 = + makeIsGroupHom λ g1 g2 → rightDistr-⌣ _ _ g1 g2 h + +characFunSpaceS¹ : ∀ {ℓ} {A : Type ℓ} → + Iso (S₊ 1 → A) (Σ[ x ∈ A ] x ≡ x) +Iso.fun characFunSpaceS¹ f = f base , cong f loop +Iso.inv characFunSpaceS¹ (x , p) base = x +Iso.inv characFunSpaceS¹ (x , p) (loop i) = p i +Iso.rightInv characFunSpaceS¹ _ = refl +Iso.leftInv characFunSpaceS¹ f = funExt λ { base → refl ; (loop i) → refl} + +characFunSpace : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + → Iso (join A B → C) + (Σ[ f ∈ (A → C) ] Σ[ g ∈ (B → C) ] + ((a : A) (b : B) → f a ≡ g b)) +Iso.fun characFunSpace f = (f ∘ inl) , ((f ∘ inr) , (λ a b → cong f (push a b))) +Iso.inv characFunSpace (f , g , p) (inl x) = f x +Iso.inv characFunSpace (f , g , p) (inr x) = g x +Iso.inv characFunSpace (f , g , p) (push a b i) = p a b i +Iso.rightInv characFunSpace (f , g , p) = refl +Iso.leftInv characFunSpace f = + funExt λ { (inl x) → refl ; (inr x) → refl ; (push a b i) → refl} + +coHomS¹-ish : (n : ℕ) → Type _ +coHomS¹-ish n = hLevelTrunc 3 (S₊ 1 → coHomK (3 +ℕ n)) + +fib : (n : ℕ) → coHomS¹-ish n → Type _ +fib n x = + trRec {B = TypeOfHLevel ℓ-zero 2} (isOfHLevelTypeOfHLevel 2) + (λ f → ∥ (Σ[ g ∈ (S₊ 3 → coHomK (3 +ℕ n)) ] + ((a : S₊ 1) (b : S₊ 3) → f a ≡ g b)) ∥₂ , squash₂) x .fst + +contrFstΣ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} + → (e : isContr A) + → Iso (Σ A B) (B (fst e)) +Iso.fun (contrFstΣ {B = B} e) (a , b) = subst B (sym (snd e a)) b +Iso.inv (contrFstΣ {B = B} e) b = (fst e) , b +Iso.rightInv (contrFstΣ {B = B} e) b = cong (λ x → subst B x b) h3 ∙ transportRefl b + where + h3 : sym (snd e (fst e)) ≡ refl + h3 = isProp→isSet (isContr→isProp e) _ _ _ _ +Iso.leftInv (contrFstΣ {B = B} e) b = + ΣPathP ((snd e (fst b)) + , (λ j → transp (λ i → B (snd e (fst b) (~ i ∨ j))) j (snd b))) + +Iso1 : (n : ℕ) → Iso (coHom (3 +ℕ n) (join S¹ (S₊ 3))) ∥ Σ (coHomS¹-ish n) (fib n) ∥₂ +Iso1 n = compIso (setTruncIso characFunSpace) Iso2 + where + Iso2 : Iso _ ∥ Σ (coHomS¹-ish n) (fib n) ∥₂ + Iso.fun Iso2 = sMap λ F → ∣ fst F ∣ , ∣ (fst (snd F)) , (snd (snd F)) ∣₂ + Iso.inv Iso2 = + sRec squash₂ + (uncurry (trElim (λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelSuc 2 squash₂) + λ f → sRec squash₂ λ p → ∣ f , p ∣₂)) + Iso.rightInv Iso2 = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + (uncurry (trElim (λ _ → isOfHLevelΠ 3 λ _ → isProp→isOfHLevelSuc 2 (squash₂ _ _)) + λ f → sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl )) + Iso.leftInv Iso2 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl + +isContrcoHomS¹-ish : (n : ℕ) → isContr (coHomS¹-ish n) +isContrcoHomS¹-ish n = isOfHLevelRetractFromIso 0 (mapCompIso characFunSpaceS¹) isContrEnd + where + isContrEnd : isContr (hLevelTrunc 3 (Σ[ x ∈ coHomK (3 +ℕ n) ] x ≡ x)) + fst isContrEnd = ∣ 0ₖ _ , refl ∣ + snd isContrEnd = + trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + (uncurry (coHomK-elim _ + (λ _ → isOfHLevelΠ (suc (suc (suc n))) + λ _ → isOfHLevelPlus' {n = n} 3 (isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _)) + (λ p → (trRec (isOfHLevelPlus' {n = n} 3 (isOfHLevelTrunc 3) _ _) + (λ p i → ∣ (0ₖ (3 +ℕ n) , p i) ∣) + (Iso.fun (PathIdTruncIso _) + (isContr→isProp + (isConnectedPath _ (isConnectedKn (2 +ℕ n)) (0ₖ _) (0ₖ _)) ∣ refl ∣ ∣ p ∣)))))) + +Iso2' : (n : ℕ) → Iso (∥ Σ (coHomS¹-ish n) (fib n) ∥₂) (fib n ∣ (λ _ → 0ₖ _) ∣) +Iso2' n = compIso (setTruncIso (contrFstΣ centerChange)) + (setTruncIdempotentIso squash₂) + where + centerChange : isContr (coHomS¹-ish n) + fst centerChange = ∣ (λ _ → 0ₖ _) ∣ + snd centerChange y = isContr→isProp (isContrcoHomS¹-ish n) _ _ + +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Relation.Nullary +ll5 : (n : ℕ) → ¬ (n ≡ 2) → isContr (fib n ∣ (λ _ → 0ₖ _) ∣) +ll5 n p = + isOfHLevelRetractFromIso 0 + (compIso + (setTruncIso + (compIso (Σ-cong-iso-snd λ f → + (compIso characFunSpaceS¹ (invIso (Σ-cong-iso-fst (iso funExt⁻ funExt (λ _ → refl) λ _ → refl))))) + (compIso ΣΣ (contrFstΣ (isContrSingl _))))) + (compIso (setTruncIso (iso funExt⁻ funExt (λ _ → refl) λ _ → refl)) + (compIso (setTruncIso (codomainIso (congIso (invIso (Iso-Kn-ΩKn+1 _))))) + ((compIso (invIso (fst (coHom≅coHomΩ _ (S₊ _)))) + (fst (Hⁿ-Sᵐ≅0 n 2 p))))))) + isContrUnit + +record ExactSeqℤ {ℓ : Level} (G : ℤ → Group ℓ) : Type ℓ where + field + maps : ∀ n → GroupHom (G n) (G (sucℤ n)) + im⊂ker : ∀ n → ∀ g → isInIm (maps n) g → isInKer (maps (sucℤ n)) g + ker⊂im : ∀ n → ∀ g → isInKer (maps (sucℤ n)) g → isInIm (maps n) g + +record ExactSeqℕ {ℓ : Level} (G : ℕ → Group ℓ) : Type ℓ where + field + maps : ∀ n → GroupHom (G n) (G (suc n)) + im⊂ker : ∀ n → ∀ g → isInIm (maps n) g → isInKer (maps (suc n)) g + ker⊂im : ∀ n → ∀ g → isInKer (maps (suc n)) g → isInIm (maps n) g + + +module _ {ℓ : Level} {A : Pointed ℓ} {n : ℕ} + where + funTyp : Type _ + funTyp = A →∙ coHomK-ptd n + + _++_ : funTyp → funTyp → funTyp + fst (f ++ g) x = fst f x +ₖ fst g x + snd (f ++ g) = cong₂ _+ₖ_ (snd f) (snd g) ∙ rUnitₖ _ (0ₖ _) + +addAgree : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (x y : funTyp {A = A} {n = n}) + → Path (fst (coHomRedGrDir n A)) + ∣ x ++ y ∣₂ + (∣ x ∣₂ +ₕ∙ ∣ y ∣₂) +addAgree {A = A} zero f g = + cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) refl) +addAgree {A = A} (suc zero) f g = + cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) refl) +addAgree {A = A} (suc (suc n)) f g = + cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) refl) + +ss : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Type ℓ''} + → isHomogeneous B + → (x y : (A →∙ B)) (f : C → x ≡ y) + → isEquiv (cong fst ∘ f) + → isEquiv f +ss homb x y f e = + isoToIsEquiv + (iso _ + (λ p → invEq (_ , e) (cong fst p)) + (λ p → →∙Homogeneous≡Path homb _ _ (secEq (_ , e) (cong fst p))) + (retEq (_ , e))) + +Whitehead : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → (f : A → B) + → isEmbedding f + → (∀ (b : B) → ∃[ a ∈ A ] f a ≡ b) + → isEquiv f +equiv-proof (Whitehead f emb p) b = + pRec isPropIsContr + (λ fib → fib , isEmbedding→hasPropFibers emb b fib) + (p b) + +Int-ind : ∀ {ℓ} (P : ℤ → Type ℓ) + → P (pos zero) → P (pos 1) + → P (negsuc zero) + → ((x y : ℤ) → P x → P y → P (x + y)) → (x : ℤ) → P x +Int-ind P z one min ind (pos zero) = z +Int-ind P z one min ind (pos (suc zero)) = one +Int-ind P z one min ind (pos (suc (suc n))) = + ind (pos (suc n)) 1 (Int-ind P z one min ind (pos (suc n))) one +Int-ind P z one min ind (negsuc zero) = min +Int-ind P z one min ind (negsuc (suc n)) = + ind (negsuc n) (negsuc zero) (Int-ind P z one min ind (negsuc n)) min + +genFunSpace : (n : ℕ) → S₊∙ n →∙ coHomK-ptd n +fst (genFunSpace zero) false = 1 +fst (genFunSpace zero) true = 0 +snd (genFunSpace zero) = refl +fst (genFunSpace (suc n)) = ∣_∣ +snd (genFunSpace (suc zero)) = refl +snd (genFunSpace (suc (suc n))) = refl + +module _ where + open import Cubical.Algebra.Monoid + open import Cubical.Algebra.Semigroup + open GroupStr + open IsGroup + open IsMonoid + open IsSemigroup + πS : (n : ℕ) → Group ℓ-zero + fst (πS n) = S₊∙ n →∙ coHomK-ptd n + 1g (snd (πS n)) = (λ _ → 0ₖ n) , refl + GroupStr._·_ (snd (πS n)) = + λ f g → (λ x → fst f x +ₖ fst g x) + , cong₂ _+ₖ_ (snd f) (snd g) ∙ rUnitₖ n (0ₖ n) + inv (snd (πS n)) f = (λ x → -ₖ fst f x) , cong -ₖ_ (snd f) ∙ -0ₖ {n = n} + is-set (isSemigroup (isMonoid (isGroup (snd (πS zero))))) = + isOfHLevelΣ 2 (isSetΠ (λ _ → isSetℤ)) + λ _ → isOfHLevelPath 2 isSetℤ _ _ + is-set (isSemigroup (isMonoid (isGroup (snd (πS (suc n)))))) = isOfHLevel↑∙' 0 n + IsSemigroup.assoc (isSemigroup (isMonoid (isGroup (snd (πS n))))) x y z = + →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ w → assocₖ n (fst x w) (fst y w) (fst z w)) + fst (identity (isMonoid (isGroup (snd (πS n)))) (f , p)) = + →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → rUnitₖ n (f x)) + snd (identity (isMonoid (isGroup (snd (πS n)))) (f , p)) = + →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → lUnitₖ n (f x)) + fst (inverse (isGroup (snd (πS n))) (f , p)) = + →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → rCancelₖ n (f x)) + snd (inverse (isGroup (snd (πS n))) (f , p)) = + →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → lCancelₖ n (f x)) + + isSetπS : (n : ℕ) → isSet (S₊∙ n →∙ coHomK-ptd n) + isSetπS = λ n → is-set (snd (πS n)) + +K : (n : ℕ) → GroupIso (coHomRedGrDir n (S₊∙ n)) (πS n) +fst (K n) = setTruncIdempotentIso (isSetπS n) +snd (K zero) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 (isSetπS 0) _ _) + λ f g → →∙Homogeneous≡ (isHomogeneousKn 0) refl) +snd (K (suc zero)) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 (isSetπS 1) _ _) + λ f g → →∙Homogeneous≡ (isHomogeneousKn 1) refl) +snd (K (suc (suc n))) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 (isSetπS _) _ _) + λ f g → →∙Homogeneous≡ (isHomogeneousKn _) refl) + +t : ∀ {ℓ} {A : Pointed ℓ} → Iso ((Bool , true) →∙ A) (typ A) +Iso.fun t f = fst f false +fst (Iso.inv t a) false = a +fst (Iso.inv (t {A = A}) a) true = pt A +snd (Iso.inv t a) = refl +Iso.rightInv t a = refl +Iso.leftInv t (f , p) = + ΣPathP ((funExt (λ { false → refl ; true → sym p})) , λ i j → p (~ i ∨ j)) + +S1' : (n : ℕ) → GroupIso (πS n) ℤGroup +fst (S1' zero) = t +snd (S1' zero) = makeIsGroupHom λ _ _ → refl +S1' (suc n) = + compGroupIso + (invGroupIso (K (suc n))) + (compGroupIso + (GroupEquiv→GroupIso (coHomGr≅coHomRedGr n (S₊∙ (suc n)))) + (Hⁿ-Sⁿ≅ℤ n)) + +S1 : (n : ℕ) → Iso (S₊∙ (suc n) →∙ coHomK-ptd (suc n)) ℤ +S1 n = compIso (invIso (setTruncIdempotentIso (isOfHLevel↑∙' 0 n))) + (compIso (equivToIso (fst (coHomGr≅coHomRedGr n (S₊∙ (suc n))))) + (fst (Hⁿ-Sⁿ≅ℤ n))) + +connFunSpace : (n i : ℕ) → (x y : S₊∙ n →∙ coHomK-ptd (suc i +' n)) → ∥ x ≡ y ∥ +connFunSpace n i f g = Iso.fun PathIdTrunc₀Iso (isContr→isProp (lem n) ∣ f ∣₂ ∣ g ∣₂) + where + open import Cubical.Relation.Nullary + natLem : (n i : ℕ) → ¬ (suc (i +ℕ n) ≡ n) + natLem zero i = snotz + natLem (suc n) i p = natLem n i (sym (+-suc i n) ∙ (cong predℕ p)) + + lem : (n : ℕ) → isContr ∥ (S₊∙ n →∙ coHomK-ptd (suc i +' n)) ∥₂ + fst (lem zero) = 0ₕ∙ (suc i) + snd (lem zero) = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → pRec (squash₂ _ _) + (λ id → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn (suc i)) + (funExt λ { false → sym id ; true → sym (snd f)}))) + (help (f .fst false)) + where + help : (x : coHomK (suc i)) → ∥ x ≡ 0ₖ _ ∥ + help = coHomK-elim _ (λ _ → isProp→isOfHLevelSuc i squash) ∣ refl ∣ + lem (suc n) = + isOfHLevelRetractFromIso 0 + (compIso (equivToIso (fst (coHomGr≅coHomRedGr (suc (i +ℕ n)) (S₊∙ (suc n))))) + (fst (Hⁿ-Sᵐ≅0 (suc (i +ℕ n)) n (natLem n i)))) + isContrUnit + +S1Pres1 : (n : ℕ) → Iso.fun (fst (S1' (suc n))) (∣_∣ , refl) ≡ pos 1 +S1Pres1 zero = refl +S1Pres1 (suc n) = cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ n))) (lem n) ∙ S1Pres1 n + where + lem : (n : ℕ) → Iso.fun (fst (suspensionAx-Sn n n)) ∣ ∣_∣ ∣₂ ≡ ∣ ∣_∣ ∣₂ + lem zero = cong ∣_∣₂ (funExt λ x → transportRefl (∣ x ∣ +ₖ (0ₖ 1)) ∙ rUnitₖ 1 ∣ x ∣) + lem (suc n) = cong ∣_∣₂ + (funExt λ x → (λ i → transportRefl ((ΩKn+1→Kn (suc (suc n)) + (transp (λ j → 0ₖ (suc (suc (suc n))) ≡ ∣ merid north (~ j ∧ ~ i) ∣) i + (λ z → ∣ compPath-filler (merid (transportRefl (transportRefl x i) i)) (sym (merid north)) i z + ∣)))) i) + ∙ Iso.leftInv (Iso-Kn-ΩKn+1 (suc (suc n))) ∣ x ∣) + +S1Pres1← : (n : ℕ) → Iso.inv (fst (S1' (suc n))) (pos 1) ≡ (∣_∣ , refl) +S1Pres1← n = sym (cong (Iso.inv (fst (S1' (suc n)))) (S1Pres1 n)) ∙ Iso.leftInv (fst (S1' (suc n))) (∣_∣ , refl) + +IsoFunSpace : (n : ℕ) → Iso (S₊∙ n →∙ coHomK-ptd n) ℤ +IsoFunSpace n = fst (S1' n) + +module g-base where + g : (n : ℕ) (i : ℕ) → coHomK i → (S₊∙ n →∙ coHomK-ptd (i +' n)) + fst (g n i x) y = x ⌣ₖ (genFunSpace n) .fst y + snd (g n i x) = cong (x ⌣ₖ_) ((genFunSpace n) .snd) ∙ ⌣ₖ-0ₖ i n x + + G : (n : ℕ) (i : ℕ) → coHomK i → (S₊∙ n →∙ coHomK-ptd (n +' i)) + fst (G n i x) y = (genFunSpace n) .fst y ⌣ₖ x + snd (G n i x) = cong (_⌣ₖ x) ((genFunSpace n) .snd) ∙ 0ₖ-⌣ₖ n i x + + eq1 : (n : ℕ) (i : ℕ) → (S₊∙ n →∙ coHomK-ptd (i +' n)) ≃ (S₊∙ n →∙ coHomK-ptd (i +' n)) + eq1 n i = isoToEquiv (iso F F FF FF) + where + lem : (i n : ℕ) → (-ₖ^ i · n) (snd (coHomK-ptd (i +' n))) ≡ 0ₖ _ + lem zero zero = refl + lem zero (suc zero) = refl + lem zero (suc (suc n)) = refl + lem (suc zero) zero = refl + lem (suc (suc i)) zero = refl + lem (suc i) (suc n) = refl + + F : S₊∙ n →∙ coHomK-ptd (i +' n) → S₊∙ n →∙ coHomK-ptd (i +' n) + fst (F f) x = (-ₖ^ i · n) (fst f x) + snd (F f) = cong (-ₖ^ i · n) (snd f) ∙ lem i n + + FF : (x : _) → F (F x) ≡ x + FF x = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ y → -ₖ-gen² i n _ _ (fst x y)) + + rCancel'' : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) → sym p ∙∙ refl ∙∙ p ≡ refl + rCancel'' p = (λ j → (λ i → p (~ i ∨ j)) ∙∙ refl ∙∙ λ i → p (i ∨ j)) ∙ sym (rUnit refl) + + transpPres0 : ∀ {k m : ℕ} (p : k ≡ m) → subst coHomK p (0ₖ k) ≡ 0ₖ m + transpPres0 {k = k} = J (λ m p → subst coHomK p (0ₖ k) ≡ 0ₖ m) (transportRefl _) + + eq2 : (n : ℕ) (i : ℕ) → (S₊∙ n →∙ coHomK-ptd (n +' i)) ≃ (S₊∙ n →∙ coHomK-ptd (i +' n)) + eq2 n i = + isoToEquiv (iso (λ f → (λ x → subst coHomK (+'-comm n i) (fst f x)) , + cong (subst coHomK (+'-comm n i)) (snd f) ∙ transpPres0 (+'-comm n i)) + (λ f → (λ x → subst coHomK (sym (+'-comm n i)) (fst f x)) + , (cong (subst coHomK (sym (+'-comm n i))) (snd f) ∙ transpPres0 (sym (+'-comm n i)))) + (λ f → →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ x → transportTransport⁻ _ (f .fst x))) + λ f → →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ x → transportTransport⁻ _ (f .fst x))) + + g≡ : (n : ℕ) (i : ℕ) → g n i ≡ λ x → fst (compEquiv (eq2 n i) (eq1 n i)) ((G n i) x) + g≡ n i = funExt (λ f → →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ y → gradedComm-⌣ₖ _ _ f (genFunSpace n .fst y))) + + glIso3-fun : (n m : ℕ) → + (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) + → (S₊ n → coHomK m) + glIso3-fun zero m (f , p) false = ΩKn+1→Kn _ (sym p ∙∙ cong f loop ∙∙ p) + glIso3-fun zero m (f , p) true = 0ₖ _ + glIso3-fun (suc n) m (f , p) x = + ΩKn+1→Kn _ (sym p ∙∙ cong f (merid x ∙ sym (merid (ptSn _))) ∙∙ p) + + glIso3-fun∙ : (n m : ℕ) → (f : _) → glIso3-fun n m f (ptSn _) ≡ 0ₖ m + glIso3-fun∙ zero m = λ _ → refl + glIso3-fun∙ (suc n) m (f , p) = + cong (ΩKn+1→Kn m) (cong (sym p ∙∙_∙∙ p) (cong (cong f) (rCancel (merid (ptSn _))))) + ∙∙ cong (ΩKn+1→Kn m) (rCancel'' p) + ∙∙ ΩKn+1→Kn-refl m + + + glIso3-inv : (n m : ℕ) → (S₊∙ n →∙ coHomK-ptd m) → (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) + fst (glIso3-inv zero m (f , p)) base = 0ₖ _ + fst (glIso3-inv zero m (f , p)) (loop i) = Kn→ΩKn+1 m (f false) i + snd (glIso3-inv zero m (f , p)) = refl + fst (glIso3-inv (suc n) m (f , p)) north = 0ₖ _ + fst (glIso3-inv (suc n) m (f , p)) south = 0ₖ _ + fst (glIso3-inv (suc n) m (f , p)) (merid a i) = Kn→ΩKn+1 m (f a) i + snd (glIso3-inv (suc n) m (f , p)) = refl + + glIso3 : (n m : ℕ) → + Iso (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) + (S₊∙ n →∙ coHomK-ptd m) + Iso.fun (glIso3 n m) f = (glIso3-fun n m f) , (glIso3-fun∙ n m f) + Iso.inv (glIso3 n m) = glIso3-inv n m + Iso.rightInv (glIso3 zero m) (f , p) = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { false → cong (ΩKn+1→Kn m) (sym (rUnit _)) ∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f false) + ; true → sym p}) + Iso.rightInv (glIso3 (suc n) m) (f , p) = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ x → (λ i → ΩKn+1→Kn m (sym (rUnit (cong-∙ (glIso3-inv (suc n) m (f , p) .fst) (merid x) (sym (merid (ptSn _))) i)) i)) + ∙∙ cong (ΩKn+1→Kn m) (cong (Kn→ΩKn+1 m (f x) ∙_) (cong sym (cong (Kn→ΩKn+1 m) p ∙ Kn→ΩKn+10ₖ m)) ∙ sym (rUnit _)) + ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f x)) + Iso.leftInv (glIso3 zero m) (f , p) = →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { base → sym p + ; (loop i) j → h3 j i}) + where + h3 : PathP (λ i → p (~ i) ≡ p (~ i)) (Kn→ΩKn+1 m (ΩKn+1→Kn m (sym p ∙∙ cong f loop ∙∙ p))) (cong f loop) + h3 = Iso.rightInv (Iso-Kn-ΩKn+1 _) _ + ◁ λ i → doubleCompPath-filler (sym p) (cong f loop) p (~ i) + Iso.leftInv (glIso3 (suc n) m) (f , p) = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { north → sym p + ; south → sym p ∙ cong f (merid (ptSn _)) + ; (merid a i) j → h3 a j i}) + where + h3 : (a : S₊ (suc n)) → PathP (λ i → p (~ i) ≡ (sym p ∙ cong f (merid (ptSn (suc n)))) i) + (Kn→ΩKn+1 m (ΩKn+1→Kn m (sym p ∙∙ cong f (merid a ∙ sym (merid (ptSn _))) ∙∙ p))) + (cong f (merid a)) + h3 a = Iso.rightInv (Iso-Kn-ΩKn+1 _) _ + ◁ λ i j → hcomp (λ k → λ { (i = i1) → (f (merid a j)) + ; (j = i0) → p (k ∧ ~ i) + ; (j = i1) → compPath-filler' (sym p) (cong f (merid (ptSn _))) k i }) + (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) + + glIsoInvHom : (n m : ℕ) (x y : coHomK n) (z : S₊ (suc m)) + → Iso.inv (glIso3 _ _) (G m n (x +ₖ y)) .fst z + ≡ Iso.inv (glIso3 _ _) (G m n x) .fst z + +ₖ Iso.inv (glIso3 _ _) (G m n y) .fst z + glIsoInvHom zero zero x y base = refl + glIsoInvHom (suc n) zero x y base = refl + glIsoInvHom zero zero x y (loop i) j = h3 j i + where + h3 : (cong (Iso.inv (glIso3 _ _) (G zero zero (x + y)) .fst) loop) + ≡ cong₂ _+ₖ_ (cong (Iso.inv (glIso3 _ _) (G zero zero x) .fst) loop) + (cong (Iso.inv (glIso3 _ _) (G zero zero y) .fst) loop) + h3 = Kn→ΩKn+1-hom 0 x y + ∙ ∙≡+₁ (cong (Iso.inv (glIso3 _ _) (G zero zero x) .fst) loop) + (cong (Iso.inv (glIso3 _ _) (G zero zero y) .fst) loop) + glIsoInvHom (suc n) zero x y (loop i) j = h3 j i + where + h3 : Kn→ΩKn+1 (suc n) ((pos (suc zero)) ·₀ (x +ₖ y)) + ≡ cong₂ _+ₖ_ (cong (Iso.inv (glIso3 zero (zero +' suc n)) (G zero (suc n) x) .fst) loop) + (cong (Iso.inv (glIso3 zero (zero +' suc n)) (G zero (suc n) y) .fst) loop) + h3 = cong (Kn→ΩKn+1 (suc n)) (lUnit⌣ₖ _ (x +ₖ y)) + ∙∙ Kn→ΩKn+1-hom (suc n) x y + ∙∙ (λ i → ∙≡+₂ n (Kn→ΩKn+1 (suc n) (lUnit⌣ₖ _ x (~ i))) + (Kn→ΩKn+1 (suc n) (lUnit⌣ₖ _ y (~ i))) i) + glIsoInvHom zero (suc m) x y north = refl + glIsoInvHom zero (suc m) x y south = refl + glIsoInvHom zero (suc m) x y (merid a i) j = h3 j i + where + h3 : Kn→ΩKn+1 (suc m) (_⌣ₖ_ {n = suc m} {m = zero} ∣ a ∣ (x + y)) + ≡ cong₂ _+ₖ_ (Kn→ΩKn+1 (suc m) (_⌣ₖ_ {n = suc m} {m = zero} ∣ a ∣ x)) + (Kn→ΩKn+1 (suc m) (_⌣ₖ_ {n = suc m} {m = zero} ∣ a ∣ y)) + h3 = cong (Kn→ΩKn+1 (suc m)) (leftDistr-⌣ₖ (suc m) 0 ∣ a ∣ x y) + ∙∙ Kn→ΩKn+1-hom (suc m) _ _ + ∙∙ ∙≡+₂ _ _ _ + glIsoInvHom (suc n) (suc m) x y north = refl + glIsoInvHom (suc n) (suc m) x y south = refl + glIsoInvHom (suc n) (suc m) x y (merid a i) j = h3 j i + where + h3 : Kn→ΩKn+1 (suc (suc (m +ℕ n))) (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ (x +ₖ y)) + ≡ cong₂ _+ₖ_ (Kn→ΩKn+1 (suc (suc (m +ℕ n))) (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ x)) + (Kn→ΩKn+1 (suc (suc (m +ℕ n))) (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ y)) + h3 = cong (Kn→ΩKn+1 (suc (suc (m +ℕ n)))) + (leftDistr-⌣ₖ (suc m) (suc n) ∣ a ∣ x y) + ∙∙ Kn→ΩKn+1-hom _ _ _ + ∙∙ ∙≡+₂ _ _ _ + + +'-suc : (n m : ℕ) → suc n +' m ≡ suc (n +' m) + +'-suc zero zero = refl + +'-suc (suc n) zero = refl + +'-suc zero (suc m) = refl + +'-suc (suc n) (suc m) = refl + + LEM : (i n : ℕ) (x : coHomK i) + → Path _ (G (suc n) i x) + (subst (λ x → S₊∙ (suc n) →∙ coHomK-ptd x) + (sym (+'-suc n i)) + ((Iso.inv (glIso3 n _)) (G n i x))) + LEM zero zero x = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ z → (λ i → x ·₀ ∣ z ∣) ∙ h3 x z ∙ sym (transportRefl _)) + where + h3 : (x : ℤ) (z : S¹) → _·₀_ {n = 1} x ∣ z ∣ ≡ fst (Iso.inv (glIso3 0 zero) (G zero zero x)) z + h3 = Int-ind _ (λ { base → refl ; (loop i) j → Kn→ΩKn+10ₖ zero (~ j) i}) + (λ { base → refl ; (loop i) j → rUnit (cong ∣_∣ₕ (lUnit loop j)) j i}) + (λ { base → refl ; (loop i) j → rUnit (cong ∣_∣ₕ (sym loop)) j i}) + λ x y inx iny z + → rightDistr-⌣ₖ 0 1 x y ∣ z ∣ + ∙∙ cong₂ _+ₖ_ (inx z) (iny z) + ∙∙ sym (glIsoInvHom zero zero x y z) + LEM (suc i) zero x = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ z → h3 z + ∙ sym (transportRefl + ((Iso.inv (glIso3 zero (suc i)) (G zero (suc i) x)) .fst z))) + where + h3 : (z : S₊ 1) → _ ≡ Iso.inv (glIso3 zero (suc i)) (G zero (suc i) x) .fst z + h3 base = refl + h3 (loop k) j = Kn→ΩKn+1 (suc i) (lUnit⌣ₖ _ x (~ j)) k + LEM zero (suc n) x = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ z → h3 x z ∙ λ i → transportRefl (Iso.inv (glIso3 (suc n) (suc n)) (G (suc n) zero x)) (~ i) .fst z) + where + +merid : rUnitₖ (suc (suc n)) ∣ south ∣ ≡ cong ∣_∣ₕ (merid (ptSn _)) + +merid = sym (lUnit _) + ∙ cong (cong ∣_∣ₕ) + λ j i → transp (λ _ → S₊ (suc (suc n))) (i ∨ j) (merid (ptSn (suc n)) i) + open import Cubical.Foundations.Path + + pp : (a : S₊ (suc n)) → PathP (λ i → ∣ merid a i ∣ₕ ≡ Kn→ΩKn+1 (suc n) (∣ a ∣ +ₖ (0ₖ _)) i) + refl (sym (rUnitₖ (suc (suc n)) ∣ south ∣)) + pp a = flipSquare ((λ i j → ∣ compPath-filler (merid a) (sym (merid (ptSn _))) i j ∣ₕ ) + ▷ cong (Kn→ΩKn+1 (suc n)) (sym (rUnitₖ (suc n) ∣ a ∣ₕ))) + ▷ sym (cong sym +merid) + + pp2 : (a : S₊ (suc n)) → (λ i → -ₖ ∣ merid a i ∣) + ≡ Kn→ΩKn+1 (suc n) (-ₖ ∣ a ∣) + pp2 a = cong (cong ∣_∣ₕ) (sym (symDistr (merid a) (sym (merid (ptSn (suc n)))))) + ∙ sym (Kn→ΩKn+1-ₖ (suc n) ∣ a ∣) + + h3 : (x : ℤ) (z : S₊ (suc (suc n))) + → _⌣ₖ_ {n = suc (suc n)} {m = 0} ∣ z ∣ x + ≡ Iso.inv (glIso3 (suc n) (suc n)) (G (suc n) zero x) .fst z + h3 = Int-ind _ + (λ { north → refl ; south → refl ; (merid a i) j → Kn→ΩKn+10ₖ (suc n) (~ j) i}) + (λ { north → refl ; south → refl + ; (merid a i) j → hcomp (λ k → λ { (i = i0) → ∣ north ∣ + ; (i = i1) → rUnitₖ (suc (suc n)) ∣ south ∣ (~ k) + ; (j = i0) → rUnitₖ (suc (suc n)) ∣ merid a i ∣ (~ k) + ; (j = i1) → pp a i k}) + ∣ merid a i ∣ₕ}) + (λ { north → refl + ; south → refl + ; (merid a i) j → pp2 a j i}) + λ x y indx indy z → leftDistr-⌣ₖ _ _ ∣ z ∣ x y + ∙ cong₂ _+ₖ_ (indx z) (indy z) + ∙ sym (glIsoInvHom _ _ _ _ _) + LEM (suc i) (suc n) x = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ z → h3 z + ∙ λ j → transportRefl ((Iso.inv (glIso3 (suc n) (suc (suc (n +ℕ i)))) + (G (suc n) (suc i) x))) (~ j) .fst z) + where + h3 : (z : S₊ (suc (suc n))) → _ + h3 north = refl + h3 south = refl + h3 (merid a i) = refl + + isEquivGzero : (i : ℕ) → isEquiv (G zero i) + isEquivGzero i = + isoToIsEquiv + (iso _ (λ f → fst f false) + (λ {(f , p) → →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { false → rUnitₖ _ (f false) ; true → sym p})}) + (lUnit⌣ₖ _)) + + isEquivG : (n i : ℕ) → isEquiv (G n i) + isEquivG zero i = + isoToIsEquiv + (iso _ (λ f → fst f false) + (λ {(f , p) → →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { false → rUnitₖ _ (f false) ; true → sym p})}) + (lUnit⌣ₖ _)) + isEquivG (suc n) i = + subst isEquiv (sym (funExt (LEM i n))) + (compEquiv (compEquiv (G n i , isEquivG n i) + (isoToEquiv (invIso (glIso3 n (n +' i))))) + (pathToEquiv (λ j → S₊∙ (suc n) →∙ coHomK-ptd (+'-suc n i (~ j)))) .snd) + + + isEquiv-g : (n i : ℕ) → isEquiv (g n i) + isEquiv-g n i = + subst isEquiv (sym (g≡ n i)) + (compEquiv (G n i , isEquivG n i) (compEquiv (eq2 n i) (eq1 n i)) .snd) + +module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) + (conB : (x y : typ B) → ∥ x ≡ y ∥) + (n : ℕ) (Q-is : Iso (typ (Q (pt B))) (S₊ n)) + (Q-is-ptd : Iso.fun Q-is (pt (Q (pt B))) ≡ snd (S₊∙ n)) + (c : (b : typ B) → (Q b →∙ coHomK-ptd n)) + (c-pt : c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) where + + g : (b : typ B) (i : ℕ) → coHomK i → (Q b →∙ coHomK-ptd (i +' n)) + fst (g b i x) y = x ⌣ₖ c b .fst y + snd (g b i x) = cong (x ⌣ₖ_) (c b .snd) ∙ ⌣ₖ-0ₖ i n x + + g' : (b : typ B) (i : ℕ) → coHomK i → coHomK i → (Q b →∙ coHomK-ptd (i +' n)) + fst (g' b i x y) z = x ⌣ₖ c b .fst z +ₖ y ⌣ₖ c b .fst z + snd (g' b i x y) = cong₂ _+ₖ_ (cong (x ⌣ₖ_) (c b .snd) ∙ ⌣ₖ-0ₖ i n x) + (cong (y ⌣ₖ_) (c b .snd) ∙ ⌣ₖ-0ₖ i n y) ∙ rUnitₖ _ (0ₖ _) + + g-hom : (b : typ B) (i : ℕ) → (x y : coHomK i) → g b i (x +ₖ y) ≡ ((g b i x) ++ (g b i y)) + g-hom b i x y = →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ z → rightDistr-⌣ₖ i n x y (c b .fst z)) + + gPathP' : (i : ℕ) → PathP (λ j → coHomK i → (isoToPath Q-is j , ua-gluePath (isoToEquiv Q-is) (Q-is-ptd) j) + →∙ coHomK-ptd (i +' n)) (g (pt B) i) (g-base.g n i) + gPathP' i = + toPathP + (funExt + λ x → →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ y → (λ i → transportRefl (transportRefl x i ⌣ₖ c (pt B) .fst (Iso.inv Q-is (transportRefl y i))) i) + ∙ cong (x ⌣ₖ_) (funExt⁻ c-pt (Iso.inv Q-is y) ∙ cong (genFunSpace n .fst) (Iso.rightInv Q-is y)))) + + g-base : (i : ℕ) → isEquiv (g (pt B) i) + g-base i = transport (λ j → isEquiv (gPathP' i (~ j))) (g-base.isEquiv-g n i) + + g-equiv : (b : typ B) (i : ℕ) → isEquiv (g b i) + g-equiv b i = + pRec (isPropIsEquiv _) + (J (λ b _ → isEquiv (g b i)) + (g-base i)) + (conB (pt B) b) + +module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) + (conB : (x y : typ B) → ∥ x ≡ y ∥₂) + (n : ℕ) (Q-is : Iso (typ (Q (pt B))) (S₊ n)) + (Q-is-ptd : Iso.fun Q-is (pt (Q (pt B))) ≡ snd (S₊∙ n)) where + + is-setQ→K : (b : typ B) → isSet (Q b →∙ coHomK-ptd n) + is-setQ→K b = + sRec (isProp→isOfHLevelSuc 1 isPropIsSet) + (J (λ b _ → isSet (Q b →∙ coHomK-ptd n)) + (subst isSet (cong (_→∙ coHomK-ptd n) + (ua∙ (isoToEquiv (invIso Q-is)) (cong (Iso.inv Q-is) (sym Q-is-ptd) ∙ Iso.leftInv Q-is _))) + (isOfHLevelRetractFromIso 2 (fst (S1' n)) isSetℤ))) + (conB (pt B) b) + + + isConnB : isConnected 3 (typ B) + fst isConnB = ∣ pt B ∣ + snd isConnB = + trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) + λ a → sRec (isOfHLevelTrunc 3 _ _) (cong ∣_∣ₕ) (conB (pt B) a) + + isPropPath : isProp (∥ pt B ≡ pt B ∥₂) + isPropPath = + isOfHLevelRetractFromIso 1 setTruncTrunc2Iso + (isContr→isProp (isConnectedPath _ isConnB (pt B) (pt B))) + + c* : Σ[ c ∈ ((b : typ B) → (Q b →∙ coHomK-ptd n)) ] + (c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) + fst c* b = + sRec (is-setQ→K b) + (J (λ b _ → Q b →∙ coHomK-ptd n) + ((λ x → genFunSpace n .fst (Iso.fun Q-is x)) + , cong (genFunSpace n .fst) Q-is-ptd ∙ genFunSpace n .snd)) (conB (pt B) b) + snd c* = + funExt λ x → (λ i → sRec (is-setQ→K (pt B)) + (J (λ b _ → Q b →∙ coHomK-ptd n) + ((λ x₁ → genFunSpace n .fst (Iso.fun Q-is x₁)) , + (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) + (isPropPath (conB (pt B) (pt B)) ∣ refl ∣₂ i) .fst x) + ∙ (λ i → transportRefl (genFunSpace n .fst (Iso.fun Q-is (transportRefl x i))) i) + + p-help : {b : fst B} (p : pt B ≡ b) → (subst (fst ∘ Q) (sym p) (snd (Q b))) ≡ (snd (Q (pt B))) + p-help {b = b} = + J (λ b p → subst (fst ∘ Q) (sym p) (snd (Q b)) ≡ snd (Q (pt B))) (transportRefl _) + + cEquiv : (b : fst B) (p : ∥ pt B ≡ b ∥₂) + → (c* .fst b) + ≡ sRec (is-setQ→K b) + (λ pp → (λ qb → genFunSpace n .fst (Iso.fun Q-is (subst (fst ∘ Q) (sym pp) qb))) + , cong (genFunSpace n .fst ∘ Iso.fun Q-is) (p-help pp) + ∙ ((λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) p + cEquiv b = + sElim (λ _ → isOfHLevelPath 2 (is-setQ→K b) _ _) + (J (λ b a → c* .fst b ≡ + sRec (is-setQ→K b) (λ pp → + (λ qb → + genFunSpace n .fst (Iso.fun Q-is (subst (fst ∘ Q) (sym pp) qb))) + , + cong (genFunSpace n .fst ∘ Iso.fun Q-is) (p-help pp) ∙ + (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd) + ∣ a ∣₂) + ((λ i → sRec (is-setQ→K (pt B)) + (J (λ b₁ _ → Q b₁ →∙ coHomK-ptd n) + ((λ x → genFunSpace n .fst (Iso.fun Q-is x)) , + (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) + (isPropPath (conB (pt B) (pt B)) ∣ refl ∣₂ i)) + ∙ →∙Homogeneous≡ (isHomogeneousKn n) + (transportRefl ((λ x → genFunSpace n .fst (Iso.fun Q-is x))) + ∙ funExt λ x → cong (genFunSpace n .fst ∘ Iso.fun Q-is) (sym (transportRefl x))))) + +module _ {ℓ ℓ'} (B : Pointed ℓ) (P : typ B → Type ℓ') where + E : Type _ + E = Σ (typ B) P + + Ẽ : Type _ + Ẽ = Pushout {A = E} (λ _ → tt) fst + + i : (n : ℕ) → (P-base : Iso (P (pt B)) (S₊ n)) → S₊ (suc n) → Ẽ + i zero P-base base = inr (pt B) + i zero P-base (loop j) = (sym (push (pt B , Iso.inv P-base false)) + ∙ push ((pt B) , Iso.inv P-base true)) j + i (suc n) P-base north = inl tt + i (suc n) P-base south = inr (pt B) + i (suc n) P-base (merid a i₁) = push (pt B , Iso.inv P-base a) i₁ + + Q : typ B → Pointed ℓ' + Q x = Susp (P x) , north + + F : Type _ + F = Σ (typ B) λ x → Q x .fst + + F̃ : Type _ + F̃ = Pushout {A = typ B} {C = F} (λ _ → tt) λ b → b , north + + invFE : Ẽ → F̃ + invFE (inl x) = inl tt + invFE (inr x) = inr (x , south) + invFE (push (x , a) i₁) = ((push x) ∙ λ i → inr (x , merid a i)) i₁ + + funFE : F̃ → Ẽ + funFE (inl x) = inl tt + funFE (inr (x , north)) = inl tt + funFE (inr (x , south)) = inr x + funFE (inr (x , merid a i₁)) = push (x , a) i₁ + funFE (push a i₁) = inl tt + + IsoFE : Iso F̃ Ẽ + Iso.fun IsoFE = funFE + Iso.inv IsoFE = invFE + Iso.rightInv IsoFE (inl x) = refl + Iso.rightInv IsoFE (inr x) = refl + Iso.rightInv IsoFE (push (x , a) i₁) k = h3 k i₁ + where + h3 : cong funFE (((push x) ∙ λ i → inr (x , merid a i))) + ≡ push (x , a) + h3 = congFunct funFE (push x) (λ i → inr (x , merid a i)) + ∙ sym (lUnit (push (x , a))) + Iso.leftInv IsoFE (inl x) = refl + Iso.leftInv IsoFE (inr (x , north)) = push x + Iso.leftInv IsoFE (inr (x , south)) = refl + Iso.leftInv IsoFE (inr (x , merid a i)) j = + compPath-filler' (push x) (λ i₁ → inr (x , merid a i₁)) (~ j) i + Iso.leftInv IsoFE (push a i₁) k = push a (i₁ ∧ k) + + + funDir1 : ∀ {ℓ} {A : Pointed ℓ} → (F̃ , inl tt) →∙ A → (b : typ B) → Q b →∙ A + fst (funDir1 {A = A , a} (f , p) b) north = f (inr (b , north)) + fst (funDir1 {A = A , a} (f , p) b) south = f (inr (b , south)) + fst (funDir1 {A = A , a} (f , p) b) (merid a₁ i₁) = f (inr (b , merid a₁ i₁)) + snd (funDir1 {A = A , a} (f , p) b) = sym (cong f (push b)) ∙ p + + funDir2 : ∀ {ℓ} {A : Pointed ℓ} → ((b : typ B) → Q b →∙ A) → (F̃ , inl tt) →∙ A + fst (funDir2 {A = A , a} f) (inl x) = a + fst (funDir2 {A = A , a} f) (inr (x , north)) = f x .fst north + fst (funDir2 {A = A , a} f) (inr (x , south)) = f x .fst south + fst (funDir2 {A = A , a} f) (inr (x , merid a₁ i₁)) = f x .fst (merid a₁ i₁) + fst (funDir2 {A = A , a} f) (push a₁ i₁) = snd (f a₁) (~ i₁) + snd (funDir2 {A = A , a} f) = refl + + funDir2-hom : (n : ℕ) → (f g : ((b : typ B) → Q b →∙ coHomK-ptd n)) + → funDir2 (λ b → f b ++ g b) ≡ (funDir2 f ++ funDir2 g) + funDir2-hom n f g = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { (inl x) → sym (rUnitₖ _ (0ₖ _)) + ; (inr (x , north)) → refl + ; (inr (x , south)) → refl + ; (inr (x , merid a i₁)) → refl + ; (push a j) i → compPath-filler (cong₂ _+ₖ_ (snd (f a)) (snd (g a))) + (rUnitₖ n (0ₖ n)) (~ i) (~ j)}) + + funDirSect : ∀ {ℓ} {A : Pointed ℓ} → (x : (b : typ B) → Q b →∙ A) (b : typ B) (q : Q b .fst) + → funDir1 (funDir2 x) b .fst q ≡ x b .fst q + funDirSect f b north = refl + funDirSect f b south = refl + funDirSect f b (merid a i₁) = refl + + funDirRetr : ∀ {ℓ} {A : Pointed ℓ} (f : F̃ → typ A) (p : _) + → (x : F̃) → fst (funDir2 {A = A} (funDir1 (f , p))) x ≡ f x + funDirRetr f p (inl x) = sym p + funDirRetr f p (inr (x , north)) = refl + funDirRetr f p (inr (x , south)) = refl + funDirRetr f p (inr (x , merid a i₁)) = refl + funDirRetr f p (push a i₁) j = compPath-filler (sym (cong f (push a))) p (~ j) (~ i₁) + + Iso2 : ∀ {ℓ} {A : Pointed ℓ} + → Iso ((F̃ , inl tt) →∙ A) + ((b : typ B) → Q b →∙ A) + Iso.fun (Iso2 {A = A , a}) = funDir1 + Iso.inv (Iso2 {A = A , a}) = funDir2 + Iso.rightInv (Iso2 {A = A , a}) f = + funExt λ b → ΣPathP (funExt (funDirSect f b) + , sym (rUnit (snd (f b)))) + Iso.leftInv (Iso2 {A = A , a}) (f , p) = + ΣPathP ((funExt (funDirRetr f p)) + , λ i j → p (~ i ∨ j)) + + ι : (k : ℕ) → Iso ((b : typ B) → Q b →∙ coHomK-ptd k) + ((Ẽ , inl tt) →∙ coHomK-ptd k) + ι k = compIso (invIso Iso2) h3 + where + h3 : Iso ((F̃ , inl tt) →∙ coHomK-ptd k) + ((Ẽ , inl tt) →∙ coHomK-ptd k) + Iso.fun h3 G = (λ x → G .fst (Iso.inv IsoFE x)) + , (snd G) + Iso.inv h3 G = (λ x → G .fst (Iso.fun IsoFE x)) + , (snd G) + Iso.rightInv h3 G = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ x → cong (G .fst) (Iso.rightInv IsoFE x)) + Iso.leftInv h3 G = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ x → cong (G .fst) (Iso.leftInv IsoFE x)) + + ι-hom : (k : ℕ) → (f g : ((b : typ B) → Q b →∙ coHomK-ptd k)) + → Iso.fun (ι k) (λ b → f b ++ g b) + ≡ (Iso.fun (ι k) f ++ Iso.fun (ι k) g) + ι-hom k f g = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ x → funExt⁻ (cong fst (funDir2-hom _ f g)) (invFE x)) + +codomainIsoDep : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : A → Type ℓ'} {C : A → Type ℓ''} + → ((a : A) → Iso (B a) (C a)) + → Iso ((a : A) → B a) ((a : A) → C a) +Iso.fun (codomainIsoDep is) f a = Iso.fun (is a) (f a) +Iso.inv (codomainIsoDep is) f a = Iso.inv (is a) (f a) +Iso.rightInv (codomainIsoDep is) f = funExt λ a → Iso.rightInv (is a) (f a) +Iso.leftInv (codomainIsoDep is) f = funExt λ a → Iso.leftInv (is a) (f a) + +module Thom {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) + (conB : (x y : typ B) → ∥ x ≡ y ∥) + (n : ℕ) (Q-is : Iso (typ (Q B P (pt B))) (S₊ n)) + (Q-is-ptd : Iso.fun Q-is (pt (Q B P (pt B))) ≡ snd (S₊∙ n)) + (c : (b : typ B) → (Q B P b →∙ coHomK-ptd n)) + (c-pt : c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) where + + ϕ : (i : ℕ) → GroupEquiv (coHomGr i (typ B)) + (coHomRedGrDir (i +' n) (Ẽ B P , inl tt)) + fst (ϕ i) = + isoToEquiv + (setTruncIso + (compIso + (codomainIsoDep + λ b → equivToIso ((g B (Q B P) conB n Q-is Q-is-ptd c c-pt b i) + , g-equiv B (Q B P) conB n Q-is Q-is-ptd c c-pt b i)) + (ι B P (i +' n)))) + snd (ϕ i) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ F G → cong ∣_∣₂ (cong (Iso.fun (ι B P (i +' n))) + (funExt (λ a → g-hom B (Q B P) conB n Q-is Q-is-ptd c c-pt a i (F a) (G a))) + ∙ ι-hom B P (i +' n) _ _) + ∙ addAgree (i +' n) _ _) + +module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) + (conB : (x y : typ B) → ∥ x ≡ y ∥₂) + (n : ℕ) (Q-is : Iso (typ (Q B P (pt B))) (S₊ n)) + (Q-is-ptd : Iso.fun Q-is (pt (Q B P (pt B))) ≡ snd (S₊∙ n)) where + + 0-connB : (x y : typ B) → ∥ x ≡ y ∥ + 0-connB x y = sRec (isProp→isSet squash) (∥_∥.∣_∣) (conB x y) + + c = (c* B (Q B P) conB n Q-is Q-is-ptd .fst) + c-ptd = (c* B (Q B P) conB n Q-is Q-is-ptd .snd) + cEq = cEquiv B (Q B P) conB n Q-is Q-is-ptd + + module w = Thom B P 0-connB n Q-is Q-is-ptd c c-ptd + + ϕ = w.ϕ + + E' = E B P + + E'̃ = Ẽ B P + + p : E' → typ B + p = fst + + e : coHom n (typ B) + e = ∣ (λ b → c b .fst south) ∣₂ + + ⌣-hom : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr (i +' n) (typ B)) + fst (⌣-hom i) x = x ⌣ e + snd (⌣-hom i) = + makeIsGroupHom λ f g → rightDistr-⌣ _ _ f g e + + p-hom : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr i E') + p-hom i = coHomMorph _ p + + E-susp : (i : ℕ) → GroupHom (coHomGr i E') (coHomRedGrDir (suc i) (E'̃ , inl tt)) + fst (E-susp i) = sMap λ f → (λ { (inl x) → 0ₖ _ + ; (inr x) → 0ₖ _ + ; (push a j) → Kn→ΩKn+1 _ (f a) j}) , refl + snd (E-susp zero) = + makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f g → + cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn 1) + (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a j) k → (Kn→ΩKn+1-hom zero (f a) (g a) + ∙ ∙≡+₁ (Kn→ΩKn+1 _ (f a)) (Kn→ΩKn+1 _ (g a))) k j}))) + snd (E-susp (suc i)) = + makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f g → + cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a j) k → (Kn→ΩKn+1-hom _ (f a) (g a) + ∙ ∙≡+₂ _ (Kn→ΩKn+1 _ (f a)) (Kn→ΩKn+1 _ (g a))) k j}))) + + module cofibSeq where + j* : (i : ℕ) → GroupHom (coHomRedGrDir i (E'̃ , (inl tt))) (coHomGr i (typ B)) + fst (j* i) = sMap λ f → λ x → fst f (inr x) + snd (j* zero) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) + snd (j* (suc zero)) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) + snd (j* (suc (suc i₁))) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) + + Im-j⊂Ker-p : (i : ℕ) (x : _) → isInIm (j* i) x → isInKer (p-hom i) x + Im-j⊂Ker-p i = + sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ f → pRec (squash₂ _ _) + (uncurry (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ g P → subst (isInKer (p-hom i)) P + (cong ∣_∣₂ (funExt λ x → cong (g .fst) (sym (push x)) ∙ g .snd)))) + + Ker-p⊂Im-j : (i : ℕ) (x : _) → isInKer (p-hom i) x → isInIm (j* i) x + Ker-p⊂Im-j i = + sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) + λ f ker → pRec squash + (λ id → ∣ ∣ (λ { (inl x) → 0ₖ _ + ; (inr x) → f x + ; (push a i₁) → funExt⁻ id a (~ i₁)}) , refl ∣₂ , refl ∣) + (Iso.fun PathIdTrunc₀Iso ker) + + Im-p⊂Ker-Susp : (i : ℕ) (x : _) → isInIm (p-hom i) x → isInKer (E-susp i) x + Im-p⊂Ker-Susp i = + sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ f → pRec (squash₂ _ _) + (uncurry (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ g y → subst (isInKer (E-susp i)) y + (cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ { (inl x) → refl + ; (inr x) → sym (Kn→ΩKn+1 _ (g x)) + ; (push a j) k → Kn→ΩKn+1 i (g (fst a)) (~ k ∧ j)}))))) + open import Cubical.Foundations.Path + Ker-Susp⊂Im-p : (i : ℕ) (x : _) → isInKer (E-susp i) x → isInIm (p-hom i) x + Ker-Susp⊂Im-p i = + sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) + λ f ker → pRec squash + (λ id → ∣ ∣ (λ x → ΩKn+1→Kn i (sym (funExt⁻ (cong fst id) (inr x)))) ∣₂ + , cong ∣_∣₂ (funExt (λ { (a , b) → + cong (ΩKn+1→Kn i) (lUnit _ ∙ cong (_∙ sym (funExt⁻ (λ i₁ → fst (id i₁)) (inr a))) (sym (flipSquare (cong snd id)) + ∙∙ (PathP→compPathR (cong (funExt⁻ (cong fst id)) (push (a , b)))) + ∙∙ assoc _ _ _ + ∙ sym (rUnit _)) + ∙ (sym (assoc _ _ _) + ∙∙ cong (Kn→ΩKn+1 i (f (a , b)) ∙_) (rCancel _) + ∙∙ sym (rUnit _))) + ∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f (a , b))})) ∣) + (Iso.fun PathIdTrunc₀Iso ker) + + Im-Susp⊂Ker-j : (i : ℕ) (x : _) → isInIm (E-susp i) x → isInKer (cofibSeq.j* (suc i)) x + Im-Susp⊂Ker-j i = + sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ g → pRec (squash₂ _ _) + (uncurry (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) + λ f id → pRec (squash₂ _ _) + (λ P → subst (isInKer (cofibSeq.j* (suc i))) (cong ∣_∣₂ P) + (cong ∣_∣₂ refl)) + (Iso.fun PathIdTrunc₀Iso id))) + + Ker-j⊂Im-Susp : (i : ℕ) (x : _) → isInKer (cofibSeq.j* (suc i)) x → isInIm (E-susp i) x + Ker-j⊂Im-Susp i = + sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) + λ f ker + → pRec squash + (λ p → ∣ ∣ (λ x → ΩKn+1→Kn _ (sym (snd f) ∙∙ cong (fst f) (push x) ∙∙ funExt⁻ p (fst x))) ∣₂ + , cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) + (funExt (λ { (inl x) → sym (snd f) + ; (inr x) → sym (funExt⁻ p x) + ; (push a j) k → h3 f p a k j}))) ∣) + (Iso.fun PathIdTrunc₀Iso ker) + where + h3 : (f : (E'̃ , inl tt) →∙ coHomK-ptd (suc i)) + → (p : (fst f ∘ inr) ≡ (λ _ → 0ₖ (suc i))) + → (a : E B P) + → PathP (λ i → snd f (~ i) ≡ p (~ i) (fst a)) + (Kn→ΩKn+1 i (ΩKn+1→Kn i (sym (snd f) ∙∙ cong (fst f) (push a) ∙∙ funExt⁻ p (fst a)))) + (cong (fst f) (push a)) + h3 f p a = Iso.rightInv (Iso-Kn-ΩKn+1 i) _ + ◁ λ i j → doubleCompPath-filler (sym (snd f)) (cong (fst f) (push a)) (funExt⁻ p (fst a)) (~ i) j + + ϕ∘j : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr (i +' n) (typ B)) + ϕ∘j i = compGroupHom (fst (fst (ϕ i)) , snd (ϕ i)) (cofibSeq.j* (i +' n)) + + +'-suc : (i n : ℕ) → (suc i +' n) ≡ suc (i +' n) + +'-suc zero zero = refl + +'-suc (suc i₁) zero = refl + +'-suc zero (suc n) = refl + +'-suc (suc i₁) (suc n) = refl + + private + h3 : ∀ {ℓ ℓ'} {n m : ℕ} (G : ℕ → Group ℓ) (H : Group ℓ') (p : n ≡ m) + → GroupEquiv (G n) H + → GroupEquiv (G m) H + h3 {n = n} G H = + J (λ m _ → GroupEquiv (G n) H → GroupEquiv (G m) H) + λ p → p + + h-ret : ∀ {ℓ ℓ'} {n m : ℕ} (G : ℕ → Group ℓ) (H : Group ℓ') + → (e : GroupEquiv (G n) H) + → (p : n ≡ m) + → (x : G m .fst) → invEq (fst e) (fst (fst (h3 G H p e)) x) ≡ subst (λ x → G x .fst) (sym p) x + h-ret G H e = + J (λ m p → ((x : G m .fst) → invEq (fst e) (fst (fst (h3 G H p e)) x) ≡ subst (λ x → G x .fst) (sym p) x)) + λ x → cong (invEq (fst e) ) + (λ i → transportRefl (transportRefl (fst (fst e) (transportRefl (transportRefl x i) i)) i) i) + ∙∙ retEq (fst e) x + ∙∙ sym (transportRefl _) + + isEquivϕ' : (i : ℕ) → GroupEquiv (coHomRedGrDir (suc (i +' n)) (E'̃ , inl tt)) + (coHomGr (suc i) (typ B)) + isEquivϕ' i = (h3 (λ n → coHomRedGrDir n (E'̃ , inl tt)) _ (+'-suc i n) + (invGroupEquiv (ϕ (suc i)))) + + ϕ' : (i : ℕ) → GroupHom (coHomRedGrDir (suc (i +' n)) (E'̃ , inl tt)) + (coHomGr (suc i) (typ B)) + ϕ' i = fst (fst (isEquivϕ' i)) , snd (isEquivϕ' i) + + susp∘ϕ : (i : ℕ) → GroupHom (coHomGr (i +' n) E') (coHomGr (suc i) (typ B)) + susp∘ϕ i = compGroupHom (E-susp (i +' n)) (ϕ' i) + + Im-ϕ∘j⊂Ker-p : (i : ℕ) (x : _) → isInIm (ϕ∘j i) x → isInKer (p-hom _) x + Im-ϕ∘j⊂Ker-p i x p = + cofibSeq.Im-j⊂Ker-p _ x + (pRec squash (uncurry (λ f p → ∣ fst (fst (ϕ _)) f , p ∣)) p) + + Ker-p⊂Im-ϕ∘j : (i : ℕ) (x : _) → isInKer (p-hom _) x → isInIm (ϕ∘j i) x + Ker-p⊂Im-ϕ∘j i x p = + pRec squash (uncurry (λ f p → ∣ (invEq (fst (ϕ _)) f) + , (cong (fst (cofibSeq.j* _)) (secEq (fst (ϕ _)) f) ∙ p) ∣)) + (cofibSeq.Ker-p⊂Im-j _ x p) + + + Im-p⊂KerSusp∘ϕ : (i : ℕ) (x : _) → isInIm (p-hom _) x → isInKer (susp∘ϕ i) x + Im-p⊂KerSusp∘ϕ i x p = cong (fst (ϕ' _)) (Im-p⊂Ker-Susp _ x p) ∙ IsGroupHom.pres1 (snd (ϕ' _)) + + KerSusp∘ϕ⊂Im-p : (i : ℕ) (x : _) → isInKer (susp∘ϕ i) x → isInIm (p-hom _) x + KerSusp∘ϕ⊂Im-p i x p = + Ker-Susp⊂Im-p _ x (sym (retEq (fst (isEquivϕ' _)) _) + ∙ (cong (invEq (fst (isEquivϕ' _))) p + ∙ IsGroupHom.pres1 (snd (invGroupEquiv (isEquivϕ' _))))) + + Im-Susp∘ϕ⊂Ker-ϕ∘j : (i : ℕ) → (x : _) → isInIm (susp∘ϕ i) x → isInKer (ϕ∘j (suc i)) x + Im-Susp∘ϕ⊂Ker-ϕ∘j i x = + pRec (squash₂ _ _) + (uncurry λ f → J (λ x p → isInKer (ϕ∘j (suc i)) x) + ((λ i → fst (cofibSeq.j* _) (fst (fst (ϕ _)) (fst (ϕ' _) (fst (E-susp _) f)))) + ∙∙ cong (fst (cofibSeq.j* _)) + ((h-ret (λ n → coHomRedGrDir n (E'̃ , inl tt)) _ + (invGroupEquiv (ϕ (suc i))) (+'-suc i n)) (fst (E-susp _) f)) + ∙∙ (natTranspLem _ (λ n → fst (cofibSeq.j* n)) (sym (+'-suc i n)) + ∙ cong (subst (λ z → coHomGr z (typ B) .fst) (sym (+'-suc i n))) + (Im-Susp⊂Ker-j _ (fst (E-susp (i +' n)) f) ∣ f , refl ∣) + ∙ tLem i n))) + where + tLem : (i n : ℕ) → subst (λ z → coHomGr z (typ B) .fst) (sym (+'-suc i n)) + (0ₕ _) ≡ 0ₕ _ + tLem zero zero = refl + tLem zero (suc n) = refl + tLem (suc i₁) zero = refl + tLem (suc i₁) (suc n) = refl + + + Ker-ϕ∘j⊂Im-Susp∘ϕ : (i : ℕ) (x : _) + → isInKer (ϕ∘j (suc i)) x → isInIm (susp∘ϕ i) x + Ker-ϕ∘j⊂Im-Susp∘ϕ i x p = + pRec squash + (uncurry (λ f p → ∣ f , cong (fst (fst (isEquivϕ' i))) p ∙ secEq (fst (isEquivϕ' _)) x ∣)) + (Ker-j⊂Im-Susp _ (invEq (fst (isEquivϕ' _)) x) + ((cong (cofibSeq.j* (suc (i +' n)) .fst ) lem2 + ∙ natTranspLem _ (λ n → cofibSeq.j* n .fst) (+'-suc i n)) + ∙∙ cong (transport (λ j → (coHomGr (+'-suc i n j) (typ B) .fst))) p + ∙∙ h2 i n)) + where + lem2 : invEq (fst (isEquivϕ' i)) x ≡ transport (λ j → coHomRed (+'-suc i n j) (E'̃ , inl tt)) (fst (fst (ϕ _)) x) + lem2 = cong (transport (λ j → coHomRed (+'-suc i n j) (E'̃ , inl tt))) + (transportRefl _ ∙ cong (fst (fst (ϕ _))) + λ i → transportRefl (transportRefl x i) i) + + h2 : (i n : ℕ) → transport (λ j → coHomGr (+'-suc i n j) (typ B) .fst) + (GroupStr.1g (coHomGr (suc i +' n) (typ B) .snd)) ≡ 0ₕ _ + h2 zero zero = refl + h2 zero (suc n) = refl + h2 (suc i₁) zero = refl + h2 (suc i₁) (suc n) = refl + + + ϕ∘j≡ : (i : ℕ) → ϕ∘j i ≡ ⌣-hom i + ϕ∘j≡ i = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ _ → refl)) + +open import Cubical.Experiments.Brunerie +open import Cubical.HITs.Hopf + +open import Cubical.HITs.Join -S³+' : S₊ 3 → S₊ 3 → S₊ 3 -S³+' north y = y -S³+' south y = y -S³+' (merid a i) north = (merid a ∙ sym (merid north)) i -S³+' (merid a i) south = (sym (merid north) ∙ merid a) i -S³+' (merid north i) (merid b j) = hcomp (λ k → λ { (i = i0) → merid b j - ; (i = i1) → merid b j - ; (j = i0) → rCancel (merid north) (~ k) i - ; (j = i1) → lCancel (merid north) (~ k) i}) - (merid b j) -S³+' (merid south i) (merid b j) = hcomp (λ k → λ { (i = i0) → merid b j - ; (i = i1) → merid b j - ; (j = i0) → (merid (merid base k) ∙ sym (merid north)) i - ; (j = i1) → (sym (merid north) ∙ merid (merid base k)) i}) - (hcomp (λ k → λ { (i = i0) → merid b j - ; (i = i1) → merid b j - ; (j = i0) → rCancel (merid north) (~ k) i - ; (j = i1) → lCancel (merid north) (~ k) i}) - (merid b j)) -S³+' (merid (merid a j) i) (merid b ik) = {!!} - -S³+ : join S¹ S¹ → join S¹ S¹ → S₊ 3 -S³+ (inl x) y = north -S³+ (inr x) y = north -S³+ (push a b i) (inl x) = north -S³+ (push a b i) (inr x) = north -S³+ (push a b i) (push c d j) = - hcomp (λ k → λ { (i = i0) → north - ; (i = i1) → north - ; (j = i0) → rCancel (merid north) k i - ; (j = i1) → rCancel (merid north) k i}) - ((merid ((merid ((a * c) * (b * c)) ∙ sym (merid base)) j) ∙ sym (merid north)) i) - -rUnitS : (x : join S¹ S¹) → S³+ x (inl base) ≡ north -rUnitS (inl x) = refl -rUnitS (inr x) = refl -rUnitS (push a b i) = refl - - - --- sss : ∥ typ ((Ω^ 3) (S₊∙ 2)) ∥₂ → ∥ typ ((Ω^ 4) (S₊∙ 3)) ∥₂ --- sss = sMap (Ω→ (Ω→ (Ω→ ((λ x → merid x ∙ sym (merid north)) , (rCancel _)))) .fst) - --- sss' : typ ((Ω^ 3) (S₊∙ 2)) → typ ((Ω^ 4) (S₊∙ 3)) --- sss' x i j k l = --- hcomp (λ r → λ { (i = i0) → rCancel (merid north) r l --- ; (i = i1) → rCancel (merid north) r l --- ; (j = i0) → rCancel (merid north) r l --- ; (j = i1) → rCancel (merid north) r l --- ; (k = i0) → rCancel (merid north) r l --- ; (k = i1) → rCancel (merid north) r l --- ; (l = i0) → north --- ; (l = i1) → north}) --- ((merid (x i j k) ∙ sym (merid north)) l) - - --- sss≡ : (x : _) → sss' x ≡ sym (sss' x) --- sss≡ x w i j k l = --- hcomp ((λ r → λ { (i = i0) → rCancel (merid north) r l --- ; (i = i1) → rCancel (merid north) r l --- ; (j = i0) → rCancel (merid north) r l --- ; (j = i1) → rCancel (merid north) r l --- ; (k = i0) → rCancel (merid north) r l --- ; (k = i1) → rCancel (merid north) r l --- ; (l = i0) → north --- ; (l = i1) → north})) --- (hcomp ((λ r → λ { (i = i0) → symDistr (merid (x i j k)) (sym (merid north)) (~ r) (~ l) --- ; (i = i1) → symDistr (merid (x i j k)) (sym (merid north)) (~ r) (~ l) --- ; (j = i0) → symDistr (merid (x i j k)) (sym (merid north)) (~ r) (~ l) --- ; (j = i1) → symDistr (merid (x i j k)) (sym (merid north)) (~ r) (~ l) --- ; (k = i0) → symDistr (merid (x i j k)) (sym (merid north)) (~ r) (~ l) --- ; (k = i1) → symDistr (merid (x i j k)) (sym (merid north)) (~ r) (~ l) --- ; (l = i0) → north --- ; (l = i1) → north --- ; (w = i0) → symDistr (merid (x i j k)) (sym (merid north)) (~ r) (~ l) -- (merid (x i j k) ∙ sym (merid north)) l --- ; (w = i1) → {!!}})) --- {!!}) -- ((merid (x i j k) ∙ sym (merid north)) ((~ l ∧ w) ∨ (l ∧ ~ w)))) - --- Π₄S³ : Type _ --- Π₄S³ = ∥ typ ((Ω^ 3) (S₊∙ 2)) ∥₂ /' λ f g → sss f ≡ sss g --- open import Cubical.Homotopy.Freudenthal - --- ll123 : Type _ --- ll123 = Pushout {A = {!!}} {B = {!!}} {!sss ∣ ? ∣₂!} (λ _ → tt) - --- isEquiv→isEquivΩ→ : {!!} --- isEquiv→isEquivΩ→ = {!!} - --- module _ {ℓ ℓ' : Level} {A : Type ℓ} {B : Type ℓ'} {x y : A} (n : ℕ) --- (f : A → B) (con : isConnectedFun (suc (suc n)) f) where --- fib≡ : (b : B) (f1 f2 : fiber f b) → Iso (f1 ≡ f2) (fiber (cong f) (snd f1 ∙ sym (snd f2))) --- fib≡ = {!!} - --- ll : (p : f x ≡ f y) → isContr (hLevelTrunc (suc n) (fiber (cong f) p)) --- ll p = subst (isContr ∘ hLevelTrunc (suc n)) --- (isoToPath (fib≡ (f x) (x , refl) (y , sym p)) --- ∙ cong (fiber (cong f)) (sym (lUnit p))) --- (isOfHLevelRetractFromIso 0 --- (invIso (PathIdTruncIso _)) --- ((isContr→isProp (con (f x)) _ _) --- , isProp→isSet (isContr→isProp (con (f x))) _ _ _)) - --- conbase : isConnectedFun 4 {A = (S₊ 2)} (σ 1 {A = S₊∙ 2} (sphereConnected 2)) --- conbase = isConnectedσ 1 (sphereConnected 2) - --- conn-cong : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A} (n : ℕ) --- → (f : A → B) --- → isConnectedFun (suc (suc n)) f --- → (p : _) → hLevelTrunc (suc n) (fiber (cong f) p) --- conn-cong {x = x} {y = y} n f con p = --- trRec (isOfHLevelTrunc (suc n)) --- (λ pp → ∣ (cong fst pp) , (cong sym (rUnit _) --- ∙∙ cong sym (cong (sym (λ i₁ → f (fst (pp i₁))) ∙_) (rUnit refl)) --- ∙∙ cong sym (PathP→compPathL (cong snd pp))) ∣) l --- where --- l : hLevelTrunc (suc n) ((x , refl) ≡ (y , sym p)) --- l = Iso.fun (PathIdTruncIso _) (isContr→isProp (con (f x)) ∣ x , refl ∣ ∣ y , sym p ∣) - --- pr : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A} (n : ℕ) --- → (f : A → B) --- → isConnectedFun (suc (suc n)) f --- → (p : _) → isProp (hLevelTrunc (suc n) (fiber {A = x ≡ y} (cong f) p)) --- pr {x = x} {y = y} n f co p = --- trElim2 (λ _ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) --- (uncurry λ p2 P → uncurry λ q Q → trRec (isOfHLevelPath' n (isOfHLevelTrunc (suc n)) _ _) --- (λ p3 → cong ∣_∣ₕ (ΣPathP ((sym (cong (cong fst) p3)) , {!cong (cong snd) p3!}))) --- (l q p2 Q P)) --- where --- l : (q p2 : x ≡ y) (Q : cong f q ≡ p) (P : cong f p2 ≡ p) --- → hLevelTrunc n (ΣPathP (q , λ j i → f (q (~ i ∧ j))) --- ≡ ΣPathP (p2 , compPathL→PathP (cong (sym (cong f p2) ∙_) (sym (rUnit refl)) --- ∙ sym (rUnit _) --- ∙ cong sym (P ∙ sym Q)))) --- l q p2 Q P = --- Iso.fun (PathIdTruncIso _) (isContr→isProp --- (isConnectedPath _ --- (co (f x)) (x , refl) (y , (cong f (sym q)))) --- ∣ ΣPathP (q , λ j i → f (q (~ i ∧ j))) ∣ --- ∣ ΣPathP (p2 , compPathL→PathP (cong (sym (cong f p2) ∙_) (sym (rUnit refl)) --- ∙ sym (rUnit _) --- ∙ cong sym (P ∙ sym Q))) ∣) - --- fib-sss : (f : ∥ typ ((Ω^ 4) (S₊∙ 3)) ∥₂) → fiber sss f --- fib-sss = sElim {!!} λ f → trRec {!!} {!!} (conn-cong {!!} --- (Ω→ (Ω→ ((λ x → merid x ∙ sym (merid north)) , (rCancel _))) .fst) --- {!conn-cong _ ? ? ?!} --- (Ω→ --- (Ω→ --- ((λ x → merid x ∙ (λ i → merid north (~ i))) , --- rCancel (merid north))) --- .snd ∙∙ f ∙∙ sym (Ω→ --- (Ω→ --- ((λ x → merid x ∙ (λ i → merid north (~ i))) , --- rCancel (merid north))) --- .snd))) - - --- -- SES→Iso : ∀ {ℓ ℓ'} {L R : Group ℓ-zero} --- -- → {G : Group ℓ} {H : Group ℓ'} --- -- → UnitGroup ≡ L --- -- → UnitGroup ≡ R --- -- → (lhom : GroupHom L G) (midhom : GroupHom G H) (rhom : GroupHom H R) --- -- → ((x : _) → isInKer midhom x → isInIm lhom x) --- -- → ((x : _) → isInKer rhom x → isInIm midhom x) --- -- → isEquiv (fst midhom) --- -- SES→Iso {R = R} {G = G} {H = H} = --- -- J (λ L _ → UnitGroup ≡ R → --- -- (lhom : GroupHom L G) (midhom : GroupHom G H) --- -- (rhom : GroupHom H R) → --- -- ((x : fst G) → isInKer midhom x → isInIm lhom x) → --- -- ((x : fst H) → isInKer rhom x → isInIm midhom x) → --- -- isEquiv (fst midhom)) --- -- ((J (λ R _ → (lhom : GroupHom UnitGroup G) (midhom : GroupHom G H) --- -- (rhom : GroupHom H R) → --- -- ((x : fst G) → isInKer midhom x → isInIm lhom x) → --- -- ((x : _) → isInKer rhom x → isInIm midhom x) → --- -- isEquiv (fst midhom)) --- -- main)) --- -- where --- -- main : (lhom : GroupHom UnitGroup G) (midhom : GroupHom G H) --- -- (rhom : GroupHom H UnitGroup) → --- -- ((x : fst G) → isInKer midhom x → isInIm lhom x) → --- -- ((x : fst H) → isInKer rhom x → isInIm midhom x) → --- -- isEquiv (fst midhom) --- -- main lhom midhom rhom lexact rexact = --- -- BijectionIsoToGroupEquiv {G = G} {H = H} --- -- bijIso' .fst .snd --- -- where --- -- bijIso' : BijectionIso G H --- -- BijectionIso.fun bijIso' = midhom --- -- BijectionIso.inj bijIso' x inker = --- -- pRec (GroupStr.is-set (snd G) _ _) --- -- (λ s → sym (snd s) ∙ IsGroupHom.pres1 (snd lhom)) (lexact _ inker) --- -- BijectionIso.surj bijIso' x = rexact x refl - --- -- HopfInvariantPush : (n : ℕ) --- -- → (f : S₊ (3 +ℕ n +ℕ n) → S₊ (2 +ℕ n)) --- -- → Type _ --- -- HopfInvariantPush n f = Pushout (λ _ → tt) f - - --- -- Hopfα : (n : ℕ) --- -- → (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) --- -- → coHom (2 +ℕ n) (HopfInvariantPush n (fst f)) --- -- Hopfα n f = ∣ (λ { (inl x) → 0ₖ _ --- -- ; (inr x) → ∣ x ∣ --- -- ; (push a i) → help a (~ i)}) ∣₂ --- -- where --- -- help : (a : S₊ (3 +ℕ n +ℕ n)) → ∣ fst f a ∣ ≡ 0ₖ (2 +ℕ n) --- -- help = sphereElim _ (λ _ → isOfHLevelPlus' {n = n} (3 +ℕ n) --- -- (isOfHLevelPath' (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n)) _ _)) (cong ∣_∣ₕ (snd f)) - --- -- Hopfβ : (n : ℕ) --- -- → (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) --- -- → coHom (4 +ℕ (n +ℕ n)) (HopfInvariantPush n (fst f)) --- -- Hopfβ n f = fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) ∣ ∣_∣ ∣₂ - --- -- module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) where --- -- module M = MV _ _ _ (λ _ → tt) (fst f) --- -- ¬lem : (n m : ℕ) → ¬ suc (suc (m +ℕ n)) ≡ suc n --- -- ¬lem zero zero p = snotz (cong predℕ p) --- -- ¬lem (suc n) zero p = ¬lem n zero (cong predℕ p) --- -- ¬lem zero (suc m) p = snotz (cong predℕ p) --- -- ¬lem (suc n) (suc m) p = --- -- ¬lem n (suc m) (cong (suc ∘ suc ) --- -- (sym (+-suc m n)) ∙ (cong predℕ p)) - --- -- GroupEquiv1 : --- -- GroupEquiv --- -- (coHomGr (3 +ℕ n +ℕ n) (fst (S₊∙ (3 +ℕ n +ℕ n)))) --- -- ((coHomGr (suc (3 +ℕ n +ℕ n)) (Pushout (λ _ → tt) (fst f)))) --- -- fst (fst GroupEquiv1) = (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) .fst --- -- snd (fst GroupEquiv1) = --- -- SES→Iso --- -- (GroupPath _ _ .fst --- -- (invGroupEquiv (isContr→≃Unit --- -- (isOfHLevelΣ 0 --- -- (isOfHLevelRetractFromIso 0 (fst (Hⁿ-Unit≅0 _)) isContrUnit) --- -- λ _ → isOfHLevelRetractFromIso 0 --- -- (fst (Hⁿ-Sᵐ≅0 _ _ (¬lem n n))) isContrUnit) --- -- , makeIsGroupHom λ _ _ → refl))) --- -- (GroupPath _ _ .fst --- -- (invGroupEquiv --- -- (isContr→≃Unit --- -- (isContrΣ --- -- (isOfHLevelRetractFromIso 0 (fst (Hⁿ-Unit≅0 _)) --- -- isContrUnit) --- -- λ _ → isOfHLevelRetractFromIso 0 --- -- (fst (Hⁿ-Sᵐ≅0 _ _ (¬lem n (suc n)))) isContrUnit) --- -- , makeIsGroupHom λ _ _ → refl))) --- -- (M.Δ (3 +ℕ n +ℕ n)) --- -- (M.d (3 +ℕ n +ℕ n)) --- -- (M.i (4 +ℕ n +ℕ n)) --- -- (M.Ker-d⊂Im-Δ _) --- -- (M.Ker-i⊂Im-d _) --- -- snd GroupEquiv1 = (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) .snd - --- -- Hopfβ-Iso : GroupIso ((coHomGr (suc (3 +ℕ n +ℕ n)) (Pushout (λ _ → tt) (fst f)))) ℤGroup --- -- Hopfβ-Iso = compGroupIso (GroupEquiv→GroupIso (invGroupEquiv GroupEquiv1)) (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) - --- -- h33 : (n : ℕ) → Iso.inv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) 1 ≡ ∣ ∣_∣ ∣₂ --- -- h33 zero = Iso.leftInv (fst (Hⁿ-Sⁿ≅ℤ (suc zero))) _ --- -- h33 (suc n) = (λ i → Iso.inv --- -- (fst --- -- (suspensionAx-Sn (suc n) (suc n))) (h33 n i)) --- -- ∙ cong ∣_∣₂ (funExt λ { north → refl --- -- ; south → cong ∣_∣ₕ (merid north) --- -- ; (merid a i) j → ∣ compPath-filler (merid a) (sym (merid north)) (~ j) i ∣}) - --- -- d-β : (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) --- -- → Iso.fun (fst (Hopfβ-Iso n f)) (Hopfβ n f) ≡ 1 --- -- d-β n f = sym (cong (Iso.fun (fst (Hopfβ-Iso n f))) h) ∙ Iso.rightInv (fst (Hopfβ-Iso n f)) (pos 1) --- -- where --- -- h : Iso.inv (fst (Hopfβ-Iso n f)) (pos 1) ≡ Hopfβ n f --- -- h = (λ i → fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) --- -- (h33 _ i)) --- -- ∙ cong ∣_∣₂ (funExt (λ { (inl x) → refl --- -- ; (inr x) → refl --- -- ; (push a i) → refl})) - --- -- module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) where --- -- 2+n = 2 +ℕ n --- -- H = HopfInvariantPush n (fst f) - --- -- ff : coHom 2+n H → coHom 2+n (S₊ (suc (suc n))) --- -- ff = sMap (_∘ inr) - --- -- grHom : IsGroupHom (snd (coHomGr 2+n H)) ff (snd (coHomGr 2+n (S₊ (suc (suc n))))) --- -- grHom = --- -- makeIsGroupHom --- -- (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) --- -- λ f g → refl) - --- -- h' : (g : (S₊ (suc (suc n)) → coHomK 2+n)) → H → coHomK (2 +ℕ n) --- -- h' g (inl x) = 0ₖ _ --- -- h' g (inr x) = g x -ₖ g north --- -- h' g (push a i) = h23 a i --- -- where --- -- h23 : (a : S₊ (suc (suc (suc (n +ℕ n))))) → 0ₖ (suc (suc n)) ≡ g (fst f a) -ₖ g north --- -- h23 = sphereElim _ (λ x → isOfHLevelPlus' {n = n} (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n) _ _)) --- -- (sym (rCancelₖ _ (g north)) ∙ cong (λ x → g x -ₖ g north) (sym (snd f))) - --- -- ff⁻ : coHom 2+n (S₊ (suc (suc n))) → coHom 2+n H --- -- ff⁻ = sMap h' - --- -- h23 : (x : _) → ∥ x ≡ 0ₖ (suc (suc n)) ∥ --- -- h23 = coHomK-elim _ (λ _ → isProp→isOfHLevelSuc (suc n) squash) ∣ refl ∣ - --- -- h : Iso (coHom (2 +ℕ n) (HopfInvariantPush n (fst f))) (coHom (2 +ℕ n) ((S₊ (suc (suc n))))) --- -- Iso.fun h = ff --- -- Iso.inv h = ff⁻ --- -- Iso.rightInv h = --- -- sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) --- -- λ g → pRec (squash₂ _ _) --- -- (λ p → cong ∣_∣₂ (funExt λ x → cong (g x +ₖ_) (cong (-ₖ_) p) ∙ rUnitₖ _ (g x))) --- -- (h23 (g north)) --- -- Iso.leftInv h = --- -- sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) --- -- λ g → pRec (squash₂ _ _) --- -- (pRec (isPropΠ (λ _ → squash₂ _ _)) --- -- (λ gn gtt → --- -- trRec (isProp→isOfHLevelSuc n (squash₂ _ _)) --- -- (λ r → cong ∣_∣₂ (funExt λ { --- -- (inl x) → sym gtt --- -- ; (inr x) → (λ i → g (inr x) -ₖ gn i) ∙ rUnitₖ _ (g (inr x)) --- -- ; (push a i) --- -- → sphereElim _ --- -- {A = λ a → PathP (λ i → h' (λ x → g (inr x)) (push a i) ≡ g (push a i)) --- -- (sym gtt) --- -- ((λ i → g (inr (fst f a)) -ₖ gn i) ∙ rUnitₖ _ (g (inr (fst f a))))} --- -- (λ _ → isOfHLevelPathP' (suc (suc (suc (n +ℕ n)))) --- -- (isOfHLevelPath (suc (suc (suc (suc (n +ℕ n))))) --- -- (isOfHLevelPlus' {n = n} (suc (suc (suc (suc n)))) --- -- (isOfHLevelTrunc (suc (suc (suc (suc n)))))) _ _) _ _) --- -- r a i})) --- -- (l g gtt gn)) --- -- (h23 (g (inr north)))) --- -- (h23 (g (inl tt))) --- -- where --- -- l : (g : HopfInvariantPush n (fst f) → coHomK (suc (suc n))) --- -- → (gtt : g (inl tt) ≡ 0ₖ (suc (suc n))) --- -- → (gn : g (inr north) ≡ 0ₖ (suc (suc n))) --- -- → hLevelTrunc (suc n) (PathP (λ i → h' (λ x → g (inr x)) (push north i) ≡ g (push north i)) --- -- (sym gtt) ((λ i → g (inr (fst f north)) -ₖ gn i) ∙ rUnitₖ _ (g (inr (fst f north))))) --- -- l g gtt gn = isConnectedPathP _ (isConnectedPath _ (isConnectedKn _) _ _) _ _ .fst - - --- -- Hopfα-Iso : GroupIso (coHomGr (2 +ℕ n) (HopfInvariantPush n (fst f))) ℤGroup --- -- Hopfα-Iso = --- -- compGroupIso --- -- (h , grHom) --- -- (Hⁿ-Sⁿ≅ℤ (suc n)) - --- -- Hopfα-Iso-α : (n : ℕ) (f : _) --- -- → Iso.fun (fst (Hopfα-Iso n f)) (Hopfα n f) --- -- ≡ 1 --- -- Hopfα-Iso-α n f = --- -- sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (h33 n)) --- -- ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) --- -- where --- -- hz : Iso.fun (h n f) (Hopfα n f) ≡ ∣ ∣_∣ ∣₂ --- -- hz = refl - --- -- _·₀ₕ_ : ∀ {ℓ} {n : ℕ} {A : Type ℓ} → ℤ → coHom n A → coHom n A --- -- _·₀ₕ_ {n = n} a = sMap λ f x → a ·₀ f x - --- -- HopfInvariant : (n : ℕ) --- -- → (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → Type _ --- -- HopfInvariant n f = --- -- Σ[ x ∈ ℤ ] --- -- Hopfα n f ⌣ Hopfα n f --- -- ≡ (x ·₀ₕ subst (λ x → coHom x (HopfInvariantPush n (fst f))) --- -- (λ i → suc (suc (suc (+-suc n n (~ i))))) --- -- (Hopfβ n f)) - --- -- pres- : (e : GroupHom ℤGroup ℤGroup) → (a : ℤ) → fst e (- a) ≡ - fst e a --- -- pres- e a = +Comm (fst e (- a)) (pos 0) --- -- ∙ cong (_+ (fst e (- a))) (sym (l (fst e a)) ∙ +Comm (fst e a) (- fst e a)) --- -- ∙∙ sym (+Assoc _ _ _) --- -- ∙∙ cong (- (fst e a) +_) (pp ∙ l (fst e a)) --- -- where --- -- l : (a : ℤ) → a + (- a) ≡ 0 --- -- l (pos zero) = refl --- -- l (pos (suc zero)) = refl --- -- l (pos (suc (suc n))) = predℤ+negsuc n (pos (suc (suc n))) ∙ l (pos (suc n)) --- -- l (negsuc zero) = refl --- -- l (negsuc (suc n)) = (cong sucℤ (sucℤ+pos n (negsuc (suc n)))) ∙ l (negsuc n) - --- -- pp : fst e a + fst e (- a) ≡ fst e a + (- fst e a) --- -- pp = sym (IsGroupHom.pres· (snd e) a (- a)) --- -- ∙∙ cong (fst e) (l a) --- -- ∙∙ (IsGroupHom.pres1 (snd e) --- -- ∙ sym (l _)) - - - --- -- GroupHomPres· : (e : GroupHom ℤGroup ℤGroup) → (a b : ℤ) → fst e (a · b) ≡ a · fst e b --- -- GroupHomPres· e (pos zero) b = IsGroupHom.pres1 (snd e) --- -- GroupHomPres· e (pos (suc n)) b = --- -- IsGroupHom.pres· (snd e) b (pos n · b) ∙ cong (fst e b +_) (GroupHomPres· e (pos n) b) --- -- GroupHomPres· e (negsuc zero) b = pres- e b --- -- GroupHomPres· e (negsuc (suc n)) b = --- -- IsGroupHom.pres· (snd e) (- b) (negsuc n · b) --- -- ∙ cong₂ _+_ (pres- e b) (GroupHomPres· e (negsuc n) b) - --- -- lUnit·₀ₕ : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → (x : coHom n A) → (pos zero) ·₀ₕ x ≡ 0ₕ _ --- -- lUnit·₀ₕ n = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl - --- -- minus≡0- : (x : ℤ) → - x ≡ (0 - x) --- -- minus≡0- x = sym (GroupStr.lid (snd ℤGroup) (- x)) - --- -- GroupHomPres·₀ : ∀ {ℓ} {A : Type ℓ} (n : ℕ) --- -- (e : GroupHom ℤGroup (coHomGr n A)) --- -- → (a b : ℤ) --- -- → fst e (a · b) --- -- ≡ (a ·₀ₕ fst e b) --- -- GroupHomPres·₀ n e (pos zero) b = IsGroupHom.pres1 (snd e) ∙ sym (lUnit·₀ₕ n (fst e b)) --- -- GroupHomPres·₀ {A = A} n e (pos (suc a)) b = --- -- IsGroupHom.pres· (snd e) b (pos a · b) --- -- ∙ (λ i → fst e b +ₕ GroupHomPres·₀ n e (pos a) b i) --- -- ∙ s (fst e b) --- -- where --- -- s : (x : coHom n A) → x +ₕ (pos a ·₀ₕ x) ≡ pos (suc a) ·₀ₕ x --- -- s = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl --- -- GroupHomPres·₀ n e (negsuc zero) b = --- -- (cong (fst e) (sym (GroupStr.lid (snd ℤGroup) (- b))) --- -- ∙ IsGroupHom.presinv (snd e) b) --- -- GroupHomPres·₀ {A = A} n e (negsuc (suc a)) b = --- -- IsGroupHom.pres· (snd e) (- b) (negsuc a · b) --- -- ∙∙ cong₂ _+ₕ_ ((cong (fst e) (sym (GroupStr.lid (snd ℤGroup) (- b))) --- -- ∙ IsGroupHom.presinv (snd e) b)) --- -- (GroupHomPres·₀ n e (negsuc a) b) --- -- ∙∙ helper (fst e b) --- -- where --- -- helper : (x : coHom n A) → (-ₕ x) +ₕ (negsuc a ·₀ₕ x) ≡ (negsuc (suc a) ·₀ₕ x) --- -- helper = --- -- sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) --- -- λ f → cong ∣_∣₂ (funExt λ x → commₖ n (-ₖ (f x)) (negsuc a ·₀ f x)) - --- -- genH : ∀ {ℓ} {A : Type ℓ} (n : ℕ) --- -- → (e : GroupIso (coHomGr n A) --- -- ℤGroup) --- -- → Σ[ x ∈ coHom n A ] --- -- ((y : coHom n A) --- -- → Σ[ a ∈ ℤ ] a ·₀ₕ x ≡ y) --- -- genH {A = A} n e = (Iso.inv (fst e) 1) --- -- , λ y → (Iso.fun (fst e) y) --- -- , (sym (GroupHomPres·₀ n (_ , snd (invGroupIso e)) (Iso.fun (fst e) y) (pos 1)) --- -- ∙∙ cong (Iso.inv (fst e)) (·Comm (Iso.fun (fst e) y) (pos 1)) --- -- ∙∙ Iso.leftInv (fst e) y) - --- -- -- open import Cubical.Data.Sum --- -- -- open import Cubical.Relation.Nullary --- -- -- l2 : (n m : ℕ) → Σ[ a ∈ ℕ ] (negsuc n · pos (suc m)) ≡ negsuc a --- -- -- l2 n zero = n , ·Comm (negsuc n) (pos 1) --- -- -- l2 n (suc m) = h _ _ .fst , --- -- -- (·Comm (negsuc n) (pos (suc (suc m))) --- -- -- ∙∙ cong (negsuc n +_) (·Comm (pos (suc m)) (negsuc n) ∙ (l2 n m .snd)) --- -- -- ∙∙ h _ _ .snd) --- -- -- where --- -- -- h : (x y : ℕ) → Σ[ a ∈ ℕ ] negsuc x + negsuc y ≡ negsuc a --- -- -- h zero zero = 1 , refl --- -- -- h zero (suc y) = (suc (suc y)) , +Comm (negsuc zero) (negsuc (suc y)) --- -- -- h (suc x) zero = (suc (suc x)) , refl --- -- -- h (suc x) (suc y) = (h (suc (suc x)) y .fst) , (predℤ+negsuc y (negsuc (suc x)) ∙ snd ((h (suc (suc x))) y)) - --- -- -- zz : (n x : ℕ) → Σ[ a ∈ ℕ ] ((pos (suc x)) · pos (suc (suc n)) ≡ pos (suc (suc a))) --- -- -- zz n zero = n , refl --- -- -- zz n (suc x) = h _ _ (zz n x) --- -- -- where --- -- -- h : (x : ℤ) (n : ℕ) → Σ[ a ∈ ℕ ] (x ≡ pos (suc (suc a))) --- -- -- → Σ[ a ∈ ℕ ] pos n + x ≡ pos (suc (suc a)) --- -- -- h x n (a , p) = n +ℕ a --- -- -- , cong (pos n +_) p ∙ cong sucℤ (sucℤ+pos a (pos n)) --- -- -- ∙ sucℤ+pos a (pos (suc n)) ∙ (sym (pos+ (suc (suc n)) a)) - --- -- -- lem22 : (n : ℕ) (x : ℤ) → ¬ pos 1 ≡ x · pos (suc (suc n)) --- -- -- lem22 n (pos zero) p = snotz (injPos p) --- -- -- lem22 n (pos (suc n₁)) p = snotz (injPos (sym (cong predℤ (snd (zz n n₁))) ∙ sym (cong predℤ p))) --- -- -- lem22 n (negsuc n₁) p = posNotnegsuc _ _ (p ∙ l2 _ _ .snd) - - --- -- -- grr : (e : GroupEquiv ℤGroup ℤGroup) (x : ℤ) → (fst (fst e) 1) ≡ x → abs (fst (fst e) 1) ≡ 1 --- -- -- grr e (pos zero) p = --- -- -- ⊥-rec (snotz (injPos (sym (retEq (fst e) 1) --- -- -- ∙∙ (cong (fst (fst (invGroupEquiv e))) p) --- -- -- ∙∙ IsGroupHom.pres1 (snd (invGroupEquiv e))))) --- -- -- grr e (pos (suc zero)) p = cong abs p --- -- -- grr e (pos (suc (suc n))) p = ⊥-rec (lem22 _ _ (h ∙ ·Comm (pos (suc (suc n))) (invEq (fst e) 1))) --- -- -- where - --- -- -- h : pos 1 ≡ _ --- -- -- h = sym (retEq (fst e) 1) --- -- -- ∙∙ (cong (fst (fst (invGroupEquiv e))) (p ∙ ·Comm 1 (pos (suc (suc n))))) --- -- -- ∙∙ GroupHomPres· (_ , snd (invGroupEquiv e)) (pos (suc (suc n))) 1 --- -- -- grr e (negsuc zero) p = cong abs p --- -- -- grr e (negsuc (suc n)) p = ⊥-rec (lem22 _ _ l33) --- -- -- where --- -- -- h : fst (fst e) (negsuc zero) ≡ pos (suc (suc n)) --- -- -- h = pres- (_ , snd e) (pos 1) ∙ cong -_ p - --- -- -- compGroup : GroupEquiv ℤGroup ℤGroup --- -- -- fst compGroup = isoToEquiv (iso -_ -_ -Involutive -Involutive) --- -- -- snd compGroup = makeIsGroupHom -Dist+ - --- -- -- compHom : GroupEquiv ℤGroup ℤGroup --- -- -- compHom = compGroupEquiv compGroup e - --- -- -- l32 : fst (fst compHom) (pos 1) ≡ pos (suc (suc n)) --- -- -- l32 = h - --- -- -- l33 : pos 1 ≡ invEq (fst compHom) (pos 1) · pos (suc (suc n)) --- -- -- l33 = sym (retEq (fst compHom) (pos 1)) --- -- -- ∙∙ cong (invEq (fst compHom)) l32 --- -- -- ∙∙ (cong (invEq (fst compHom)) (·Comm (pos 1) (pos (suc (suc n)))) --- -- -- ∙ GroupHomPres· (_ , (snd (invGroupEquiv compHom))) (pos (suc (suc n))) (pos 1) --- -- -- ∙ ·Comm (pos (suc (suc n))) (invEq (fst compHom) (pos 1))) - --- -- -- abs→⊎ : (x : ℤ) (n : ℕ) → abs x ≡ n → (x ≡ pos n) ⊎ (x ≡ - pos n) --- -- -- abs→⊎ x n = J (λ n _ → (x ≡ pos n) ⊎ (x ≡ - pos n)) (help x) --- -- -- where --- -- -- help : (x : ℤ) → (x ≡ pos (abs x)) ⊎ (x ≡ - pos (abs x)) --- -- -- help (pos n) = inl refl --- -- -- help (negsuc n) = inr refl - --- -- -- JGroupEquiv : ∀ {ℓ ℓ'} {G : Group ℓ} (P : (H : Group ℓ) → GroupEquiv G H → Type ℓ') --- -- -- → P G idGroupEquiv --- -- -- → ∀ {H} e → P H e --- -- -- JGroupEquiv {G = G} P p {H} e = --- -- -- transport (λ i → P (GroupPath G H .fst e i) --- -- -- (transp (λ j → GroupEquiv G (GroupPath G H .fst e (i ∨ ~ j))) i e)) --- -- -- (subst (P G) (sym l) p) --- -- -- where --- -- -- l : (transp (λ j → GroupEquiv G (GroupPath G H .fst e (~ j))) i0 e) ≡ idGroupEquiv --- -- -- l = Σ≡Prop (λ _ → isPropIsGroupHom _ _) --- -- -- (Σ≡Prop (λ _ → isPropIsEquiv _) --- -- -- (funExt λ x → (λ i → fst (fst (fst e .snd .equiv-proof --- -- -- (transportRefl (fst (fst e) (transportRefl x i)) i)))) --- -- -- ∙ retEq (fst e) x)) - --- -- -- sesIsoPresGen : --- -- -- ∀ (G : Group ℓ-zero) --- -- -- → (iso : GroupEquiv ℤGroup G) --- -- -- → (H : Group ℓ-zero) --- -- -- → (iso2 : GroupEquiv ℤGroup H) --- -- -- → (e : fst G) --- -- -- → invEq (fst iso) e ≡ 1 --- -- -- → (hom : GroupEquiv G H) --- -- -- → abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1 --- -- -- sesIsoPresGen G = JGroupEquiv (λ G iso → --- -- -- (H : Group ℓ-zero) --- -- -- (iso2 : GroupEquiv ℤGroup H) --- -- -- → (e : fst G) --- -- -- → invEq (fst iso) e ≡ 1 --- -- -- → (hom : GroupEquiv G H) --- -- -- → abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1) --- -- -- λ H → JGroupEquiv (λ H iso2 → (e : ℤ) → e ≡ pos 1 → --- -- -- (hom : GroupEquiv ℤGroup H) → --- -- -- abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1) --- -- -- λ e p hom → cong (abs ∘ (fst (fst hom))) p ∙ autoPres1 hom --- -- -- where --- -- -- autoPres1 : (e : GroupEquiv ℤGroup ℤGroup) --- -- -- → abs (fst (fst e) 1) ≡ 1 --- -- -- autoPres1 e = grr e _ refl - --- -- -- characFunSpaceS¹ : ∀ {ℓ} {A : Type ℓ} → --- -- -- Iso (S₊ 1 → A) (Σ[ x ∈ A ] x ≡ x) --- -- -- Iso.fun characFunSpaceS¹ f = f base , cong f loop --- -- -- Iso.inv characFunSpaceS¹ (x , p) base = x --- -- -- Iso.inv characFunSpaceS¹ (x , p) (loop i) = p i --- -- -- Iso.rightInv characFunSpaceS¹ _ = refl --- -- -- Iso.leftInv characFunSpaceS¹ f = funExt λ { base → refl ; (loop i) → refl} - --- -- -- characFunSpace : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} --- -- -- → Iso (join A B → C) --- -- -- (Σ[ f ∈ (A → C) ] Σ[ g ∈ (B → C) ] --- -- -- ((a : A) (b : B) → f a ≡ g b)) --- -- -- Iso.fun characFunSpace f = (f ∘ inl) , ((f ∘ inr) , (λ a b → cong f (push a b))) --- -- -- Iso.inv characFunSpace (f , g , p) (inl x) = f x --- -- -- Iso.inv characFunSpace (f , g , p) (inr x) = g x --- -- -- Iso.inv characFunSpace (f , g , p) (push a b i) = p a b i --- -- -- Iso.rightInv characFunSpace (f , g , p) = refl --- -- -- Iso.leftInv characFunSpace f = --- -- -- funExt λ { (inl x) → refl ; (inr x) → refl ; (push a b i) → refl} - --- -- -- coHomS¹-ish : (n : ℕ) → Type _ --- -- -- coHomS¹-ish n = hLevelTrunc 3 (S₊ 1 → coHomK (3 +ℕ n)) - --- -- -- fib : (n : ℕ) → coHomS¹-ish n → Type _ --- -- -- fib n x = --- -- -- trRec {B = TypeOfHLevel ℓ-zero 2} (isOfHLevelTypeOfHLevel 2) --- -- -- (λ f → ∥ (Σ[ g ∈ (S₊ 3 → coHomK (3 +ℕ n)) ] --- -- -- ((a : S₊ 1) (b : S₊ 3) → f a ≡ g b)) ∥₂ , squash₂) x .fst - --- -- -- contrFstΣ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} --- -- -- → (e : isContr A) --- -- -- → Iso (Σ A B) (B (fst e)) --- -- -- Iso.fun (contrFstΣ {B = B} e) (a , b) = subst B (sym (snd e a)) b --- -- -- Iso.inv (contrFstΣ {B = B} e) b = (fst e) , b --- -- -- Iso.rightInv (contrFstΣ {B = B} e) b = cong (λ x → subst B x b) h ∙ transportRefl b --- -- -- where --- -- -- h : sym (snd e (fst e)) ≡ refl --- -- -- h = isProp→isSet (isContr→isProp e) _ _ _ _ --- -- -- Iso.leftInv (contrFstΣ {B = B} e) b = --- -- -- ΣPathP ((snd e (fst b)) --- -- -- , (λ j → transp (λ i → B (snd e (fst b) (~ i ∨ j))) j (snd b))) - --- -- -- Iso1 : (n : ℕ) → Iso (coHom (3 +ℕ n) (join S¹ (S₊ 3))) ∥ Σ (coHomS¹-ish n) (fib n) ∥₂ --- -- -- Iso1 n = compIso (setTruncIso characFunSpace) Iso2 --- -- -- where --- -- -- Iso2 : Iso _ ∥ Σ (coHomS¹-ish n) (fib n) ∥₂ --- -- -- Iso.fun Iso2 = sMap λ F → ∣ fst F ∣ , ∣ (fst (snd F)) , (snd (snd F)) ∣₂ --- -- -- Iso.inv Iso2 = --- -- -- sRec squash₂ --- -- -- (uncurry (trElim (λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelSuc 2 squash₂) --- -- -- λ f → sRec squash₂ λ p → ∣ f , p ∣₂)) --- -- -- Iso.rightInv Iso2 = --- -- -- sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) --- -- -- (uncurry (trElim (λ _ → isOfHLevelΠ 3 λ _ → isProp→isOfHLevelSuc 2 (squash₂ _ _)) --- -- -- λ f → sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl )) --- -- -- Iso.leftInv Iso2 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl - --- -- -- isContrcoHomS¹-ish : (n : ℕ) → isContr (coHomS¹-ish n) --- -- -- isContrcoHomS¹-ish n = isOfHLevelRetractFromIso 0 (mapCompIso characFunSpaceS¹) isContrEnd --- -- -- where --- -- -- isContrEnd : isContr (hLevelTrunc 3 (Σ[ x ∈ coHomK (3 +ℕ n) ] x ≡ x)) --- -- -- fst isContrEnd = ∣ 0ₖ _ , refl ∣ --- -- -- snd isContrEnd = --- -- -- trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) --- -- -- (uncurry (coHomK-elim _ --- -- -- (λ _ → isOfHLevelΠ (suc (suc (suc n))) --- -- -- λ _ → isOfHLevelPlus' {n = n} 3 (isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _)) --- -- -- (λ p → (trRec (isOfHLevelPlus' {n = n} 3 (isOfHLevelTrunc 3) _ _) --- -- -- (λ p i → ∣ (0ₖ (3 +ℕ n) , p i) ∣) --- -- -- (Iso.fun (PathIdTruncIso _) --- -- -- (isContr→isProp --- -- -- (isConnectedPath _ (isConnectedKn (2 +ℕ n)) (0ₖ _) (0ₖ _)) ∣ refl ∣ ∣ p ∣)))))) - --- -- -- Iso2' : (n : ℕ) → Iso (∥ Σ (coHomS¹-ish n) (fib n) ∥₂) (fib n ∣ (λ _ → 0ₖ _) ∣) --- -- -- Iso2' n = compIso (setTruncIso (contrFstΣ centerChange)) --- -- -- (setTruncIdempotentIso squash₂) --- -- -- where --- -- -- centerChange : isContr (coHomS¹-ish n) --- -- -- fst centerChange = ∣ (λ _ → 0ₖ _) ∣ --- -- -- snd centerChange y = isContr→isProp (isContrcoHomS¹-ish n) _ _ - --- -- -- open import Cubical.Foundations.Equiv.HalfAdjoint --- -- -- open import Cubical.Relation.Nullary --- -- -- ll5 : (n : ℕ) → ¬ (n ≡ 2) → isContr (fib n ∣ (λ _ → 0ₖ _) ∣) --- -- -- ll5 n p = --- -- -- isOfHLevelRetractFromIso 0 --- -- -- (compIso --- -- -- (setTruncIso --- -- -- (compIso (Σ-cong-iso-snd λ f → --- -- -- (compIso characFunSpaceS¹ (invIso (Σ-cong-iso-fst (iso funExt⁻ funExt (λ _ → refl) λ _ → refl))))) --- -- -- (compIso ΣΣ (contrFstΣ (isContrSingl _))))) --- -- -- (compIso (setTruncIso (iso funExt⁻ funExt (λ _ → refl) λ _ → refl)) --- -- -- (compIso (setTruncIso (codomainIso (congIso (invIso (Iso-Kn-ΩKn+1 _))))) --- -- -- ((compIso (invIso (fst (coHom≅coHomΩ _ (S₊ _)))) --- -- -- (fst (Hⁿ-Sᵐ≅0 n 2 p))))))) --- -- -- isContrUnit - --- -- -- record ExactSeqℤ {ℓ : Level} (G : ℤ → Group ℓ) : Type ℓ where --- -- -- field --- -- -- maps : ∀ n → GroupHom (G n) (G (sucℤ n)) --- -- -- im⊂ker : ∀ n → ∀ g → isInIm (maps n) g → isInKer (maps (sucℤ n)) g --- -- -- ker⊂im : ∀ n → ∀ g → isInKer (maps (sucℤ n)) g → isInIm (maps n) g - --- -- -- record ExactSeqℕ {ℓ : Level} (G : ℕ → Group ℓ) : Type ℓ where --- -- -- field --- -- -- maps : ∀ n → GroupHom (G n) (G (suc n)) --- -- -- im⊂ker : ∀ n → ∀ g → isInIm (maps n) g → isInKer (maps (suc n)) g --- -- -- ker⊂im : ∀ n → ∀ g → isInKer (maps (suc n)) g → isInIm (maps n) g - - --- -- -- module _ {ℓ : Level} {A : Pointed ℓ} {n : ℕ} --- -- -- where --- -- -- funTyp : Type _ --- -- -- funTyp = A →∙ coHomK-ptd n - --- -- -- _++_ : funTyp → funTyp → funTyp --- -- -- fst (f ++ g) x = fst f x +ₖ fst g x --- -- -- snd (f ++ g) = cong₂ _+ₖ_ (snd f) (snd g) ∙ rUnitₖ _ (0ₖ _) - --- -- -- addAgree : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (x y : funTyp {A = A} {n = n}) --- -- -- → Path (fst (coHomRedGrDir n A)) --- -- -- ∣ x ++ y ∣₂ --- -- -- (∣ x ∣₂ +ₕ∙ ∣ y ∣₂) --- -- -- addAgree {A = A} zero f g = --- -- -- cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) refl) --- -- -- addAgree {A = A} (suc zero) f g = --- -- -- cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) refl) --- -- -- addAgree {A = A} (suc (suc n)) f g = --- -- -- cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) refl) - --- -- -- ss : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Type ℓ''} --- -- -- → isHomogeneous B --- -- -- → (x y : (A →∙ B)) (f : C → x ≡ y) --- -- -- → isEquiv (cong fst ∘ f) --- -- -- → isEquiv f --- -- -- ss homb x y f e = --- -- -- isoToIsEquiv --- -- -- (iso _ --- -- -- (λ p → invEq (_ , e) (cong fst p)) --- -- -- (λ p → →∙Homogeneous≡Path homb _ _ (secEq (_ , e) (cong fst p))) --- -- -- (retEq (_ , e))) - --- -- -- Whitehead : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} --- -- -- → (f : A → B) --- -- -- → isEmbedding f --- -- -- → (∀ (b : B) → ∃[ a ∈ A ] f a ≡ b) --- -- -- → isEquiv f --- -- -- equiv-proof (Whitehead f emb p) b = --- -- -- pRec isPropIsContr --- -- -- (λ fib → fib , isEmbedding→hasPropFibers emb b fib) --- -- -- (p b) - --- -- -- Int-ind : ∀ {ℓ} (P : ℤ → Type ℓ) --- -- -- → P (pos zero) → P (pos 1) --- -- -- → P (negsuc zero) --- -- -- → ((x y : ℤ) → P x → P y → P (x + y)) → (x : ℤ) → P x --- -- -- Int-ind P z one min ind (pos zero) = z --- -- -- Int-ind P z one min ind (pos (suc zero)) = one --- -- -- Int-ind P z one min ind (pos (suc (suc n))) = --- -- -- ind (pos (suc n)) 1 (Int-ind P z one min ind (pos (suc n))) one --- -- -- Int-ind P z one min ind (negsuc zero) = min --- -- -- Int-ind P z one min ind (negsuc (suc n)) = --- -- -- ind (negsuc n) (negsuc zero) (Int-ind P z one min ind (negsuc n)) min - --- -- -- genFunSpace : (n : ℕ) → S₊∙ n →∙ coHomK-ptd n --- -- -- fst (genFunSpace zero) false = 1 --- -- -- fst (genFunSpace zero) true = 0 --- -- -- snd (genFunSpace zero) = refl --- -- -- fst (genFunSpace (suc n)) = ∣_∣ --- -- -- snd (genFunSpace (suc zero)) = refl --- -- -- snd (genFunSpace (suc (suc n))) = refl - --- -- -- module _ where --- -- -- open import Cubical.Algebra.Monoid --- -- -- open import Cubical.Algebra.Semigroup --- -- -- open GroupStr --- -- -- open IsGroup --- -- -- open IsMonoid --- -- -- open IsSemigroup --- -- -- πS : (n : ℕ) → Group ℓ-zero --- -- -- fst (πS n) = S₊∙ n →∙ coHomK-ptd n --- -- -- 1g (snd (πS n)) = (λ _ → 0ₖ n) , refl --- -- -- GroupStr._·_ (snd (πS n)) = --- -- -- λ f g → (λ x → fst f x +ₖ fst g x) --- -- -- , cong₂ _+ₖ_ (snd f) (snd g) ∙ rUnitₖ n (0ₖ n) --- -- -- inv (snd (πS n)) f = (λ x → -ₖ fst f x) , cong -ₖ_ (snd f) ∙ -0ₖ {n = n} --- -- -- is-set (isSemigroup (isMonoid (isGroup (snd (πS zero))))) = --- -- -- isOfHLevelΣ 2 (isSetΠ (λ _ → isSetℤ)) --- -- -- λ _ → isOfHLevelPath 2 isSetℤ _ _ --- -- -- is-set (isSemigroup (isMonoid (isGroup (snd (πS (suc n)))))) = isOfHLevel↑∙' 0 n --- -- -- IsSemigroup.assoc (isSemigroup (isMonoid (isGroup (snd (πS n))))) x y z = --- -- -- →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ w → assocₖ n (fst x w) (fst y w) (fst z w)) --- -- -- fst (identity (isMonoid (isGroup (snd (πS n)))) (f , p)) = --- -- -- →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → rUnitₖ n (f x)) --- -- -- snd (identity (isMonoid (isGroup (snd (πS n)))) (f , p)) = --- -- -- →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → lUnitₖ n (f x)) --- -- -- fst (inverse (isGroup (snd (πS n))) (f , p)) = --- -- -- →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → rCancelₖ n (f x)) --- -- -- snd (inverse (isGroup (snd (πS n))) (f , p)) = --- -- -- →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → lCancelₖ n (f x)) - --- -- -- isSetπS : (n : ℕ) → isSet (S₊∙ n →∙ coHomK-ptd n) --- -- -- isSetπS = λ n → is-set (snd (πS n)) - --- -- -- K : (n : ℕ) → GroupIso (coHomRedGrDir n (S₊∙ n)) (πS n) --- -- -- fst (K n) = setTruncIdempotentIso (isSetπS n) --- -- -- snd (K zero) = --- -- -- makeIsGroupHom --- -- -- (sElim2 (λ _ _ → isOfHLevelPath 2 (isSetπS 0) _ _) --- -- -- λ f g → →∙Homogeneous≡ (isHomogeneousKn 0) refl) --- -- -- snd (K (suc zero)) = --- -- -- makeIsGroupHom --- -- -- (sElim2 (λ _ _ → isOfHLevelPath 2 (isSetπS 1) _ _) --- -- -- λ f g → →∙Homogeneous≡ (isHomogeneousKn 1) refl) --- -- -- snd (K (suc (suc n))) = --- -- -- makeIsGroupHom --- -- -- (sElim2 (λ _ _ → isOfHLevelPath 2 (isSetπS _) _ _) --- -- -- λ f g → →∙Homogeneous≡ (isHomogeneousKn _) refl) - --- -- -- t : ∀ {ℓ} {A : Pointed ℓ} → Iso ((Bool , true) →∙ A) (typ A) --- -- -- Iso.fun t f = fst f false --- -- -- fst (Iso.inv t a) false = a --- -- -- fst (Iso.inv (t {A = A}) a) true = pt A --- -- -- snd (Iso.inv t a) = refl --- -- -- Iso.rightInv t a = refl --- -- -- Iso.leftInv t (f , p) = --- -- -- ΣPathP ((funExt (λ { false → refl ; true → sym p})) , λ i j → p (~ i ∨ j)) - --- -- -- S1' : (n : ℕ) → GroupIso (πS n) ℤGroup --- -- -- fst (S1' zero) = t --- -- -- snd (S1' zero) = makeIsGroupHom λ _ _ → refl --- -- -- S1' (suc n) = --- -- -- compGroupIso --- -- -- (invGroupIso (K (suc n))) --- -- -- (compGroupIso --- -- -- (GroupEquiv→GroupIso (coHomGr≅coHomRedGr n (S₊∙ (suc n)))) --- -- -- (Hⁿ-Sⁿ≅ℤ n)) - --- -- -- S1 : (n : ℕ) → Iso (S₊∙ (suc n) →∙ coHomK-ptd (suc n)) ℤ --- -- -- S1 n = compIso (invIso (setTruncIdempotentIso (isOfHLevel↑∙' 0 n))) --- -- -- (compIso (equivToIso (fst (coHomGr≅coHomRedGr n (S₊∙ (suc n))))) --- -- -- (fst (Hⁿ-Sⁿ≅ℤ n))) - --- -- -- connFunSpace : (n i : ℕ) → (x y : S₊∙ n →∙ coHomK-ptd (suc i +' n)) → ∥ x ≡ y ∥ --- -- -- connFunSpace n i f g = Iso.fun PathIdTrunc₀Iso (isContr→isProp (lem n) ∣ f ∣₂ ∣ g ∣₂) --- -- -- where --- -- -- open import Cubical.Relation.Nullary --- -- -- natLem : (n i : ℕ) → ¬ (suc (i +ℕ n) ≡ n) --- -- -- natLem zero i = snotz --- -- -- natLem (suc n) i p = natLem n i (sym (+-suc i n) ∙ (cong predℕ p)) - --- -- -- lem : (n : ℕ) → isContr ∥ (S₊∙ n →∙ coHomK-ptd (suc i +' n)) ∥₂ --- -- -- fst (lem zero) = 0ₕ∙ (suc i) --- -- -- snd (lem zero) = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) --- -- -- λ f → pRec (squash₂ _ _) --- -- -- (λ id → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn (suc i)) --- -- -- (funExt λ { false → sym id ; true → sym (snd f)}))) --- -- -- (help (f .fst false)) --- -- -- where --- -- -- help : (x : coHomK (suc i)) → ∥ x ≡ 0ₖ _ ∥ --- -- -- help = coHomK-elim _ (λ _ → isProp→isOfHLevelSuc i squash) ∣ refl ∣ --- -- -- lem (suc n) = --- -- -- isOfHLevelRetractFromIso 0 --- -- -- (compIso (equivToIso (fst (coHomGr≅coHomRedGr (suc (i +ℕ n)) (S₊∙ (suc n))))) --- -- -- (fst (Hⁿ-Sᵐ≅0 (suc (i +ℕ n)) n (natLem n i)))) --- -- -- isContrUnit - --- -- -- S1Pres1 : (n : ℕ) → Iso.fun (fst (S1' (suc n))) (∣_∣ , refl) ≡ pos 1 --- -- -- S1Pres1 zero = refl --- -- -- S1Pres1 (suc n) = cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ n))) (lem n) ∙ S1Pres1 n --- -- -- where --- -- -- lem : (n : ℕ) → Iso.fun (fst (suspensionAx-Sn n n)) ∣ ∣_∣ ∣₂ ≡ ∣ ∣_∣ ∣₂ --- -- -- lem zero = cong ∣_∣₂ (funExt λ x → transportRefl (∣ x ∣ +ₖ (0ₖ 1)) ∙ rUnitₖ 1 ∣ x ∣) --- -- -- lem (suc n) = cong ∣_∣₂ --- -- -- (funExt λ x → (λ i → transportRefl ((ΩKn+1→Kn (suc (suc n)) --- -- -- (transp (λ j → 0ₖ (suc (suc (suc n))) ≡ ∣ merid north (~ j ∧ ~ i) ∣) i --- -- -- (λ z → ∣ compPath-filler (merid (transportRefl (transportRefl x i) i)) (sym (merid north)) i z --- -- -- ∣)))) i) --- -- -- ∙ Iso.leftInv (Iso-Kn-ΩKn+1 (suc (suc n))) ∣ x ∣) - --- -- -- S1Pres1← : (n : ℕ) → Iso.inv (fst (S1' (suc n))) (pos 1) ≡ (∣_∣ , refl) --- -- -- S1Pres1← n = sym (cong (Iso.inv (fst (S1' (suc n)))) (S1Pres1 n)) ∙ Iso.leftInv (fst (S1' (suc n))) (∣_∣ , refl) - --- -- -- IsoFunSpace : (n : ℕ) → Iso (S₊∙ n →∙ coHomK-ptd n) ℤ --- -- -- IsoFunSpace n = fst (S1' n) - --- -- -- module g-base where --- -- -- g : (n : ℕ) (i : ℕ) → coHomK i → (S₊∙ n →∙ coHomK-ptd (i +' n)) --- -- -- fst (g n i x) y = x ⌣ₖ (genFunSpace n) .fst y --- -- -- snd (g n i x) = cong (x ⌣ₖ_) ((genFunSpace n) .snd) ∙ ⌣ₖ-0ₖ i n x - --- -- -- G : (n : ℕ) (i : ℕ) → coHomK i → (S₊∙ n →∙ coHomK-ptd (n +' i)) --- -- -- fst (G n i x) y = (genFunSpace n) .fst y ⌣ₖ x --- -- -- snd (G n i x) = cong (_⌣ₖ x) ((genFunSpace n) .snd) ∙ 0ₖ-⌣ₖ n i x - --- -- -- eq1 : (n : ℕ) (i : ℕ) → (S₊∙ n →∙ coHomK-ptd (i +' n)) ≃ (S₊∙ n →∙ coHomK-ptd (i +' n)) --- -- -- eq1 n i = isoToEquiv (iso F F FF FF) --- -- -- where --- -- -- lem : (i n : ℕ) → (-ₖ^ i · n) (snd (coHomK-ptd (i +' n))) ≡ 0ₖ _ --- -- -- lem zero zero = refl --- -- -- lem zero (suc zero) = refl --- -- -- lem zero (suc (suc n)) = refl --- -- -- lem (suc zero) zero = refl --- -- -- lem (suc (suc i)) zero = refl --- -- -- lem (suc i) (suc n) = refl - --- -- -- F : S₊∙ n →∙ coHomK-ptd (i +' n) → S₊∙ n →∙ coHomK-ptd (i +' n) --- -- -- fst (F f) x = (-ₖ^ i · n) (fst f x) --- -- -- snd (F f) = cong (-ₖ^ i · n) (snd f) ∙ lem i n - --- -- -- FF : (x : _) → F (F x) ≡ x --- -- -- FF x = --- -- -- →∙Homogeneous≡ (isHomogeneousKn _) --- -- -- (funExt λ y → -ₖ-gen² i n _ _ (fst x y)) - --- -- -- rCancel'' : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) → sym p ∙∙ refl ∙∙ p ≡ refl --- -- -- rCancel'' p = (λ j → (λ i → p (~ i ∨ j)) ∙∙ refl ∙∙ λ i → p (i ∨ j)) ∙ sym (rUnit refl) - --- -- -- transpPres0 : ∀ {k m : ℕ} (p : k ≡ m) → subst coHomK p (0ₖ k) ≡ 0ₖ m --- -- -- transpPres0 {k = k} = J (λ m p → subst coHomK p (0ₖ k) ≡ 0ₖ m) (transportRefl _) - --- -- -- eq2 : (n : ℕ) (i : ℕ) → (S₊∙ n →∙ coHomK-ptd (n +' i)) ≃ (S₊∙ n →∙ coHomK-ptd (i +' n)) --- -- -- eq2 n i = --- -- -- isoToEquiv (iso (λ f → (λ x → subst coHomK (+'-comm n i) (fst f x)) , --- -- -- cong (subst coHomK (+'-comm n i)) (snd f) ∙ transpPres0 (+'-comm n i)) --- -- -- (λ f → (λ x → subst coHomK (sym (+'-comm n i)) (fst f x)) --- -- -- , (cong (subst coHomK (sym (+'-comm n i))) (snd f) ∙ transpPres0 (sym (+'-comm n i)))) --- -- -- (λ f → →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ x → transportTransport⁻ _ (f .fst x))) --- -- -- λ f → →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ x → transportTransport⁻ _ (f .fst x))) - --- -- -- g≡ : (n : ℕ) (i : ℕ) → g n i ≡ λ x → fst (compEquiv (eq2 n i) (eq1 n i)) ((G n i) x) --- -- -- g≡ n i = funExt (λ f → →∙Homogeneous≡ (isHomogeneousKn _) --- -- -- (funExt λ y → gradedComm-⌣ₖ _ _ f (genFunSpace n .fst y))) - --- -- -- glIso3-fun : (n m : ℕ) → --- -- -- (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) --- -- -- → (S₊ n → coHomK m) --- -- -- glIso3-fun zero m (f , p) false = ΩKn+1→Kn _ (sym p ∙∙ cong f loop ∙∙ p) --- -- -- glIso3-fun zero m (f , p) true = 0ₖ _ --- -- -- glIso3-fun (suc n) m (f , p) x = --- -- -- ΩKn+1→Kn _ (sym p ∙∙ cong f (merid x ∙ sym (merid (ptSn _))) ∙∙ p) - --- -- -- glIso3-fun∙ : (n m : ℕ) → (f : _) → glIso3-fun n m f (ptSn _) ≡ 0ₖ m --- -- -- glIso3-fun∙ zero m = λ _ → refl --- -- -- glIso3-fun∙ (suc n) m (f , p) = --- -- -- cong (ΩKn+1→Kn m) (cong (sym p ∙∙_∙∙ p) (cong (cong f) (rCancel (merid (ptSn _))))) --- -- -- ∙∙ cong (ΩKn+1→Kn m) (rCancel'' p) --- -- -- ∙∙ ΩKn+1→Kn-refl m - - --- -- -- glIso3-inv : (n m : ℕ) → (S₊∙ n →∙ coHomK-ptd m) → (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) --- -- -- fst (glIso3-inv zero m (f , p)) base = 0ₖ _ --- -- -- fst (glIso3-inv zero m (f , p)) (loop i) = Kn→ΩKn+1 m (f false) i --- -- -- snd (glIso3-inv zero m (f , p)) = refl --- -- -- fst (glIso3-inv (suc n) m (f , p)) north = 0ₖ _ --- -- -- fst (glIso3-inv (suc n) m (f , p)) south = 0ₖ _ --- -- -- fst (glIso3-inv (suc n) m (f , p)) (merid a i) = Kn→ΩKn+1 m (f a) i --- -- -- snd (glIso3-inv (suc n) m (f , p)) = refl - --- -- -- glIso3 : (n m : ℕ) → --- -- -- Iso (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) --- -- -- (S₊∙ n →∙ coHomK-ptd m) --- -- -- Iso.fun (glIso3 n m) f = (glIso3-fun n m f) , (glIso3-fun∙ n m f) --- -- -- Iso.inv (glIso3 n m) = glIso3-inv n m --- -- -- Iso.rightInv (glIso3 zero m) (f , p) = --- -- -- →∙Homogeneous≡ (isHomogeneousKn _) --- -- -- (funExt λ { false → cong (ΩKn+1→Kn m) (sym (rUnit _)) ∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f false) --- -- -- ; true → sym p}) --- -- -- Iso.rightInv (glIso3 (suc n) m) (f , p) = --- -- -- →∙Homogeneous≡ (isHomogeneousKn _) --- -- -- (funExt λ x → (λ i → ΩKn+1→Kn m (sym (rUnit (cong-∙ (glIso3-inv (suc n) m (f , p) .fst) (merid x) (sym (merid (ptSn _))) i)) i)) --- -- -- ∙∙ cong (ΩKn+1→Kn m) (cong (Kn→ΩKn+1 m (f x) ∙_) (cong sym (cong (Kn→ΩKn+1 m) p ∙ Kn→ΩKn+10ₖ m)) ∙ sym (rUnit _)) --- -- -- ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f x)) --- -- -- Iso.leftInv (glIso3 zero m) (f , p) = →∙Homogeneous≡ (isHomogeneousKn _) --- -- -- (funExt λ { base → sym p --- -- -- ; (loop i) j → h j i}) --- -- -- where --- -- -- h : PathP (λ i → p (~ i) ≡ p (~ i)) (Kn→ΩKn+1 m (ΩKn+1→Kn m (sym p ∙∙ cong f loop ∙∙ p))) (cong f loop) --- -- -- h = Iso.rightInv (Iso-Kn-ΩKn+1 _) _ --- -- -- ◁ λ i → doubleCompPath-filler (sym p) (cong f loop) p (~ i) --- -- -- Iso.leftInv (glIso3 (suc n) m) (f , p) = --- -- -- →∙Homogeneous≡ (isHomogeneousKn _) --- -- -- (funExt λ { north → sym p --- -- -- ; south → sym p ∙ cong f (merid (ptSn _)) --- -- -- ; (merid a i) j → h a j i}) --- -- -- where --- -- -- h : (a : S₊ (suc n)) → PathP (λ i → p (~ i) ≡ (sym p ∙ cong f (merid (ptSn (suc n)))) i) --- -- -- (Kn→ΩKn+1 m (ΩKn+1→Kn m (sym p ∙∙ cong f (merid a ∙ sym (merid (ptSn _))) ∙∙ p))) --- -- -- (cong f (merid a)) --- -- -- h a = Iso.rightInv (Iso-Kn-ΩKn+1 _) _ --- -- -- ◁ λ i j → hcomp (λ k → λ { (i = i1) → (f (merid a j)) --- -- -- ; (j = i0) → p (k ∧ ~ i) --- -- -- ; (j = i1) → compPath-filler' (sym p) (cong f (merid (ptSn _))) k i }) --- -- -- (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) - --- -- -- glIsoInvHom : (n m : ℕ) (x y : coHomK n) (z : S₊ (suc m)) --- -- -- → Iso.inv (glIso3 _ _) (G m n (x +ₖ y)) .fst z --- -- -- ≡ Iso.inv (glIso3 _ _) (G m n x) .fst z --- -- -- +ₖ Iso.inv (glIso3 _ _) (G m n y) .fst z --- -- -- glIsoInvHom zero zero x y base = refl --- -- -- glIsoInvHom (suc n) zero x y base = refl --- -- -- glIsoInvHom zero zero x y (loop i) j = h j i --- -- -- where --- -- -- h : (cong (Iso.inv (glIso3 _ _) (G zero zero (x + y)) .fst) loop) --- -- -- ≡ cong₂ _+ₖ_ (cong (Iso.inv (glIso3 _ _) (G zero zero x) .fst) loop) --- -- -- (cong (Iso.inv (glIso3 _ _) (G zero zero y) .fst) loop) --- -- -- h = Kn→ΩKn+1-hom 0 x y --- -- -- ∙ ∙≡+₁ (cong (Iso.inv (glIso3 _ _) (G zero zero x) .fst) loop) --- -- -- (cong (Iso.inv (glIso3 _ _) (G zero zero y) .fst) loop) --- -- -- glIsoInvHom (suc n) zero x y (loop i) j = h j i --- -- -- where --- -- -- h : Kn→ΩKn+1 (suc n) ((pos (suc zero)) ·₀ (x +ₖ y)) --- -- -- ≡ cong₂ _+ₖ_ (cong (Iso.inv (glIso3 zero (zero +' suc n)) (G zero (suc n) x) .fst) loop) --- -- -- (cong (Iso.inv (glIso3 zero (zero +' suc n)) (G zero (suc n) y) .fst) loop) --- -- -- h = cong (Kn→ΩKn+1 (suc n)) (lUnit⌣ₖ _ (x +ₖ y)) --- -- -- ∙∙ Kn→ΩKn+1-hom (suc n) x y --- -- -- ∙∙ (λ i → ∙≡+₂ n (Kn→ΩKn+1 (suc n) (lUnit⌣ₖ _ x (~ i))) --- -- -- (Kn→ΩKn+1 (suc n) (lUnit⌣ₖ _ y (~ i))) i) --- -- -- glIsoInvHom zero (suc m) x y north = refl --- -- -- glIsoInvHom zero (suc m) x y south = refl --- -- -- glIsoInvHom zero (suc m) x y (merid a i) j = h j i --- -- -- where --- -- -- h : Kn→ΩKn+1 (suc m) (_⌣ₖ_ {n = suc m} {m = zero} ∣ a ∣ (x + y)) --- -- -- ≡ cong₂ _+ₖ_ (Kn→ΩKn+1 (suc m) (_⌣ₖ_ {n = suc m} {m = zero} ∣ a ∣ x)) --- -- -- (Kn→ΩKn+1 (suc m) (_⌣ₖ_ {n = suc m} {m = zero} ∣ a ∣ y)) --- -- -- h = cong (Kn→ΩKn+1 (suc m)) (leftDistr-⌣ₖ (suc m) 0 ∣ a ∣ x y) --- -- -- ∙∙ Kn→ΩKn+1-hom (suc m) _ _ --- -- -- ∙∙ ∙≡+₂ _ _ _ --- -- -- glIsoInvHom (suc n) (suc m) x y north = refl --- -- -- glIsoInvHom (suc n) (suc m) x y south = refl --- -- -- glIsoInvHom (suc n) (suc m) x y (merid a i) j = h j i --- -- -- where --- -- -- h : Kn→ΩKn+1 (suc (suc (m +ℕ n))) (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ (x +ₖ y)) --- -- -- ≡ cong₂ _+ₖ_ (Kn→ΩKn+1 (suc (suc (m +ℕ n))) (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ x)) --- -- -- (Kn→ΩKn+1 (suc (suc (m +ℕ n))) (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ y)) --- -- -- h = cong (Kn→ΩKn+1 (suc (suc (m +ℕ n)))) --- -- -- (leftDistr-⌣ₖ (suc m) (suc n) ∣ a ∣ x y) --- -- -- ∙∙ Kn→ΩKn+1-hom _ _ _ --- -- -- ∙∙ ∙≡+₂ _ _ _ - --- -- -- +'-suc : (n m : ℕ) → suc n +' m ≡ suc (n +' m) --- -- -- +'-suc zero zero = refl --- -- -- +'-suc (suc n) zero = refl --- -- -- +'-suc zero (suc m) = refl --- -- -- +'-suc (suc n) (suc m) = refl - --- -- -- LEM : (i n : ℕ) (x : coHomK i) --- -- -- → Path _ (G (suc n) i x) --- -- -- (subst (λ x → S₊∙ (suc n) →∙ coHomK-ptd x) --- -- -- (sym (+'-suc n i)) --- -- -- ((Iso.inv (glIso3 n _)) (G n i x))) --- -- -- LEM zero zero x = --- -- -- →∙Homogeneous≡ (isHomogeneousKn _) --- -- -- (funExt λ z → (λ i → x ·₀ ∣ z ∣) ∙ h x z ∙ sym (transportRefl _)) --- -- -- where --- -- -- h : (x : ℤ) (z : S¹) → _·₀_ {n = 1} x ∣ z ∣ ≡ fst (Iso.inv (glIso3 0 zero) (G zero zero x)) z --- -- -- h = Int-ind _ (λ { base → refl ; (loop i) j → Kn→ΩKn+10ₖ zero (~ j) i}) --- -- -- (λ { base → refl ; (loop i) j → rUnit (cong ∣_∣ₕ (lUnit loop j)) j i}) --- -- -- (λ { base → refl ; (loop i) j → rUnit (cong ∣_∣ₕ (sym loop)) j i}) --- -- -- λ x y inx iny z --- -- -- → rightDistr-⌣ₖ 0 1 x y ∣ z ∣ --- -- -- ∙∙ cong₂ _+ₖ_ (inx z) (iny z) --- -- -- ∙∙ sym (glIsoInvHom zero zero x y z) --- -- -- LEM (suc i) zero x = --- -- -- →∙Homogeneous≡ (isHomogeneousKn _) --- -- -- (funExt λ z → h z --- -- -- ∙ sym (transportRefl --- -- -- ((Iso.inv (glIso3 zero (suc i)) (G zero (suc i) x)) .fst z))) --- -- -- where --- -- -- h : (z : S₊ 1) → _ ≡ Iso.inv (glIso3 zero (suc i)) (G zero (suc i) x) .fst z --- -- -- h base = refl --- -- -- h (loop k) j = Kn→ΩKn+1 (suc i) (lUnit⌣ₖ _ x (~ j)) k --- -- -- LEM zero (suc n) x = --- -- -- →∙Homogeneous≡ (isHomogeneousKn _) --- -- -- (funExt λ z → h x z ∙ λ i → transportRefl (Iso.inv (glIso3 (suc n) (suc n)) (G (suc n) zero x)) (~ i) .fst z) --- -- -- where --- -- -- +merid : rUnitₖ (suc (suc n)) ∣ south ∣ ≡ cong ∣_∣ₕ (merid (ptSn _)) --- -- -- +merid = sym (lUnit _) --- -- -- ∙ cong (cong ∣_∣ₕ) --- -- -- λ j i → transp (λ _ → S₊ (suc (suc n))) (i ∨ j) (merid (ptSn (suc n)) i) --- -- -- open import Cubical.Foundations.Path - --- -- -- pp : (a : S₊ (suc n)) → PathP (λ i → ∣ merid a i ∣ₕ ≡ Kn→ΩKn+1 (suc n) (∣ a ∣ +ₖ (0ₖ _)) i) --- -- -- refl (sym (rUnitₖ (suc (suc n)) ∣ south ∣)) --- -- -- pp a = flipSquare ((λ i j → ∣ compPath-filler (merid a) (sym (merid (ptSn _))) i j ∣ₕ ) --- -- -- ▷ cong (Kn→ΩKn+1 (suc n)) (sym (rUnitₖ (suc n) ∣ a ∣ₕ))) --- -- -- ▷ sym (cong sym +merid) - --- -- -- pp2 : (a : S₊ (suc n)) → (λ i → -ₖ ∣ merid a i ∣) --- -- -- ≡ Kn→ΩKn+1 (suc n) (-ₖ ∣ a ∣) --- -- -- pp2 a = cong (cong ∣_∣ₕ) (sym (symDistr (merid a) (sym (merid (ptSn (suc n)))))) --- -- -- ∙ sym (Kn→ΩKn+1-ₖ (suc n) ∣ a ∣) - --- -- -- h : (x : ℤ) (z : S₊ (suc (suc n))) --- -- -- → _⌣ₖ_ {n = suc (suc n)} {m = 0} ∣ z ∣ x --- -- -- ≡ Iso.inv (glIso3 (suc n) (suc n)) (G (suc n) zero x) .fst z --- -- -- h = Int-ind _ --- -- -- (λ { north → refl ; south → refl ; (merid a i) j → Kn→ΩKn+10ₖ (suc n) (~ j) i}) --- -- -- (λ { north → refl ; south → refl --- -- -- ; (merid a i) j → hcomp (λ k → λ { (i = i0) → ∣ north ∣ --- -- -- ; (i = i1) → rUnitₖ (suc (suc n)) ∣ south ∣ (~ k) --- -- -- ; (j = i0) → rUnitₖ (suc (suc n)) ∣ merid a i ∣ (~ k) --- -- -- ; (j = i1) → pp a i k}) --- -- -- ∣ merid a i ∣ₕ}) --- -- -- (λ { north → refl --- -- -- ; south → refl --- -- -- ; (merid a i) j → pp2 a j i}) --- -- -- λ x y indx indy z → leftDistr-⌣ₖ _ _ ∣ z ∣ x y --- -- -- ∙ cong₂ _+ₖ_ (indx z) (indy z) --- -- -- ∙ sym (glIsoInvHom _ _ _ _ _) --- -- -- LEM (suc i) (suc n) x = --- -- -- →∙Homogeneous≡ (isHomogeneousKn _) --- -- -- (funExt λ z → h z --- -- -- ∙ λ j → transportRefl ((Iso.inv (glIso3 (suc n) (suc (suc (n +ℕ i)))) --- -- -- (G (suc n) (suc i) x))) (~ j) .fst z) --- -- -- where --- -- -- h : (z : S₊ (suc (suc n))) → _ --- -- -- h north = refl --- -- -- h south = refl --- -- -- h (merid a i) = refl - --- -- -- isEquivGzero : (i : ℕ) → isEquiv (G zero i) --- -- -- isEquivGzero i = --- -- -- isoToIsEquiv --- -- -- (iso _ (λ f → fst f false) --- -- -- (λ {(f , p) → →∙Homogeneous≡ (isHomogeneousKn _) --- -- -- (funExt λ { false → rUnitₖ _ (f false) ; true → sym p})}) --- -- -- (lUnit⌣ₖ _)) - --- -- -- isEquivG : (n i : ℕ) → isEquiv (G n i) --- -- -- isEquivG zero i = --- -- -- isoToIsEquiv --- -- -- (iso _ (λ f → fst f false) --- -- -- (λ {(f , p) → →∙Homogeneous≡ (isHomogeneousKn _) --- -- -- (funExt λ { false → rUnitₖ _ (f false) ; true → sym p})}) --- -- -- (lUnit⌣ₖ _)) --- -- -- isEquivG (suc n) i = --- -- -- subst isEquiv (sym (funExt (LEM i n))) --- -- -- (compEquiv (compEquiv (G n i , isEquivG n i) --- -- -- (isoToEquiv (invIso (glIso3 n (n +' i))))) --- -- -- (pathToEquiv (λ j → S₊∙ (suc n) →∙ coHomK-ptd (+'-suc n i (~ j)))) .snd) - - --- -- -- isEquiv-g : (n i : ℕ) → isEquiv (g n i) --- -- -- isEquiv-g n i = --- -- -- subst isEquiv (sym (g≡ n i)) --- -- -- (compEquiv (G n i , isEquivG n i) (compEquiv (eq2 n i) (eq1 n i)) .snd) - --- -- -- module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) --- -- -- (conB : (x y : typ B) → ∥ x ≡ y ∥) --- -- -- (n : ℕ) (Q-is : Iso (typ (Q (pt B))) (S₊ n)) --- -- -- (Q-is-ptd : Iso.fun Q-is (pt (Q (pt B))) ≡ snd (S₊∙ n)) --- -- -- (c : (b : typ B) → (Q b →∙ coHomK-ptd n)) --- -- -- (c-pt : c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) where - --- -- -- g : (b : typ B) (i : ℕ) → coHomK i → (Q b →∙ coHomK-ptd (i +' n)) --- -- -- fst (g b i x) y = x ⌣ₖ c b .fst y --- -- -- snd (g b i x) = cong (x ⌣ₖ_) (c b .snd) ∙ ⌣ₖ-0ₖ i n x - --- -- -- g' : (b : typ B) (i : ℕ) → coHomK i → coHomK i → (Q b →∙ coHomK-ptd (i +' n)) --- -- -- fst (g' b i x y) z = x ⌣ₖ c b .fst z +ₖ y ⌣ₖ c b .fst z --- -- -- snd (g' b i x y) = cong₂ _+ₖ_ (cong (x ⌣ₖ_) (c b .snd) ∙ ⌣ₖ-0ₖ i n x) --- -- -- (cong (y ⌣ₖ_) (c b .snd) ∙ ⌣ₖ-0ₖ i n y) ∙ rUnitₖ _ (0ₖ _) - --- -- -- g-hom : (b : typ B) (i : ℕ) → (x y : coHomK i) → g b i (x +ₖ y) ≡ ((g b i x) ++ (g b i y)) --- -- -- g-hom b i x y = →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ z → rightDistr-⌣ₖ i n x y (c b .fst z)) - --- -- -- gPathP' : (i : ℕ) → PathP (λ j → coHomK i → (isoToPath Q-is j , ua-gluePath (isoToEquiv Q-is) (Q-is-ptd) j) --- -- -- →∙ coHomK-ptd (i +' n)) (g (pt B) i) (g-base.g n i) --- -- -- gPathP' i = --- -- -- toPathP --- -- -- (funExt --- -- -- λ x → →∙Homogeneous≡ (isHomogeneousKn _) --- -- -- (funExt λ y → (λ i → transportRefl (transportRefl x i ⌣ₖ c (pt B) .fst (Iso.inv Q-is (transportRefl y i))) i) --- -- -- ∙ cong (x ⌣ₖ_) (funExt⁻ c-pt (Iso.inv Q-is y) ∙ cong (genFunSpace n .fst) (Iso.rightInv Q-is y)))) - --- -- -- g-base : (i : ℕ) → isEquiv (g (pt B) i) --- -- -- g-base i = transport (λ j → isEquiv (gPathP' i (~ j))) (g-base.isEquiv-g n i) - --- -- -- g-equiv : (b : typ B) (i : ℕ) → isEquiv (g b i) --- -- -- g-equiv b i = --- -- -- pRec (isPropIsEquiv _) --- -- -- (J (λ b _ → isEquiv (g b i)) --- -- -- (g-base i)) --- -- -- (conB (pt B) b) - --- -- -- module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) --- -- -- (conB : (x y : typ B) → ∥ x ≡ y ∥₂) --- -- -- (n : ℕ) (Q-is : Iso (typ (Q (pt B))) (S₊ n)) --- -- -- (Q-is-ptd : Iso.fun Q-is (pt (Q (pt B))) ≡ snd (S₊∙ n)) where - --- -- -- is-setQ→K : (b : typ B) → isSet (Q b →∙ coHomK-ptd n) --- -- -- is-setQ→K b = --- -- -- sRec (isProp→isOfHLevelSuc 1 isPropIsSet) --- -- -- (J (λ b _ → isSet (Q b →∙ coHomK-ptd n)) --- -- -- (subst isSet (cong (_→∙ coHomK-ptd n) --- -- -- (ua∙ (isoToEquiv (invIso Q-is)) (cong (Iso.inv Q-is) (sym Q-is-ptd) ∙ Iso.leftInv Q-is _))) --- -- -- (isOfHLevelRetractFromIso 2 (fst (S1' n)) isSetℤ))) --- -- -- (conB (pt B) b) - - --- -- -- isConnB : isConnected 3 (typ B) --- -- -- fst isConnB = ∣ pt B ∣ --- -- -- snd isConnB = --- -- -- trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) --- -- -- λ a → sRec (isOfHLevelTrunc 3 _ _) (cong ∣_∣ₕ) (conB (pt B) a) - --- -- -- isPropPath : isProp (∥ pt B ≡ pt B ∥₂) --- -- -- isPropPath = --- -- -- isOfHLevelRetractFromIso 1 setTruncTrunc2Iso --- -- -- (isContr→isProp (isConnectedPath _ isConnB (pt B) (pt B))) - --- -- -- c* : Σ[ c ∈ ((b : typ B) → (Q b →∙ coHomK-ptd n)) ] --- -- -- (c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) --- -- -- fst c* b = --- -- -- sRec (is-setQ→K b) --- -- -- (J (λ b _ → Q b →∙ coHomK-ptd n) --- -- -- ((λ x → genFunSpace n .fst (Iso.fun Q-is x)) --- -- -- , cong (genFunSpace n .fst) Q-is-ptd ∙ genFunSpace n .snd)) (conB (pt B) b) --- -- -- snd c* = --- -- -- funExt λ x → (λ i → sRec (is-setQ→K (pt B)) --- -- -- (J (λ b _ → Q b →∙ coHomK-ptd n) --- -- -- ((λ x₁ → genFunSpace n .fst (Iso.fun Q-is x₁)) , --- -- -- (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) --- -- -- (isPropPath (conB (pt B) (pt B)) ∣ refl ∣₂ i) .fst x) --- -- -- ∙ (λ i → transportRefl (genFunSpace n .fst (Iso.fun Q-is (transportRefl x i))) i) - --- -- -- p-help : {b : fst B} (p : pt B ≡ b) → (subst (fst ∘ Q) (sym p) (snd (Q b))) ≡ (snd (Q (pt B))) --- -- -- p-help {b = b} = --- -- -- J (λ b p → subst (fst ∘ Q) (sym p) (snd (Q b)) ≡ snd (Q (pt B))) (transportRefl _) - --- -- -- cEquiv : (b : fst B) (p : ∥ pt B ≡ b ∥₂) --- -- -- → (c* .fst b) --- -- -- ≡ sRec (is-setQ→K b) --- -- -- (λ pp → (λ qb → genFunSpace n .fst (Iso.fun Q-is (subst (fst ∘ Q) (sym pp) qb))) --- -- -- , cong (genFunSpace n .fst ∘ Iso.fun Q-is) (p-help pp) --- -- -- ∙ ((λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) p --- -- -- cEquiv b = --- -- -- sElim (λ _ → isOfHLevelPath 2 (is-setQ→K b) _ _) --- -- -- (J (λ b a → c* .fst b ≡ --- -- -- sRec (is-setQ→K b) (λ pp → --- -- -- (λ qb → --- -- -- genFunSpace n .fst (Iso.fun Q-is (subst (fst ∘ Q) (sym pp) qb))) --- -- -- , --- -- -- cong (genFunSpace n .fst ∘ Iso.fun Q-is) (p-help pp) ∙ --- -- -- (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd) --- -- -- ∣ a ∣₂) --- -- -- ((λ i → sRec (is-setQ→K (pt B)) --- -- -- (J (λ b₁ _ → Q b₁ →∙ coHomK-ptd n) --- -- -- ((λ x → genFunSpace n .fst (Iso.fun Q-is x)) , --- -- -- (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) --- -- -- (isPropPath (conB (pt B) (pt B)) ∣ refl ∣₂ i)) --- -- -- ∙ →∙Homogeneous≡ (isHomogeneousKn n) --- -- -- (transportRefl ((λ x → genFunSpace n .fst (Iso.fun Q-is x))) --- -- -- ∙ funExt λ x → cong (genFunSpace n .fst ∘ Iso.fun Q-is) (sym (transportRefl x))))) - --- -- -- module _ {ℓ ℓ'} (B : Pointed ℓ) (P : typ B → Type ℓ') where --- -- -- E : Type _ --- -- -- E = Σ (typ B) P - --- -- -- Ẽ : Type _ --- -- -- Ẽ = Pushout {A = E} (λ _ → tt) fst - --- -- -- i : (n : ℕ) → (P-base : Iso (P (pt B)) (S₊ n)) → S₊ (suc n) → Ẽ --- -- -- i zero P-base base = inr (pt B) --- -- -- i zero P-base (loop j) = (sym (push (pt B , Iso.inv P-base false)) --- -- -- ∙ push ((pt B) , Iso.inv P-base true)) j --- -- -- i (suc n) P-base north = inl tt --- -- -- i (suc n) P-base south = inr (pt B) --- -- -- i (suc n) P-base (merid a i₁) = push (pt B , Iso.inv P-base a) i₁ - --- -- -- Q : typ B → Pointed ℓ' --- -- -- Q x = Susp (P x) , north - --- -- -- F : Type _ --- -- -- F = Σ (typ B) λ x → Q x .fst - --- -- -- F̃ : Type _ --- -- -- F̃ = Pushout {A = typ B} {C = F} (λ _ → tt) λ b → b , north - --- -- -- invFE : Ẽ → F̃ --- -- -- invFE (inl x) = inl tt --- -- -- invFE (inr x) = inr (x , south) --- -- -- invFE (push (x , a) i₁) = ((push x) ∙ λ i → inr (x , merid a i)) i₁ - --- -- -- funFE : F̃ → Ẽ --- -- -- funFE (inl x) = inl tt --- -- -- funFE (inr (x , north)) = inl tt --- -- -- funFE (inr (x , south)) = inr x --- -- -- funFE (inr (x , merid a i₁)) = push (x , a) i₁ --- -- -- funFE (push a i₁) = inl tt - --- -- -- IsoFE : Iso F̃ Ẽ --- -- -- Iso.fun IsoFE = funFE --- -- -- Iso.inv IsoFE = invFE --- -- -- Iso.rightInv IsoFE (inl x) = refl --- -- -- Iso.rightInv IsoFE (inr x) = refl --- -- -- Iso.rightInv IsoFE (push (x , a) i₁) k = h k i₁ --- -- -- where --- -- -- h : cong funFE (((push x) ∙ λ i → inr (x , merid a i))) --- -- -- ≡ push (x , a) --- -- -- h = congFunct funFE (push x) (λ i → inr (x , merid a i)) --- -- -- ∙ sym (lUnit (push (x , a))) --- -- -- Iso.leftInv IsoFE (inl x) = refl --- -- -- Iso.leftInv IsoFE (inr (x , north)) = push x --- -- -- Iso.leftInv IsoFE (inr (x , south)) = refl --- -- -- Iso.leftInv IsoFE (inr (x , merid a i)) j = --- -- -- compPath-filler' (push x) (λ i₁ → inr (x , merid a i₁)) (~ j) i --- -- -- Iso.leftInv IsoFE (push a i₁) k = push a (i₁ ∧ k) - - --- -- -- funDir1 : ∀ {ℓ} {A : Pointed ℓ} → (F̃ , inl tt) →∙ A → (b : typ B) → Q b →∙ A --- -- -- fst (funDir1 {A = A , a} (f , p) b) north = f (inr (b , north)) --- -- -- fst (funDir1 {A = A , a} (f , p) b) south = f (inr (b , south)) --- -- -- fst (funDir1 {A = A , a} (f , p) b) (merid a₁ i₁) = f (inr (b , merid a₁ i₁)) --- -- -- snd (funDir1 {A = A , a} (f , p) b) = sym (cong f (push b)) ∙ p - --- -- -- funDir2 : ∀ {ℓ} {A : Pointed ℓ} → ((b : typ B) → Q b →∙ A) → (F̃ , inl tt) →∙ A --- -- -- fst (funDir2 {A = A , a} f) (inl x) = a --- -- -- fst (funDir2 {A = A , a} f) (inr (x , north)) = f x .fst north --- -- -- fst (funDir2 {A = A , a} f) (inr (x , south)) = f x .fst south --- -- -- fst (funDir2 {A = A , a} f) (inr (x , merid a₁ i₁)) = f x .fst (merid a₁ i₁) --- -- -- fst (funDir2 {A = A , a} f) (push a₁ i₁) = snd (f a₁) (~ i₁) --- -- -- snd (funDir2 {A = A , a} f) = refl - --- -- -- funDir2-hom : (n : ℕ) → (f g : ((b : typ B) → Q b →∙ coHomK-ptd n)) --- -- -- → funDir2 (λ b → f b ++ g b) ≡ (funDir2 f ++ funDir2 g) --- -- -- funDir2-hom n f g = --- -- -- →∙Homogeneous≡ (isHomogeneousKn _) --- -- -- (funExt λ { (inl x) → sym (rUnitₖ _ (0ₖ _)) --- -- -- ; (inr (x , north)) → refl --- -- -- ; (inr (x , south)) → refl --- -- -- ; (inr (x , merid a i₁)) → refl --- -- -- ; (push a j) i → compPath-filler (cong₂ _+ₖ_ (snd (f a)) (snd (g a))) --- -- -- (rUnitₖ n (0ₖ n)) (~ i) (~ j)}) - --- -- -- funDirSect : ∀ {ℓ} {A : Pointed ℓ} → (x : (b : typ B) → Q b →∙ A) (b : typ B) (q : Q b .fst) --- -- -- → funDir1 (funDir2 x) b .fst q ≡ x b .fst q --- -- -- funDirSect f b north = refl --- -- -- funDirSect f b south = refl --- -- -- funDirSect f b (merid a i₁) = refl - --- -- -- funDirRetr : ∀ {ℓ} {A : Pointed ℓ} (f : F̃ → typ A) (p : _) --- -- -- → (x : F̃) → fst (funDir2 {A = A} (funDir1 (f , p))) x ≡ f x --- -- -- funDirRetr f p (inl x) = sym p --- -- -- funDirRetr f p (inr (x , north)) = refl --- -- -- funDirRetr f p (inr (x , south)) = refl --- -- -- funDirRetr f p (inr (x , merid a i₁)) = refl --- -- -- funDirRetr f p (push a i₁) j = compPath-filler (sym (cong f (push a))) p (~ j) (~ i₁) - --- -- -- Iso2 : ∀ {ℓ} {A : Pointed ℓ} --- -- -- → Iso ((F̃ , inl tt) →∙ A) --- -- -- ((b : typ B) → Q b →∙ A) --- -- -- Iso.fun (Iso2 {A = A , a}) = funDir1 --- -- -- Iso.inv (Iso2 {A = A , a}) = funDir2 --- -- -- Iso.rightInv (Iso2 {A = A , a}) f = --- -- -- funExt λ b → ΣPathP (funExt (funDirSect f b) --- -- -- , sym (rUnit (snd (f b)))) --- -- -- Iso.leftInv (Iso2 {A = A , a}) (f , p) = --- -- -- ΣPathP ((funExt (funDirRetr f p)) --- -- -- , λ i j → p (~ i ∨ j)) - --- -- -- ι : (k : ℕ) → Iso ((b : typ B) → Q b →∙ coHomK-ptd k) --- -- -- ((Ẽ , inl tt) →∙ coHomK-ptd k) --- -- -- ι k = compIso (invIso Iso2) h --- -- -- where --- -- -- h : Iso ((F̃ , inl tt) →∙ coHomK-ptd k) --- -- -- ((Ẽ , inl tt) →∙ coHomK-ptd k) --- -- -- Iso.fun h G = (λ x → G .fst (Iso.inv IsoFE x)) --- -- -- , (snd G) --- -- -- Iso.inv h G = (λ x → G .fst (Iso.fun IsoFE x)) --- -- -- , (snd G) --- -- -- Iso.rightInv h G = --- -- -- →∙Homogeneous≡ (isHomogeneousKn _) --- -- -- (funExt λ x → cong (G .fst) (Iso.rightInv IsoFE x)) --- -- -- Iso.leftInv h G = --- -- -- →∙Homogeneous≡ (isHomogeneousKn _) --- -- -- (funExt λ x → cong (G .fst) (Iso.leftInv IsoFE x)) - --- -- -- ι-hom : (k : ℕ) → (f g : ((b : typ B) → Q b →∙ coHomK-ptd k)) --- -- -- → Iso.fun (ι k) (λ b → f b ++ g b) --- -- -- ≡ (Iso.fun (ι k) f ++ Iso.fun (ι k) g) --- -- -- ι-hom k f g = --- -- -- →∙Homogeneous≡ (isHomogeneousKn _) --- -- -- (funExt λ x → funExt⁻ (cong fst (funDir2-hom _ f g)) (invFE x)) - --- -- -- codomainIsoDep : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : A → Type ℓ'} {C : A → Type ℓ''} --- -- -- → ((a : A) → Iso (B a) (C a)) --- -- -- → Iso ((a : A) → B a) ((a : A) → C a) --- -- -- Iso.fun (codomainIsoDep is) f a = Iso.fun (is a) (f a) --- -- -- Iso.inv (codomainIsoDep is) f a = Iso.inv (is a) (f a) --- -- -- Iso.rightInv (codomainIsoDep is) f = funExt λ a → Iso.rightInv (is a) (f a) --- -- -- Iso.leftInv (codomainIsoDep is) f = funExt λ a → Iso.leftInv (is a) (f a) - --- -- -- module Thom {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) --- -- -- (conB : (x y : typ B) → ∥ x ≡ y ∥) --- -- -- (n : ℕ) (Q-is : Iso (typ (Q B P (pt B))) (S₊ n)) --- -- -- (Q-is-ptd : Iso.fun Q-is (pt (Q B P (pt B))) ≡ snd (S₊∙ n)) --- -- -- (c : (b : typ B) → (Q B P b →∙ coHomK-ptd n)) --- -- -- (c-pt : c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) where - --- -- -- ϕ : (i : ℕ) → GroupEquiv (coHomGr i (typ B)) --- -- -- (coHomRedGrDir (i +' n) (Ẽ B P , inl tt)) --- -- -- fst (ϕ i) = --- -- -- isoToEquiv --- -- -- (setTruncIso --- -- -- (compIso --- -- -- (codomainIsoDep --- -- -- λ b → equivToIso ((g B (Q B P) conB n Q-is Q-is-ptd c c-pt b i) --- -- -- , g-equiv B (Q B P) conB n Q-is Q-is-ptd c c-pt b i)) --- -- -- (ι B P (i +' n)))) --- -- -- snd (ϕ i) = --- -- -- makeIsGroupHom --- -- -- (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) --- -- -- λ F G → cong ∣_∣₂ (cong (Iso.fun (ι B P (i +' n))) --- -- -- (funExt (λ a → g-hom B (Q B P) conB n Q-is Q-is-ptd c c-pt a i (F a) (G a))) --- -- -- ∙ ι-hom B P (i +' n) _ _) --- -- -- ∙ addAgree (i +' n) _ _) - --- -- -- module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) --- -- -- (conB : (x y : typ B) → ∥ x ≡ y ∥₂) --- -- -- (n : ℕ) (Q-is : Iso (typ (Q B P (pt B))) (S₊ n)) --- -- -- (Q-is-ptd : Iso.fun Q-is (pt (Q B P (pt B))) ≡ snd (S₊∙ n)) where - --- -- -- 0-connB : (x y : typ B) → ∥ x ≡ y ∥ --- -- -- 0-connB x y = sRec (isProp→isSet squash) (∥_∥.∣_∣) (conB x y) - --- -- -- c = (c* B (Q B P) conB n Q-is Q-is-ptd .fst) --- -- -- c-ptd = (c* B (Q B P) conB n Q-is Q-is-ptd .snd) --- -- -- cEq = cEquiv B (Q B P) conB n Q-is Q-is-ptd - --- -- -- module w = Thom B P 0-connB n Q-is Q-is-ptd c c-ptd - --- -- -- ϕ = w.ϕ - --- -- -- E' = E B P - --- -- -- E'̃ = Ẽ B P - --- -- -- p : E' → typ B --- -- -- p = fst - --- -- -- e : coHom n (typ B) --- -- -- e = ∣ (λ b → c b .fst south) ∣₂ - --- -- -- ⌣-hom : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr (i +' n) (typ B)) --- -- -- fst (⌣-hom i) x = x ⌣ e --- -- -- snd (⌣-hom i) = --- -- -- makeIsGroupHom λ f g → rightDistr-⌣ _ _ f g e - --- -- -- p-hom : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr i E') --- -- -- p-hom i = coHomMorph _ p - --- -- -- E-susp : (i : ℕ) → GroupHom (coHomGr i E') (coHomRedGrDir (suc i) (E'̃ , inl tt)) --- -- -- fst (E-susp i) = sMap λ f → (λ { (inl x) → 0ₖ _ --- -- -- ; (inr x) → 0ₖ _ --- -- -- ; (push a j) → Kn→ΩKn+1 _ (f a) j}) , refl --- -- -- snd (E-susp zero) = --- -- -- makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) --- -- -- λ f g → --- -- -- cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn 1) --- -- -- (funExt λ { (inl x) → refl --- -- -- ; (inr x) → refl --- -- -- ; (push a j) k → (Kn→ΩKn+1-hom zero (f a) (g a) --- -- -- ∙ ∙≡+₁ (Kn→ΩKn+1 _ (f a)) (Kn→ΩKn+1 _ (g a))) k j}))) --- -- -- snd (E-susp (suc i)) = --- -- -- makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) --- -- -- λ f g → --- -- -- cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) --- -- -- (funExt λ { (inl x) → refl --- -- -- ; (inr x) → refl --- -- -- ; (push a j) k → (Kn→ΩKn+1-hom _ (f a) (g a) --- -- -- ∙ ∙≡+₂ _ (Kn→ΩKn+1 _ (f a)) (Kn→ΩKn+1 _ (g a))) k j}))) - --- -- -- module cofibSeq where --- -- -- j* : (i : ℕ) → GroupHom (coHomRedGrDir i (E'̃ , (inl tt))) (coHomGr i (typ B)) --- -- -- fst (j* i) = sMap λ f → λ x → fst f (inr x) --- -- -- snd (j* zero) = --- -- -- makeIsGroupHom --- -- -- (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) --- -- -- snd (j* (suc zero)) = --- -- -- makeIsGroupHom --- -- -- (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) --- -- -- snd (j* (suc (suc i₁))) = --- -- -- makeIsGroupHom --- -- -- (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl) - --- -- -- Im-j⊂Ker-p : (i : ℕ) (x : _) → isInIm (j* i) x → isInKer (p-hom i) x --- -- -- Im-j⊂Ker-p i = --- -- -- sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) --- -- -- λ f → pRec (squash₂ _ _) --- -- -- (uncurry (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) --- -- -- λ g P → subst (isInKer (p-hom i)) P --- -- -- (cong ∣_∣₂ (funExt λ x → cong (g .fst) (sym (push x)) ∙ g .snd)))) - --- -- -- Ker-p⊂Im-j : (i : ℕ) (x : _) → isInKer (p-hom i) x → isInIm (j* i) x --- -- -- Ker-p⊂Im-j i = --- -- -- sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) --- -- -- λ f ker → pRec squash --- -- -- (λ id → ∣ ∣ (λ { (inl x) → 0ₖ _ --- -- -- ; (inr x) → f x --- -- -- ; (push a i₁) → funExt⁻ id a (~ i₁)}) , refl ∣₂ , refl ∣) --- -- -- (Iso.fun PathIdTrunc₀Iso ker) - --- -- -- Im-p⊂Ker-Susp : (i : ℕ) (x : _) → isInIm (p-hom i) x → isInKer (E-susp i) x --- -- -- Im-p⊂Ker-Susp i = --- -- -- sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) --- -- -- λ f → pRec (squash₂ _ _) --- -- -- (uncurry (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) --- -- -- λ g y → subst (isInKer (E-susp i)) y --- -- -- (cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) --- -- -- (funExt λ { (inl x) → refl --- -- -- ; (inr x) → sym (Kn→ΩKn+1 _ (g x)) --- -- -- ; (push a j) k → Kn→ΩKn+1 i (g (fst a)) (~ k ∧ j)}))))) --- -- -- open import Cubical.Foundations.Path --- -- -- Ker-Susp⊂Im-p : (i : ℕ) (x : _) → isInKer (E-susp i) x → isInIm (p-hom i) x --- -- -- Ker-Susp⊂Im-p i = --- -- -- sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) --- -- -- λ f ker → pRec squash --- -- -- (λ id → ∣ ∣ (λ x → ΩKn+1→Kn i (sym (funExt⁻ (cong fst id) (inr x)))) ∣₂ --- -- -- , cong ∣_∣₂ (funExt (λ { (a , b) → --- -- -- cong (ΩKn+1→Kn i) (lUnit _ ∙ cong (_∙ sym (funExt⁻ (λ i₁ → fst (id i₁)) (inr a))) (sym (flipSquare (cong snd id)) --- -- -- ∙∙ (PathP→compPathR (cong (funExt⁻ (cong fst id)) (push (a , b)))) --- -- -- ∙∙ assoc _ _ _ --- -- -- ∙ sym (rUnit _)) --- -- -- ∙ (sym (assoc _ _ _) --- -- -- ∙∙ cong (Kn→ΩKn+1 i (f (a , b)) ∙_) (rCancel _) --- -- -- ∙∙ sym (rUnit _))) --- -- -- ∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f (a , b))})) ∣) --- -- -- (Iso.fun PathIdTrunc₀Iso ker) - --- -- -- Im-Susp⊂Ker-j : (i : ℕ) (x : _) → isInIm (E-susp i) x → isInKer (cofibSeq.j* (suc i)) x --- -- -- Im-Susp⊂Ker-j i = --- -- -- sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) --- -- -- λ g → pRec (squash₂ _ _) --- -- -- (uncurry (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) --- -- -- λ f id → pRec (squash₂ _ _) --- -- -- (λ P → subst (isInKer (cofibSeq.j* (suc i))) (cong ∣_∣₂ P) --- -- -- (cong ∣_∣₂ refl)) --- -- -- (Iso.fun PathIdTrunc₀Iso id))) - --- -- -- Ker-j⊂Im-Susp : (i : ℕ) (x : _) → isInKer (cofibSeq.j* (suc i)) x → isInIm (E-susp i) x --- -- -- Ker-j⊂Im-Susp i = --- -- -- sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) --- -- -- λ f ker --- -- -- → pRec squash --- -- -- (λ p → ∣ ∣ (λ x → ΩKn+1→Kn _ (sym (snd f) ∙∙ cong (fst f) (push x) ∙∙ funExt⁻ p (fst x))) ∣₂ --- -- -- , cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) --- -- -- (funExt (λ { (inl x) → sym (snd f) --- -- -- ; (inr x) → sym (funExt⁻ p x) --- -- -- ; (push a j) k → h f p a k j}))) ∣) --- -- -- (Iso.fun PathIdTrunc₀Iso ker) --- -- -- where --- -- -- h : (f : (E'̃ , inl tt) →∙ coHomK-ptd (suc i)) --- -- -- → (p : (fst f ∘ inr) ≡ (λ _ → 0ₖ (suc i))) --- -- -- → (a : E B P) --- -- -- → PathP (λ i → snd f (~ i) ≡ p (~ i) (fst a)) --- -- -- (Kn→ΩKn+1 i (ΩKn+1→Kn i (sym (snd f) ∙∙ cong (fst f) (push a) ∙∙ funExt⁻ p (fst a)))) --- -- -- (cong (fst f) (push a)) --- -- -- h f p a = Iso.rightInv (Iso-Kn-ΩKn+1 i) _ --- -- -- ◁ λ i j → doubleCompPath-filler (sym (snd f)) (cong (fst f) (push a)) (funExt⁻ p (fst a)) (~ i) j - --- -- -- ϕ∘j : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr (i +' n) (typ B)) --- -- -- ϕ∘j i = compGroupHom (fst (fst (ϕ i)) , snd (ϕ i)) (cofibSeq.j* (i +' n)) - --- -- -- +'-suc : (i n : ℕ) → (suc i +' n) ≡ suc (i +' n) --- -- -- +'-suc zero zero = refl --- -- -- +'-suc (suc i₁) zero = refl --- -- -- +'-suc zero (suc n) = refl --- -- -- +'-suc (suc i₁) (suc n) = refl - --- -- -- private --- -- -- h : ∀ {ℓ ℓ'} {n m : ℕ} (G : ℕ → Group ℓ) (H : Group ℓ') (p : n ≡ m) --- -- -- → GroupEquiv (G n) H --- -- -- → GroupEquiv (G m) H --- -- -- h {n = n} G H = --- -- -- J (λ m _ → GroupEquiv (G n) H → GroupEquiv (G m) H) --- -- -- λ p → p - --- -- -- h-ret : ∀ {ℓ ℓ'} {n m : ℕ} (G : ℕ → Group ℓ) (H : Group ℓ') --- -- -- → (e : GroupEquiv (G n) H) --- -- -- → (p : n ≡ m) --- -- -- → (x : G m .fst) → invEq (fst e) (fst (fst (h G H p e)) x) ≡ subst (λ x → G x .fst) (sym p) x --- -- -- h-ret G H e = --- -- -- J (λ m p → ((x : G m .fst) → invEq (fst e) (fst (fst (h G H p e)) x) ≡ subst (λ x → G x .fst) (sym p) x)) --- -- -- λ x → cong (invEq (fst e) ) --- -- -- (λ i → transportRefl (transportRefl (fst (fst e) (transportRefl (transportRefl x i) i)) i) i) --- -- -- ∙∙ retEq (fst e) x --- -- -- ∙∙ sym (transportRefl _) - --- -- -- isEquivϕ' : (i : ℕ) → GroupEquiv (coHomRedGrDir (suc (i +' n)) (E'̃ , inl tt)) --- -- -- (coHomGr (suc i) (typ B)) --- -- -- isEquivϕ' i = (h (λ n → coHomRedGrDir n (E'̃ , inl tt)) _ (+'-suc i n) --- -- -- (invGroupEquiv (ϕ (suc i)))) - --- -- -- ϕ' : (i : ℕ) → GroupHom (coHomRedGrDir (suc (i +' n)) (E'̃ , inl tt)) --- -- -- (coHomGr (suc i) (typ B)) --- -- -- ϕ' i = fst (fst (isEquivϕ' i)) , snd (isEquivϕ' i) - --- -- -- susp∘ϕ : (i : ℕ) → GroupHom (coHomGr (i +' n) E') (coHomGr (suc i) (typ B)) --- -- -- susp∘ϕ i = compGroupHom (E-susp (i +' n)) (ϕ' i) - --- -- -- Im-ϕ∘j⊂Ker-p : (i : ℕ) (x : _) → isInIm (ϕ∘j i) x → isInKer (p-hom _) x --- -- -- Im-ϕ∘j⊂Ker-p i x p = --- -- -- cofibSeq.Im-j⊂Ker-p _ x --- -- -- (pRec squash (uncurry (λ f p → ∣ fst (fst (ϕ _)) f , p ∣)) p) - --- -- -- Ker-p⊂Im-ϕ∘j : (i : ℕ) (x : _) → isInKer (p-hom _) x → isInIm (ϕ∘j i) x --- -- -- Ker-p⊂Im-ϕ∘j i x p = --- -- -- pRec squash (uncurry (λ f p → ∣ (invEq (fst (ϕ _)) f) --- -- -- , (cong (fst (cofibSeq.j* _)) (secEq (fst (ϕ _)) f) ∙ p) ∣)) --- -- -- (cofibSeq.Ker-p⊂Im-j _ x p) - - --- -- -- Im-p⊂KerSusp∘ϕ : (i : ℕ) (x : _) → isInIm (p-hom _) x → isInKer (susp∘ϕ i) x --- -- -- Im-p⊂KerSusp∘ϕ i x p = cong (fst (ϕ' _)) (Im-p⊂Ker-Susp _ x p) ∙ IsGroupHom.pres1 (snd (ϕ' _)) - --- -- -- KerSusp∘ϕ⊂Im-p : (i : ℕ) (x : _) → isInKer (susp∘ϕ i) x → isInIm (p-hom _) x --- -- -- KerSusp∘ϕ⊂Im-p i x p = --- -- -- Ker-Susp⊂Im-p _ x (sym (retEq (fst (isEquivϕ' _)) _) --- -- -- ∙ (cong (invEq (fst (isEquivϕ' _))) p --- -- -- ∙ IsGroupHom.pres1 (snd (invGroupEquiv (isEquivϕ' _))))) - --- -- -- Im-Susp∘ϕ⊂Ker-ϕ∘j : (i : ℕ) → (x : _) → isInIm (susp∘ϕ i) x → isInKer (ϕ∘j (suc i)) x --- -- -- Im-Susp∘ϕ⊂Ker-ϕ∘j i x = --- -- -- pRec (squash₂ _ _) --- -- -- (uncurry λ f → J (λ x p → isInKer (ϕ∘j (suc i)) x) --- -- -- ((λ i → fst (cofibSeq.j* _) (fst (fst (ϕ _)) (fst (ϕ' _) (fst (E-susp _) f)))) --- -- -- ∙∙ cong (fst (cofibSeq.j* _)) --- -- -- ((h-ret (λ n → coHomRedGrDir n (E'̃ , inl tt)) _ --- -- -- (invGroupEquiv (ϕ (suc i))) (+'-suc i n)) (fst (E-susp _) f)) --- -- -- ∙∙ (natTranspLem _ (λ n → fst (cofibSeq.j* n)) (sym (+'-suc i n)) --- -- -- ∙ cong (subst (λ z → coHomGr z (typ B) .fst) (sym (+'-suc i n))) --- -- -- (Im-Susp⊂Ker-j _ (fst (E-susp (i +' n)) f) ∣ f , refl ∣) --- -- -- ∙ tLem i n))) --- -- -- where --- -- -- tLem : (i n : ℕ) → subst (λ z → coHomGr z (typ B) .fst) (sym (+'-suc i n)) --- -- -- (0ₕ _) ≡ 0ₕ _ --- -- -- tLem zero zero = refl --- -- -- tLem zero (suc n) = refl --- -- -- tLem (suc i₁) zero = refl --- -- -- tLem (suc i₁) (suc n) = refl - - --- -- -- Ker-ϕ∘j⊂Im-Susp∘ϕ : (i : ℕ) (x : _) --- -- -- → isInKer (ϕ∘j (suc i)) x → isInIm (susp∘ϕ i) x --- -- -- Ker-ϕ∘j⊂Im-Susp∘ϕ i x p = --- -- -- pRec squash --- -- -- (uncurry (λ f p → ∣ f , cong (fst (fst (isEquivϕ' i))) p ∙ secEq (fst (isEquivϕ' _)) x ∣)) --- -- -- (Ker-j⊂Im-Susp _ (invEq (fst (isEquivϕ' _)) x) --- -- -- ((cong (cofibSeq.j* (suc (i +' n)) .fst ) lem2 --- -- -- ∙ natTranspLem _ (λ n → cofibSeq.j* n .fst) (+'-suc i n)) --- -- -- ∙∙ cong (transport (λ j → (coHomGr (+'-suc i n j) (typ B) .fst))) p --- -- -- ∙∙ h2 i n)) --- -- -- where --- -- -- lem2 : invEq (fst (isEquivϕ' i)) x ≡ transport (λ j → coHomRed (+'-suc i n j) (E'̃ , inl tt)) (fst (fst (ϕ _)) x) --- -- -- lem2 = cong (transport (λ j → coHomRed (+'-suc i n j) (E'̃ , inl tt))) --- -- -- (transportRefl _ ∙ cong (fst (fst (ϕ _))) --- -- -- λ i → transportRefl (transportRefl x i) i) - --- -- -- h2 : (i n : ℕ) → transport (λ j → coHomGr (+'-suc i n j) (typ B) .fst) --- -- -- (GroupStr.1g (coHomGr (suc i +' n) (typ B) .snd)) ≡ 0ₕ _ --- -- -- h2 zero zero = refl --- -- -- h2 zero (suc n) = refl --- -- -- h2 (suc i₁) zero = refl --- -- -- h2 (suc i₁) (suc n) = refl - - --- -- -- ϕ∘j≡ : (i : ℕ) → ϕ∘j i ≡ ⌣-hom i --- -- -- ϕ∘j≡ i = --- -- -- Σ≡Prop (λ _ → isPropIsGroupHom _ _) --- -- -- (funExt (sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) --- -- -- λ _ → refl)) - --- -- -- open import Cubical.Experiments.Brunerie --- -- -- open import Cubical.HITs.Hopf - --- -- -- open import Cubical.HITs.Join - --- -- -- module fibS1 = hopfBase S1-AssocHSpace (sphereElim2 _ (λ _ _ → squash) ∣ refl ∣) - --- -- -- S¹Hopf : S₊ 2 → Type --- -- -- S¹Hopf = fibS1.Hopf - --- -- -- TotalHopf' : Type _ --- -- -- TotalHopf' = Σ (S₊ 2) S¹Hopf - --- -- -- CP2 : Type _ --- -- -- CP2 = fibS1.megaPush - --- -- -- fibr : CP2 → Type _ --- -- -- fibr = fibS1.P - --- -- -- hopf : join S¹ S¹ → S₊ 2 --- -- -- hopf x = fst (JoinS¹S¹→TotalHopf x) - --- -- -- E* : Type _ --- -- -- E* = fibS1.totalSpaceMegaPush - --- -- -- IsoE' : Iso E* (join S¹ (join S¹ S¹)) --- -- -- IsoE' = fibS1.IsoJoin₁ - --- -- -- IsoE2 : (join S¹ (join S¹ S¹)) ≡ join S¹ (S₊ 3) --- -- -- IsoE2 = cong (join S¹) (sym S³≡joinS¹S¹ ∙ isoToPath IsoS³S3) - --- -- -- IsoTotalHopf' : Iso fibS1.TotalSpaceHopf' (join S¹ S¹) --- -- -- IsoTotalHopf' = fibS1.joinIso₁ - --- -- -- CP' : Type _ --- -- -- CP' = Pushout (λ _ → tt) hopf - --- -- -- conCP2 : (x y : CP2) → ∥ x ≡ y ∥₂ --- -- -- conCP2 x y = sRec2 squash₂ (λ p q → ∣ p ∙ sym q ∣₂) (conCP2' x) (conCP2' y) --- -- -- where --- -- -- conCP2' : (x : CP2) → ∥ x ≡ inl tt ∥₂ --- -- -- conCP2' (inl x) = ∣ refl ∣₂ --- -- -- conCP2' (inr x) = sphereElim 1 {A = λ x → ∥ inr x ≡ inl tt ∥₂} (λ _ → squash₂) ∣ sym (push (inl base)) ∣₂ x --- -- -- conCP2' (push a i₁) = ll a i₁ --- -- -- where --- -- -- h2 : ∀ {ℓ} {A : fibS1.TotalSpaceHopf' → Type ℓ} → ((a : _) → isProp (A a)) --- -- -- → A (inl base) --- -- -- → ((a : fibS1.TotalSpaceHopf') → A a) --- -- -- h2 {A = A} p b = --- -- -- PushoutToProp p (sphereElim 0 (λ _ → p _) b) --- -- -- (sphereElim 0 (λ _ → p _) (subst A (push (base , base)) b)) - --- -- -- ll : (a : fibS1.TotalSpaceHopf') → PathP (λ i → ∥ Path CP2 (push a i) (inl tt) ∥₂) (conCP2' (inl tt)) (conCP2' (inr (fibS1.induced a))) --- -- -- ll = h2 (λ _ → isOfHLevelPathP' 1 squash₂ _ _) λ j → ∣ (λ i → push (inl base) (~ i ∧ j)) ∣₂ - --- -- -- module GysinS1 = Gysin (CP2 , inl tt) fibr conCP2 2 idIso refl - --- -- -- PushoutReplaceBase : --- -- -- ∀ {ℓ ℓ' ℓ''} {A B : Type ℓ} {C : Type ℓ'} {D : Type ℓ''} {f : A → C} {g : A → D} --- -- -- (e : B ≃ A) → Pushout (f ∘ fst e) (g ∘ fst e) ≡ Pushout f g --- -- -- PushoutReplaceBase {f = f} {g = g} = --- -- -- EquivJ (λ _ e → Pushout (f ∘ fst e) (g ∘ fst e) ≡ Pushout f g) --- -- -- refl - --- -- -- isContrH³E : isContr (coHom 3 (GysinS1.E')) --- -- -- isContrH³E = --- -- -- subst isContr --- -- -- (sym (isoToPath (compIso (Iso1 0) (Iso2' 0))) --- -- -- ∙ cong (coHom 3) (sym (isoToPath IsoE' ∙ IsoE2))) --- -- -- (ll5 0 (snotz ∘ sym)) - --- -- -- isContrH⁴E : isContr (coHom 4 (GysinS1.E')) --- -- -- isContrH⁴E = --- -- -- subst isContr --- -- -- (sym (isoToPath (compIso (Iso1 1) (Iso2' 1))) --- -- -- ∙ cong (coHom 4) (sym (isoToPath IsoE' ∙ IsoE2))) --- -- -- (ll5 1 λ p → snotz (sym (cong predℕ p))) - --- -- -- genH2 = GysinS1.e - --- -- -- S³→Groupoid : ∀ {ℓ} (P : join S¹ S¹ → Type ℓ) --- -- -- → ((x : _) → isGroupoid (P x)) --- -- -- → P (inl base) --- -- -- → (x : _) → P x --- -- -- S³→Groupoid P grp b = --- -- -- transport (λ i → (x : (sym (isoToPath S³IsojoinS¹S¹) ∙ isoToPath IsoS³S3) (~ i)) --- -- -- → P (transp (λ j → (sym (isoToPath S³IsojoinS¹S¹) ∙ isoToPath IsoS³S3) (~ i ∧ ~ j)) i x)) --- -- -- (sphereElim _ (λ _ → grp _) b) - --- -- -- TotalHopf→Gpd : ∀ {ℓ} (P : fibS1.TotalSpaceHopf' → Type ℓ) --- -- -- → ((x : _) → isGroupoid (P x)) --- -- -- → P (inl base) --- -- -- → (x : _) → P x --- -- -- TotalHopf→Gpd P grp b = --- -- -- transport (λ i → (x : sym (isoToPath IsoTotalHopf') i) --- -- -- → P (transp (λ j → isoToPath IsoTotalHopf' (~ i ∧ ~ j)) i x)) --- -- -- (S³→Groupoid _ (λ _ → grp _) b) - --- -- -- TotalHopf→Gpd' : ∀ {ℓ} (P : Σ (S₊ 2) fibS1.Hopf → Type ℓ) --- -- -- → ((x : _) → isGroupoid (P x)) --- -- -- → P (north , base) --- -- -- → (x : _) → P x --- -- -- TotalHopf→Gpd' P grp b = --- -- -- transport (λ i → (x : ua (_ , fibS1.isEquivTotalSpaceHopf'→TotalSpace) i) --- -- -- → P (transp (λ j → ua (_ , fibS1.isEquivTotalSpaceHopf'→TotalSpace) (i ∨ j)) i x)) --- -- -- (TotalHopf→Gpd _ (λ _ → grp _) b) - --- -- -- CP2→Groupoid : ∀ {ℓ} {P : CP2 → Type ℓ} --- -- -- → ((x : _) → is2Groupoid (P x)) --- -- -- → (b : P (inl tt)) --- -- -- → (s2 : ((x : _) → P (inr x))) --- -- -- → PathP (λ i → P (push (inl base) i)) b (s2 north) --- -- -- → (x : _) → P x --- -- -- CP2→Groupoid {P = P} grp b s2 pp (inl x) = b --- -- -- CP2→Groupoid {P = P} grp b s2 pp (inr x) = s2 x --- -- -- CP2→Groupoid {P = P} grp b s2 pp (push a i₁) = h23 a i₁ --- -- -- where --- -- -- h23 : (a : fibS1.TotalSpaceHopf') → PathP (λ i → P (push a i)) b (s2 _) --- -- -- h23 = TotalHopf→Gpd _ (λ _ → isOfHLevelPathP' 3 (grp _) _ _) pp - --- -- -- mm : (S₊ 2 → coHomK 2) → (CP2 → coHomK 2) --- -- -- mm f = λ { (inl x) → 0ₖ _ --- -- -- ; (inr x) → f x -ₖ f north --- -- -- ; (push a i) → TotalHopf→Gpd (λ x → 0ₖ 2 ≡ f (fibS1.TotalSpaceHopf'→TotalSpace x .fst) -ₖ f north) --- -- -- (λ _ → isOfHLevelTrunc 4 _ _) --- -- -- (sym (rCancelₖ 2 (f north))) --- -- -- a i} - --- -- -- e-inv : coHomGr 2 (S₊ 2) .fst → coHomGr 2 CP2 .fst --- -- -- e-inv = sMap mm - --- -- -- cancel : (f : CP2 → coHomK 2) → ∥ mm (f ∘ inr) ≡ f ∥ --- -- -- cancel f = --- -- -- pRec squash --- -- -- (λ p → pRec squash --- -- -- (λ f → ∣ funExt f ∣) --- -- -- (zz2 p)) --- -- -- (h1 (f (inl tt))) --- -- -- where --- -- -- h1 : (x : coHomK 2) → ∥ x ≡ 0ₖ 2 ∥ --- -- -- h1 = coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash) ∣ refl ∣ - --- -- -- zz2 : (p : f (inl tt) ≡ 0ₖ 2) → ∥ ((x : CP2) → mm (f ∘ inr) x ≡ f x) ∥ --- -- -- zz2 p = trRec squash (λ pp → --- -- -- ∣ CP2→Groupoid (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) --- -- -- (sym p) --- -- -- (λ x → (λ i → f (inr x) -ₖ f (push (inl base) (~ i))) --- -- -- ∙∙ (λ i → f (inr x) -ₖ p i) --- -- -- ∙∙ rUnitₖ 2 (f (inr x))) --- -- -- pp ∣) --- -- -- l --- -- -- where --- -- -- l : hLevelTrunc 1 (PathP ((λ i₁ → mm (f ∘ inr) (push (inl base) i₁) ≡ f (push (inl base) i₁))) --- -- -- (sym p) --- -- -- (((λ i₁ → f (inr north) -ₖ f (push (inl base) (~ i₁))) ∙∙ --- -- -- (λ i₁ → f (inr north) -ₖ p i₁) ∙∙ rUnitₖ 2 (f (inr north))))) --- -- -- l = isConnectedPathP 1 (isConnectedPath 2 (isConnectedKn 1) _ _) _ _ .fst - --- -- -- e' : GroupIso (coHomGr 2 CP2) (coHomGr 2 (S₊ 2)) --- -- -- Iso.fun (fst e') = sMap (_∘ inr) --- -- -- Iso.inv (fst e') = e-inv --- -- -- Iso.rightInv (fst e') = --- -- -- sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) --- -- -- λ f → trRec {B = Iso.fun (fst e') (Iso.inv (fst e') ∣ f ∣₂) ≡ ∣ f ∣₂} --- -- -- (isOfHLevelPath 2 squash₂ _ _) --- -- -- (λ p → cong ∣_∣₂ (funExt λ x → cong (λ y → (f x) -ₖ y) p ∙ rUnitₖ 2 (f x))) --- -- -- (Iso.fun (PathIdTruncIso _) (isContr→isProp (isConnectedKn 1) ∣ f north ∣ ∣ 0ₖ 2 ∣)) --- -- -- Iso.leftInv (fst e') = --- -- -- sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) --- -- -- λ f → pRec (squash₂ _ _) (cong ∣_∣₂) (cancel f) --- -- -- snd e' = --- -- -- makeIsGroupHom --- -- -- (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ f g → refl) - --- -- -- grrg : GroupIso (coHomGr 2 CP2) ℤGroup --- -- -- grrg = compGroupIso e' (Hⁿ-Sⁿ≅ℤ 1) - --- -- -- ⌣hom : GroupEquiv (coHomGr 2 CP2) (coHomGr 4 CP2) --- -- -- fst (fst ⌣hom) = GysinS1.⌣-hom 2 .fst --- -- -- snd (fst ⌣hom) = subst isEquiv (cong fst (GysinS1.ϕ∘j≡ 2)) h23 --- -- -- where --- -- -- h23 : isEquiv (GysinS1.ϕ∘j 2 .fst) --- -- -- h23 = --- -- -- SES→Iso --- -- -- (GroupPath _ _ .fst (invGroupEquiv --- -- -- (isContr→≃Unit isContrH³E --- -- -- , makeIsGroupHom λ _ _ → refl))) --- -- -- (GroupPath _ _ .fst (invGroupEquiv --- -- -- (isContr→≃Unit isContrH⁴E --- -- -- , makeIsGroupHom λ _ _ → refl))) --- -- -- (GysinS1.susp∘ϕ 1) --- -- -- (GysinS1.ϕ∘j 2) --- -- -- (GysinS1.p-hom 4) --- -- -- (GysinS1.Ker-ϕ∘j⊂Im-Susp∘ϕ _) --- -- -- (GysinS1.Ker-p⊂Im-ϕ∘j _) --- -- -- snd ⌣hom = GysinS1.⌣-hom 2 .snd - --- -- -- isGenerator≃ℤ : ∀ {ℓ} (G : Group ℓ) (g : fst G) --- -- -- → Type _ --- -- -- isGenerator≃ℤ G g = --- -- -- ∃[ e ∈ GroupIso G ℤGroup ] abs (Iso.fun (fst e) g) ≡ 1 - - - --- -- -- -- ⌣-pres1 : ∀ {ℓ} (G : Type ℓ) (n : ℕ) → Type _ --- -- -- -- ⌣-pres1 G n = --- -- -- -- Σ[ x ∈ coHomGr n G .fst ] --- -- -- -- isGenerator≃ℤ (coHomGr n G) x × isGenerator≃ℤ (coHomGr (n +' n) G) (x ⌣ x) - --- -- -- -- genr : coHom 2 CP2 --- -- -- -- genr = ∣ CP2→Groupoid (λ _ → isOfHLevelTrunc 4) --- -- -- -- (0ₖ _) --- -- -- -- ∣_∣ --- -- -- -- refl ∣₂ - --- -- -- -- ≡CP2 : (f g : CP2 → coHomK 2) → ∥ f ∘ inr ≡ g ∘ inr ∥ → Path (coHom 2 CP2) ∣ f ∣₂ ∣ g ∣₂ --- -- -- -- ≡CP2 f g = pRec (squash₂ _ _) --- -- -- -- (λ p → pRec (squash₂ _ _) (λ id → trRec (squash₂ _ _) --- -- -- -- (λ pp → cong ∣_∣₂ --- -- -- -- (funExt (CP2→Groupoid (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) --- -- -- -- id --- -- -- -- (funExt⁻ p) --- -- -- -- (compPathR→PathP pp)))) --- -- -- -- (pp2 (f (inl tt)) (g (inl tt)) id --- -- -- -- (cong f (push (inl base)) ∙ (funExt⁻ p north) ∙ cong g (sym (push (inl base)))))) --- -- -- -- (conn (f (inl tt)) (g (inl tt)))) - --- -- -- -- where --- -- -- -- conn : (x y : coHomK 2) → ∥ x ≡ y ∥ --- -- -- -- conn = coHomK-elim _ (λ _ → isSetΠ λ _ → isOfHLevelSuc 1 squash) --- -- -- -- (coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash) ∣ refl ∣) - --- -- -- -- pp2 : (x y : coHomK 2) (p q : x ≡ y) → hLevelTrunc 1 (p ≡ q) --- -- -- -- pp2 x y = λ p q → Iso.fun (PathIdTruncIso _) (isContr→isProp (isConnectedPath _ (isConnectedKn 1) x y) ∣ p ∣ ∣ q ∣) - --- -- -- -- rUnit* : (x : S¹) → x * base ≡ x --- -- -- -- rUnit* base = refl --- -- -- -- rUnit* (loop i₁) = refl - --- -- -- -- meridP : (a x : S¹) → Path (Path (coHomK 2) _ _) (cong ∣_∣ₕ (merid (a * x) ∙ sym (merid base))) --- -- -- -- ((cong ∣_∣ₕ (merid a ∙ sym (merid base))) ∙ (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) --- -- -- -- meridP = wedgeconFun _ _ (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) --- -- -- -- (λ x → lUnit _ ∙ cong (_∙ cong ∣_∣ₕ (merid x ∙ sym (merid base))) (cong (cong ∣_∣ₕ) (sym (rCancel (merid base))))) --- -- -- -- (λ x → (λ i → cong ∣_∣ₕ (merid (rUnit* x i) ∙ sym (merid base))) --- -- -- -- ∙∙ rUnit _ --- -- -- -- ∙∙ cong (cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙_) --- -- -- -- (cong (cong ∣_∣ₕ) (sym (rCancel (merid base))))) --- -- -- -- (sym (l (cong ∣_∣ₕ (merid base ∙ sym (merid base))) --- -- -- -- (cong (cong ∣_∣ₕ) (sym (rCancel (merid base)))))) --- -- -- -- where --- -- -- -- l : ∀ {ℓ} {A : Type ℓ} {x : A} (p : x ≡ x) (P : refl ≡ p) --- -- -- -- → lUnit p ∙ cong (_∙ p) P ≡ rUnit p ∙ cong (p ∙_) P --- -- -- -- l p = J (λ p P → lUnit p ∙ cong (_∙ p) P ≡ rUnit p ∙ cong (p ∙_) P) refl - --- -- -- -- lemmie : (x : S¹) → ptSn 1 ≡ x * (invLooper x) --- -- -- -- lemmie base = refl --- -- -- -- lemmie (loop i) j = --- -- -- -- hcomp (λ r → λ {(i = i0) → base ; (i = i1) → base ; (j = i0) → base}) --- -- -- -- base - --- -- -- -- mmInv : (x : S¹) → Path (Path (coHomK 2) _ _) --- -- -- -- (cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base))) --- -- -- -- (cong ∣_∣ₕ (sym (merid x ∙ sym (merid base)))) --- -- -- -- mmInv x = (lUnit _ --- -- -- -- ∙∙ cong (_∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base))) (sym (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base))))) ∙∙ sym (assoc _ _ _)) --- -- -- -- ∙∙ cong (sym (cong ∣_∣ₕ (merid x ∙ sym (merid base))) ∙_) hh --- -- -- -- ∙∙ (assoc _ _ _ --- -- -- -- ∙∙ cong (_∙ (cong ∣_∣ₕ (sym (merid x ∙ sym (merid base))))) (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) --- -- -- -- ∙∙ sym (lUnit _)) --- -- -- -- where --- -- -- -- hh : cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base)) --- -- -- -- ≡ cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙ cong ∣_∣ₕ (sym (merid x ∙ sym (merid base))) --- -- -- -- hh = sym (meridP x (invLooper x)) ∙ ((λ i → cong ∣_∣ₕ (merid (lemmie x (~ i)) ∙ sym (merid base))) ∙ cong (cong ∣_∣ₕ) (rCancel (merid base))) ∙ sym (rCancel _) - --- -- -- -- commS1 : (a x : S¹) → a * x ≡ x * a --- -- -- -- commS1 = wedgeconFun _ _ (λ _ _ → isGroupoidS¹ _ _) --- -- -- -- (sym ∘ rUnit*) --- -- -- -- rUnit* --- -- -- -- refl - --- -- -- -- s233 : (a x : S¹) → (invEq (fibS1.μ-eq a) x) * a ≡ (invLooper a * x) * a --- -- -- -- s233 a x = secEq (fibS1.μ-eq a) x ∙∙ cong (_* x) (lemmie a) ∙∙ assocHSpace.μ-assoc S1-AssocHSpace a (invLooper a) x ∙ commS1 _ _ - --- -- -- -- ss23 : (a x : S¹) → invEq (fibS1.μ-eq a) x ≡ invLooper a * x --- -- -- -- ss23 a x = sym (retEq (fibS1.μ-eq a) (invEq (fibS1.μ-eq a) x)) --- -- -- -- ∙∙ cong (invEq (fibS1.μ-eq a)) (s233 a x) --- -- -- -- ∙∙ retEq (fibS1.μ-eq a) (invLooper a * x) - --- -- -- -- ll : GysinS1.e ≡ genr --- -- -- -- ll = ≡CP2 _ _ ∣ funExt (λ x → funExt⁻ (cong fst (ll32 x)) south) ∣ --- -- -- -- where --- -- -- -- kr : (x : Σ (S₊ 2) fibS1.Hopf) → Path (hLevelTrunc 4 _) ∣ fst x ∣ ∣ north ∣ --- -- -- -- kr = uncurry λ { north → λ y → cong ∣_∣ₕ (merid base ∙ sym (merid y)) --- -- -- -- ; south → λ y → cong ∣_∣ₕ (sym (merid y)) --- -- -- -- ; (merid a i) → lem a i} --- -- -- -- where --- -- -- -- lem : (a : S¹) → PathP (λ i → (y : fibS1.Hopf (merid a i)) → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a i ∣ ∣ north ∣) --- -- -- -- (λ y → cong ∣_∣ₕ (merid base ∙ sym (merid y))) --- -- -- -- λ y → cong ∣_∣ₕ (sym (merid y)) --- -- -- -- lem a = toPathP (funExt λ x → cong (transport ((λ i₁ → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a i₁ ∣ ∣ north ∣))) --- -- -- -- ((λ i → (λ z → cong ∣_∣ₕ (merid base --- -- -- -- ∙ sym (merid (transport (λ j → uaInvEquiv (fibS1.μ-eq a) (~ i) j) x))) z)) --- -- -- -- ∙ λ i → cong ∣_∣ₕ (merid base --- -- -- -- ∙ sym (merid (transportRefl (invEq (fibS1.μ-eq a) x) i)))) --- -- -- -- ∙∙ (λ i → transp (λ i₁ → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a (i₁ ∨ i) ∣ ∣ north ∣) i --- -- -- -- (compPath-filler' (cong ∣_∣ₕ (sym (merid a))) --- -- -- -- (cong ∣_∣ₕ (merid base ∙ sym (merid (ss23 a x i)))) i)) --- -- -- -- ∙∙ cong ((cong ∣_∣ₕ) (sym (merid a)) ∙_) (cong (cong ∣_∣ₕ) (cong sym (symDistr (merid base) (sym (merid (invLooper a * x))))) --- -- -- -- ∙ cong sym (meridP (invLooper a) x) --- -- -- -- ∙ symDistr ((cong ∣_∣ₕ) (merid (invLooper a) ∙ sym (merid base))) --- -- -- -- ((cong ∣_∣ₕ) (merid x ∙ sym (merid base))) --- -- -- -- ∙ isCommΩK 2 (sym (λ i₁ → ∣ (merid x ∙ (λ i₂ → merid base (~ i₂))) i₁ ∣)) --- -- -- -- (sym (λ i₁ → ∣ (merid (invLooper a) ∙ (λ i₂ → merid base (~ i₂))) i₁ ∣)) --- -- -- -- ∙ cong₂ _∙_ (cong sym (mmInv a) ∙ cong-∙ ∣_∣ₕ (merid a) (sym (merid base))) --- -- -- -- (cong (cong ∣_∣ₕ) (symDistr (merid x) (sym (merid base))) ∙ cong-∙ ∣_∣ₕ (merid base) (sym (merid x)))) --- -- -- -- ∙∙ (λ j → (λ i₁ → ∣ merid a (~ i₁ ∨ j) ∣) ∙ --- -- -- -- ((λ i₁ → ∣ merid a (i₁ ∨ j) ∣) ∙ --- -- -- -- (λ i₁ → ∣ merid base (~ i₁ ∨ j) ∣)) --- -- -- -- ∙ --- -- -- -- (λ i₁ → ∣ merid base (i₁ ∨ j) ∣) ∙ --- -- -- -- (λ i₁ → ∣ merid x (~ i₁) ∣ₕ)) --- -- -- -- ∙∙ sym (lUnit _) --- -- -- -- ∙∙ sym (assoc _ _ _) --- -- -- -- ∙∙ (sym (lUnit _) ∙ sym (lUnit _) ∙ sym (lUnit _))) - --- -- -- -- psst : (x : S₊ 2) → Q (CP2 , inl tt) fibr (inr x) →∙ coHomK-ptd 2 --- -- -- -- fst (psst x) north = ∣ north ∣ --- -- -- -- fst (psst x) south = ∣ x ∣ --- -- -- -- fst (psst x) (merid a i₁) = kr (x , a) (~ i₁) --- -- -- -- snd (psst x) = refl - --- -- -- -- ll32-fst : GysinS1.c (inr north) .fst ≡ psst north .fst --- -- -- -- ll32-fst = cong fst (GysinS1.cEq (inr north) ∣ push (inl base) ∣₂) --- -- -- -- ∙ (funExt l) --- -- -- -- where --- -- -- -- l : (qb : _) → --- -- -- -- ∣ (subst (fst ∘ Q (CP2 , inl tt) fibr) (sym (push (inl base))) qb) ∣ --- -- -- -- ≡ psst north .fst qb --- -- -- -- l north = refl --- -- -- -- l south = cong ∣_∣ₕ (sym (merid base)) --- -- -- -- l (merid a i) j = --- -- -- -- hcomp (λ k → λ { (i = i0) → ∣ merid a (~ k ∧ j) ∣ --- -- -- -- ; (i = i1) → ∣ merid base (~ j) ∣ --- -- -- -- ; (j = i0) → ∣ transportRefl (merid a i) (~ k) ∣ --- -- -- -- ; (j = i1) → ∣ compPath-filler (merid base) (sym (merid a)) k (~ i) ∣ₕ}) --- -- -- -- (hcomp (λ k → λ { (i = i0) → ∣ merid a (j ∨ ~ k) ∣ --- -- -- -- ; (i = i1) → ∣ merid base (~ j ∨ ~ k) ∣ --- -- -- -- ; (j = i0) → ∣ merid a (~ k ∨ i) ∣ --- -- -- -- ; (j = i1) → ∣ merid base (~ i ∨ ~ k) ∣ₕ}) --- -- -- -- ∣ south ∣) - --- -- -- -- is-setHepl : (x : S₊ 2) → isSet (Q (CP2 , inl tt) fibr (inr x) →∙ coHomK-ptd 2) --- -- -- -- is-setHepl = sphereElim _ (λ _ → isProp→isOfHLevelSuc 1 (isPropIsOfHLevel 2)) --- -- -- -- (isOfHLevel↑∙' 0 1) - --- -- -- -- ll32 : (x : S₊ 2) → (GysinS1.c (inr x) ≡ psst x) --- -- -- -- ll32 = sphereElim _ (λ x → isOfHLevelPath 2 (is-setHepl x) _ _) --- -- -- -- (→∙Homogeneous≡ (isHomogeneousKn _) ll32-fst) - --- -- -- -- isGenerator≃ℤ-e : isGenerator≃ℤ (coHomGr 2 CP2) GysinS1.e --- -- -- -- isGenerator≃ℤ-e = --- -- -- -- subst (isGenerator≃ℤ (coHomGr 2 CP2)) (sym ll) --- -- -- -- ∣ grrg , refl ∣ - --- -- -- -- ⌣-pres1-CP2 : ⌣-pres1 CP2 2 --- -- -- -- ⌣-pres1-CP2 = --- -- -- -- GysinS1.e --- -- -- -- , (isGenerator≃ℤ-e --- -- -- -- , ∣ compGroupIso (GroupEquiv→GroupIso (invGroupEquiv ⌣hom)) grrg --- -- -- -- , {!!} ∣) - - --- -- -- -- {- --- -- -- -- Goal: snd (v' (pt A) (push a i₁)) ≡ --- -- -- -- ua-gluePt (μ-eq (snd a)) i₁ (fst a) --- -- -- -- ———— Boundary —————————————————————————————————————————————— --- -- -- -- i₁ = i0 ⊢ HSpace.μₗ e (fst a) --- -- -- -- i₁ = i1 ⊢ HSpace.μₗ e (HSpace.μ e (fst a) (snd a)) --- -- -- -- -} +module fibS1 = hopfBase S1-AssocHSpace (sphereElim2 _ (λ _ _ → squash) ∣ refl ∣) + +S¹Hopf : S₊ 2 → Type +S¹Hopf = fibS1.Hopf + +TotalHopf' : Type _ +TotalHopf' = Σ (S₊ 2) S¹Hopf + +CP2 : Type _ +CP2 = fibS1.megaPush + +fibr : CP2 → Type _ +fibr = fibS1.P + +hopf : join S¹ S¹ → S₊ 2 +hopf x = fst (JoinS¹S¹→TotalHopf x) + +E* : Type _ +E* = fibS1.totalSpaceMegaPush + +IsoE' : Iso E* (join S¹ (join S¹ S¹)) +IsoE' = fibS1.IsoJoin₁ + +IsoE2 : (join S¹ (join S¹ S¹)) ≡ join S¹ (S₊ 3) +IsoE2 = cong (join S¹) (sym S³≡joinS¹S¹ ∙ isoToPath IsoS³S3) + +IsoTotalHopf' : Iso fibS1.TotalSpaceHopf' (join S¹ S¹) +IsoTotalHopf' = fibS1.joinIso₁ + +CP' : Type _ +CP' = Pushout (λ _ → tt) hopf + +conCP2 : (x y : CP2) → ∥ x ≡ y ∥₂ +conCP2 x y = sRec2 squash₂ (λ p q → ∣ p ∙ sym q ∣₂) (conCP2' x) (conCP2' y) + where + conCP2' : (x : CP2) → ∥ x ≡ inl tt ∥₂ + conCP2' (inl x) = ∣ refl ∣₂ + conCP2' (inr x) = sphereElim 1 {A = λ x → ∥ inr x ≡ inl tt ∥₂} (λ _ → squash₂) ∣ sym (push (inl base)) ∣₂ x + conCP2' (push a i₁) = ll a i₁ + where + h2 : ∀ {ℓ} {A : fibS1.TotalSpaceHopf' → Type ℓ} → ((a : _) → isProp (A a)) + → A (inl base) + → ((a : fibS1.TotalSpaceHopf') → A a) + h2 {A = A} p b = + PushoutToProp p (sphereElim 0 (λ _ → p _) b) + (sphereElim 0 (λ _ → p _) (subst A (push (base , base)) b)) + + ll : (a : fibS1.TotalSpaceHopf') → PathP (λ i → ∥ Path CP2 (push a i) (inl tt) ∥₂) (conCP2' (inl tt)) (conCP2' (inr (fibS1.induced a))) + ll = h2 (λ _ → isOfHLevelPathP' 1 squash₂ _ _) λ j → ∣ (λ i → push (inl base) (~ i ∧ j)) ∣₂ + +module GysinS1 = Gysin (CP2 , inl tt) fibr conCP2 2 idIso refl + +PushoutReplaceBase : + ∀ {ℓ ℓ' ℓ''} {A B : Type ℓ} {C : Type ℓ'} {D : Type ℓ''} {f : A → C} {g : A → D} + (e : B ≃ A) → Pushout (f ∘ fst e) (g ∘ fst e) ≡ Pushout f g +PushoutReplaceBase {f = f} {g = g} = + EquivJ (λ _ e → Pushout (f ∘ fst e) (g ∘ fst e) ≡ Pushout f g) + refl + +isContrH³E : isContr (coHom 3 (GysinS1.E')) +isContrH³E = + subst isContr + (sym (isoToPath (compIso (Iso1 0) (Iso2' 0))) + ∙ cong (coHom 3) (sym (isoToPath IsoE' ∙ IsoE2))) + (ll5 0 (snotz ∘ sym)) + +isContrH⁴E : isContr (coHom 4 (GysinS1.E')) +isContrH⁴E = + subst isContr + (sym (isoToPath (compIso (Iso1 1) (Iso2' 1))) + ∙ cong (coHom 4) (sym (isoToPath IsoE' ∙ IsoE2))) + (ll5 1 λ p → snotz (sym (cong predℕ p))) + +genH2 = GysinS1.e + +S³→Groupoid : ∀ {ℓ} (P : join S¹ S¹ → Type ℓ) + → ((x : _) → isGroupoid (P x)) + → P (inl base) + → (x : _) → P x +S³→Groupoid P grp b = + transport (λ i → (x : (sym (isoToPath S³IsojoinS¹S¹) ∙ isoToPath IsoS³S3) (~ i)) + → P (transp (λ j → (sym (isoToPath S³IsojoinS¹S¹) ∙ isoToPath IsoS³S3) (~ i ∧ ~ j)) i x)) + (sphereElim _ (λ _ → grp _) b) + +TotalHopf→Gpd : ∀ {ℓ} (P : fibS1.TotalSpaceHopf' → Type ℓ) + → ((x : _) → isGroupoid (P x)) + → P (inl base) + → (x : _) → P x +TotalHopf→Gpd P grp b = + transport (λ i → (x : sym (isoToPath IsoTotalHopf') i) + → P (transp (λ j → isoToPath IsoTotalHopf' (~ i ∧ ~ j)) i x)) + (S³→Groupoid _ (λ _ → grp _) b) + +TotalHopf→Gpd' : ∀ {ℓ} (P : Σ (S₊ 2) fibS1.Hopf → Type ℓ) + → ((x : _) → isGroupoid (P x)) + → P (north , base) + → (x : _) → P x +TotalHopf→Gpd' P grp b = + transport (λ i → (x : ua (_ , fibS1.isEquivTotalSpaceHopf'→TotalSpace) i) + → P (transp (λ j → ua (_ , fibS1.isEquivTotalSpaceHopf'→TotalSpace) (i ∨ j)) i x)) + (TotalHopf→Gpd _ (λ _ → grp _) b) + +CP2→Groupoid : ∀ {ℓ} {P : CP2 → Type ℓ} + → ((x : _) → is2Groupoid (P x)) + → (b : P (inl tt)) + → (s2 : ((x : _) → P (inr x))) + → PathP (λ i → P (push (inl base) i)) b (s2 north) + → (x : _) → P x +CP2→Groupoid {P = P} grp b s2 pp (inl x) = b +CP2→Groupoid {P = P} grp b s2 pp (inr x) = s2 x +CP2→Groupoid {P = P} grp b s2 pp (push a i₁) = h3 a i₁ + where + h3 : (a : fibS1.TotalSpaceHopf') → PathP (λ i → P (push a i)) b (s2 _) + h3 = TotalHopf→Gpd _ (λ _ → isOfHLevelPathP' 3 (grp _) _ _) pp + +mm : (S₊ 2 → coHomK 2) → (CP2 → coHomK 2) +mm f = λ { (inl x) → 0ₖ _ + ; (inr x) → f x -ₖ f north + ; (push a i) → TotalHopf→Gpd (λ x → 0ₖ 2 ≡ f (fibS1.TotalSpaceHopf'→TotalSpace x .fst) -ₖ f north) + (λ _ → isOfHLevelTrunc 4 _ _) + (sym (rCancelₖ 2 (f north))) + a i} + +e-inv : coHomGr 2 (S₊ 2) .fst → coHomGr 2 CP2 .fst +e-inv = sMap mm + +cancel : (f : CP2 → coHomK 2) → ∥ mm (f ∘ inr) ≡ f ∥ +cancel f = + pRec squash + (λ p → pRec squash + (λ f → ∣ funExt f ∣) + (zz2 p)) + (h1 (f (inl tt))) + where + h1 : (x : coHomK 2) → ∥ x ≡ 0ₖ 2 ∥ + h1 = coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash) ∣ refl ∣ + + zz2 : (p : f (inl tt) ≡ 0ₖ 2) → ∥ ((x : CP2) → mm (f ∘ inr) x ≡ f x) ∥ + zz2 p = trRec squash (λ pp → + ∣ CP2→Groupoid (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) + (sym p) + (λ x → (λ i → f (inr x) -ₖ f (push (inl base) (~ i))) + ∙∙ (λ i → f (inr x) -ₖ p i) + ∙∙ rUnitₖ 2 (f (inr x))) + pp ∣) + l + where + l : hLevelTrunc 1 (PathP ((λ i₁ → mm (f ∘ inr) (push (inl base) i₁) ≡ f (push (inl base) i₁))) + (sym p) + (((λ i₁ → f (inr north) -ₖ f (push (inl base) (~ i₁))) ∙∙ + (λ i₁ → f (inr north) -ₖ p i₁) ∙∙ rUnitₖ 2 (f (inr north))))) + l = isConnectedPathP 1 (isConnectedPath 2 (isConnectedKn 1) _ _) _ _ .fst + +e' : GroupIso (coHomGr 2 CP2) (coHomGr 2 (S₊ 2)) +Iso.fun (fst e') = sMap (_∘ inr) +Iso.inv (fst e') = e-inv +Iso.rightInv (fst e') = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → trRec {B = Iso.fun (fst e') (Iso.inv (fst e') ∣ f ∣₂) ≡ ∣ f ∣₂} + (isOfHLevelPath 2 squash₂ _ _) + (λ p → cong ∣_∣₂ (funExt λ x → cong (λ y → (f x) -ₖ y) p ∙ rUnitₖ 2 (f x))) + (Iso.fun (PathIdTruncIso _) (isContr→isProp (isConnectedKn 1) ∣ f north ∣ ∣ 0ₖ 2 ∣)) +Iso.leftInv (fst e') = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → pRec (squash₂ _ _) (cong ∣_∣₂) (cancel f) +snd e' = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ f g → refl) + +grrg : GroupIso (coHomGr 2 CP2) ℤGroup +grrg = compGroupIso e' (Hⁿ-Sⁿ≅ℤ 1) + +⌣hom : GroupEquiv (coHomGr 2 CP2) (coHomGr 4 CP2) +fst (fst ⌣hom) = GysinS1.⌣-hom 2 .fst +snd (fst ⌣hom) = subst isEquiv (cong fst (GysinS1.ϕ∘j≡ 2)) h3 + where + h3 : isEquiv (GysinS1.ϕ∘j 2 .fst) + h3 = + SES→Iso + (GroupPath _ _ .fst (invGroupEquiv + (isContr→≃Unit isContrH³E + , makeIsGroupHom λ _ _ → refl))) + (GroupPath _ _ .fst (invGroupEquiv + (isContr→≃Unit isContrH⁴E + , makeIsGroupHom λ _ _ → refl))) + (GysinS1.susp∘ϕ 1) + (GysinS1.ϕ∘j 2) + (GysinS1.p-hom 4) + (GysinS1.Ker-ϕ∘j⊂Im-Susp∘ϕ _) + (GysinS1.Ker-p⊂Im-ϕ∘j _) +snd ⌣hom = GysinS1.⌣-hom 2 .snd + +isGenerator≃ℤ : ∀ {ℓ} (G : Group ℓ) (g : fst G) + → Type _ +isGenerator≃ℤ G g = + Σ[ e ∈ GroupIso G ℤGroup ] abs (Iso.fun (fst e) g) ≡ 1 + + +⌣-pres1 : ∀ {ℓ} (G : Type ℓ) (n : ℕ) → Type _ +⌣-pres1 G n = + Σ[ x ∈ coHomGr n G .fst ] + isGenerator≃ℤ (coHomGr n G) x × isGenerator≃ℤ (coHomGr (n +' n) G) (x ⌣ x) + +genr : coHom 2 CP2 +genr = ∣ CP2→Groupoid (λ _ → isOfHLevelTrunc 4) + (0ₖ _) + ∣_∣ + refl ∣₂ + +≡CP2 : (f g : CP2 → coHomK 2) → ∥ f ∘ inr ≡ g ∘ inr ∥ → Path (coHom 2 CP2) ∣ f ∣₂ ∣ g ∣₂ +≡CP2 f g = pRec (squash₂ _ _) + (λ p → pRec (squash₂ _ _) (λ id → trRec (squash₂ _ _) + (λ pp → cong ∣_∣₂ + (funExt (CP2→Groupoid (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) + id + (funExt⁻ p) + (compPathR→PathP pp)))) + (pp2 (f (inl tt)) (g (inl tt)) id + (cong f (push (inl base)) ∙ (funExt⁻ p north) ∙ cong g (sym (push (inl base)))))) + (conn (f (inl tt)) (g (inl tt)))) + + where + conn : (x y : coHomK 2) → ∥ x ≡ y ∥ + conn = coHomK-elim _ (λ _ → isSetΠ λ _ → isOfHLevelSuc 1 squash) + (coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash) ∣ refl ∣) + + pp2 : (x y : coHomK 2) (p q : x ≡ y) → hLevelTrunc 1 (p ≡ q) + pp2 x y = λ p q → Iso.fun (PathIdTruncIso _) (isContr→isProp (isConnectedPath _ (isConnectedKn 1) x y) ∣ p ∣ ∣ q ∣) + +rUnit* : (x : S¹) → x * base ≡ x +rUnit* base = refl +rUnit* (loop i₁) = refl + +meridP : (a x : S¹) → Path (Path (coHomK 2) _ _) (cong ∣_∣ₕ (merid (a * x) ∙ sym (merid base))) + ((cong ∣_∣ₕ (merid a ∙ sym (merid base))) ∙ (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) +meridP = wedgeconFun _ _ (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) + (λ x → lUnit _ ∙ cong (_∙ cong ∣_∣ₕ (merid x ∙ sym (merid base))) (cong (cong ∣_∣ₕ) (sym (rCancel (merid base))))) + (λ x → (λ i → cong ∣_∣ₕ (merid (rUnit* x i) ∙ sym (merid base))) + ∙∙ rUnit _ + ∙∙ cong (cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙_) + (cong (cong ∣_∣ₕ) (sym (rCancel (merid base))))) + (sym (l (cong ∣_∣ₕ (merid base ∙ sym (merid base))) + (cong (cong ∣_∣ₕ) (sym (rCancel (merid base)))))) + where + l : ∀ {ℓ} {A : Type ℓ} {x : A} (p : x ≡ x) (P : refl ≡ p) + → lUnit p ∙ cong (_∙ p) P ≡ rUnit p ∙ cong (p ∙_) P + l p = J (λ p P → lUnit p ∙ cong (_∙ p) P ≡ rUnit p ∙ cong (p ∙_) P) refl + +lemmie : (x : S¹) → ptSn 1 ≡ x * (invLooper x) +lemmie base = refl +lemmie (loop i) j = + hcomp (λ r → λ {(i = i0) → base ; (i = i1) → base ; (j = i0) → base}) + base + +mmInv : (x : S¹) → Path (Path (coHomK 2) _ _) + (cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base))) + (cong ∣_∣ₕ (sym (merid x ∙ sym (merid base)))) +mmInv x = (lUnit _ + ∙∙ cong (_∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base))) (sym (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base))))) ∙∙ sym (assoc _ _ _)) + ∙∙ cong (sym (cong ∣_∣ₕ (merid x ∙ sym (merid base))) ∙_) hh + ∙∙ (assoc _ _ _ + ∙∙ cong (_∙ (cong ∣_∣ₕ (sym (merid x ∙ sym (merid base))))) (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) + ∙∙ sym (lUnit _)) + where + hh : cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base)) + ≡ cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙ cong ∣_∣ₕ (sym (merid x ∙ sym (merid base))) + hh = sym (meridP x (invLooper x)) ∙ ((λ i → cong ∣_∣ₕ (merid (lemmie x (~ i)) ∙ sym (merid base))) ∙ cong (cong ∣_∣ₕ) (rCancel (merid base))) ∙ sym (rCancel _) + +commS1 : (a x : S¹) → a * x ≡ x * a +commS1 = wedgeconFun _ _ (λ _ _ → isGroupoidS¹ _ _) + (sym ∘ rUnit*) + rUnit* + refl + +s233 : (a x : S¹) → (invEq (fibS1.μ-eq a) x) * a ≡ (invLooper a * x) * a +s233 a x = secEq (fibS1.μ-eq a) x ∙∙ cong (_* x) (lemmie a) ∙∙ assocHSpace.μ-assoc S1-AssocHSpace a (invLooper a) x ∙ commS1 _ _ + +ss23 : (a x : S¹) → invEq (fibS1.μ-eq a) x ≡ invLooper a * x +ss23 a x = sym (retEq (fibS1.μ-eq a) (invEq (fibS1.μ-eq a) x)) + ∙∙ cong (invEq (fibS1.μ-eq a)) (s233 a x) + ∙∙ retEq (fibS1.μ-eq a) (invLooper a * x) + +ll : GysinS1.e ≡ genr +ll = ≡CP2 _ _ ∣ funExt (λ x → funExt⁻ (cong fst (ll32 x)) south) ∣ + where + kr : (x : Σ (S₊ 2) fibS1.Hopf) → Path (hLevelTrunc 4 _) ∣ fst x ∣ ∣ north ∣ + kr = uncurry λ { north → λ y → cong ∣_∣ₕ (merid base ∙ sym (merid y)) + ; south → λ y → cong ∣_∣ₕ (sym (merid y)) + ; (merid a i) → lem a i} + where + lem : (a : S¹) → PathP (λ i → (y : fibS1.Hopf (merid a i)) → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a i ∣ ∣ north ∣) + (λ y → cong ∣_∣ₕ (merid base ∙ sym (merid y))) + λ y → cong ∣_∣ₕ (sym (merid y)) + lem a = toPathP (funExt λ x → cong (transport ((λ i₁ → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a i₁ ∣ ∣ north ∣))) + ((λ i → (λ z → cong ∣_∣ₕ (merid base + ∙ sym (merid (transport (λ j → uaInvEquiv (fibS1.μ-eq a) (~ i) j) x))) z)) + ∙ λ i → cong ∣_∣ₕ (merid base + ∙ sym (merid (transportRefl (invEq (fibS1.μ-eq a) x) i)))) + ∙∙ (λ i → transp (λ i₁ → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a (i₁ ∨ i) ∣ ∣ north ∣) i + (compPath-filler' (cong ∣_∣ₕ (sym (merid a))) + (cong ∣_∣ₕ (merid base ∙ sym (merid (ss23 a x i)))) i)) + ∙∙ cong ((cong ∣_∣ₕ) (sym (merid a)) ∙_) (cong (cong ∣_∣ₕ) (cong sym (symDistr (merid base) (sym (merid (invLooper a * x))))) + ∙ cong sym (meridP (invLooper a) x) + ∙ symDistr ((cong ∣_∣ₕ) (merid (invLooper a) ∙ sym (merid base))) + ((cong ∣_∣ₕ) (merid x ∙ sym (merid base))) + ∙ isCommΩK 2 (sym (λ i₁ → ∣ (merid x ∙ (λ i₂ → merid base (~ i₂))) i₁ ∣)) + (sym (λ i₁ → ∣ (merid (invLooper a) ∙ (λ i₂ → merid base (~ i₂))) i₁ ∣)) + ∙ cong₂ _∙_ (cong sym (mmInv a) ∙ cong-∙ ∣_∣ₕ (merid a) (sym (merid base))) + (cong (cong ∣_∣ₕ) (symDistr (merid x) (sym (merid base))) ∙ cong-∙ ∣_∣ₕ (merid base) (sym (merid x)))) + ∙∙ (λ j → (λ i₁ → ∣ merid a (~ i₁ ∨ j) ∣) ∙ + ((λ i₁ → ∣ merid a (i₁ ∨ j) ∣) ∙ + (λ i₁ → ∣ merid base (~ i₁ ∨ j) ∣)) + ∙ + (λ i₁ → ∣ merid base (i₁ ∨ j) ∣) ∙ + (λ i₁ → ∣ merid x (~ i₁) ∣ₕ)) + ∙∙ sym (lUnit _) + ∙∙ sym (assoc _ _ _) + ∙∙ (sym (lUnit _) ∙ sym (lUnit _) ∙ sym (lUnit _))) + + psst : (x : S₊ 2) → Q (CP2 , inl tt) fibr (inr x) →∙ coHomK-ptd 2 + fst (psst x) north = ∣ north ∣ + fst (psst x) south = ∣ x ∣ + fst (psst x) (merid a i₁) = kr (x , a) (~ i₁) + snd (psst x) = refl + + ll32-fst : GysinS1.c (inr north) .fst ≡ psst north .fst + ll32-fst = cong fst (GysinS1.cEq (inr north) ∣ push (inl base) ∣₂) + ∙ (funExt l) + where + l : (qb : _) → + ∣ (subst (fst ∘ Q (CP2 , inl tt) fibr) (sym (push (inl base))) qb) ∣ + ≡ psst north .fst qb + l north = refl + l south = cong ∣_∣ₕ (sym (merid base)) + l (merid a i) j = + hcomp (λ k → λ { (i = i0) → ∣ merid a (~ k ∧ j) ∣ + ; (i = i1) → ∣ merid base (~ j) ∣ + ; (j = i0) → ∣ transportRefl (merid a i) (~ k) ∣ + ; (j = i1) → ∣ compPath-filler (merid base) (sym (merid a)) k (~ i) ∣ₕ}) + (hcomp (λ k → λ { (i = i0) → ∣ merid a (j ∨ ~ k) ∣ + ; (i = i1) → ∣ merid base (~ j ∨ ~ k) ∣ + ; (j = i0) → ∣ merid a (~ k ∨ i) ∣ + ; (j = i1) → ∣ merid base (~ i ∨ ~ k) ∣ₕ}) + ∣ south ∣) + + is-setHepl : (x : S₊ 2) → isSet (Q (CP2 , inl tt) fibr (inr x) →∙ coHomK-ptd 2) + is-setHepl = sphereElim _ (λ _ → isProp→isOfHLevelSuc 1 (isPropIsOfHLevel 2)) + (isOfHLevel↑∙' 0 1) + + ll32 : (x : S₊ 2) → (GysinS1.c (inr x) ≡ psst x) + ll32 = sphereElim _ (λ x → isOfHLevelPath 2 (is-setHepl x) _ _) + (→∙Homogeneous≡ (isHomogeneousKn _) ll32-fst) + +isGenerator≃ℤ-e : isGenerator≃ℤ (coHomGr 2 CP2) GysinS1.e +isGenerator≃ℤ-e = + subst (isGenerator≃ℤ (coHomGr 2 CP2)) (sym ll) + (grrg , refl) + +hopfFun : S₊ 3 → S₊ 2 +hopfFun x = fibS1.TotalSpaceHopf'→TotalSpace + (Iso.inv IsoTotalHopf' + (Iso.fun S³IsojoinS¹S¹ (Iso.inv IsoS³S3 x))) .fst + +CP2' : Type _ +CP2' = Pushout {A = S₊ 3} (λ _ → tt) hopfFun + +CP2≡CP2' : CP2' ≡ CP2 +CP2≡CP2' = + PushoutReplaceBase + (isoToEquiv (compIso (invIso IsoS³S3) + (compIso S³IsojoinS¹S¹ (invIso IsoTotalHopf')))) + + +{- +⌣-pres1-CP2 : ⌣-pres1 CP2 2 +fst ⌣-pres1-CP2 = GysinS1.e +fst (snd ⌣-pres1-CP2) = isGenerator≃ℤ-e +snd (snd ⌣-pres1-CP2) = ∣ p ∣ + where + p : Σ _ _ + fst p = compGroupIso (GroupEquiv→GroupIso (invGroupEquiv ⌣hom)) grrg + snd p = {!!} + ∙ sesIsoPresGen _ (invGroupEquiv (GroupIso→GroupEquiv grrg)) _ + (compGroupEquiv (invGroupEquiv (GroupIso→GroupEquiv grrg)) (invGroupEquiv (invGroupEquiv ⌣hom))) + GysinS1.e {!!} ⌣hom +{- + GysinS1.e + , (isGenerator≃ℤ-e + , ∣ compGroupIso (GroupEquiv→GroupIso (invGroupEquiv ⌣hom)) grrg + , {!!} ∣) + -} + +⌣-pres1-CP2' : ⌣-pres1 CP2' 2 +⌣-pres1-CP2' = subst (λ x → ⌣-pres1 x 2) (sym CP2≡CP2') ⌣-pres1-CP2 + -} +-- {- +-- Goal: snd (v' (pt A) (push a i₁)) ≡ +-- ua-gluePt (μ-eq (snd a)) i₁ (fst a) +-- ———— Boundary —————————————————————————————————————————————— +-- i₁ = i0 ⊢ HSpace.μₗ e (fst a) +-- i₁ = i1 ⊢ HSpace.μₗ e (HSpace.μ e (fst a) (snd a)) +-- -} --- -- -- -- -- help : (x : _) → (v' (pt A)) x ≡ TotalSpaceHopf'→TotalSpace x --- -- -- -- -- help (inl x) = ΣPathP (refl , HSpace.μₗ e x) --- -- -- -- -- help (inr x) = ΣPathP (refl , (HSpace.μₗ e x)) --- -- -- -- -- help (push (x , y) i) j = --- -- -- -- -- comp (λ _ → Σ (Susp (typ A)) Hopf) --- -- -- -- -- (λ k → λ {(i = i0) → merid y i , HSpace.μₗ e x j --- -- -- -- -- ; (i = i1) → merid y i , assocHSpace.μ-assoc-filler e-ass x y j k --- -- -- -- -- ; (j = i0) → merid y i , hfill --- -- -- -- -- (λ j → λ { (i = i0) → HSpace.μ e (pt A) x --- -- -- -- -- ; (i = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j --- -- -- -- -- }) --- -- -- -- -- (inS (ua-gluePt (μ-eq y) i (HSpace.μ e (pt A) x))) --- -- -- -- -- k --- -- -- -- -- ; (j = i1) → merid y i , ua-gluePt (μ-eq y) i x}) --- -- -- -- -- (merid y i , ua-gluePt (μ-eq y) i (HSpace.μₗ e x j)) --- -- -- -- -- where --- -- -- -- -- open import Cubical.Foundations.Path - --- -- -- -- -- PPΣ : ∀ {ℓ} {A : Type ℓ} {f : A ≃ A} (p : f ≡ f) → {!!} --- -- -- -- -- PPΣ = {!!} - --- -- -- -- -- V : PathP (λ i₁ → hcomp --- -- -- -- -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x --- -- -- -- -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j --- -- -- -- -- }) --- -- -- -- -- (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x)) ≡ --- -- -- -- -- ua-gluePt (μ-eq y) i₁ x) --- -- -- -- -- (HSpace.μₗ e x) --- -- -- -- -- (HSpace.μₗ e (HSpace.μ e x y)) -- (HSpace.μₗ e (HSpace.μ e (fst a) (snd a))) --- -- -- -- -- V = transport (λ z → {!PathP (λ i₁ → hfill --- -- -- -- -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x --- -- -- -- -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j --- -- -- -- -- }) --- -- -- -- -- (inS (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x))) z ≡ ua-gluePt (μ-eq y) i₁ x) --- -- -- -- -- ? ?!}) --- -- -- -- -- {!hfill --- -- -- -- -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x --- -- -- -- -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j --- -- -- -- -- }) --- -- -- -- -- (inS (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x))) ?!} -- toPathP ({!!} ∙∙ {!!} ∙∙ {!!}) -- toPathP (flipSquare {!!}) -- hcomp {!!} {!!} - --- -- -- -- -- P : Pushout {A = TotalSpaceHopf'} (λ _ → tt) induced → Type _ --- -- -- -- -- P (inl x) = typ A --- -- -- -- -- P (inr x) = Hopf x --- -- -- -- -- P (push a i₁) = ua (v a) i₁ +-- -- help : (x : _) → (v' (pt A)) x ≡ TotalSpaceHopf'→TotalSpace x +-- -- help (inl x) = ΣPathP (refl , HSpace.μₗ e x) +-- -- help (inr x) = ΣPathP (refl , (HSpace.μₗ e x)) +-- -- help (push (x , y) i) j = +-- -- comp (λ _ → Σ (Susp (typ A)) Hopf) +-- -- (λ k → λ {(i = i0) → merid y i , HSpace.μₗ e x j +-- -- ; (i = i1) → merid y i , assocHSpace.μ-assoc-filler e-ass x y j k +-- -- ; (j = i0) → merid y i , hfill +-- -- (λ j → λ { (i = i0) → HSpace.μ e (pt A) x +-- -- ; (i = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j +-- -- }) +-- -- (inS (ua-gluePt (μ-eq y) i (HSpace.μ e (pt A) x))) +-- -- k +-- -- ; (j = i1) → merid y i , ua-gluePt (μ-eq y) i x}) +-- -- (merid y i , ua-gluePt (μ-eq y) i (HSpace.μₗ e x j)) +-- -- where +-- -- open import Cubical.Foundations.Path + +-- -- PPΣ : ∀ {ℓ} {A : Type ℓ} {f : A ≃ A} (p : f ≡ f) → {!!} +-- -- PPΣ = {!!} + +-- -- V : PathP (λ i₁ → hcomp +-- -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x +-- -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j +-- -- }) +-- -- (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x)) ≡ +-- -- ua-gluePt (μ-eq y) i₁ x) +-- -- (HSpace.μₗ e x) +-- -- (HSpace.μₗ e (HSpace.μ e x y)) -- (HSpace.μₗ e (HSpace.μ e (fst a) (snd a))) +-- -- V = transport (λ z → {!PathP (λ i₁ → hfill +-- -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x +-- -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j +-- -- }) +-- -- (inS (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x))) z ≡ ua-gluePt (μ-eq y) i₁ x) +-- -- ? ?!}) +-- -- {!hfill +-- -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x +-- -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j +-- -- }) +-- -- (inS (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x))) ?!} -- toPathP ({!!} ∙∙ {!!} ∙∙ {!!}) -- toPathP (flipSquare {!!}) -- hcomp {!!} {!!} + +-- -- P : Pushout {A = TotalSpaceHopf'} (λ _ → tt) induced → Type _ +-- -- P (inl x) = typ A +-- -- P (inr x) = Hopf x +-- -- P (push a i₁) = ua (v a) i₁ diff --git a/Cubical/ZCohomology/Gysin2.agda b/Cubical/ZCohomology/Gysin2.agda new file mode 100644 index 000000000..b986144f4 --- /dev/null +++ b/Cubical/ZCohomology/Gysin2.agda @@ -0,0 +1,1221 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.ZCohomology.Gysin2 where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Groups.Connected +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Groups.Wedge +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.MayerVietorisUnreduced +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.RingStructure.CupProduct +open import Cubical.ZCohomology.RingStructure.RingLaws +open import Cubical.ZCohomology.RingStructure.GradedCommutativity + +open import Cubical.Data.Sum +open import Cubical.Relation.Nullary + +open import Cubical.Functions.Embedding + +open import Cubical.Data.Fin + +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Path +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Function hiding (case_of_) +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Pointed.Homogeneous + +open import Cubical.Foundations.Univalence + +open import Cubical.Relation.Nullary + +open import Cubical.Data.Empty renaming (rec to ⊥-rec) +open import Cubical.Data.Sigma +open import Cubical.Data.Int hiding (_+'_) +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_ ; _·_ to _·ℕ_) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Unit +open import Cubical.Data.Bool +open import Cubical.Algebra.Group + renaming (ℤ to ℤGroup ; Unit to UnitGroup) hiding (Bool) + +open import Cubical.HITs.Pushout.Flattening +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.EilenbergSteenrod +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 renaming (_·_ to _*_) +open import Cubical.HITs.Truncation + renaming (rec to trRec ; elim to trElim ; elim2 to trElim2) +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap) +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; elim to pElim) + +open import Cubical.Algebra.AbGroup + +open import Cubical.Homotopy.Loopspace + +open import Cubical.HITs.Join + +open import Cubical.Homotopy.Hopf + +open import Cubical.HITs.SetQuotients renaming (_/_ to _/'_) +open import Cubical.ZCohomology.Gysin + + +rUnitℤ· : ∀ {ℓ} (G : Group ℓ) → (x : ℤ) → (x ℤ[ G ]· GroupStr.1g (snd G)) + ≡ GroupStr.1g (snd G) +rUnitℤ· G (pos zero) = refl +rUnitℤ· G (pos (suc n)) = + cong (GroupStr._·_ (snd G) (GroupStr.1g (snd G))) + (rUnitℤ· G (pos n)) + ∙ GroupStr.lid (snd G) (GroupStr.1g (snd G)) +rUnitℤ· G (negsuc zero) = GroupTheory.inv1g G +rUnitℤ· G (negsuc (suc n)) = + cong₂ (GroupStr._·_ (snd G)) (GroupTheory.inv1g G) (rUnitℤ· G (negsuc n)) + ∙ GroupStr.lid (snd G) (GroupStr.1g (snd G)) + +rUnitℤ·ℤ : (x : ℤ) → (x ℤ[ ℤGroup ]· 1) ≡ x +rUnitℤ·ℤ (pos zero) = refl +rUnitℤ·ℤ (pos (suc n)) = cong (pos 1 +_) (rUnitℤ·ℤ (pos n)) ∙ sym (pos+ 1 n) +rUnitℤ·ℤ (negsuc zero) = refl +rUnitℤ·ℤ (negsuc (suc n)) = cong (-1 +_) (rUnitℤ·ℤ (negsuc n)) ∙ +Comm (negsuc 0) (negsuc n) + +private + precommℤ : ∀ {ℓ} (G : Group ℓ) → (g : fst G) (y : ℤ) → (snd G GroupStr.· (y ℤ[ G ]· g)) g + ≡ (sucℤ y ℤ[ G ]· g) + precommℤ G g (pos zero) = GroupStr.lid (snd G) g + ∙ sym (GroupStr.rid (snd G) g) + precommℤ G g (pos (suc n)) = + sym (GroupStr.assoc (snd G) _ _ _) + ∙ cong ((snd G GroupStr.· g)) (precommℤ G g (pos n)) + precommℤ G g (negsuc zero) = + GroupStr.invl (snd G) g + precommℤ G g (negsuc (suc n)) = + sym (GroupStr.assoc (snd G) _ _ _) + ∙ cong ((snd G GroupStr.· GroupStr.inv (snd G) g)) (precommℤ G g (negsuc n)) + ∙ negsucLem n + where + negsucLem : (n : ℕ) → (snd G GroupStr.· GroupStr.inv (snd G) g) + (sucℤ (negsuc n) ℤ[ G ]· g) + ≡ (sucℤ (negsuc (suc n)) ℤ[ G ]· g) + negsucLem zero = (GroupStr.rid (snd G) _) + negsucLem (suc n) = refl + +distrℤ· : ∀ {ℓ} (G : Group ℓ) → (g : fst G) (x y : ℤ) + → ((x + y) ℤ[ G ]· g) + ≡ GroupStr._·_ (snd G) (x ℤ[ G ]· g) (y ℤ[ G ]· g) +distrℤ· G' g (pos zero) y = cong (_ℤ[ G' ]· g) (+Comm 0 y) + ∙ sym (GroupStr.lid (snd G') _) +distrℤ· G' g (pos (suc n)) (pos n₁) = + cong (_ℤ[ G' ]· g) (sym (pos+ (suc n) n₁)) + ∙ cong (GroupStr._·_ (snd G') g) (cong (_ℤ[ G' ]· g) (pos+ n n₁) ∙ distrℤ· G' g (pos n) (pos n₁)) + ∙ GroupStr.assoc (snd G') _ _ _ +distrℤ· G' g (pos (suc n)) (negsuc zero) = + distrℤ· G' g (pos n) 0 + ∙ ((GroupStr.rid (snd G') (pos n ℤ[ G' ]· g) ∙ sym (GroupStr.lid (snd G') (pos n ℤ[ G' ]· g))) + ∙ cong (λ x → GroupStr._·_ (snd G') x (pos n ℤ[ G' ]· g)) + (sym (GroupStr.invl (snd G') g)) ∙ sym (GroupStr.assoc (snd G') _ _ _)) + ∙ (GroupStr.assoc (snd G') _ _ _) + ∙ cong (λ x → GroupStr._·_ (snd G') x (pos n ℤ[ G' ]· g)) (GroupStr.invl (snd G') g) + ∙ GroupStr.lid (snd G') _ + ∙ sym (GroupStr.rid (snd G') _) + ∙ (cong (GroupStr._·_ (snd G') (pos n ℤ[ G' ]· g)) (sym (GroupStr.invr (snd G') g)) + ∙ GroupStr.assoc (snd G') _ _ _) + ∙ cong (λ x → GroupStr._·_ (snd G') x (GroupStr.inv (snd G') g)) + (precommℤ G' g (pos n)) +distrℤ· G' g (pos (suc n)) (negsuc (suc n₁)) = + cong (_ℤ[ G' ]· g) (predℤ+negsuc n₁ (pos (suc n))) + ∙∙ distrℤ· G' g (pos n) (negsuc n₁) + ∙∙ (cong (λ x → GroupStr._·_ (snd G') x (negsuc n₁ ℤ[ G' ]· g)) + ((sym (GroupStr.rid (snd G') (pos n ℤ[ G' ]· g)) + ∙ cong (GroupStr._·_ (snd G') (pos n ℤ[ G' ]· g)) (sym (GroupStr.invr (snd G') g))) + ∙∙ GroupStr.assoc (snd G') _ _ _ + ∙∙ cong (λ x → GroupStr._·_ (snd G') x (GroupStr.inv (snd G') g)) (precommℤ G' g (pos n) )) + ∙ sym (GroupStr.assoc (snd G') _ _ _)) +distrℤ· G' g (negsuc zero) y = + cong (_ℤ[ G' ]· g) (+Comm -1 y) ∙ lol1 y + module _ where + lol1 : (y : ℤ) → (predℤ y ℤ[ G' ]· g) ≡ (snd G' GroupStr.· GroupStr.inv (snd G') g) (y ℤ[ G' ]· g) + lol1 (pos zero) = sym (GroupStr.rid (snd G') _) + lol1 (pos (suc n)) = + sym (GroupStr.lid (snd G') ((pos n ℤ[ G' ]· g))) + ∙∙ cong (λ x → GroupStr._·_ (snd G') x (pos n ℤ[ G' ]· g)) (sym (GroupStr.invl (snd G') g)) + ∙∙ sym (GroupStr.assoc (snd G') _ _ _) + lol1 (negsuc n) = refl +distrℤ· G' g (negsuc (suc n)) y = + cong (_ℤ[ G' ]· g) (+Comm (negsuc (suc n)) y) + ∙∙ lol1 G' g 0 (y +negsuc n) + ∙∙ cong (snd G' GroupStr.· GroupStr.inv (snd G') g) + (cong (_ℤ[ G' ]· g) (+Comm y (negsuc n)) ∙ distrℤ· G' g (negsuc n) y) + ∙ (GroupStr.assoc (snd G') _ _ _) + + + +HopfInvariantPushElim : ∀ {ℓ} n → (f : _) → {P : HopfInvariantPush n f → Type ℓ} + → (isOfHLevel (suc (suc (suc (suc (n +ℕ n))))) (P (inl tt))) + → (e : P (inl tt)) + (g : (x : _) → P (inr x)) + (r : PathP (λ i → P (push north i)) e (g (f north))) + → (x : _) → P x +HopfInvariantPushElim n f {P = P} hlev e g r (inl x) = e +HopfInvariantPushElim n f {P = P} hlev e g r (inr x) = g x +HopfInvariantPushElim n f {P = P} hlev e g r (push a i₁) = help a i₁ + where + help : (a : Susp (Susp (S₊ (suc (n +ℕ n))))) → PathP (λ i → P (push a i)) e (g (f a)) + help = sphereElim _ (sphereElim _ (λ _ → isProp→isOfHLevelSuc ((suc (suc (n +ℕ n)))) (isPropIsOfHLevel _)) + (isOfHLevelPathP' (suc (suc (suc (n +ℕ n)))) + (subst (isOfHLevel (suc (suc (suc (suc (n +ℕ n)))))) + (cong P (push north)) + hlev) _ _)) r + +transportCohomIso : ∀ {ℓ} {A : Type ℓ} {n m : ℕ} + → (p : n ≡ m) + → GroupIso (coHomGr n A) (coHomGr m A) +Iso.fun (fst (transportCohomIso {A = A} p)) = subst (λ n → coHom n A) p +Iso.inv (fst (transportCohomIso {A = A} p)) = subst (λ n → coHom n A) (sym p) +Iso.rightInv (fst (transportCohomIso p)) = transportTransport⁻ _ +Iso.leftInv (fst (transportCohomIso p)) = transportTransport⁻ _ +snd (transportCohomIso {A = A} {n = n} {m = m} p) = + makeIsGroupHom (λ x y → help x y p) + where + help : (x y : coHom n A) (p : n ≡ m) → subst (λ n → coHom n A) p (x +ₕ y) + ≡ subst (λ n → coHom n A) p x +ₕ subst (λ n → coHom n A) p y + help x y = J (λ m p → subst (λ n → coHom n A) p (x +ₕ y) + ≡ subst (λ n → coHom n A) p x +ₕ subst (λ n → coHom n A) p y) + (transportRefl (x +ₕ y) ∙ cong₂ _+ₕ_ (sym (transportRefl x)) (sym (transportRefl y))) + +α-hopfMap : abs (HopfFunction zero (hopfFun , λ _ → hopfFun (snd (S₊∙ 3)))) ≡ suc zero +α-hopfMap = cong abs (cong (Iso.fun (fst (Hopfβ-Iso zero (hopfFun , refl)))) + (transportRefl (⌣-α 0 (hopfFun , refl)))) ∙ l3 (sym CP2≡CP2') GysinS1.e (Hopfα zero (hopfFun , refl)) + (l isGenerator≃ℤ-e) + (GroupIso→GroupEquiv (Hopfα-Iso 0 (hopfFun , refl)) , refl) + (snd (fst ⌣hom)) + (GroupIso→GroupEquiv (Hopfβ-Iso zero (hopfFun , refl))) + where + l : Σ-syntax (GroupIso (coHomGr 2 CP2) ℤGroup) + (λ ϕ → abs (Iso.fun (fst ϕ) GysinS1.e) ≡ 1) + → Σ-syntax (GroupEquiv (coHomGr 2 CP2) ℤGroup) + (λ ϕ → abs (fst (fst ϕ) GysinS1.e) ≡ 1) + l p = (GroupIso→GroupEquiv (fst p)) , (snd p) + +·Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (S₊∙ n →∙ A) + → (S₊∙ n →∙ A) + → (S₊∙ n →∙ A) +·Π {A = A} {n = zero} p q = (λ _ → pt A) , refl +fst (·Π {A = A} {n = suc zero} (f , p) (g , q)) base = pt A +fst (·Π {A = A} {n = suc zero} (f , p) (g , q)) (loop j) = + ((sym p ∙∙ cong f loop ∙∙ p) ∙ (sym q ∙∙ cong g loop ∙∙ q)) j +snd (·Π {A = A} {n = suc zero} (f , p) (g , q)) = refl +fst (·Π {A = A} {n = suc (suc n)} (f , p) (g , q)) north = pt A +fst (·Π {A = A} {n = suc (suc n)} (f , p) (g , q)) south = pt A +fst (·Π {A = A} {n = suc (suc n)} (f , p) (g , q)) (merid a j) = + ((sym p ∙∙ cong f (merid a ∙ sym (merid (ptSn (suc n)))) ∙∙ p) + ∙ (sym q ∙∙ cong g (merid a ∙ sym (merid (ptSn (suc n)))) ∙∙ q)) j +snd (·Π {A = A} {n = suc (suc n)} (f , p) (g , q)) = refl + +flipLoopSpace : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → Iso (fst ((Ω^ (suc n)) A)) (fst ((Ω^ n) (Ω A))) +flipLoopSpace n = {!!} + +module _ {ℓ : Level} {A : Pointed ℓ} where + convertLoopFun→ : (n : ℕ) → (S₊∙ (suc n) →∙ Ω A) → (S₊∙ (suc (suc n)) →∙ A) + fst (convertLoopFun→ n f) north = pt A + fst (convertLoopFun→ n f) south = pt A + fst (convertLoopFun→ n f) (merid a i₁) = fst f a i₁ + snd (convertLoopFun→ n f) = refl + + myFiller : (n : ℕ) → (S₊∙ (suc (suc n)) →∙ A) → I → I → I → fst A + myFiller n f i j k = + hfill (λ k → λ { (i = i0) → doubleCompPath-filler (sym (snd f)) (cong (fst f) (merid (ptSn _) ∙ sym (merid (ptSn _)))) (snd f) k j + ; (i = i1) → snd f k + ; (j = i0) → snd f k + ; (j = i1) → snd f k}) + (inS ((fst f) (rCancel (merid (ptSn _)) i j))) + k + + convertLoopFun← : (n : ℕ) → (S₊∙ (suc (suc n)) →∙ A) → (S₊∙ (suc n) →∙ Ω A) + fst (convertLoopFun← n f) a = sym (snd f) ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f + snd (convertLoopFun← n f) i j = myFiller n f i j i1 + + convertLoopFun : (n : ℕ) → Iso (S₊∙ (suc n) →∙ Ω A) (S₊∙ (suc (suc n)) →∙ A) + Iso.fun (convertLoopFun n) f = convertLoopFun→ n f + Iso.inv (convertLoopFun n) f = convertLoopFun← n f + Iso.rightInv (convertLoopFun n) (f , p) = + ΣPathP (funExt help , λ i j → p (~ i ∨ j)) + where + help : (x : _) → fst (convertLoopFun→ n (convertLoopFun← n (f , p))) x ≡ f x + help north = sym p + help south = sym p ∙ cong f (merid (ptSn _)) + help (merid a j) i = + hcomp (λ k → λ { (i = i1) → f (merid a j) + ; (j = i0) → p (~ i ∧ k) + ; (j = i1) → compPath-filler' (sym p) (cong f (merid (ptSn _))) k i}) + (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) + Iso.leftInv (convertLoopFun n) f = + ΣPathP (funExt help34 , help2) + where + + sillyFill₂ : (x : S₊ (suc n)) → I → I → I → fst A + sillyFill₂ x i j k = + hfill (λ k → λ { (i = i0) → convertLoopFun→ n f .fst (compPath-filler' (merid x) (sym (merid (ptSn _))) k j) + ; (i = i1) → snd A + ; (j = i0) → fst f x (i ∨ ~ k) + ; (j = i1) → snd A}) + (inS (snd f i (~ j))) + k + + sillyFill : (x : S₊ (suc n)) → I → I → I → fst A + sillyFill x i j k = + hfill (λ k → λ { (i = i0) → doubleCompPath-filler refl (cong (convertLoopFun→ n f .fst) (merid x ∙ sym (merid (ptSn _)))) refl k j + ; (i = i1) → fst f x (j ∨ ~ k) + ; (j = i0) → fst f x (i ∧ ~ k) + ; (j = i1) → snd A}) + (inS (sillyFill₂ x i j i1)) + k + + + + help34 : (x : _) → fst (convertLoopFun← n (convertLoopFun→ n f)) x ≡ fst f x + help34 x i j = sillyFill x i j i1 + + help2 : PathP _ _ _ + help2 i j k = + hcomp (λ r → λ {(i = i0) → myFiller n (convertLoopFun→ n f) j k r + ; (i = i1) → snd f j (k ∨ ~ r) + ; (j = i0) → sillyFill (ptSn _) i k r + ; (j = i1) → snd A + ; (k = i0) → snd f j (i ∧ ~ r) + ; (k = i1) → snd A}) + (hcomp ((λ r → λ {(i = i0) → {!!} + ; (i = i1) → {!!} + ; (j = i0) → {!!} + ; (j = i1) → {!!} + ; (k = i0) → {!!} + ; (k = i1) → {!!}})) + {!!}) + + {- + i = i0 ⊢ snd (convertLoopFun← n (Iso.fun (convertLoopFun n) f)) j k +i = i1 ⊢ snd f j k +j = i0 ⊢ funExt help i (snd (S₊∙ (suc n))) k +j = i1 ⊢ snd (Ω A) k +k = i0 ⊢ snd A +k = i1 ⊢ snd A + -} + +→' : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) → (S₊∙ (suc n) →∙ A) → (fst ((Ω^ (suc n)) A)) +→' zero (f , p) = sym p ∙∙ cong f loop ∙∙ p +→' {A = A} (suc n) (f , p) = Iso.inv (flipLoopSpace _) (→' {A = Ω A} n (Iso.inv (convertLoopFun _) (f , p))) + +←' : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) → (fst ((Ω^ (suc n)) A)) → (S₊∙ (suc n) →∙ A) +fst (←' {A = A} zero f) base = pt A +fst (←' {A = A} zero f) (loop i₁) = f i₁ +snd (←' {A = A} zero f) = refl +←' {A = A} (suc n) f = Iso.fun (convertLoopFun _) (←' n (Iso.fun (flipLoopSpace _) f)) + +cancel-r : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (x : fst ((Ω^ (suc n)) A)) + → →' n (←' n x) ≡ x +cancel-r zero x = sym (rUnit _) +cancel-r (suc n) x = + cong (Iso.inv (flipLoopSpace (suc n)) ∘ →' n) + (Iso.leftInv (convertLoopFun _) + (←' n (Iso.fun (flipLoopSpace (suc n)) x))) + ∙∙ cong (Iso.inv (flipLoopSpace (suc n))) + (cancel-r n (Iso.fun (flipLoopSpace (suc n)) x)) + ∙∙ Iso.leftInv (flipLoopSpace (suc n)) x + +cancel-l : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (x : (S₊∙ (suc n) →∙ A)) + → ←' n (→' n x) ≡ x +cancel-l zero (f , p) = ΣPathP (funExt fstp , λ i j → p (~ i ∨ j)) + where + fstp : (x : _) → _ ≡ f x + fstp base = sym p + fstp (loop i) j = doubleCompPath-filler (sym p) (cong f loop) p (~ j) i +cancel-l (suc n) f = + cong (Iso.fun (convertLoopFun n) ∘ ←' n) (Iso.rightInv (flipLoopSpace (suc n)) + (→' n (Iso.inv (convertLoopFun n) f))) + ∙∙ cong (Iso.fun (convertLoopFun n)) (cancel-l n (Iso.inv (convertLoopFun n) f)) + ∙∙ Iso.rightInv (convertLoopFun _) _ + +→∙Ω : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (S₊∙ n →∙ A) (fst ((Ω^ n) A)) +Iso.fun (→∙Ω {A = A} zero) (f , p) = f false +fst (Iso.inv (→∙Ω {A = A} zero) a) false = a +fst (Iso.inv (→∙Ω {A = A} zero) a) true = pt A +snd (Iso.inv (→∙Ω {A = A} zero) a) = refl +Iso.rightInv (→∙Ω {A = A} zero) a = refl +Iso.leftInv (→∙Ω {A = A} zero) (f , p) = ΣPathP ((funExt (λ { false → refl ; true → sym p})) , λ i j → p (~ i ∨ j)) +→∙Ω {A = A} (suc n) = {!!} + +-- 0Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → (S₊∙ n →∙ A) +-- fst (0Π {A = A}) _ = pt A +-- snd (0Π {A = A}) = refl + +-- comm·Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} +-- → (f g : S₊∙ (suc (suc n)) →∙ A) +-- → ·Π f g ≡ ·Π g f +-- fst (comm·Π {A = A} {n = n} f g i) north = pt A +-- fst (comm·Π {A = A} {n = n} f g i) south = pt A +-- fst (comm·Π {A = A} {n = zero} f g i) (merid base j) = {!!} +-- fst (comm·Π {A = A} {n = zero} f g i) (merid (loop k) j) = {!!} +-- fst (comm·Π {A = A} {n = suc n} f g i) (merid a j) = {!!} +-- snd (comm·Π {A = A} {n = n} f g i) = refl + +-- rUnit·Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} +-- → (f : S₊∙ (suc n) →∙ A) +-- → ·Π f 0Π ≡ f +-- fst (rUnit·Π {A = A} {n = zero} f i) base = snd f (~ i) +-- fst (rUnit·Π {A = A} {n = zero} f i) (loop j) = help i j +-- where +-- help : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) +-- (((sym (snd f)) ∙∙ (cong (fst f) loop) ∙∙ snd f) +-- ∙ (refl ∙ refl)) +-- (cong (fst f) loop) +-- help = (cong ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) ∙_) (sym (rUnit refl)) ∙ sym (rUnit _)) +-- ◁ λ i j → doubleCompPath-filler (sym (snd f)) (cong (fst f) loop) (snd f) (~ i) j +-- snd (rUnit·Π {A = A} {n = zero} f i) j = snd f (~ i ∨ j) +-- fst (rUnit·Π {A = A} {n = suc n} f i) north = snd f (~ i) +-- fst (rUnit·Π {A = A} {n = suc n} f i) south = (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i +-- fst (rUnit·Π {A = A} {n = suc n} f i) (merid a j) = help i j +-- where +-- help : PathP (λ i → snd f (~ i) ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) +-- (((sym (snd f)) ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) ∙∙ snd f) +-- ∙ (refl ∙ refl)) +-- (cong (fst f) (merid a)) +-- help = (cong (((sym (snd f)) ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) ∙∙ snd f) ∙_) +-- (sym (rUnit refl)) +-- ∙ sym (rUnit _)) +-- ◁ λ i j → hcomp (λ k → λ { (j = i0) → snd f (~ i ∧ k) +-- ; (j = i1) → compPath-filler' (sym (snd f)) (cong (fst f) (merid (ptSn (suc n)))) k i +-- ; (i = i0) → doubleCompPath-filler (sym (snd f)) +-- (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) +-- (snd f) k j +-- ; (i = i1) → fst f (merid a j)}) +-- (fst f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) +-- snd (rUnit·Π {A = A} {n = suc n} f i) j = snd f (~ i ∨ j) + +-- -- _·π_ : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} +-- -- → ∥ S₊∙ n →∙ A ∥₂ +-- -- → ∥ S₊∙ n →∙ A ∥₂ +-- -- → ∥ S₊∙ n →∙ A ∥₂ +-- -- _·π_ = sRec2 squash₂ λ f g → ∣ ·Π f g ∣₂ + +-- -- open import Cubical.HITs.Wedge +-- -- _∨→_ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} +-- -- → (f : A →∙ C) (g : B →∙ C) +-- -- → A ⋁ B → fst C +-- -- (f ∨→ g) (inl x) = fst f x +-- -- (f ∨→ g) (inr x) = fst g x +-- -- (f ∨→ g) (push a i₁) = (snd f ∙ sym (snd g)) i₁ + + +-- -- add2 : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → (f g : S₊∙ n →∙ A) → ((S₊∙ n) ⋁ (S₊∙ n) , inl (ptSn _)) →∙ A +-- -- fst (add2 {A = A} f g) = f ∨→ g +-- -- snd (add2 {A = A} f g) = snd f + +-- -- C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Type _ +-- -- C* n f g = Pushout (λ _ → tt) (fst (add2 f g)) + +-- -- θ : ∀ {ℓ} {A : Pointed ℓ} → Susp (fst A) → (Susp (fst A) , north) ⋁ (Susp (fst A) , north) +-- -- θ north = inl north +-- -- θ south = inr north +-- -- θ {A = A} (merid a i₁) = +-- -- ((λ i → inl ((merid a ∙ sym (merid (pt A))) i)) +-- -- ∙∙ push tt +-- -- ∙∙ λ i → inr ((merid a ∙ sym (merid (pt A))) i)) i₁ + + +-- -- help3 : ∀ {ℓ} {A : Type ℓ} {x y z w : A} (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) +-- -- → p ∙∙ q ∙∙ r ≡ p ∙ q ∙ r +-- -- help3 {x = x} {y = y} {z = z} {w = w} = +-- -- J (λ y p → (q : y ≡ z) (r : z ≡ w) → +-- -- (p ∙∙ q ∙∙ r) ≡ p ∙ q ∙ r) +-- -- λ q r → lUnit (q ∙ r) + +-- -- +≡ : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → (x : _) → ·Π f g .fst x ≡ (f ∨→ g) (θ {A = S₊∙ _} x) +-- -- +≡ n f g north = sym (snd f) +-- -- +≡ n f g south = sym (snd g) +-- -- +≡ n f g (merid a i₁) j = main j i₁ +-- -- where + +-- -- help : cong (f ∨→ g) (cong (θ {A = S₊∙ _}) (merid a)) +-- -- ≡ (refl ∙∙ cong (fst f) (merid a ∙ sym (merid north)) ∙∙ snd f) +-- -- ∙ (sym (snd g) ∙∙ cong (fst g) (merid a ∙ sym (merid north)) ∙∙ refl) +-- -- help = cong-∙∙ (f ∨→ g) ((λ i → inl ((merid a ∙ sym (merid north)) i))) +-- -- (push tt) +-- -- (λ i → inr ((merid a ∙ sym (merid north)) i)) +-- -- ∙∙ help3 _ _ _ +-- -- ∙∙ cong (cong (f ∨→ g) +-- -- (λ i₂ → inl ((merid a ∙ (λ i₃ → merid north (~ i₃))) i₂)) ∙_) +-- -- (sym (assoc _ _ _)) +-- -- ∙∙ assoc _ _ _ +-- -- ∙∙ cong₂ _∙_ refl (compPath≡compPath' _ _) + +-- -- main : PathP (λ i → snd f (~ i) ≡ snd g (~ i)) (λ i₁ → ·Π f g .fst (merid a i₁)) +-- -- λ i₁ → (f ∨→ g) (θ {A = S₊∙ _} (merid a i₁)) +-- -- main = (λ i → ((λ j → snd f (~ i ∧ ~ j)) ∙∙ cong (fst f) (merid a ∙ sym (merid north)) ∙∙ snd f) +-- -- ∙ (sym (snd g) ∙∙ cong (fst g) (merid a ∙ sym (merid north)) ∙∙ λ j → snd g (~ i ∧ j))) +-- -- ▷ sym help + +-- -- +≡' : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → ·Π f g ≡ ((f ∨→ g) ∘ (θ {A = S₊∙ _}) , snd f) +-- -- +≡' n f g = ΣPathP ((funExt (+≡ n f g)) , (λ j i → snd f (i ∨ ~ j))) + +-- -- WedgeElim : (n : ℕ) → ∀ {ℓ} {P : S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) → Type ℓ} +-- -- → ((x : _) → isOfHLevel (3 +ℕ n) (P x)) +-- -- → P (inl north) +-- -- → (x : _) → P x +-- -- WedgeElim n {P = P} x s (inl x₁) = +-- -- sphereElim _ {A = P ∘ inl} +-- -- (λ _ → isOfHLevelPlus' {n = n} (3 +ℕ n) (x _)) s x₁ +-- -- WedgeElim n {P = P} x s (inr x₁) = +-- -- sphereElim _ {A = P ∘ inr} +-- -- (sphereElim _ (λ _ → isProp→isOfHLevelSuc ((suc (suc (n +ℕ n)))) +-- -- (isPropIsOfHLevel (suc (suc (suc (n +ℕ n)))))) +-- -- (subst (isOfHLevel ((3 +ℕ n) +ℕ n)) (cong P (push tt)) +-- -- (isOfHLevelPlus' {n = n} (3 +ℕ n) (x _)))) +-- -- (subst P (push tt) s) x₁ +-- -- WedgeElim n {P = P} x s (push a j) = transp (λ i → P (push tt (i ∧ j))) (~ j) s + + +-- -- H²-C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → GroupIso (coHomGr (2 +ℕ n) (C* n f g)) ℤGroup +-- -- H²-C* n f g = compGroupIso groupIso1 (Hⁿ-Sⁿ≅ℤ (suc n)) +-- -- where +-- -- help : (coHom (2 +ℕ n) (C* n f g)) → coHom (2 +ℕ n) (S₊ (2 +ℕ n)) +-- -- help = sMap λ f → f ∘ inr + + +-- -- invMapPrim : (S₊ (2 +ℕ n) → coHomK (2 +ℕ n)) → C* n f g → coHomK (2 +ℕ n) +-- -- invMapPrim h (inl x) = h (ptSn _) +-- -- invMapPrim h (inr x) = h x +-- -- invMapPrim h (push a i₁) = WedgeElim n {P = λ a → h north ≡ h (fst (add2 f g) a)} +-- -- (λ _ → isOfHLevelTrunc (4 +ℕ n) _ _) +-- -- (cong h (sym (snd f))) a i₁ + +-- -- invMap : coHom (2 +ℕ n) (S₊ (2 +ℕ n)) → (coHom (2 +ℕ n) (C* n f g)) +-- -- invMap = sMap invMapPrim + +-- -- ret1 : (x : coHom (2 +ℕ n) (S₊ (2 +ℕ n))) → help (invMap x) ≡ x +-- -- ret1 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) +-- -- λ a → refl + +-- -- ret2 : (x : (coHom (2 +ℕ n) (C* n f g))) → invMap (help x) ≡ x +-- -- ret2 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) +-- -- λ h → cong ∣_∣₂ (funExt λ { (inl x) → cong h ((λ i → inr ((snd f) (~ i))) +-- -- ∙ sym (push (inl north))) +-- -- ; (inr x) → refl +-- -- ; (push a i₁) → help2 h a i₁}) +-- -- where +-- -- help2 : (h : C* n f g → coHomK (2 +ℕ n)) +-- -- → (a : _) → PathP (λ i → invMapPrim (h ∘ inr) (push a i) ≡ h (push a i)) +-- -- (cong h ((λ i → inr ((snd f) (~ i))) ∙ sym (push (inl north)))) +-- -- refl +-- -- help2 h = WedgeElim n (λ _ → isOfHLevelPathP (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n) _ _) _ _) +-- -- help4 + +-- -- where +-- -- help4 : PathP (λ i → invMapPrim (h ∘ inr) (push (inl north) i) ≡ h (push (inl north) i)) +-- -- (cong h ((λ i → inr ((snd f) (~ i))) ∙ sym (push (inl north)))) +-- -- refl +-- -- help4 i j = h (hcomp (λ k → λ { (i = i1) → inr (fst f north) +-- -- ; (j = i0) → inr (snd f (~ i)) +-- -- ; (j = i1) → push (inl north) (i ∨ ~ k)}) +-- -- (inr (snd f (~ i ∧ ~ j)))) + +-- -- groupIso1 : GroupIso ((coHomGr (2 +ℕ n) (C* n f g))) (coHomGr (2 +ℕ n) (S₊ (2 +ℕ n))) +-- -- Iso.fun (fst groupIso1) = help +-- -- Iso.inv (fst groupIso1) = invMap +-- -- Iso.rightInv (fst groupIso1) = ret1 +-- -- Iso.leftInv (fst groupIso1) = ret2 +-- -- snd groupIso1 = makeIsGroupHom +-- -- (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) +-- -- λ f g → refl) + + +-- -- C+ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Type _ +-- -- C+ n f g = Pushout (λ _ → tt) (·Π f g .fst) + +-- -- H⁴-C∨ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → GroupIso (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C+ n f g)) +-- -- ℤGroup +-- -- H⁴-C∨ n f g = compGroupIso +-- -- (transportCohomIso (cong (3 +ℕ_) (+-suc n n))) (Hopfβ-Iso n (·Π f g)) + +-- -- lol : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) (x : ℤ) +-- -- → Iso.inv (fst (H⁴-C∨ n f g)) x +-- -- ≡ subst (λ m → coHom m (C+ n f g)) (sym (cong (3 +ℕ_) (+-suc n n))) (Iso.inv (fst (Hopfβ-Iso n (·Π f g))) x) +-- -- lol n f g = λ _ → refl + +-- -- H²-C∨ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → GroupIso (coHomGr (2 +ℕ n) (C+ n f g)) +-- -- ℤGroup +-- -- H²-C∨ n f g = Hopfα-Iso n (·Π f g) + + +-- -- H⁴-C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → GroupIso (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) +-- -- (DirProd ℤGroup ℤGroup) +-- -- H⁴-C* n f g = compGroupIso (GroupEquiv→GroupIso (invGroupEquiv fstIso)) +-- -- (compGroupIso (transportCohomIso (cong (2 +ℕ_) (+-suc n n))) +-- -- (compGroupIso (Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) _) +-- -- (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ _) (Hⁿ-Sⁿ≅ℤ _)))) +-- -- where +-- -- module Ms = MV _ _ _ (λ _ → tt) (fst (add2 f g)) +-- -- fstIso : GroupEquiv (coHomGr ((suc (suc (n +ℕ suc n)))) (S₊∙ (3 +ℕ (n +ℕ n)) ⋁ S₊∙ (3 +ℕ (n +ℕ n)))) +-- -- (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) +-- -- fst (fst fstIso) = Ms.d (suc (suc (n +ℕ suc n))) .fst +-- -- snd (fst fstIso) = +-- -- SES→Iso +-- -- (GroupPath _ _ .fst (compGroupEquiv (invEquiv (isContr→≃Unit ((tt , tt) , (λ _ → refl))) , makeIsGroupHom λ _ _ → refl) +-- -- (GroupEquivDirProd +-- -- (GroupIso→GroupEquiv (invGroupIso (Hⁿ-Unit≅0 _))) +-- -- (GroupIso→GroupEquiv +-- -- (invGroupIso +-- -- (Hⁿ-Sᵐ≅0 _ _ λ p → ¬lem 0 ((λ _ → north) , refl) n n (cong suc (sym (+-suc n n)) ∙ p))))))) +-- -- (GroupPath _ _ .fst +-- -- (compGroupEquiv ((invEquiv (isContr→≃Unit ((tt , tt) , (λ _ → refl))) , makeIsGroupHom λ _ _ → refl)) +-- -- ((GroupEquivDirProd +-- -- (GroupIso→GroupEquiv (invGroupIso (Hⁿ-Unit≅0 _))) +-- -- (GroupIso→GroupEquiv +-- -- (invGroupIso (Hⁿ-Sᵐ≅0 _ _ λ p → ¬lem 0 ((λ _ → north) , refl) n (suc n) (cong (2 +ℕ_) (sym (+-suc n n)) ∙ p)))))))) +-- -- (Ms.Δ (suc (suc (n +ℕ suc n)))) +-- -- (Ms.d (suc (suc (n +ℕ suc n)))) +-- -- (Ms.i (suc (suc (suc (n +ℕ suc n))))) +-- -- (Ms.Ker-d⊂Im-Δ _) +-- -- (Ms.Ker-i⊂Im-d _) +-- -- snd fstIso = Ms.d (suc (suc (n +ℕ suc n))) .snd + +-- -- Hopfie : {n : ℕ} → ∥ S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n) ∥₂ → ℤ +-- -- Hopfie = sRec isSetℤ (HopfFunction _) + +-- -- subber : (n : ℕ) (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → _ +-- -- subber n f = subst (λ x → coHom x (HopfInvariantPush n (fst f))) +-- -- (λ i → suc (suc (suc (+-suc n n i)))) + +-- -- module _ (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) where + +-- -- C+' = C+ n f g + +-- -- βₗ : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) +-- -- βₗ = Iso.inv (fst (H⁴-C* n f g)) (1 , 0) + +-- -- βᵣ : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) +-- -- βᵣ = Iso.inv (fst (H⁴-C* n f g)) (0 , 1) + +-- -- βᵣ'-fun : C* n f g → coHomK ((4 +ℕ (n +ℕ n))) +-- -- βᵣ'-fun (inl x) = 0ₖ _ +-- -- βᵣ'-fun (inr x) = 0ₖ _ +-- -- βᵣ'-fun (push (inl x) i₁) = 0ₖ _ +-- -- βᵣ'-fun (push (inr x) i₁) = Kn→ΩKn+1 _ ∣ x ∣ i₁ +-- -- βᵣ'-fun (push (push a i₂) i₁) = Kn→ΩKn+10ₖ _ (~ i₂) i₁ + + +-- -- βₗ'-fun : C* n f g → coHomK (4 +ℕ (n +ℕ n)) +-- -- βₗ'-fun (inl x) = 0ₖ _ +-- -- βₗ'-fun (inr x) = 0ₖ _ +-- -- βₗ'-fun (push (inl x) i₁) = Kn→ΩKn+1 _ ∣ x ∣ i₁ +-- -- βₗ'-fun (push (inr x) i₁) = 0ₖ _ +-- -- βₗ'-fun (push (push a i₂) i₁) = Kn→ΩKn+10ₖ _ i₂ i₁ + +-- -- βₗ''-fun : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) +-- -- βₗ''-fun = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) ∣ βₗ'-fun ∣₂ + +-- -- βᵣ''-fun : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) +-- -- βᵣ''-fun = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) ∣ βᵣ'-fun ∣₂ + +-- -- α∨ : coHom (2 +ℕ n) (C* n f g) +-- -- α∨ = Iso.inv (fst (H²-C* n f g)) 1 + +-- -- private +-- -- pp : (a : _) → 0ₖ (suc (suc n)) ≡ ∣ (f ∨→ g) a ∣ +-- -- pp = WedgeElim _ (λ _ → isOfHLevelTrunc (4 +ℕ n) _ _) +-- -- (cong ∣_∣ₕ (sym (snd f))) + +-- -- pp-inr : pp (inr north) ≡ cong ∣_∣ₕ (sym (snd g)) +-- -- pp-inr = (λ j → transport (λ i → 0ₖ (2 +ℕ n) ≡ ∣ compPath-filler' (snd f) (sym (snd g)) (~ j) i ∣ₕ) +-- -- λ i → ∣ snd f (~ i ∨ j) ∣ₕ) +-- -- ∙ λ j → transp (λ i → 0ₖ (2 +ℕ n) ≡ ∣ snd g (~ i ∧ ~ j) ∣ₕ) j λ i → ∣ snd g (~ i ∨ ~ j) ∣ₕ + +-- -- α∨'-fun : C* n f g → coHomK (2 +ℕ n) +-- -- α∨'-fun (inl x) = 0ₖ _ +-- -- α∨'-fun (inr x) = ∣ x ∣ +-- -- α∨'-fun (push a i₁) = pp a i₁ + +-- -- α∨' : coHom (2 +ℕ n) (C* n f g) +-- -- α∨' = ∣ α∨'-fun ∣₂ + + +-- -- β+ : coHom ((2 +ℕ n) +' (2 +ℕ n)) C+' +-- -- β+ = Iso.inv (fst (H⁴-C∨ n f g)) 1 + +-- -- q : C+' → C* n f g +-- -- q (inl x) = inl x +-- -- q (inr x) = inr x +-- -- q (push a i₁) = (push (θ a) ∙ λ i → inr (+≡ n f g a (~ i))) i₁ + +-- -- jₗ : HopfInvariantPush n (fst f) → C* n f g +-- -- jₗ (inl x) = inl x +-- -- jₗ (inr x) = inr x +-- -- jₗ (push a i₁) = push (inl a) i₁ + +-- -- jᵣ : HopfInvariantPush n (fst g) → C* n f g +-- -- jᵣ (inl x) = inl x +-- -- jᵣ (inr x) = inr x +-- -- jᵣ (push a i₁) = push (inr a) i₁ + +-- -- α-∨→1 : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → Iso.fun (fst (H²-C* n f g)) (α∨' n f g) ≡ 1 +-- -- α-∨→1 n f g = sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (h33 n)) +-- -- ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) + +-- -- α-∨-lem : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → α∨ n f g ≡ α∨' n f g +-- -- α-∨-lem n f g = sym (Iso.leftInv (fst (H²-C* n f g)) _) +-- -- ∙∙ cong (Iso.inv (fst (H²-C* n f g))) help +-- -- ∙∙ Iso.leftInv (fst (H²-C* n f g)) _ +-- -- where +-- -- help : Iso.fun (fst (H²-C* n f g)) (α∨ n f g) ≡ Iso.fun (fst (H²-C* n f g)) (α∨' n f g) +-- -- help = (Iso.rightInv (fst (H²-C* n f g)) (pos 1)) ∙ sym (α-∨→1 n f g) + +-- -- q-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → coHomFun _ (q n f g) (α∨ n f g) ≡ Hopfα n (·Π f g) +-- -- q-α n f g = (λ i → coHomFun _ (q n f g) (α-∨-lem n f g i)) +-- -- ∙ sym (Iso.leftInv is _) +-- -- ∙∙ cong (Iso.inv is) help +-- -- ∙∙ Iso.leftInv is _ +-- -- where +-- -- is = fst (Hopfα-Iso n (·Π f g)) + +-- -- help : Iso.fun is (coHomFun _ (q n f g) (α∨' n f g)) +-- -- ≡ Iso.fun is (Hopfα n (·Π f g)) +-- -- help = sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (h33 n)) +-- -- ∙∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) +-- -- ∙∙ sym (Hopfα-Iso-α n (·Π f g)) + + +-- -- βₗ≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → βₗ n f g ≡ βₗ''-fun n f g +-- -- βₗ≡ n f g = cong (FF ∘ subber2) +-- -- (cong (Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) +-- -- (S₊∙ (suc (suc (suc (n +ℕ n))))) +-- -- (((suc (suc (n +ℕ n))))))))) help +-- -- ∙ help2) +-- -- ∙ funExt⁻ FF∘subber ∣ wedgeGen ∣₂ +-- -- ∙ cong subber3 (sym βₗ'-fun≡) +-- -- where +-- -- FF = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (n +ℕ suc n))) .fst +-- -- FF' = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (suc (n +ℕ n)))) .fst + +-- -- help : Iso.inv (fst (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (1 , 0) +-- -- ≡ (∣ ∣_∣ ∣₂ , 0ₕ _) +-- -- help = ΣPathP ((h33 _) , IsGroupHom.pres1 (snd (invGroupIso (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) + +-- -- subber3 = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + +-- -- subber2 = (subst (λ m → coHom m (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) +-- -- (sym (cong (2 +ℕ_) (+-suc n n)))) + +-- -- FF∘subber : FF ∘ subber2 ≡ subber3 ∘ FF' +-- -- FF∘subber = +-- -- funExt (λ a → (λ j → transp (λ i → coHom ((suc ∘ suc ∘ suc) (+-suc n n (~ i ∧ j))) (C* n f g)) (~ j) +-- -- (MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) ((suc (suc (+-suc n n j)))) .fst +-- -- (transp (λ i → coHom (2 +ℕ (+-suc n n (j ∨ ~ i))) +-- -- (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) j +-- -- a)))) + +-- -- wedgeGen : (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) → +-- -- coHomK (suc (suc (suc (n +ℕ n))))) +-- -- wedgeGen (inl x) = ∣ x ∣ +-- -- wedgeGen (inr x) = 0ₖ _ +-- -- wedgeGen (push a i₁) = 0ₖ _ + +-- -- βₗ'-fun≡ : ∣ βₗ'-fun n f g ∣₂ ≡ FF' ∣ wedgeGen ∣₂ +-- -- βₗ'-fun≡ = cong ∣_∣₂ (funExt λ { (inl x) → refl +-- -- ; (inr x) → refl +-- -- ; (push (inl x) i₁) → refl +-- -- ; (push (inr x) i₁) j → Kn→ΩKn+10ₖ _ (~ j) i₁ +-- -- ; (push (push a i₂) i₁) j → Kn→ΩKn+10ₖ _ (~ j ∧ i₂) i₁}) + +-- -- help2 : Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) (((suc (suc (n +ℕ n)))))))) +-- -- (∣ ∣_∣ ∣₂ , 0ₕ _) +-- -- ≡ ∣ wedgeGen ∣₂ +-- -- help2 = cong ∣_∣₂ (funExt λ { (inl x) → rUnitₖ (suc (suc (suc (n +ℕ n)))) ∣ x ∣ +-- -- ; (inr x) → refl +-- -- ; (push a i₁) → refl}) + +-- -- βᵣ≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → βᵣ n f g ≡ βᵣ''-fun n f g +-- -- βᵣ≡ n f g = +-- -- cong (FF ∘ subber2) +-- -- (cong (Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) +-- -- (S₊∙ (suc (suc (suc (n +ℕ n))))) +-- -- (((suc (suc (n +ℕ n))))))))) +-- -- help +-- -- ∙ help2) +-- -- ∙ funExt⁻ FF∘subber ∣ wedgeGen ∣₂ +-- -- ∙ cong (subber3) (sym βᵣ'-fun≡) +-- -- where +-- -- FF = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (n +ℕ suc n))) .fst +-- -- FF' = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (suc (n +ℕ n)))) .fst + +-- -- help : Iso.inv (fst (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (0 , 1) +-- -- ≡ (0ₕ _ , ∣ ∣_∣ ∣₂) +-- -- help = ΣPathP (IsGroupHom.pres1 (snd (invGroupIso (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) , (h33 _)) + +-- -- subber3 = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + +-- -- subber2 = (subst (λ m → coHom m (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) +-- -- (sym (cong (2 +ℕ_) (+-suc n n)))) + +-- -- FF∘subber : FF ∘ subber2 ≡ subber3 ∘ FF' +-- -- FF∘subber = +-- -- funExt (λ a → (λ j → transp (λ i → coHom ((suc ∘ suc ∘ suc) (+-suc n n (~ i ∧ j))) (C* n f g)) (~ j) +-- -- (MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) ((suc (suc (+-suc n n j)))) .fst +-- -- (transp (λ i → coHom (2 +ℕ (+-suc n n (j ∨ ~ i))) +-- -- (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) j +-- -- a)))) + +-- -- wedgeGen : (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) → +-- -- coHomK (suc (suc (suc (n +ℕ n))))) +-- -- wedgeGen (inl x) = 0ₖ _ +-- -- wedgeGen (inr x) = ∣ x ∣ +-- -- wedgeGen (push a i₁) = 0ₖ _ + + +-- -- help2 : Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) (((suc (suc (n +ℕ n)))))))) +-- -- (0ₕ _ , ∣ ∣_∣ ∣₂) +-- -- ≡ ∣ wedgeGen ∣₂ +-- -- help2 = cong ∣_∣₂ (funExt λ { (inl x) → refl +-- -- ; (inr x) → lUnitₖ (suc (suc (suc (n +ℕ n)))) ∣ x ∣ +-- -- ; (push a i₁) → refl}) + +-- -- βᵣ'-fun≡ : ∣ βᵣ'-fun n f g ∣₂ ≡ FF' ∣ wedgeGen ∣₂ +-- -- βᵣ'-fun≡ = cong ∣_∣₂ (funExt λ { (inl x) → refl +-- -- ; (inr x) → refl +-- -- ; (push (inl x) i₁) j → Kn→ΩKn+10ₖ _ (~ j) i₁ +-- -- ; (push (inr x) i₁) → refl +-- -- ; (push (push a i₂) i₁) j → Kn→ΩKn+10ₖ _ (~ j ∧ ~ i₂) i₁}) + +-- -- q-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → coHomFun _ (q n f g) (βₗ n f g) +-- -- ≡ β+ n f g +-- -- q-βₗ n f g = cong (coHomFun _ (q n f g)) (βₗ≡ n f g) +-- -- ∙ transportLem +-- -- ∙ cong (subst (λ m → coHom m (C+' n f g)) +-- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) +-- -- (sym (Iso.leftInv (fst (Hopfβ-Iso n (·Π f g))) (Hopfβ n (·Π f g))) +-- -- ∙ cong (Iso.inv ((fst (Hopfβ-Iso n (·Π f g))))) (d-β n (·Π f g))) +-- -- where +-- -- transportLem : coHomFun (suc (suc (suc (n +ℕ suc n)))) (q n f g) +-- -- (βₗ''-fun n f g) +-- -- ≡ subst (λ m → coHom m (C+' n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- -- (Hopfβ n (·Π f g)) +-- -- transportLem = +-- -- natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (q n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- -- ∙ cong (subst (λ m → coHom m (C+' n f g)) +-- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) +-- -- (cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → loll a i₁})) +-- -- where +-- -- open import Cubical.Foundations.Path +-- -- loll : (x : fst (S₊∙ (3 +ℕ (n +ℕ n)))) → +-- -- PathP +-- -- (λ i₁ → +-- -- βₗ'-fun n f g ((push (θ x) ∙ λ i → inr (+≡ n f g x (~ i))) i₁) ≡ +-- -- MV.d-pre Unit (fst (S₊∙ (2 +ℕ n))) (fst (S₊∙ (3 +ℕ n +ℕ n))) +-- -- (λ _ → tt) (fst (·Π f g)) (3 +ℕ n +ℕ n) ∣_∣ (push x i₁)) +-- -- (λ _ → βₗ'-fun n f g (q n f g (inl tt))) +-- -- (λ _ → βₗ'-fun n f g (q n f g (inr (·Π f g .fst x)))) +-- -- loll x = +-- -- flipSquare (cong-∙ (βₗ'-fun n f g) (push (θ x)) (λ i → inr (+≡ n f g x (~ i))) +-- -- ∙∙ sym (rUnit _) +-- -- ∙∙ lem₁ x) + +-- -- where +-- -- lem₁ : (x : _) → cong (βₗ'-fun n f g) (push (θ {A = S₊∙ _} x)) ≡ Kn→ΩKn+1 _ ∣ x ∣ +-- -- lem₁ north = refl +-- -- lem₁ south = sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)) +-- -- lem₁ (merid a j) k = lala k j +-- -- where +-- -- lala : PathP (λ k → Kn→ΩKn+1 _ ∣ north ∣ₕ ≡ (sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north))) k) +-- -- (λ j → cong (βₗ'-fun n f g) (push (θ {A = S₊∙ _} (merid a j)))) +-- -- (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) +-- -- lala = (cong-∙∙ (cong (βₗ'-fun n f g) ∘ push) +-- -- ((λ i₁ → inl ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁))) +-- -- (push tt) +-- -- ((λ i₁ → inr ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁))) +-- -- ∙ sym (compPath≡compPath' _ _) +-- -- ∙ (λ _ → (λ j → Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∣ (merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) j ∣ₕ) +-- -- ∙ Kn→ΩKn+10ₖ _) +-- -- ∙ cong (_∙ Kn→ΩKn+10ₖ _) (cong-∙ ((Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ)) +-- -- (merid a) (sym (merid north))) +-- -- ∙ sym (assoc _ _ _) +-- -- ∙ cong (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a) ∙_) +-- -- (sym (symDistr (sym ((Kn→ΩKn+10ₖ _))) (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)))))) +-- -- ◁ λ i j → compPath-filler (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) +-- -- (sym (sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) +-- -- (cong ∣_∣ₕ (merid north)))) +-- -- (~ i) j + +-- -- q-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → coHomFun _ (q n f g) (βᵣ n f g) +-- -- ≡ β+ n f g +-- -- q-βᵣ n f g = cong (coHomFun _ (q n f g)) (βᵣ≡ n f g) +-- -- ∙ transportLem +-- -- ∙ cong (subst (λ m → coHom m (C+' n f g)) +-- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) +-- -- (sym (Iso.leftInv (fst (Hopfβ-Iso n (·Π f g))) (Hopfβ n (·Π f g))) +-- -- ∙ cong (Iso.inv ((fst (Hopfβ-Iso n (·Π f g))))) (d-β n (·Π f g))) +-- -- where +-- -- transportLem : coHomFun (suc (suc (suc (n +ℕ suc n)))) (q n f g) +-- -- (βᵣ''-fun n f g) +-- -- ≡ subst (λ m → coHom m (C+' n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- -- (Hopfβ n (·Π f g)) +-- -- transportLem = +-- -- natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (q n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- -- ∙ cong (subst (λ m → coHom m (C+' n f g)) +-- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) +-- -- (cong ∣_∣₂ (funExt λ { (inl x) → refl +-- -- ; (inr x) → refl +-- -- ; (push a i₁) → loll a i₁})) +-- -- where +-- -- open import Cubical.Foundations.Path +-- -- loll : (x : fst (S₊∙ (3 +ℕ (n +ℕ n)))) → +-- -- PathP +-- -- (λ i₁ → +-- -- βᵣ'-fun n f g ((push (θ x) ∙ λ i → inr (+≡ n f g x (~ i))) i₁) ≡ +-- -- MV.d-pre Unit (fst (S₊∙ (2 +ℕ n))) (fst (S₊∙ (3 +ℕ n +ℕ n))) +-- -- (λ _ → tt) (fst (·Π f g)) (3 +ℕ n +ℕ n) ∣_∣ (push x i₁)) +-- -- (λ _ → βᵣ'-fun n f g (q n f g (inl tt))) +-- -- (λ _ → βᵣ'-fun n f g (q n f g (inr (·Π f g .fst x)))) +-- -- loll x = flipSquare (cong-∙ (βᵣ'-fun n f g) (push (θ x)) (λ i → inr (+≡ n f g x (~ i))) +-- -- ∙∙ sym (rUnit _) +-- -- ∙∙ lem₁ x) +-- -- where +-- -- lem₁ : (x : _) → cong (βᵣ'-fun n f g) (push (θ {A = S₊∙ _} x)) ≡ Kn→ΩKn+1 _ ∣ x ∣ +-- -- lem₁ north = sym (Kn→ΩKn+10ₖ _) +-- -- lem₁ south = cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)) +-- -- lem₁ (merid a j) k = lala k j -- lala k j +-- -- where +-- -- lala : PathP (λ k → (Kn→ΩKn+10ₖ _) (~ k) ≡ (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north))) k) +-- -- (λ j → cong (βᵣ'-fun n f g) (push (θ {A = S₊∙ _} (merid a j)))) +-- -- (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) +-- -- lala = (cong-∙∙ (cong (βᵣ'-fun n f g) ∘ push) +-- -- (λ i₁ → inl ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁)) +-- -- (push tt) +-- -- (λ i₁ → inr ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁)) +-- -- ∙∙ (cong (sym (Kn→ΩKn+10ₖ _) ∙_) (cong-∙ (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a) (sym (merid (ptSn _))))) +-- -- ∙∙ sym (help3 _ _ _)) +-- -- ◁ symP (doubleCompPath-filler +-- -- (sym (Kn→ΩKn+10ₖ _)) +-- -- (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) +-- -- (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (sym (merid north)))) +-- -- open import Cubical.Foundations.Path +-- -- jₗ-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → coHomFun _ (jₗ n f g) (α∨ n f g) +-- -- ≡ Hopfα n f +-- -- jₗ-α n f g = +-- -- cong (coHomFun _ (jₗ n f g)) (α-∨-lem n f g) +-- -- ∙ cong ∣_∣₂ (funExt (HopfInvariantPushElim n (fst f) +-- -- (isOfHLevelPath ((suc (suc (suc (suc (n +ℕ n)))))) +-- -- (isOfHLevelPlus' {n = n} (4 +ℕ n) (isOfHLevelTrunc (4 +ℕ n))) _ _) +-- -- refl +-- -- (λ _ → refl) +-- -- λ j → refl)) + +-- -- jₗ-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → coHomFun _ (jₗ n f g) (βₗ n f g) +-- -- ≡ subst (λ m → coHom m (HopfInvariantPush n (fst f))) +-- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- -- (Hopfβ n f) +-- -- jₗ-βₗ n f g = +-- -- cong (coHomFun _ (jₗ n f g)) (βₗ≡ n f g) +-- -- ∙∙ natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (jₗ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- -- ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) +-- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) +-- -- (cong ∣_∣₂ +-- -- (funExt λ { (inl x) → refl +-- -- ; (inr x) → refl +-- -- ; (push a i₁) → refl})) + +-- -- jₗ-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → coHomFun _ (jₗ n f g) (βᵣ n f g) +-- -- ≡ 0ₕ _ +-- -- jₗ-βᵣ n f g = +-- -- cong (coHomFun _ (jₗ n f g)) (βᵣ≡ n f g) +-- -- ∙∙ natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (jₗ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- -- ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) +-- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) +-- -- cool +-- -- where +-- -- cool : coHomFun (suc (suc (suc (suc (n +ℕ n))))) (jₗ n f g) +-- -- ∣ βᵣ'-fun n f g ∣₂ ≡ 0ₕ _ +-- -- cool = cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → refl}) + +-- -- jᵣ-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → coHomFun _ (jᵣ n f g) (α∨ n f g) +-- -- ≡ Hopfα n g +-- -- jᵣ-α n f g = cong (coHomFun _ (jᵣ n f g)) (α-∨-lem n f g) +-- -- ∙ cong ∣_∣₂ (funExt (HopfInvariantPushElim n (fst g) +-- -- (isOfHLevelPath ((suc (suc (suc (suc (n +ℕ n)))))) +-- -- (isOfHLevelPlus' {n = n} (4 +ℕ n) (isOfHLevelTrunc (4 +ℕ n))) _ _) +-- -- refl +-- -- (λ _ → refl) +-- -- (flipSquare (pp-inr n f g)))) + +-- -- jᵣ-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → coHomFun _ (jᵣ n f g) (βₗ n f g) ≡ 0ₕ _ +-- -- jᵣ-βₗ n f g = +-- -- cong (coHomFun _ (jᵣ n f g)) (βₗ≡ n f g) +-- -- ∙∙ natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (jᵣ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- -- ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) +-- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) +-- -- cool +-- -- where +-- -- cool : coHomFun (suc (suc (suc (suc (n +ℕ n))))) (jᵣ n f g) +-- -- ∣ βₗ'-fun n f g ∣₂ ≡ 0ₕ _ +-- -- cool = cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → refl}) + + +-- -- jᵣ-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → coHomFun _ (jᵣ n f g) (βᵣ n f g) +-- -- ≡ subst (λ m → coHom m (HopfInvariantPush n (fst g))) +-- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- -- (Hopfβ n g) +-- -- jᵣ-βᵣ n f g = +-- -- cong (coHomFun _ (jᵣ n f g)) (βᵣ≡ n f g) +-- -- ∙∙ natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (jᵣ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- -- ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst g))) +-- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) +-- -- (cong ∣_∣₂ +-- -- (funExt λ { (inl x) → refl +-- -- ; (inr x) → refl +-- -- ; (push a i₁) → refl})) + +-- -- gen₂ℤ×ℤ : gen₂-by (DirProd ℤGroup ℤGroup) (1 , 0) (0 , 1) +-- -- fst (gen₂ℤ×ℤ (x , y)) = x , y +-- -- snd (gen₂ℤ×ℤ (x , y)) = +-- -- ΣPathP ((cong₂ _+_ ((·Comm 1 x) ∙ cong fst (sym (distrLem 1 0 x))) ((·Comm 0 y) ∙ cong fst (sym (distrLem 0 1 y)))) +-- -- , +Comm y 0 ∙ cong₂ _+_ (·Comm 0 x ∙ cong snd (sym (distrLem 1 0 x))) (·Comm 1 y ∙ cong snd (sym (distrLem 0 1 y)))) +-- -- where +-- -- ℤ×ℤ = DirProd ℤGroup ℤGroup +-- -- _+''_ = GroupStr._·_ (snd ℤ×ℤ) + +-- -- ll3 : (x : ℤ) → - x ≡ 0 - x +-- -- ll3 (pos zero) = refl +-- -- ll3 (pos (suc zero)) = refl +-- -- ll3 (pos (suc (suc n))) = +-- -- cong predℤ (ll3 (pos (suc n))) +-- -- ll3 (negsuc zero) = refl +-- -- ll3 (negsuc (suc n)) = cong sucℤ (ll3 (negsuc n)) + +-- -- distrLem : (x y : ℤ) (z : ℤ) +-- -- → Path (ℤ × ℤ) (z ℤ[ ℤ×ℤ ]· (x , y)) (z · x , z · y) +-- -- distrLem x y (pos zero) = refl +-- -- distrLem x y (pos (suc n)) = +-- -- (cong ((x , y) +''_) (distrLem x y (pos n))) +-- -- distrLem x y (negsuc zero) = ΣPathP (sym (ll3 x) , sym (ll3 y)) +-- -- distrLem x y (negsuc (suc n)) = +-- -- cong₂ _+''_ (ΣPathP (sym (ll3 x) , sym (ll3 y))) +-- -- (distrLem x y (negsuc n)) + +-- -- genH²ⁿC* : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → gen₂-by (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) +-- -- (βₗ n f g) +-- -- (βᵣ n f g) +-- -- genH²ⁿC* n f g = +-- -- Iso-pres-gen₂ (DirProd ℤGroup ℤGroup) +-- -- (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) +-- -- (1 , 0) +-- -- (0 , 1) +-- -- gen₂ℤ×ℤ +-- -- (invGroupIso (H⁴-C* n f g)) + +-- -- private + +-- -- H⁴C* : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Group _ +-- -- H⁴C* n f g = coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) + +-- -- X : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → ℤ +-- -- X n f g = (genH²ⁿC* n f g) (α∨ n f g ⌣ α∨ n f g) .fst .fst +-- -- Y : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → ℤ +-- -- Y n f g = (genH²ⁿC* n f g) (α∨ n f g ⌣ α∨ n f g) .fst .snd + +-- -- α∨≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → α∨ n f g ⌣ α∨ n f g ≡ ((X n f g) ℤ[ H⁴C* n f g ]· βₗ n f g) +-- -- +ₕ ((Y n f g) ℤ[ H⁴C* n f g ]· βᵣ n f g) +-- -- α∨≡ n f g = (genH²ⁿC* n f g) (α∨ n f g ⌣ α∨ n f g) .snd + + +-- -- eq₁ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → (Hopfα n (·Π f g)) ⌣ (Hopfα n (·Π f g)) +-- -- ≡ ((X n f g + Y n f g) ℤ[ coHomGr _ _ ]· (β+ n f g)) +-- -- eq₁ n f g = +-- -- cong (λ x → x ⌣ x) (sym (q-α n f g) ∙ cong (coHomFun (suc (suc n)) (q n f g)) (α-∨-lem n f g)) +-- -- ∙ cong ((coHomFun _) (q _ f g)) (cong (λ x → x ⌣ x) (sym (α-∨-lem n f g)) +-- -- ∙ α∨≡ n f g) +-- -- ∙ IsGroupHom.pres· (coHomMorph _ (q n f g) .snd) _ _ +-- -- ∙ cong₂ _+ₕ_ (hompres· (coHomMorph _ (q n f g)) (βₗ n f g) (X n f g) +-- -- ∙ cong (λ z → (X n f g) ℤ[ coHomGr _ _ ]· z) +-- -- (q-βₗ n f g)) +-- -- (hompres· (coHomMorph _ (q n f g)) (βᵣ n f g) (Y n f g) +-- -- ∙ cong (λ z → (Y n f g) ℤ[ coHomGr _ _ ]· z) +-- -- (q-βᵣ n f g)) +-- -- ∙ sym (distrℤ· (coHomGr _ _) (β+ n f g) (X n f g) (Y n f g)) + +-- -- coHomFun⌣ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) +-- -- → (n : ℕ) (x y : coHom n B) +-- -- → coHomFun _ f (x ⌣ y) ≡ coHomFun _ f x ⌣ coHomFun _ f y +-- -- coHomFun⌣ f n = sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl + +-- -- eq₂ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → Hopfα n f ⌣ Hopfα n f +-- -- ≡ (X n f g ℤ[ coHomGr _ _ ]· subst (λ m → coHom m (HopfInvariantPush n (fst f))) +-- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- -- (Hopfβ n f)) +-- -- eq₂ n f g = +-- -- cong (λ x → x ⌣ x) (sym (jₗ-α n f g)) +-- -- ∙∙ sym (coHomFun⌣ (jₗ n f g) _ _ _) +-- -- ∙∙ cong (coHomFun _ (jₗ n f g)) (α∨≡ n f g) +-- -- ∙∙ IsGroupHom.pres· (snd (coHomMorph _ (jₗ n f g))) _ _ +-- -- ∙∙ cong₂ _+ₕ_ (hompres· (coHomMorph _ (jₗ n f g)) (βₗ n f g) (X n f g)) +-- -- (hompres· (coHomMorph _ (jₗ n f g)) (βᵣ n f g) (Y n f g) +-- -- ∙ cong (λ z → (Y n f g) ℤ[ coHomGr _ _ ]· z) +-- -- (jₗ-βᵣ n f g)) +-- -- ∙∙ cong₂ _+ₕ_ refl (rUnitℤ· (coHomGr _ _) (Y n f g)) +-- -- ∙∙ (rUnitₕ _ _ +-- -- ∙ cong (X n f g ℤ[ coHomGr _ _ ]·_) (jₗ-βₗ n f g)) + +-- -- eq₃ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- -- → Hopfα n g ⌣ Hopfα n g +-- -- ≡ (Y n f g ℤ[ coHomGr _ _ ]· subst (λ m → coHom m (HopfInvariantPush n (fst f))) +-- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- -- (Hopfβ n g)) +-- -- eq₃ n f g = +-- -- cong (λ x → x ⌣ x) (sym (jᵣ-α n f g)) +-- -- ∙∙ sym (coHomFun⌣ (jᵣ n f g) _ _ _) +-- -- ∙∙ cong (coHomFun _ (jᵣ n f g)) (α∨≡ n f g) +-- -- ∙∙ IsGroupHom.pres· (snd (coHomMorph _ (jᵣ n f g))) _ _ +-- -- ∙∙ cong₂ _+ₕ_ (hompres· (coHomMorph _ (jᵣ n f g)) (βₗ n f g) (X n f g) +-- -- ∙ cong (λ z → (X n f g) ℤ[ coHomGr _ _ ]· z) +-- -- (jᵣ-βₗ n f g)) +-- -- (hompres· (coHomMorph _ (jᵣ n f g)) (βᵣ n f g) (Y n f g)) +-- -- ∙∙ cong₂ _+ₕ_ (rUnitℤ· (coHomGr _ _) (X n f g)) refl +-- -- ∙∙ ((lUnitₕ _ _) ∙ cong (Y n f g ℤ[ coHomGr _ _ ]·_) (jᵣ-βᵣ n f g)) + +-- -- eq₂-eq₂ : (n : ℕ) (f g : S₊∙ (suc (suc (suc (n +ℕ n)))) →∙ S₊∙ (suc (suc n))) +-- -- → HopfFunction n (·Π f g) ≡ HopfFunction n f + HopfFunction n g +-- -- eq₂-eq₂ n f g = +-- -- mainL +-- -- ∙ sym (cong₂ _+_ (help n f g) (helpg n f g)) +-- -- where +-- -- transpLem : ∀ {ℓ} {G : ℕ → Group ℓ} +-- -- → (n m : ℕ) +-- -- → (x : ℤ) +-- -- → (p : m ≡ n) +-- -- → (g : fst (G n)) +-- -- → subst (fst ∘ G) p (x ℤ[ G m ]· subst (fst ∘ G) (sym p) g) ≡ (x ℤ[ G n ]· g) +-- -- transpLem {G = G} n m x = +-- -- J (λ n p → (g : fst (G n)) → subst (fst ∘ G) p (x ℤ[ G m ]· subst (fst ∘ G) (sym p) g) +-- -- ≡ (x ℤ[ G n ]· g)) +-- -- λ g → transportRefl _ ∙ cong (x ℤ[ G m ]·_) (transportRefl _) + +-- -- mainL : HopfFunction n (·Π f g) ≡ X n f g + Y n f g +-- -- mainL = cong (Iso.fun (fst (Hopfβ-Iso n (·Π f g)))) +-- -- (cong (subst (λ x → coHom x (HopfInvariantPush n (fst (·Π f g)))) +-- -- λ i₁ → suc (suc (suc (+-suc n n i₁)))) +-- -- (eq₁ n f g)) +-- -- ∙∙ cong (Iso.fun (fst (Hopfβ-Iso n (·Π f g)))) +-- -- (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst (·Π f g)))} _ _ +-- -- (X n f g + Y n f g) (λ i₁ → suc (suc (suc (+-suc n n i₁)))) +-- -- (Iso.inv (fst (Hopfβ-Iso n (·Π f g))) 1)) +-- -- ∙∙ hompres· (_ , snd (Hopfβ-Iso n (·Π f g))) (Iso.inv (fst (Hopfβ-Iso n (·Π f g))) 1) +-- -- (X n f g + Y n f g) +-- -- ∙∙ cong ((X n f g + Y n f g) ℤ[ ℤ , ℤGroup .snd ]·_) +-- -- (Iso.rightInv ((fst (Hopfβ-Iso n (·Π f g)))) 1) +-- -- ∙∙ rUnitℤ·ℤ (X n f g + Y n f g) + + +-- -- help : (n : ℕ) (f g : _) → HopfFunction n f ≡ X n f g +-- -- help n f g = cong (Iso.fun (fst (Hopfβ-Iso n f))) +-- -- (cong (subst (λ x → coHom x (HopfInvariantPush n (fst f))) +-- -- (λ i₁ → suc (suc (suc (+-suc n n i₁))))) (eq₂ n f g)) +-- -- ∙ cong (Iso.fun (fst (Hopfβ-Iso n f))) +-- -- (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst f))} _ _ +-- -- (X n f g) +-- -- ((cong (suc ∘ suc ∘ suc) (+-suc n n))) +-- -- (Hopfβ n f)) +-- -- ∙ hompres· (_ , snd (Hopfβ-Iso n f)) (Hopfβ n f) (X n f g) +-- -- ∙ cong (X n f g ℤ[ ℤ , ℤGroup .snd ]·_) (d-β n f) +-- -- ∙ rUnitℤ·ℤ (X n f g) + +-- -- helpg : (n : ℕ) (f g : _) → HopfFunction n g ≡ Y n f g +-- -- helpg n f g = +-- -- cong (Iso.fun (fst (Hopfβ-Iso n g))) +-- -- (cong (subst (λ x → coHom x (HopfInvariantPush n (fst g))) +-- -- (λ i₁ → suc (suc (suc (+-suc n n i₁))))) +-- -- (eq₃ n f g)) +-- -- ∙∙ cong (Iso.fun (fst (Hopfβ-Iso n g))) +-- -- (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst g))} _ _ +-- -- (Y n f g) +-- -- ((cong (suc ∘ suc ∘ suc) (+-suc n n))) +-- -- (Hopfβ n g)) +-- -- ∙∙ hompres· (_ , snd (Hopfβ-Iso n g)) (Hopfβ n g) (Y n f g) +-- -- ∙∙ cong (Y n f g ℤ[ ℤ , ℤGroup .snd ]·_) (d-β n g) +-- -- ∙∙ rUnitℤ·ℤ (Y n f g) + + +-- -- {- +-- -- q-α : coHomFun _ q α∨ ≡ Hopfα n (·Π f g) +-- -- q-α with n +-- -- ... | zero = ? +-- -- ... | (suc n) = ? +-- -- where +-- -- helpMe! : Iso.fun (fst (Hopfα-Iso _ (·Π f g))) (Hopfα n (·Π f g)) +-- -- ≡ Iso.fun (fst (Hopfα-Iso _ (·Π f g))) (coHomFun _ q α∨) +-- -- helpMe! = {!!} + +-- -- q-βₗ : coHomFun _ q βₗ ≡ {!(Hopfβ n ?)!} +-- -- q-βₗ = {!!} + +-- -- q-βᵣ : coHomFun _ q βᵣ ≡ {!!} +-- -- q-βᵣ = {!!} -} + +-- -- -- helpi : {!(f : ? →∙ ?) (HopfInvariantPush n (fst f) → ?)!} +-- -- -- helpi = {!!} + +-- -- -- lal : {!!} +-- -- -- lal = {!!} + +-- -- -- isHomHopf : ∀ n → (x y : ∥ S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n) ∥₂) +-- -- -- → Hopfie (x ·π y) ≡ Hopfie x + Hopfie y +-- -- -- isHomHopf n = +-- -- -- sElim2 (λ _ _ → isOfHLevelPath 2 isSetℤ _ _) +-- -- -- λ f g → (λ i → Iso.fun (fst (Hopfβ-Iso n (·Π f g))) (subber n (·Π f g) (⌣-α n (·Π f g)))) +-- -- -- ∙ cong (Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) +-- -- -- {!!} +-- -- -- ∙∙ IsGroupHom.pres· (((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) .snd) +-- -- -- (invEq (fst (GroupEquiv1 n f)) (subber n f (⌣-α n f))) +-- -- -- (invEq (fst (GroupEquiv1 n g)) (subber n g (⌣-α n g))) +-- -- -- ∙∙ λ i → Iso.fun (fst (Hopfβ-Iso n f)) (subber n f (⌣-α n f)) +-- -- -- + Iso.fun (fst (Hopfβ-Iso n g)) (subber n g (⌣-α n g)) +-- -- -- where +-- -- -- cool : (f g : _) → Iso.fun (fst (Hopfβ-Iso n (·Π f g))) (subber n (·Π f g) (⌣-α n f)) +-- -- -- ≡ Iso.fun (fst (Hopfβ-Iso n f)) (subber n f (⌣-α n f)) +-- -- -- cool f g = {!!} diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index 97ef86156..e2392008f 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -52,893 +52,603 @@ private variable ℓ ℓ' : Level -cup-pure : {!!} -cup-pure = {!!} -open import Cubical.HITs.S1 renaming (_·_ to _·'_) - -S₊3 : S₊ 3 → S₊ 3 → S₊ 3 -S₊3 north y = north -S₊3 south y = north -S₊3 (merid north i) north = north -S₊3 (merid south i) north = north -S₊3 (merid (merid base i) j) north = north -S₊3 (merid (merid (loop i₁) i) j) north = north -S₊3 (merid north i) south = north -S₊3 (merid south i) south = north -S₊3 (merid (merid base j) i) south = north -S₊3 (merid (merid (loop k) i) j) south = {!!} -S₊3 (merid north i) (merid north j) = north -S₊3 (merid north i) (merid south j) = north -S₊3 (merid north i) (merid (merid a i₁) j) = {!!} -S₊3 (merid south i) (merid north j) = north -S₊3 (merid south i) (merid south j) = {!north!} -S₊3 (merid south i) (merid (merid a i₁) j) = {!!} -S₊3 (merid (merid a i₁) i) (merid north j) = {!!} -S₊3 (merid (merid a i₁) i) (merid south j) = {!!} -S₊3 (merid (merid a i) j) (merid (merid a₁ k) l) = {!!} - --- open import Cubical.HITs.Join --- rofl : join S¹ S¹ → Type _ --- rofl (inl x) = {!!} --- rofl (inr x) = {!!} --- rofl (push a b i) = {!!} --- where --- zz : (x : S₊ 2) → hLevelTrunc 5 (S₊ 3) → hLevelTrunc 5 (S₊ 3) --- zz x = {!trElim ? ?!} - --- r4 : hLevelTrunc 7 {!!} → {!!} → {!!} --- r4 = {!!} - --- r3 : hLevelTrunc 6 (S₊ 2) → S₊ 4 → hLevelTrunc 6 (S₊ 2) --- r3 = trRec (isOfHLevelΠ 6 (λ _ → isOfHLevelTrunc 6)) --- (wedgeconFun 1 3 (λ _ _ → isOfHLevelTrunc 6) --- {!!} ∣_∣ --- {!!}) - --- t : hLevelTrunc 6 (S₊ 3) → hLevelTrunc 5 (S₊ 2) → hLevelTrunc 6 (S₊ 3) --- t = trRec {!!} {!!} --- where --- pp : S₊ 3 → hLevelTrunc 5 (S₊ 2) → hLevelTrunc 6 (S₊ 3) --- pp north = λ _ → ∣ north ∣ --- pp south = λ _ → ∣ north ∣ --- pp (merid a i) x = {!!} --- where --- h : (x : hLevelTrunc 5 (S₊ 2)) → Path (hLevelTrunc 6 (S₊ 3)) ∣ north ∣ ∣ north ∣ --- h = trElim (λ _ → isOfHLevelTrunc 6 _ _) --- λ b → {!!} - --- S³∥ : Type _ --- S³∥ = hLevelTrunc 6 (S₊ 3) - --- _+3_ : S³∥ → S³∥ → S³∥ --- _+3_ = elim2 (λ _ _ → isOfHLevelTrunc 6) (wedgeconFun _ _ (λ _ _ → isOfHLevelTrunc 6) --- ∣_∣ --- ∣_∣ --- refl) - --- 0₃ : hLevelTrunc 6 (S₊ 3) --- 0₃ = ∣ north ∣ - --- rUnit3 : ∀ x → x +3 0₃ ≡ x --- rUnit3 = trElim (λ _ → isOfHLevelPath 6 (isOfHLevelTrunc 6) _ _) --- (wedgeconRight 2 2 {λ _ _ → S³∥} --- (λ _ _ → isOfHLevelTrunc 6) ∣_∣ₕ ∣_∣ₕ refl) - --- lUnit3 : ∀ x → 0₃ +3 x ≡ x --- lUnit3 = trElim (λ _ → isOfHLevelPath 6 (isOfHLevelTrunc 6) _ _) --- λ _ → refl - --- min3 : hLevelTrunc 6 (S₊ 3) → hLevelTrunc 6 (S₊ 3) --- min3 = trMap m --- where --- m : S₊ 3 → S₊ 3 --- m north = north --- m south = north --- m (merid a i) = (merid a ∙ sym (merid north)) i - --- comm3 : ∀ x y → x +3 y ≡ y +3 x --- comm3 = elim2 (λ _ _ → isOfHLevelPath 6 (isOfHLevelTrunc 6) _ _) --- (wedgeconFun 2 2 --- (λ _ _ → isOfHLevelPath 6 (isOfHLevelTrunc 6) _ _) --- (sym ∘ rUnit3 ∘ ∣_∣ₕ) --- (rUnit3 ∘ ∣_∣ₕ) --- refl) - --- rCancel3 : (x : hLevelTrunc 6 (S₊ 3)) → x +3 min3 x ≡ 0₃ --- rCancel3 = --- trElim {!!} --- λ { north → refl --- ; south → refl --- ; (merid a i) → {!!}} --- where --- p+ : ∀ {x : S³∥} (p : 0₃ ≡ x) → cong₂ _+3_ p (sym p) ≡ lUnit3 x ∙ sym (rUnit3 x) --- p+ = J (λ x p → cong₂ _+3_ p (sym p) ≡ lUnit3 x ∙ sym (rUnit3 x)) (rUnit refl) - --- p+merid : ∀ (x : S₊ 2) → cong₂ _+3_ (cong ∣_∣ₕ (merid x)) (sym (cong ∣_∣ₕ (merid x))) --- ≡ {!!} --- p+merid = {!!} - --- π-iso : Iso (∥ (S₊∙ 4 →∙ (S³∥ , ∣ north ∣)) ∥₂) (∥ (S₊ 4 → S³∥) ∥₂) --- fun π-iso = map fst --- inv' π-iso = map λ f → (λ x → (f x) +3 min3 (f north)) , rCancel3 (f north) --- rightInv π-iso = sElim {!!} {!!} --- leftInv π-iso = {!!} - --- p2 : Iso ∥ (S₊ 3 → typ (Ω (S³∥ , ∣ north ∣))) ∥₂ --- ∥ (S₊ 2 → typ ((Ω^ 2) (S³∥ , ∣ north ∣))) ∥₂ --- p2 = {!!} - --- data RP2 : Type ℓ-zero where --- b : RP2 --- l1 l2 : b ≡ b --- sq : l1 ≡ sym l2 - --- rar : S₊ 2 → RP2 --- rar north = b --- rar south = b --- rar (merid base i) = (l1 ∙ l2) i --- rar (merid (loop j) i) = --- hcomp (λ k → λ { (i = i0) → b --- ; (i = i1) → b --- ; (j = i0) → (sym (lCancel l2) ∙ cong (_∙ l2) (sym sq)) k i --- ; (j = i1) → (sym (lCancel l2) ∙ cong (_∙ l2) (sym sq)) k i}) --- b - - - --- jmap : join S¹ S¹ → S₊ 2 --- jmap (inl x) = north --- jmap (inr x) = south --- jmap (push a b₁ i) = merid (a ·' b₁) i - --- gl : Susp (join S¹ S¹) → S₊ 3 --- gl north = north --- gl south = south --- gl (merid a i) = merid (jmap a) i - --- _+j_ : (Susp (join S¹ S¹) , north) →∙ S₊∙ 3 --- → (Susp (join S¹ S¹) , north) →∙ S₊∙ 3 --- → (Susp (join S¹ S¹) , north) →∙ S₊∙ 3 --- fst (f +j g) north = north --- fst (f +j g) south = north --- fst (f +j g) (merid a i) = --- ((sym (snd f) ∙∙ cong (fst f) (merid a ∙ sym (merid (inl base))) ∙∙ snd f) --- ∙ ((sym (snd g) ∙∙ cong (fst g) (merid a ∙ sym (merid (inl base))) ∙∙ snd g))) i --- snd (f +j g) = refl - --- ss123 : {!!} --- ss123 = {!!} - --- thegen : {!(Susp (join S¹ S¹) , north) →∙ S₊∙ 3!} --- thegen = {!!} - --- thegen2 : (Susp (join S¹ S¹) , north) →∙ (S³∥ , ∣ north ∣) --- fst thegen2 north = ∣ north ∣ --- fst thegen2 south = ∣ north ∣ --- fst thegen2 (merid a i) = --- ∣ (cong gl (merid a ∙ sym (merid (inl base))) --- ∙ cong gl (merid a ∙ sym (merid (inl base)))) i ∣ --- snd thegen2 = refl - --- s : join S¹ S¹ → {!!} --- s = {!!} - --- rs : (x : _) → fst thegen2 x ≡ ∣ north ∣ --- rs north = refl --- rs south = refl --- rs (merid a i) j = pp a j i --- where --- open import Cubical.HITs.Wedge --- wedge→gpd : ∀ {ℓ} (P : (S₊∙ 1) ⋁ S₊∙ 1 → Type ℓ) --- → ((s : (S₊∙ 1) ⋁ S₊∙ 1) → isGroupoid (P s)) --- → {!!} --- → {!!} --- wedge→gpd = {!!} - --- pp3 : (a : join S¹ S¹) → --- cong (∣_∣ₕ {n = 6}) (cong gl (merid a ∙ sym (merid (inl base)))) ≡ --- sym (cong (∣_∣ₕ {n = 6}) (cong gl (merid a ∙ sym (merid (inl base))))) -- 4 --- pp3 (inl x) = {!!} --- pp3 (inr x) = {!!} --- pp3 (push a b₁ i) = {!!} -- 3 - --- pp4 : (a : join S¹ S¹) → cong (∣_∣ₕ {n = 6}) (merid (jmap a) ∙ sym (merid north)) --- ≡ sym (cong (∣_∣ₕ {n = 6}) (merid (jmap a) ∙ sym (merid north))) --- -- 4-type --- pp4 (inl x) = cong (cong (∣_∣ₕ {n = 6})) (rCancel (merid north)) --- ∙ cong sym (cong (cong (∣_∣ₕ {n = 6})) (sym (rCancel (merid north)))) --- pp4 (inr x) = (cong (cong (∣_∣ₕ {n = 6})) λ i → merid (merid base (~ i)) ∙ sym (merid north)) --- ∙∙ (cong (cong (∣_∣ₕ {n = 6})) (rCancel (merid north)) --- ∙ cong sym (cong (cong (∣_∣ₕ {n = 6})) (sym (rCancel (merid north))))) --- ∙∙ cong sym (cong (cong (∣_∣ₕ {n = 6})) --- λ i → merid (merid base i) ∙ sym (merid north)) --- -- 3-type --- pp4 (push a c i) j k = {!!} --- where --- c' : ∥ cong merid (merid a) ≡ cong merid (merid base) ∥ 1 --- c' = isConnectedPath _ (isConnectedPath 2 (isConnectedPath 3 (sphereConnected 3) _ _) _ _) _ _ .fst - --- pr : (a c : S¹) → cong merid (merid (a ·' c)) ≡ cong merid (merid a ∙ sym (merid c)) --- ∙ (cong merid (merid base)) --- pr base base = {!!} --- pr base (loop i) = {!!} --- pr (loop i) base = {!cong ∣_∣ₕ (merid (cong jmap (push a c)) ∙ sym (merid north))!} --- pr (loop i) (loop j) = {!!} - --- hej : Cube (λ j → cong ∣_∣ₕ (merid (merid (a ·' c) j) ∙ sym (merid north))) -- i j k --- (λ j → cong ∣_∣ₕ (merid (merid a j) ∙ sym (merid (merid c j)))) --- {!!} --- {!!} --- {!!} --- {!!} --- hej = {!!} --- pp : (a : join S¹ S¹) → --- cong (∣_∣ₕ {n = 6}) ((cong gl (merid a ∙ sym (merid (inl base))) --- ∙ cong gl (merid a ∙ sym (merid (inl base))))) ≡ refl --- pp a = {!merid a!} ∙ {!!} --- +gen : ℤ → (Susp (join S¹ S¹) , north) →∙ S₊∙ 3 --- +gen (pos zero) = (λ x → north) , refl --- fst (+gen (pos (suc n))) north = north --- fst (+gen (pos (suc n))) south = north --- fst (+gen (pos (suc n))) (merid a i) = --- {!cong (fst (+gen (pos n))) (merid a ∙ sym (merid (inl base)))!} --- snd (+gen (pos (suc n))) = {!!} --- +gen (negsuc zero) = {!!} --- +gen (negsuc (suc n)) = {!!} - --- _+∙_ : {!!} -- {n m : ℕ} → (S₊∙ n → S₊∙ m) → (S₊∙ n → S₊∙ m) → (S₊∙ n → S₊∙ m) --- _+∙_ = {!!} - --- sn : Iso ∥ (typ ((Ω^ 4) (S₊∙ 3))) ∥₂ --- (∥ (Σ[ x ∈ typ ((Ω^ 3) (S₊∙ 3)) ] --- x ≡ sym x) ∥₂) --- fun sn = map λ p → (λ i j k → p (i ∨ j) (j ∨ k) (k ∨ i) (~ k ∨ ~ i ∨ ~ j)) --- , {!!} --- inv' sn = {!!} --- rightInv sn = {!!} --- leftInv sn = {!!} - --- gaza : ∥ {!!} ≃ {!!} ∥₂ --- gaza = {!!} - --- Σ- : ∀ {ℓ} {A : Type ℓ} → Susp A → Susp A --- Σ- north = south --- Σ- south = north --- Σ- (merid a i) = merid a (~ i) - --- rr : Iso (∥ (Σ[ f ∈ S₊∙ 3 →∙ (S³∥ , ∣ north ∣) ] ((x : S₊ 3) → fst f x ≡ fst f x)) ∥₂) --- (∥ (Σ[ f ∈ S₊∙ 3 →∙ coHomK-ptd 3 ] ((x : S₊ 3) → fst f x +ₖ fst f x ≡ 0ₖ 3)) ∥₂) --- fun rr = {!sMap ?!} --- inv' rr = {!!} --- rightInv rr = {!!} --- leftInv rr = {!!} - --- open import Cubical.HITs.Pushout --- RP3 : Type ℓ-zero --- RP3 = Pushout rar λ _ → tt - --- RP4 : S₊ 3 → RP3 --- RP4 north = inl b --- RP4 south = inr tt --- RP4 (merid north i) = push north i --- RP4 (merid south i) = push south i --- RP4 (merid (merid base j) i) = push {!rar!} i --- RP4 (merid (merid (loop k) j) i) = {!!} - --- ll : hLevelTrunc 5 (S₊ 2) → hLevelTrunc 5 (S₊ 2) → {!!} --- ll = {!!} - --- ss : Iso ∥ (S₊ 1 → (S₊∙ 1 →∙ ((Ω^ 2) (S³∥ , ∣ north ∣)))) ∥₂ --- ∥ ((S₊ 1 → (S₊ 1 → typ ((Ω^ 2) (S³∥ , ∣ north ∣))))) ∥₂ --- fun ss = map λ f → λ x y → f x .fst y --- inv' ss = map λ f x → {!f!} --- rightInv ss = {!!} --- leftInv ss = {!!} - --- p3 : Iso ∥ (S₊ 2 → coHomK-ptd 1 →∙ (coHomK-ptd 1 →∙ S³∥ , ∣ north ∣ ∙)) ∥₂ --- ∥ {!!} ∥₂ --- p3 = {!S₊∙ 3 !} - --- cong2 : {!!} --- cong2 = {!!} - --- -- ------------------- Connectedness --------------------- --- -- is2ConnectedKn : (n : ℕ) → isConnected 2 (coHomK (suc n)) --- -- is2ConnectedKn zero = ∣ ∣ base ∣ ∣ --- -- , trElim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) --- -- (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 (isOfHLevelTrunc 2)) _ _) --- -- (toPropElim (λ _ → isOfHLevelTrunc 2 _ _) refl)) --- -- is2ConnectedKn (suc n) = ∣ ∣ north ∣ ∣ --- -- , trElim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) --- -- (trElim (λ _ → isProp→isOfHLevelSuc (3 + n) (isOfHLevelTrunc 2 _ _)) --- -- (suspToPropElim (ptSn (suc n)) (λ _ → isOfHLevelTrunc 2 _ _) refl)) - --- -- isConnectedKn : (n : ℕ) → isConnected (2 + n) (coHomK (suc n)) --- -- isConnectedKn n = isOfHLevelRetractFromIso 0 (invIso (truncOfTruncIso (2 + n) 1)) (sphereConnected (suc n)) - --- -- -- direct proof of connectedness of ΩKₙ₊₁ not relying on the equivalence ∥ a ≡ b ∥ₙ ≃ (∣ a ∣ₙ₊₁ ≡ ∣ b ∣ₙ₊₁) --- -- isConnectedPathKn : (n : ℕ) (x y : (coHomK (suc n))) → isConnected (suc n) (x ≡ y) --- -- isConnectedPathKn n = --- -- trElim (λ _ → isProp→isOfHLevelSuc (2 + n) (isPropΠ λ _ → isPropIsContr)) --- -- (sphereElim _ (λ _ → isProp→isOfHLevelSuc n (isPropΠ λ _ → isPropIsContr)) --- -- λ y → isContrRetractOfConstFun --- -- {B = (hLevelTrunc (suc n) (ptSn (suc n) ≡ ptSn (suc n)))} ∣ refl ∣ --- -- (fun⁻ n y --- -- , trElim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) --- -- (J (λ y p → fun⁻ n y _ ≡ _) (funExt⁻ (fun⁻Id n) ∣ refl ∣)))) --- -- where --- -- fun⁻ : (n : ℕ) → (y : coHomK (suc n)) → --- -- hLevelTrunc (suc n) (ptSn (suc n) ≡ ptSn (suc n)) --- -- → hLevelTrunc (suc n) (∣ ptSn (suc n) ∣ ≡ y) --- -- fun⁻ n = --- -- trElim (λ _ → isOfHLevelΠ (3 + n) λ _ → isOfHLevelSuc (2 + n) (isOfHLevelSuc (suc n) (isOfHLevelTrunc (suc n)))) --- -- (sphereElim n (λ _ → isOfHLevelΠ (suc n) λ _ → isOfHLevelTrunc (suc n)) λ _ → ∣ refl ∣) - --- -- fun⁻Id : (n : ℕ) → fun⁻ n ∣ ptSn (suc n) ∣ ≡ λ _ → ∣ refl ∣ --- -- fun⁻Id zero = refl --- -- fun⁻Id (suc n) = refl - --- -- ------------------- --- -- -- Induction principles for cohomology groups (n ≥ 1) --- -- -- If we want to show a proposition about some x : Hⁿ(A), it suffices to show it under the --- -- -- assumption that x = ∣ f ∣₂ for some f : A → Kₙ and that f is pointed - --- -- coHomPointedElim : {A : Type ℓ} (n : ℕ) (a : A) {B : coHom (suc n) A → Type ℓ'} --- -- → ((x : coHom (suc n) A) → isProp (B x)) --- -- → ((f : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → B ∣ f ∣₂) --- -- → (x : coHom (suc n) A) → B x --- -- coHomPointedElim {ℓ' = ℓ'} {A = A} n a isprop indp = --- -- sElim (λ _ → isOfHLevelSuc 1 (isprop _)) --- -- λ f → helper n isprop indp f (f a) refl --- -- where --- -- helper : (n : ℕ) {B : coHom (suc n) A → Type ℓ'} --- -- → ((x : coHom (suc n) A) → isProp (B x)) --- -- → ((f : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → B ∣ f ∣₂) --- -- → (f : A → coHomK (suc n)) --- -- → (x : coHomK (suc n)) --- -- → f a ≡ x → B ∣ f ∣₂ --- -- -- pattern matching a bit extra to avoid isOfHLevelPlus' --- -- helper zero isprop ind f = --- -- trElim (λ _ → isOfHLevelPlus {n = 1} 2 (isPropΠ λ _ → isprop _)) --- -- (toPropElim (λ _ → isPropΠ λ _ → isprop _) (ind f)) --- -- helper (suc zero) isprop ind f = --- -- trElim (λ _ → isOfHLevelPlus {n = 1} 3 (isPropΠ λ _ → isprop _)) --- -- (suspToPropElim base (λ _ → isPropΠ λ _ → isprop _) (ind f)) --- -- helper (suc (suc zero)) isprop ind f = --- -- trElim (λ _ → isOfHLevelPlus {n = 1} 4 (isPropΠ λ _ → isprop _)) --- -- (suspToPropElim north (λ _ → isPropΠ λ _ → isprop _) (ind f)) --- -- helper (suc (suc (suc n))) isprop ind f = --- -- trElim (λ _ → isOfHLevelPlus' {n = 5 + n} 1 (isPropΠ λ _ → isprop _)) --- -- (suspToPropElim north (λ _ → isPropΠ λ _ → isprop _) (ind f)) - - --- -- coHomPointedElim2 : {A : Type ℓ} (n : ℕ) (a : A) {B : coHom (suc n) A → coHom (suc n) A → Type ℓ'} --- -- → ((x y : coHom (suc n) A) → isProp (B x y)) --- -- → ((f g : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → g a ≡ coHom-pt (suc n) → B ∣ f ∣₂ ∣ g ∣₂) --- -- → (x y : coHom (suc n) A) → B x y --- -- coHomPointedElim2 {ℓ' = ℓ'} {A = A} n a isprop indp = sElim2 (λ _ _ → isOfHLevelSuc 1 (isprop _ _)) --- -- λ f g → helper n a isprop indp f g (f a) (g a) refl refl --- -- where --- -- helper : (n : ℕ) (a : A) {B : coHom (suc n) A → coHom (suc n) A → Type ℓ'} --- -- → ((x y : coHom (suc n) A) → isProp (B x y)) --- -- → ((f g : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → g a ≡ coHom-pt (suc n) → B ∣ f ∣₂ ∣ g ∣₂) --- -- → (f g : A → coHomK (suc n)) --- -- → (x y : coHomK (suc n)) --- -- → f a ≡ x → g a ≡ y --- -- → B ∣ f ∣₂ ∣ g ∣₂ --- -- helper zero a isprop indp f g = --- -- elim2 (λ _ _ → isOfHLevelPlus {n = 1} 2 (isPropΠ2 λ _ _ → isprop _ _)) --- -- (toPropElim2 (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) --- -- helper (suc zero) a isprop indp f g = --- -- elim2 (λ _ _ → isOfHLevelPlus {n = 1} 3 (isPropΠ2 λ _ _ → isprop _ _)) --- -- (suspToPropElim2 base (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) --- -- helper (suc (suc zero)) a isprop indp f g = --- -- elim2 (λ _ _ → isOfHLevelPlus {n = 1} 4 (isPropΠ2 λ _ _ → isprop _ _)) --- -- (suspToPropElim2 north (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) --- -- helper (suc (suc (suc n))) a isprop indp f g = --- -- elim2 (λ _ _ → isOfHLevelPlus' {n = 5 + n} 1 (isPropΠ2 λ _ _ → isprop _ _)) --- -- (suspToPropElim2 north (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) - --- -- coHomK-elim : ∀ {ℓ} (n : ℕ) {B : coHomK (suc n) → Type ℓ} --- -- → ((x : _) → isOfHLevel (suc n) (B x)) --- -- → B (0ₖ (suc n)) --- -- → (x : _) → B x --- -- coHomK-elim n {B = B } hlev b = --- -- trElim (λ _ → isOfHLevelPlus {n = (suc n)} 2 (hlev _)) --- -- (sphereElim _ (hlev ∘ ∣_∣) b) - --- -- {- Equivalence between cohomology of A and reduced cohomology of (A + 1) -} --- -- coHomRed+1Equiv : (n : ℕ) → --- -- (A : Type ℓ) → --- -- (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) --- -- coHomRed+1Equiv zero A i = ∥ helpLemma {C = (ℤ , pos 0)} i ∥₂ --- -- module coHomRed+1 where --- -- helpLemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →∙ C))) --- -- helpLemma {C = C} = isoToPath (iso map1 --- -- map2 --- -- (λ b → linvPf b) --- -- (λ _ → refl)) --- -- where --- -- map1 : (A → typ C) → ((((A ⊎ Unit) , inr (tt)) →∙ C)) --- -- map1 f = map1' , refl --- -- module helpmap where --- -- map1' : A ⊎ Unit → fst C --- -- map1' (inl x) = f x --- -- map1' (inr x) = pt C - --- -- map2 : ((((A ⊎ Unit) , inr (tt)) →∙ C)) → (A → typ C) --- -- map2 (g , pf) x = g (inl x) - --- -- linvPf : (b :((((A ⊎ Unit) , inr (tt)) →∙ C))) → map1 (map2 b) ≡ b --- -- linvPf (f , snd) i = (λ x → helper x i) , λ j → snd ((~ i) ∨ j) --- -- where --- -- helper : (x : A ⊎ Unit) → ((helpmap.map1') (map2 (f , snd)) x) ≡ f x --- -- helper (inl x) = refl --- -- helper (inr tt) = sym snd --- -- coHomRed+1Equiv (suc zero) A i = ∥ coHomRed+1.helpLemma A i {C = (coHomK 1 , ∣ base ∣)} i ∥₂ --- -- coHomRed+1Equiv (suc (suc n)) A i = ∥ coHomRed+1.helpLemma A i {C = (coHomK (2 + n) , ∣ north ∣)} i ∥₂ - --- -- Iso-coHom-coHomRed : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → Iso (coHomRed (suc n) A) (coHom (suc n) (typ A)) --- -- fun (Iso-coHom-coHomRed {A = A , a} n) = map fst --- -- inv' (Iso-coHom-coHomRed {A = A , a} n) = map λ f → (λ x → f x -ₖ f a) , rCancelₖ _ _ --- -- rightInv (Iso-coHom-coHomRed {A = A , a} n) = --- -- sElim (λ _ → isOfHLevelPath 2 § _ _) --- -- λ f → trRec (isProp→isOfHLevelSuc _ (§ _ _)) --- -- (λ p → cong ∣_∣₂ (funExt λ x → cong (λ y → f x +ₖ y) (cong -ₖ_ p ∙ -0ₖ) ∙ rUnitₖ _ (f x))) --- -- (Iso.fun (PathIdTruncIso (suc n)) (isContr→isProp (isConnectedKn n) ∣ f a ∣ ∣ 0ₖ _ ∣)) --- -- leftInv (Iso-coHom-coHomRed {A = A , a} n) = --- -- sElim (λ _ → isOfHLevelPath 2 § _ _) --- -- λ {(f , p) → cong ∣_∣₂ (ΣPathP (((funExt λ x → (cong (λ y → f x -ₖ y) p --- -- ∙∙ cong (λ y → f x +ₖ y) -0ₖ --- -- ∙∙ rUnitₖ _ (f x)) ∙ refl)) --- -- , helper n (f a) (sym p)))} --- -- where --- -- path : (n : ℕ) (x : coHomK (suc n)) (p : 0ₖ _ ≡ x) → _ --- -- path n x p = (cong (λ y → x -ₖ y) (sym p) ∙∙ cong (λ y → x +ₖ y) -0ₖ ∙∙ rUnitₖ _ x) ∙ refl - --- -- helper : (n : ℕ) (x : coHomK (suc n)) (p : 0ₖ _ ≡ x) --- -- → PathP (λ i → path n x p i ≡ 0ₖ _) (rCancelₖ _ x) (sym p) --- -- helper zero x = --- -- J (λ x p → PathP (λ i → path 0 x p i ≡ 0ₖ _) --- -- (rCancelₖ _ x) (sym p)) --- -- λ i j → rUnit (rUnit (λ _ → 0ₖ 1) (~ j)) (~ j) i --- -- helper (suc n) x = --- -- J (λ x p → PathP (λ i → path (suc n) x p i ≡ 0ₖ _) (rCancelₖ _ x) (sym p)) --- -- λ i j → rCancelₖ (suc (suc n)) (0ₖ (suc (suc n))) (~ i ∧ ~ j) - --- -- +∙≡+ : (n : ℕ) {A : Pointed ℓ} (x y : coHomRed (suc n) A) --- -- → Iso.fun (Iso-coHom-coHomRed n) (x +ₕ∙ y) --- -- ≡ Iso.fun (Iso-coHom-coHomRed n) x +ₕ Iso.fun (Iso-coHom-coHomRed n) y - --- -- +∙≡+ zero = sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl --- -- +∙≡+ (suc n) = sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl - --- -- private --- -- homhelp : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) (x y : coHom (suc n) (typ A)) --- -- → Iso.inv (Iso-coHom-coHomRed {A = A} n) (x +ₕ y) --- -- ≡ Iso.inv (Iso-coHom-coHomRed n) x +ₕ∙ Iso.inv (Iso-coHom-coHomRed n) y --- -- homhelp n A = morphLemmas.isMorphInv _+ₕ∙_ _+ₕ_ --- -- (Iso.fun (Iso-coHom-coHomRed n)) (+∙≡+ n) _ --- -- (Iso.rightInv (Iso-coHom-coHomRed n)) (Iso.leftInv (Iso-coHom-coHomRed n)) - --- -- coHomGr≅coHomRedGr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) --- -- → GroupEquiv (coHomRedGrDir (suc n) A) (coHomGr (suc n) (typ A)) --- -- fst (coHomGr≅coHomRedGr n A) = isoToEquiv (Iso-coHom-coHomRed n) --- -- snd (coHomGr≅coHomRedGr n A) = makeIsGroupHom (+∙≡+ n) - --- -- coHomRedGroup : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → AbGroup ℓ --- -- coHomRedGroup zero A = coHomRedGroupDir zero A --- -- coHomRedGroup (suc n) A = --- -- InducedAbGroup (coHomGroup (suc n) (typ A)) --- -- _+ₕ∙_ --- -- (isoToEquiv (invIso (Iso-coHom-coHomRed n))) --- -- (homhelp n A) - --- -- abstract --- -- coHomGroup≡coHomRedGroup : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) --- -- → coHomGroup (suc n) (typ A) ≡ coHomRedGroup (suc n) A --- -- coHomGroup≡coHomRedGroup n A = --- -- InducedAbGroupPath (coHomGroup (suc n) (typ A)) --- -- _+ₕ∙_ --- -- (isoToEquiv (invIso (Iso-coHom-coHomRed n))) --- -- (homhelp n A) - --- -- ------------------- Kₙ ≃ ΩKₙ₊₁ --------------------- --- -- -- This proof uses the encode-decode method rather than Freudenthal - --- -- -- We define the map σ : Kₙ → ΩKₙ₊₁ and prove that it is a morphism --- -- private --- -- module _ (n : ℕ) where --- -- σ : {n : ℕ} → coHomK (suc n) → Path (coHomK (2 + n)) ∣ north ∣ ∣ north ∣ --- -- σ {n = n} = trRec (isOfHLevelTrunc (4 + n) _ _) --- -- λ a → cong ∣_∣ (merid a ∙ sym (merid (ptSn (suc n)))) - --- -- σ-hom-helper : ∀ {ℓ} {A : Type ℓ} {a : A} (p : a ≡ a) (r : refl ≡ p) --- -- → lUnit p ∙ cong (_∙ p) r ≡ rUnit p ∙ cong (p ∙_) r --- -- σ-hom-helper p = J (λ p r → lUnit p ∙ cong (_∙ p) r ≡ rUnit p ∙ cong (p ∙_) r) refl - --- -- σ-hom : {n : ℕ} (x y : coHomK (suc n)) → σ (x +ₖ y) ≡ σ x ∙ σ y --- -- σ-hom {n = zero} = --- -- elim2 (λ _ _ → isOfHLevelPath 3 (isOfHLevelTrunc 4 _ _) _ _) --- -- (wedgeconFun _ _ --- -- (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) --- -- (λ x → lUnit _ --- -- ∙ cong (_∙ σ ∣ x ∣) (cong (cong ∣_∣) (sym (rCancel (merid base))))) --- -- (λ y → cong σ (rUnitₖ 1 ∣ y ∣) --- -- ∙∙ rUnit _ --- -- ∙∙ cong (σ ∣ y ∣ ∙_) (cong (cong ∣_∣) (sym (rCancel (merid base))))) --- -- (sym (σ-hom-helper (σ ∣ base ∣) (cong (cong ∣_∣) (sym (rCancel (merid base))))))) --- -- σ-hom {n = suc n} = --- -- elim2 (λ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (5 + n) _ _) _ _) --- -- (wedgeconFun _ _ (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (wedgeConHLev' n) _ _) --- -- (λ x → lUnit _ --- -- ∙ cong (_∙ σ ∣ x ∣) (cong (cong ∣_∣) (sym (rCancel (merid north))))) --- -- (λ y → cong σ (rUnitₖ (2 + n) ∣ y ∣) --- -- ∙∙ rUnit _ --- -- ∙∙ cong (σ ∣ y ∣ ∙_) (cong (cong ∣_∣) (sym (rCancel (merid north))))) --- -- (sym (σ-hom-helper (σ ∣ north ∣) (cong (cong ∣_∣) (sym (rCancel (merid north))))))) - --- -- -- We will need to following lemma --- -- σ-minusDistr : {n : ℕ} (x y : coHomK (suc n)) → σ (x -ₖ y) ≡ σ x ∙ sym (σ y) --- -- σ-minusDistr {n = n} = --- -- morphLemmas.distrMinus' --- -- _+ₖ_ _∙_ --- -- σ σ-hom ∣ (ptSn (suc n)) ∣ refl --- -- -ₖ_ sym --- -- (λ x → sym (lUnit x)) (λ x → sym (rUnit x)) --- -- (rUnitₖ (suc n)) --- -- (lCancelₖ (suc n)) rCancel --- -- (assocₖ (suc n)) assoc∙ --- -- (cong (cong ∣_∣) (rCancel (merid (ptSn (suc n))))) - --- -- -- we define the code using addIso --- -- Code : (n : ℕ) → coHomK (2 + n) → Type₀ --- -- Code n x = (trRec {B = TypeOfHLevel ℓ-zero (3 + n)} (isOfHLevelTypeOfHLevel (3 + n)) --- -- λ a → Code' a , hLevCode' a) x .fst --- -- where --- -- Code' : (S₊ (2 + n)) → Type₀ --- -- Code' north = coHomK (suc n) --- -- Code' south = coHomK (suc n) --- -- Code' (merid a i) = isoToPath (addIso (suc n) ∣ a ∣) i - --- -- hLevCode' : (x : S₊ (2 + n)) → isOfHLevel (3 + n) (Code' x) --- -- hLevCode' = suspToPropElim (ptSn (suc n)) (λ _ → isPropIsOfHLevel (3 + n)) (isOfHLevelTrunc (3 + n)) - --- -- symMeridLem : (n : ℕ) → (x : S₊ (suc n)) (y : coHomK (suc n)) --- -- → subst (Code n) (cong ∣_∣ (sym (merid x))) y ≡ y -ₖ ∣ x ∣ --- -- symMeridLem n x = trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) --- -- (λ y → cong (_-ₖ ∣ x ∣) (transportRefl ∣ y ∣)) - --- -- decode : {n : ℕ} (x : coHomK (2 + n)) → Code n x → ∣ north ∣ ≡ x --- -- decode {n = n} = trElim (λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) --- -- decode-elim --- -- where --- -- north≡merid : (a : S₊ (suc n)) --- -- → Path (coHomK (2 + n)) ∣ north ∣ ∣ north ∣ --- -- ≡ (Path (coHomK (2 + n)) ∣ north ∣ ∣ south ∣) --- -- north≡merid a i = Path (coHomK (2 + n)) ∣ north ∣ ∣ merid a i ∣ - --- -- decode-elim : (a : S₊ (2 + n)) → Code n ∣ a ∣ → Path (coHomK (2 + n)) ∣ north ∣ ∣ a ∣ --- -- decode-elim north = σ --- -- decode-elim south = trRec (isOfHLevelTrunc (4 + n) _ _) --- -- λ a → cong ∣_∣ (merid a) --- -- decode-elim (merid a i) = --- -- hcomp (λ k → λ { (i = i0) → σ --- -- ; (i = i1) → mainPath a k}) --- -- (funTypeTransp (Code n) (λ x → ∣ north ∣ ≡ x) (cong ∣_∣ (merid a)) σ i) --- -- where --- -- mainPath : (a : (S₊ (suc n))) → --- -- transport (north≡merid a) ∘ σ ∘ transport (λ i → Code n ∣ merid a (~ i) ∣) --- -- ≡ trRec (isOfHLevelTrunc (4 + n) _ _) λ a → cong ∣_∣ (merid a) --- -- mainPath a = funExt (trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (4 + n) _ _) _ _) --- -- (λ x → (λ i → transport (north≡merid a) (σ (symMeridLem n a ∣ x ∣ i))) --- -- ∙∙ cong (transport (north≡merid a)) (-distrHelp x) --- -- ∙∙ (substAbove x))) --- -- where --- -- -distrHelp : (x : S₊ (suc n)) → σ (∣ x ∣ -ₖ ∣ a ∣) ≡ cong ∣_∣ (merid x) ∙ cong ∣_∣ (sym (merid a)) --- -- -distrHelp x = --- -- σ-minusDistr ∣ x ∣ ∣ a ∣ --- -- ∙ (λ i → (cong ∣_∣ (compPath-filler (merid x) (λ j → merid (ptSn (suc n)) (~ j ∨ i)) (~ i))) --- -- ∙ (cong ∣_∣ (sym (compPath-filler (merid a) (λ j → merid (ptSn (suc n)) (~ j ∨ i)) (~ i))))) - --- -- substAbove : (x : S₊ (suc n)) → transport (north≡merid a) (cong ∣_∣ (merid x) ∙ cong ∣_∣ (sym (merid a))) --- -- ≡ cong ∣_∣ (merid x) --- -- substAbove x i = transp (λ j → north≡merid a (i ∨ j)) i --- -- (compPath-filler (cong ∣_∣ (merid x)) (λ j → ∣ merid a (~ j ∨ i) ∣) (~ i)) - --- -- encode : {n : ℕ} {x : coHomK (2 + n)} → Path (coHomK (2 + n)) ∣ north ∣ x → Code n x --- -- encode {n = n} p = transport (cong (Code n) p) ∣ (ptSn (suc n)) ∣ - --- -- decode-encode : {n : ℕ} {x : coHomK (2 + n)} (p : Path (coHomK (2 + n)) ∣ north ∣ x) --- -- → decode _ (encode p) ≡ p --- -- decode-encode {n = n} = --- -- J (λ y p → decode _ (encode p) ≡ p) --- -- (cong (decode ∣ north ∣) (transportRefl ∣ ptSn (suc n) ∣) --- -- ∙ cong (cong ∣_∣) (rCancel (merid (ptSn (suc n))))) - --- -- -- We define an addition operation on Code which we can use in order to show that encode is a --- -- -- morphism (in a very loose sense) --- -- hLevCode : {n : ℕ} (x : coHomK (2 + n)) → isOfHLevel (3 + n) (Code n x) --- -- hLevCode {n = n} = --- -- trElim (λ _ → isProp→isOfHLevelSuc (3 + n) (isPropIsOfHLevel (3 + n))) --- -- (sphereToPropElim _ --- -- (λ _ → (isPropIsOfHLevel (3 + n))) (isOfHLevelTrunc (3 + n))) - --- -- Code-add' : {n : ℕ} (x : _) → Code n ∣ north ∣ → Code n ∣ x ∣ → Code n ∣ x ∣ --- -- Code-add' {n = n} north = _+ₖ_ --- -- Code-add' {n = n} south = _+ₖ_ --- -- Code-add' {n = n} (merid a i) = helper n a i --- -- where --- -- help : (n : ℕ) → (x y a : S₊ (suc n)) --- -- → transport (λ i → Code n ∣ north ∣ → Code n ∣ merid a i ∣ → Code n ∣ merid a i ∣) --- -- (_+ₖ_) ∣ x ∣ ∣ y ∣ --- -- ≡ ∣ x ∣ +ₖ ∣ y ∣ --- -- help n x y a = --- -- (λ i → transportRefl ((∣ transportRefl x i ∣ +ₖ (∣ transportRefl y i ∣ -ₖ ∣ a ∣)) +ₖ ∣ a ∣) i) --- -- ∙∙ cong (_+ₖ ∣ a ∣) (assocₖ _ ∣ x ∣ ∣ y ∣ (-ₖ ∣ a ∣)) --- -- ∙∙ sym (assocₖ _ (∣ x ∣ +ₖ ∣ y ∣) (-ₖ ∣ a ∣) ∣ a ∣) --- -- ∙∙ cong ((∣ x ∣ +ₖ ∣ y ∣) +ₖ_) (lCancelₖ _ ∣ a ∣) --- -- ∙∙ rUnitₖ _ _ - --- -- helper : (n : ℕ) (a : S₊ (suc n)) --- -- → PathP (λ i → Code n ∣ north ∣ → Code n ∣ merid a i ∣ → Code n ∣ merid a i ∣) _+ₖ_ _+ₖ_ --- -- helper n a = --- -- toPathP (funExt --- -- (trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelΠ (3 + n) (λ _ → isOfHLevelTrunc (3 + n))) _ _) --- -- λ x → funExt --- -- (trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) --- -- λ y → help n x y a))) - --- -- Code-add : {n : ℕ} (x : _) → Code n ∣ north ∣ → Code n x → Code n x --- -- Code-add {n = n} = --- -- trElim (λ x → isOfHLevelΠ (4 + n) λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelSuc (3 + n) (hLevCode {n = n} x)) --- -- Code-add' - --- -- encode-hom : {n : ℕ} {x : _} (q : 0ₖ _ ≡ 0ₖ _) (p : 0ₖ _ ≡ x) --- -- → encode (q ∙ p) ≡ Code-add {n = n} x (encode q) (encode p) --- -- encode-hom {n = n} q = J (λ x p → encode (q ∙ p) ≡ Code-add {n = n} x (encode q) (encode p)) --- -- (cong encode (sym (rUnit q)) --- -- ∙∙ sym (rUnitₖ _ (encode q)) --- -- ∙∙ cong (encode q +ₖ_) (cong ∣_∣ (sym (transportRefl _)))) - --- -- stabSpheres : (n : ℕ) → Iso (coHomK (suc n)) (typ (Ω (coHomK-ptd (2 + n)))) --- -- fun (stabSpheres n) = decode _ --- -- inv' (stabSpheres n) = encode --- -- rightInv (stabSpheres n) p = decode-encode p --- -- leftInv (stabSpheres n) = --- -- trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) --- -- λ a → cong encode (congFunct ∣_∣ (merid a) (sym (merid (ptSn (suc n))))) --- -- ∙∙ (λ i → transport (congFunct (Code n) (cong ∣_∣ (merid a)) --- -- (cong ∣_∣ (sym (merid (ptSn (suc n))))) i) ∣ ptSn (suc n) ∣) --- -- ∙∙ (substComposite (λ x → x) --- -- (cong (Code n) (cong ∣_∣ (merid a))) --- -- (cong (Code n) (cong ∣_∣ (sym (merid (ptSn (suc n)))))) ∣ ptSn (suc n) ∣ --- -- ∙∙ cong (transport (λ i → Code n ∣ merid (ptSn (suc n)) (~ i) ∣)) --- -- (transportRefl (∣ (ptSn (suc n)) ∣ +ₖ ∣ a ∣) ∙ lUnitₖ (suc n) ∣ a ∣) --- -- ∙∙ symMeridLem n (ptSn (suc n)) ∣ a ∣ --- -- ∙∙ cong (∣ a ∣ +ₖ_) -0ₖ --- -- ∙∙ rUnitₖ (suc n) ∣ a ∣) - --- -- Iso-Kn-ΩKn+1 : (n : HLevel) → Iso (coHomK n) (typ (Ω (coHomK-ptd (suc n)))) --- -- Iso-Kn-ΩKn+1 zero = invIso (compIso (congIso (truncIdempotentIso _ isGroupoidS¹)) ΩS¹Isoℤ) --- -- Iso-Kn-ΩKn+1 (suc n) = stabSpheres n - --- -- Kn≃ΩKn+1 : {n : ℕ} → coHomK n ≃ typ (Ω (coHomK-ptd (suc n))) --- -- Kn≃ΩKn+1 {n = n} = isoToEquiv (Iso-Kn-ΩKn+1 n) - --- -- -- Some properties of the Iso --- -- Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) --- -- Kn→ΩKn+1 n = Iso.fun (Iso-Kn-ΩKn+1 n) - --- -- ΩKn+1→Kn : (n : ℕ) → typ (Ω (coHomK-ptd (suc n))) → coHomK n --- -- ΩKn+1→Kn n = Iso.inv (Iso-Kn-ΩKn+1 n) - --- -- Kn→ΩKn+10ₖ : (n : ℕ) → Kn→ΩKn+1 n (0ₖ n) ≡ refl --- -- Kn→ΩKn+10ₖ zero = sym (rUnit refl) --- -- Kn→ΩKn+10ₖ (suc n) i j = ∣ (rCancel (merid (ptSn (suc n))) i j) ∣ - --- -- ΩKn+1→Kn-refl : (n : ℕ) → ΩKn+1→Kn n refl ≡ 0ₖ n --- -- ΩKn+1→Kn-refl zero = refl --- -- ΩKn+1→Kn-refl (suc zero) = refl --- -- ΩKn+1→Kn-refl (suc (suc n)) = refl - --- -- Kn≃ΩKn+1∙ : {n : ℕ} → coHomK-ptd n ≡ (Ω (coHomK-ptd (suc n))) --- -- Kn≃ΩKn+1∙ {n = n} = ua∙ Kn≃ΩKn+1 (Kn→ΩKn+10ₖ n) - --- -- Kn→ΩKn+1-hom : (n : ℕ) (x y : coHomK n) → Kn→ΩKn+1 n (x +[ n ]ₖ y) ≡ Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y --- -- Kn→ΩKn+1-hom zero x y = (λ j i → hfill (doubleComp-faces (λ i₁ → ∣ base ∣) (λ _ → ∣ base ∣) i) --- -- (inS (∣ intLoop (x ℤ+ y) i ∣)) (~ j)) --- -- ∙∙ (λ j i → ∣ intLoop-hom x y (~ j) i ∣) --- -- ∙∙ (congFunct ∣_∣ (intLoop x) (intLoop y) --- -- ∙ cong₂ _∙_ (λ j i → hfill (doubleComp-faces (λ i₁ → ∣ base ∣) (λ _ → ∣ base ∣) i) --- -- (inS (∣ intLoop x i ∣)) j) --- -- λ j i → hfill (doubleComp-faces (λ i₁ → ∣ base ∣) (λ _ → ∣ base ∣) i) --- -- (inS (∣ intLoop y i ∣)) j) --- -- Kn→ΩKn+1-hom (suc n) = σ-hom - --- -- ΩKn+1→Kn-hom : (n : ℕ) (x y : Path (coHomK (suc n)) (0ₖ _) (0ₖ _)) --- -- → ΩKn+1→Kn n (x ∙ y) ≡ ΩKn+1→Kn n x +[ n ]ₖ ΩKn+1→Kn n y --- -- ΩKn+1→Kn-hom zero p q = --- -- cong winding (congFunct (trRec isGroupoidS¹ (λ x → x)) p q) --- -- ∙ winding-hom (cong (trRec isGroupoidS¹ (λ x → x)) p) (cong (trRec isGroupoidS¹ (λ x → x)) q) --- -- ΩKn+1→Kn-hom (suc n) = encode-hom - --- -- Kn→ΩKn+1-ₖ : (n : ℕ) (x : coHomK n) → Kn→ΩKn+1 n (-ₖ x) ≡ sym (Kn→ΩKn+1 n x) --- -- Kn→ΩKn+1-ₖ n x = --- -- lUnit _ --- -- ∙∙ cong (_∙ Kn→ΩKn+1 n (-ₖ x)) (sym (lCancel _)) --- -- ∙∙ sym (assoc∙ _ _ _) --- -- ∙∙ cong (sym (Kn→ΩKn+1 n x) ∙_) help --- -- ∙∙ sym (rUnit _) --- -- where --- -- help : Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n (-ₖ x) ≡ refl --- -- help = sym (Kn→ΩKn+1-hom n x (-ₖ x)) ∙∙ cong (Kn→ΩKn+1 n) (rCancelₖ n x) ∙∙ Kn→ΩKn+10ₖ n - --- -- isHomogeneousKn : (n : HLevel) → isHomogeneous (coHomK-ptd n) --- -- isHomogeneousKn n = --- -- subst isHomogeneous --- -- (sym (ΣPathP (ua Kn≃ΩKn+1 , ua-gluePath _ (Kn→ΩKn+10ₖ n)))) --- -- (isHomogeneousPath _ _) - --- -- -- With the equivalence Kn≃ΩKn+1, we get that the two definitions of cohomology groups agree --- -- open IsGroupHom - --- -- coHom≅coHomΩ : ∀ {ℓ} (n : ℕ) (A : Type ℓ) → GroupIso (coHomGr n A) (coHomGrΩ n A) --- -- fun (fst (coHom≅coHomΩ n A)) = map λ f a → Kn→ΩKn+1 n (f a) --- -- inv' (fst (coHom≅coHomΩ n A)) = map λ f a → ΩKn+1→Kn n (f a) --- -- rightInv (fst (coHom≅coHomΩ n A)) = --- -- sElim (λ _ → isOfHLevelPath 2 § _ _) --- -- λ f → cong ∣_∣₂ (funExt λ x → rightInv (Iso-Kn-ΩKn+1 n) (f x)) --- -- leftInv (fst (coHom≅coHomΩ n A)) = --- -- sElim (λ _ → isOfHLevelPath 2 § _ _) --- -- λ f → cong ∣_∣₂ (funExt λ x → leftInv (Iso-Kn-ΩKn+1 n) (f x)) --- -- snd (coHom≅coHomΩ n A) = --- -- makeIsGroupHom --- -- (sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) --- -- λ f g → cong ∣_∣₂ (funExt λ x → Kn→ΩKn+1-hom n (f x) (g x))) - --- -- module lockedKnIso (key : Unit') where --- -- Kn→ΩKn+1' : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) --- -- Kn→ΩKn+1' n = lock key (Iso.fun (Iso-Kn-ΩKn+1 n)) - --- -- ΩKn+1→Kn' : (n : ℕ) → typ (Ω (coHomK-ptd (suc n))) → coHomK n --- -- ΩKn+1→Kn' n = lock key (Iso.inv (Iso-Kn-ΩKn+1 n)) - --- -- ΩKn+1→Kn→ΩKn+1 : (n : ℕ) → (x : typ (Ω (coHomK-ptd (suc n)))) → Kn→ΩKn+1' n (ΩKn+1→Kn' n x) ≡ x --- -- ΩKn+1→Kn→ΩKn+1 n x = pm key --- -- where --- -- pm : (key : Unit') → lock key (Iso.fun (Iso-Kn-ΩKn+1 n)) (lock key (Iso.inv (Iso-Kn-ΩKn+1 n)) x) ≡ x --- -- pm unlock = Iso.rightInv (Iso-Kn-ΩKn+1 n) x - --- -- Kn→ΩKn+1→Kn : (n : ℕ) → (x : coHomK n) → ΩKn+1→Kn' n (Kn→ΩKn+1' n x) ≡ x --- -- Kn→ΩKn+1→Kn n x = pm key --- -- where --- -- pm : (key : Unit') → lock key (Iso.inv (Iso-Kn-ΩKn+1 n)) (lock key (Iso.fun (Iso-Kn-ΩKn+1 n)) x) ≡ x --- -- pm unlock = Iso.leftInv (Iso-Kn-ΩKn+1 n) x - --- -- -distrLemma : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n m : ℕ) (f : GroupHom (coHomGr n A) (coHomGr m B)) --- -- (x y : coHom n A) --- -- → fst f (x -[ n ]ₕ y) ≡ fst f x -[ m ]ₕ fst f y --- -- -distrLemma n m f' x y = sym (-cancelRₕ m (f y) (f (x -[ n ]ₕ y))) --- -- ∙∙ cong (λ x → x -[ m ]ₕ f y) (sym (f' .snd .pres· (x -[ n ]ₕ y) y)) --- -- ∙∙ cong (λ x → x -[ m ]ₕ f y) ( cong f (-+cancelₕ n _ _)) --- -- where --- -- f = fst f' - --- -- -- HLevel stuff for cup product --- -- isContr-↓∙ : (n : ℕ) → isContr (coHomK-ptd (suc n) →∙ coHomK-ptd n) --- -- fst (isContr-↓∙ zero) = (λ _ → 0) , refl --- -- snd (isContr-↓∙ zero) (f , p) = --- -- Σ≡Prop (λ f → isSetℤ _ _) --- -- (funExt (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 isSetℤ) _ _) --- -- (toPropElim (λ _ → isSetℤ _ _) (sym p)))) --- -- fst (isContr-↓∙ (suc n)) = (λ _ → 0ₖ _) , refl --- -- fst (snd (isContr-↓∙ (suc n)) f i) x = --- -- trElim {B = λ x → 0ₖ (suc n) ≡ fst f x} --- -- (λ _ → isOfHLevelPath (4 + n) (isOfHLevelSuc (3 + n) (isOfHLevelTrunc (3 + n))) _ _) --- -- (sphereElim _ (λ _ → isOfHLevelTrunc (3 + n) _ _) --- -- (sym (snd f))) x i --- -- snd (snd (isContr-↓∙ (suc n)) f i) j = snd f (~ i ∨ j) - --- -- isContr-↓∙' : (n : ℕ) → isContr (S₊∙ (suc n) →∙ coHomK-ptd n) --- -- fst (isContr-↓∙' zero) = (λ _ → 0) , refl --- -- snd (isContr-↓∙' zero) (f , p) = --- -- Σ≡Prop (λ f → isSetℤ _ _) --- -- (funExt (toPropElim (λ _ → isSetℤ _ _) (sym p))) --- -- fst (isContr-↓∙' (suc n)) = (λ _ → 0ₖ _) , refl --- -- fst (snd (isContr-↓∙' (suc n)) f i) x = --- -- sphereElim _ {A = λ x → 0ₖ (suc n) ≡ fst f x} --- -- (λ _ → isOfHLevelTrunc (3 + n) _ _) --- -- (sym (snd f)) x i --- -- snd (snd (isContr-↓∙' (suc n)) f i) j = snd f (~ i ∨ j) - --- -- isOfHLevel→∙Kn : {A : Pointed₀} (n m : ℕ) --- -- → isOfHLevel (suc m) (A →∙ coHomK-ptd n) → isOfHLevel (suc (suc m)) (A →∙ coHomK-ptd (suc n)) --- -- isOfHLevel→∙Kn {A = A} n m hlev = step₃ --- -- where --- -- step₁ : isOfHLevel (suc m) (A →∙ Ω (coHomK-ptd (suc n))) --- -- step₁ = --- -- subst (isOfHLevel (suc m)) --- -- (λ i → A →∙ ua∙ {A = Ω (coHomK-ptd (suc n))} {B = coHomK-ptd n} --- -- (invEquiv Kn≃ΩKn+1) --- -- (ΩKn+1→Kn-refl n) (~ i)) hlev - --- -- step₂ : isOfHLevel (suc m) (typ (Ω (A →∙ coHomK-ptd (suc n) ∙))) --- -- step₂ = isOfHLevelRetractFromIso (suc m) (invIso (invIso (ΩfunExtIso _ _))) step₁ - --- -- step₃ : isOfHLevel (suc (suc m)) (A →∙ coHomK-ptd (suc n)) --- -- step₃ = --- -- isOfHLevelΩ→isOfHLevel m --- -- λ f → subst (λ x → isOfHLevel (suc m) (typ (Ω x))) --- -- (isHomogeneous→∙ (isHomogeneousKn (suc n)) f) --- -- step₂ - --- -- isOfHLevel↑∙ : ∀ n m → isOfHLevel (2 + n) (coHomK-ptd (suc m) →∙ coHomK-ptd (suc (n + m))) --- -- isOfHLevel↑∙ zero m = isOfHLevel→∙Kn m 0 (isContr→isProp (isContr-↓∙ m)) --- -- isOfHLevel↑∙ (suc n) m = isOfHLevel→∙Kn (suc (n + m)) (suc n) (isOfHLevel↑∙ n m) - --- -- isOfHLevel↑∙' : ∀ n m → isOfHLevel (2 + n) (S₊∙ (suc m) →∙ coHomK-ptd (suc (n + m))) --- -- isOfHLevel↑∙' zero m = isOfHLevel→∙Kn m 0 (isContr→isProp (isContr-↓∙' m)) --- -- isOfHLevel↑∙' (suc n) m = isOfHLevel→∙Kn (suc (n + m)) (suc n) (isOfHLevel↑∙' n m) - --- -- →∙KnPath : (A : Pointed₀) (n : ℕ) → Ω (A →∙ coHomK-ptd (suc n) ∙) ≡ (A →∙ coHomK-ptd n ∙) --- -- →∙KnPath A n = --- -- ua∙ (isoToEquiv (ΩfunExtIso A (coHomK-ptd (suc n)))) refl --- -- ∙ (λ i → (A →∙ Kn≃ΩKn+1∙ {n = n} (~ i) ∙)) - --- -- contr∙-lem : (n m : ℕ) --- -- → isContr (coHomK-ptd (suc n) --- -- →∙ (coHomK-ptd (suc m) --- -- →∙ coHomK-ptd (suc (n + m)) ∙)) --- -- fst (contr∙-lem n m) = (λ _ → (λ _ → 0ₖ _) , refl) , refl --- -- snd (contr∙-lem n m) (f , p) = →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousKn _)) (sym (funExt help)) --- -- where --- -- help : (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) --- -- help = --- -- trElim (λ _ → isOfHLevelPath (3 + n) --- -- (subst (λ x → isOfHLevel x (coHomK-ptd (suc m) →∙ coHomK-ptd (suc (n + m)))) (λ i → suc (suc (+-comm n 1 i))) --- -- (isOfHLevelPlus' {n = 1} (2 + n) (isOfHLevel↑∙ n m))) _ _) --- -- (sphereElim _ (λ _ → isOfHLevel↑∙ n m _ _) p) - --- -- isOfHLevel↑∙∙ : ∀ n m l --- -- → isOfHLevel (2 + l) (coHomK-ptd (suc n) --- -- →∙ (coHomK-ptd (suc m) --- -- →∙ coHomK-ptd (suc (suc (l + n + m))) ∙)) --- -- isOfHLevel↑∙∙ n m zero = --- -- isOfHLevelΩ→isOfHLevel 0 λ f --- -- → subst --- -- isProp (cong (λ x → typ (Ω x)) --- -- (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousKn _)) f)) --- -- (isOfHLevelRetractFromIso 1 (ΩfunExtIso _ _) h) --- -- where --- -- h : isProp (coHomK-ptd (suc n) --- -- →∙ (Ω (coHomK-ptd (suc m) --- -- →∙ coHomK-ptd (suc (suc (n + m))) ∙))) --- -- h = --- -- subst isProp (λ i → coHomK-ptd (suc n) →∙ (→∙KnPath (coHomK-ptd (suc m)) (suc (n + m)) (~ i))) --- -- (isContr→isProp (contr∙-lem n m)) --- -- isOfHLevel↑∙∙ n m (suc l) = --- -- isOfHLevelΩ→isOfHLevel (suc l) --- -- λ f → --- -- subst --- -- (isOfHLevel (2 + l)) (cong (λ x → typ (Ω x)) --- -- (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousKn _)) f)) --- -- (isOfHLevelRetractFromIso (2 + l) (ΩfunExtIso _ _) h) --- -- where --- -- h : isOfHLevel (2 + l) --- -- (coHomK-ptd (suc n) --- -- →∙ (Ω (coHomK-ptd (suc m) --- -- →∙ coHomK-ptd (suc (suc (suc (l + n + m)))) ∙))) --- -- h = --- -- subst (isOfHLevel (2 + l)) --- -- (λ i → coHomK-ptd (suc n) →∙ →∙KnPath (coHomK-ptd (suc m)) (suc (suc (l + n + m))) (~ i)) --- -- (isOfHLevel↑∙∙ n m l) - --- -- -- Misc. --- -- isoType→isoCohom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) --- -- → Iso A B --- -- → GroupIso (coHomGr n A) (coHomGr n B) --- -- fst (isoType→isoCohom n is) = setTruncIso (domIso is) --- -- snd (isoType→isoCohom n is) = --- -- makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) --- -- (λ _ _ → refl)) +------------------- Connectedness --------------------- +is2ConnectedKn : (n : ℕ) → isConnected 2 (coHomK (suc n)) +is2ConnectedKn zero = ∣ ∣ base ∣ ∣ + , trElim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 (isOfHLevelTrunc 2)) _ _) + (toPropElim (λ _ → isOfHLevelTrunc 2 _ _) refl)) +is2ConnectedKn (suc n) = ∣ ∣ north ∣ ∣ + , trElim (λ _ → isOfHLevelPath 2 (isOfHLevelTrunc 2) _ _) + (trElim (λ _ → isProp→isOfHLevelSuc (3 + n) (isOfHLevelTrunc 2 _ _)) + (suspToPropElim (ptSn (suc n)) (λ _ → isOfHLevelTrunc 2 _ _) refl)) + +isConnectedKn : (n : ℕ) → isConnected (2 + n) (coHomK (suc n)) +isConnectedKn n = isOfHLevelRetractFromIso 0 (invIso (truncOfTruncIso (2 + n) 1)) (sphereConnected (suc n)) + +-- direct proof of connectedness of ΩKₙ₊₁ not relying on the equivalence ∥ a ≡ b ∥ₙ ≃ (∣ a ∣ₙ₊₁ ≡ ∣ b ∣ₙ₊₁) +isConnectedPathKn : (n : ℕ) (x y : (coHomK (suc n))) → isConnected (suc n) (x ≡ y) +isConnectedPathKn n = + trElim (λ _ → isProp→isOfHLevelSuc (2 + n) (isPropΠ λ _ → isPropIsContr)) + (sphereElim _ (λ _ → isProp→isOfHLevelSuc n (isPropΠ λ _ → isPropIsContr)) + λ y → isContrRetractOfConstFun + {B = (hLevelTrunc (suc n) (ptSn (suc n) ≡ ptSn (suc n)))} ∣ refl ∣ + (fun⁻ n y + , trElim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) + (J (λ y p → fun⁻ n y _ ≡ _) (funExt⁻ (fun⁻Id n) ∣ refl ∣)))) + where + fun⁻ : (n : ℕ) → (y : coHomK (suc n)) → + hLevelTrunc (suc n) (ptSn (suc n) ≡ ptSn (suc n)) + → hLevelTrunc (suc n) (∣ ptSn (suc n) ∣ ≡ y) + fun⁻ n = + trElim (λ _ → isOfHLevelΠ (3 + n) λ _ → isOfHLevelSuc (2 + n) (isOfHLevelSuc (suc n) (isOfHLevelTrunc (suc n)))) + (sphereElim n (λ _ → isOfHLevelΠ (suc n) λ _ → isOfHLevelTrunc (suc n)) λ _ → ∣ refl ∣) + + fun⁻Id : (n : ℕ) → fun⁻ n ∣ ptSn (suc n) ∣ ≡ λ _ → ∣ refl ∣ + fun⁻Id zero = refl + fun⁻Id (suc n) = refl + +------------------- +-- Induction principles for cohomology groups (n ≥ 1) +-- If we want to show a proposition about some x : Hⁿ(A), it suffices to show it under the +-- assumption that x = ∣ f ∣₂ for some f : A → Kₙ and that f is pointed + +coHomPointedElim : {A : Type ℓ} (n : ℕ) (a : A) {B : coHom (suc n) A → Type ℓ'} + → ((x : coHom (suc n) A) → isProp (B x)) + → ((f : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → B ∣ f ∣₂) + → (x : coHom (suc n) A) → B x +coHomPointedElim {ℓ' = ℓ'} {A = A} n a isprop indp = + sElim (λ _ → isOfHLevelSuc 1 (isprop _)) + λ f → helper n isprop indp f (f a) refl + where + helper : (n : ℕ) {B : coHom (suc n) A → Type ℓ'} + → ((x : coHom (suc n) A) → isProp (B x)) + → ((f : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → B ∣ f ∣₂) + → (f : A → coHomK (suc n)) + → (x : coHomK (suc n)) + → f a ≡ x → B ∣ f ∣₂ + -- pattern matching a bit extra to avoid isOfHLevelPlus' + helper zero isprop ind f = + trElim (λ _ → isOfHLevelPlus {n = 1} 2 (isPropΠ λ _ → isprop _)) + (toPropElim (λ _ → isPropΠ λ _ → isprop _) (ind f)) + helper (suc zero) isprop ind f = + trElim (λ _ → isOfHLevelPlus {n = 1} 3 (isPropΠ λ _ → isprop _)) + (suspToPropElim base (λ _ → isPropΠ λ _ → isprop _) (ind f)) + helper (suc (suc zero)) isprop ind f = + trElim (λ _ → isOfHLevelPlus {n = 1} 4 (isPropΠ λ _ → isprop _)) + (suspToPropElim north (λ _ → isPropΠ λ _ → isprop _) (ind f)) + helper (suc (suc (suc n))) isprop ind f = + trElim (λ _ → isOfHLevelPlus' {n = 5 + n} 1 (isPropΠ λ _ → isprop _)) + (suspToPropElim north (λ _ → isPropΠ λ _ → isprop _) (ind f)) + + +coHomPointedElim2 : {A : Type ℓ} (n : ℕ) (a : A) {B : coHom (suc n) A → coHom (suc n) A → Type ℓ'} + → ((x y : coHom (suc n) A) → isProp (B x y)) + → ((f g : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → g a ≡ coHom-pt (suc n) → B ∣ f ∣₂ ∣ g ∣₂) + → (x y : coHom (suc n) A) → B x y +coHomPointedElim2 {ℓ' = ℓ'} {A = A} n a isprop indp = sElim2 (λ _ _ → isOfHLevelSuc 1 (isprop _ _)) + λ f g → helper n a isprop indp f g (f a) (g a) refl refl + where + helper : (n : ℕ) (a : A) {B : coHom (suc n) A → coHom (suc n) A → Type ℓ'} + → ((x y : coHom (suc n) A) → isProp (B x y)) + → ((f g : A → coHomK (suc n)) → f a ≡ coHom-pt (suc n) → g a ≡ coHom-pt (suc n) → B ∣ f ∣₂ ∣ g ∣₂) + → (f g : A → coHomK (suc n)) + → (x y : coHomK (suc n)) + → f a ≡ x → g a ≡ y + → B ∣ f ∣₂ ∣ g ∣₂ + helper zero a isprop indp f g = + elim2 (λ _ _ → isOfHLevelPlus {n = 1} 2 (isPropΠ2 λ _ _ → isprop _ _)) + (toPropElim2 (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) + helper (suc zero) a isprop indp f g = + elim2 (λ _ _ → isOfHLevelPlus {n = 1} 3 (isPropΠ2 λ _ _ → isprop _ _)) + (suspToPropElim2 base (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) + helper (suc (suc zero)) a isprop indp f g = + elim2 (λ _ _ → isOfHLevelPlus {n = 1} 4 (isPropΠ2 λ _ _ → isprop _ _)) + (suspToPropElim2 north (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) + helper (suc (suc (suc n))) a isprop indp f g = + elim2 (λ _ _ → isOfHLevelPlus' {n = 5 + n} 1 (isPropΠ2 λ _ _ → isprop _ _)) + (suspToPropElim2 north (λ _ _ → isPropΠ2 λ _ _ → isprop _ _) (indp f g)) + +coHomK-elim : ∀ {ℓ} (n : ℕ) {B : coHomK (suc n) → Type ℓ} + → ((x : _) → isOfHLevel (suc n) (B x)) + → B (0ₖ (suc n)) + → (x : _) → B x +coHomK-elim n {B = B } hlev b = + trElim (λ _ → isOfHLevelPlus {n = (suc n)} 2 (hlev _)) + (sphereElim _ (hlev ∘ ∣_∣) b) + +{- Equivalence between cohomology of A and reduced cohomology of (A + 1) -} +coHomRed+1Equiv : (n : ℕ) → + (A : Type ℓ) → + (coHom n A) ≡ (coHomRed n ((A ⊎ Unit , inr (tt)))) +coHomRed+1Equiv zero A i = ∥ helpLemma {C = (ℤ , pos 0)} i ∥₂ + module coHomRed+1 where + helpLemma : {C : Pointed ℓ} → ( (A → (typ C)) ≡ ((((A ⊎ Unit) , inr (tt)) →∙ C))) + helpLemma {C = C} = isoToPath (iso map1 + map2 + (λ b → linvPf b) + (λ _ → refl)) + where + map1 : (A → typ C) → ((((A ⊎ Unit) , inr (tt)) →∙ C)) + map1 f = map1' , refl + module helpmap where + map1' : A ⊎ Unit → fst C + map1' (inl x) = f x + map1' (inr x) = pt C + + map2 : ((((A ⊎ Unit) , inr (tt)) →∙ C)) → (A → typ C) + map2 (g , pf) x = g (inl x) + + linvPf : (b :((((A ⊎ Unit) , inr (tt)) →∙ C))) → map1 (map2 b) ≡ b + linvPf (f , snd) i = (λ x → helper x i) , λ j → snd ((~ i) ∨ j) + where + helper : (x : A ⊎ Unit) → ((helpmap.map1') (map2 (f , snd)) x) ≡ f x + helper (inl x) = refl + helper (inr tt) = sym snd +coHomRed+1Equiv (suc zero) A i = ∥ coHomRed+1.helpLemma A i {C = (coHomK 1 , ∣ base ∣)} i ∥₂ +coHomRed+1Equiv (suc (suc n)) A i = ∥ coHomRed+1.helpLemma A i {C = (coHomK (2 + n) , ∣ north ∣)} i ∥₂ + +Iso-coHom-coHomRed : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → Iso (coHomRed (suc n) A) (coHom (suc n) (typ A)) +fun (Iso-coHom-coHomRed {A = A , a} n) = map fst +inv' (Iso-coHom-coHomRed {A = A , a} n) = map λ f → (λ x → f x -ₖ f a) , rCancelₖ _ _ +rightInv (Iso-coHom-coHomRed {A = A , a} n) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ f → trRec (isProp→isOfHLevelSuc _ (§ _ _)) + (λ p → cong ∣_∣₂ (funExt λ x → cong (λ y → f x +ₖ y) (cong -ₖ_ p ∙ -0ₖ) ∙ rUnitₖ _ (f x))) + (Iso.fun (PathIdTruncIso (suc n)) (isContr→isProp (isConnectedKn n) ∣ f a ∣ ∣ 0ₖ _ ∣)) +leftInv (Iso-coHom-coHomRed {A = A , a} n) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ {(f , p) → cong ∣_∣₂ (ΣPathP (((funExt λ x → (cong (λ y → f x -ₖ y) p + ∙∙ cong (λ y → f x +ₖ y) -0ₖ + ∙∙ rUnitₖ _ (f x)) ∙ refl)) + , helper n (f a) (sym p)))} + where + path : (n : ℕ) (x : coHomK (suc n)) (p : 0ₖ _ ≡ x) → _ + path n x p = (cong (λ y → x -ₖ y) (sym p) ∙∙ cong (λ y → x +ₖ y) -0ₖ ∙∙ rUnitₖ _ x) ∙ refl + + helper : (n : ℕ) (x : coHomK (suc n)) (p : 0ₖ _ ≡ x) + → PathP (λ i → path n x p i ≡ 0ₖ _) (rCancelₖ _ x) (sym p) + helper zero x = + J (λ x p → PathP (λ i → path 0 x p i ≡ 0ₖ _) + (rCancelₖ _ x) (sym p)) + λ i j → rUnit (rUnit (λ _ → 0ₖ 1) (~ j)) (~ j) i + helper (suc n) x = + J (λ x p → PathP (λ i → path (suc n) x p i ≡ 0ₖ _) (rCancelₖ _ x) (sym p)) + λ i j → rCancelₖ (suc (suc n)) (0ₖ (suc (suc n))) (~ i ∧ ~ j) + ++∙≡+ : (n : ℕ) {A : Pointed ℓ} (x y : coHomRed (suc n) A) + → Iso.fun (Iso-coHom-coHomRed n) (x +ₕ∙ y) + ≡ Iso.fun (Iso-coHom-coHomRed n) x +ₕ Iso.fun (Iso-coHom-coHomRed n) y + ++∙≡+ zero = sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl ++∙≡+ (suc n) = sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) λ _ _ → refl + +private + homhelp : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) (x y : coHom (suc n) (typ A)) + → Iso.inv (Iso-coHom-coHomRed {A = A} n) (x +ₕ y) + ≡ Iso.inv (Iso-coHom-coHomRed n) x +ₕ∙ Iso.inv (Iso-coHom-coHomRed n) y + homhelp n A = morphLemmas.isMorphInv _+ₕ∙_ _+ₕ_ + (Iso.fun (Iso-coHom-coHomRed n)) (+∙≡+ n) _ + (Iso.rightInv (Iso-coHom-coHomRed n)) (Iso.leftInv (Iso-coHom-coHomRed n)) + +coHomGr≅coHomRedGr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) + → GroupEquiv (coHomRedGrDir (suc n) A) (coHomGr (suc n) (typ A)) +fst (coHomGr≅coHomRedGr n A) = isoToEquiv (Iso-coHom-coHomRed n) +snd (coHomGr≅coHomRedGr n A) = makeIsGroupHom (+∙≡+ n) + +coHomRedGroup : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → AbGroup ℓ +coHomRedGroup zero A = coHomRedGroupDir zero A +coHomRedGroup (suc n) A = + InducedAbGroup (coHomGroup (suc n) (typ A)) + _+ₕ∙_ + (isoToEquiv (invIso (Iso-coHom-coHomRed n))) + (homhelp n A) + +abstract + coHomGroup≡coHomRedGroup : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) + → coHomGroup (suc n) (typ A) ≡ coHomRedGroup (suc n) A + coHomGroup≡coHomRedGroup n A = + InducedAbGroupPath (coHomGroup (suc n) (typ A)) + _+ₕ∙_ + (isoToEquiv (invIso (Iso-coHom-coHomRed n))) + (homhelp n A) + +------------------- Kₙ ≃ ΩKₙ₊₁ --------------------- +-- This proof uses the encode-decode method rather than Freudenthal + +-- We define the map σ : Kₙ → ΩKₙ₊₁ and prove that it is a morphism +private + module _ (n : ℕ) where + σ : {n : ℕ} → coHomK (suc n) → Path (coHomK (2 + n)) ∣ north ∣ ∣ north ∣ + σ {n = n} = trRec (isOfHLevelTrunc (4 + n) _ _) + λ a → cong ∣_∣ (merid a ∙ sym (merid (ptSn (suc n)))) + + σ-hom-helper : ∀ {ℓ} {A : Type ℓ} {a : A} (p : a ≡ a) (r : refl ≡ p) + → lUnit p ∙ cong (_∙ p) r ≡ rUnit p ∙ cong (p ∙_) r + σ-hom-helper p = J (λ p r → lUnit p ∙ cong (_∙ p) r ≡ rUnit p ∙ cong (p ∙_) r) refl + + σ-hom : {n : ℕ} (x y : coHomK (suc n)) → σ (x +ₖ y) ≡ σ x ∙ σ y + σ-hom {n = zero} = + elim2 (λ _ _ → isOfHLevelPath 3 (isOfHLevelTrunc 4 _ _) _ _) + (wedgeconFun _ _ + (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) + (λ x → lUnit _ + ∙ cong (_∙ σ ∣ x ∣) (cong (cong ∣_∣) (sym (rCancel (merid base))))) + (λ y → cong σ (rUnitₖ 1 ∣ y ∣) + ∙∙ rUnit _ + ∙∙ cong (σ ∣ y ∣ ∙_) (cong (cong ∣_∣) (sym (rCancel (merid base))))) + (sym (σ-hom-helper (σ ∣ base ∣) (cong (cong ∣_∣) (sym (rCancel (merid base))))))) + σ-hom {n = suc n} = + elim2 (λ _ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (5 + n) _ _) _ _) + (wedgeconFun _ _ (λ _ _ → isOfHLevelPath ((2 + n) + (2 + n)) (wedgeConHLev' n) _ _) + (λ x → lUnit _ + ∙ cong (_∙ σ ∣ x ∣) (cong (cong ∣_∣) (sym (rCancel (merid north))))) + (λ y → cong σ (rUnitₖ (2 + n) ∣ y ∣) + ∙∙ rUnit _ + ∙∙ cong (σ ∣ y ∣ ∙_) (cong (cong ∣_∣) (sym (rCancel (merid north))))) + (sym (σ-hom-helper (σ ∣ north ∣) (cong (cong ∣_∣) (sym (rCancel (merid north))))))) + + -- We will need to following lemma + σ-minusDistr : {n : ℕ} (x y : coHomK (suc n)) → σ (x -ₖ y) ≡ σ x ∙ sym (σ y) + σ-minusDistr {n = n} = + morphLemmas.distrMinus' + _+ₖ_ _∙_ + σ σ-hom ∣ (ptSn (suc n)) ∣ refl + -ₖ_ sym + (λ x → sym (lUnit x)) (λ x → sym (rUnit x)) + (rUnitₖ (suc n)) + (lCancelₖ (suc n)) rCancel + (assocₖ (suc n)) assoc∙ + (cong (cong ∣_∣) (rCancel (merid (ptSn (suc n))))) + + -- we define the code using addIso + Code : (n : ℕ) → coHomK (2 + n) → Type₀ + Code n x = (trRec {B = TypeOfHLevel ℓ-zero (3 + n)} (isOfHLevelTypeOfHLevel (3 + n)) + λ a → Code' a , hLevCode' a) x .fst + where + Code' : (S₊ (2 + n)) → Type₀ + Code' north = coHomK (suc n) + Code' south = coHomK (suc n) + Code' (merid a i) = isoToPath (addIso (suc n) ∣ a ∣) i + + hLevCode' : (x : S₊ (2 + n)) → isOfHLevel (3 + n) (Code' x) + hLevCode' = suspToPropElim (ptSn (suc n)) (λ _ → isPropIsOfHLevel (3 + n)) (isOfHLevelTrunc (3 + n)) + + symMeridLem : (n : ℕ) → (x : S₊ (suc n)) (y : coHomK (suc n)) + → subst (Code n) (cong ∣_∣ (sym (merid x))) y ≡ y -ₖ ∣ x ∣ + symMeridLem n x = trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) + (λ y → cong (_-ₖ ∣ x ∣) (transportRefl ∣ y ∣)) + + decode : {n : ℕ} (x : coHomK (2 + n)) → Code n x → ∣ north ∣ ≡ x + decode {n = n} = trElim (λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelPath (4 + n) (isOfHLevelTrunc (4 + n)) _ _) + decode-elim + where + north≡merid : (a : S₊ (suc n)) + → Path (coHomK (2 + n)) ∣ north ∣ ∣ north ∣ + ≡ (Path (coHomK (2 + n)) ∣ north ∣ ∣ south ∣) + north≡merid a i = Path (coHomK (2 + n)) ∣ north ∣ ∣ merid a i ∣ + + decode-elim : (a : S₊ (2 + n)) → Code n ∣ a ∣ → Path (coHomK (2 + n)) ∣ north ∣ ∣ a ∣ + decode-elim north = σ + decode-elim south = trRec (isOfHLevelTrunc (4 + n) _ _) + λ a → cong ∣_∣ (merid a) + decode-elim (merid a i) = + hcomp (λ k → λ { (i = i0) → σ + ; (i = i1) → mainPath a k}) + (funTypeTransp (Code n) (λ x → ∣ north ∣ ≡ x) (cong ∣_∣ (merid a)) σ i) + where + mainPath : (a : (S₊ (suc n))) → + transport (north≡merid a) ∘ σ ∘ transport (λ i → Code n ∣ merid a (~ i) ∣) + ≡ trRec (isOfHLevelTrunc (4 + n) _ _) λ a → cong ∣_∣ (merid a) + mainPath a = funExt (trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (4 + n) _ _) _ _) + (λ x → (λ i → transport (north≡merid a) (σ (symMeridLem n a ∣ x ∣ i))) + ∙∙ cong (transport (north≡merid a)) (-distrHelp x) + ∙∙ (substAbove x))) + where + -distrHelp : (x : S₊ (suc n)) → σ (∣ x ∣ -ₖ ∣ a ∣) ≡ cong ∣_∣ (merid x) ∙ cong ∣_∣ (sym (merid a)) + -distrHelp x = + σ-minusDistr ∣ x ∣ ∣ a ∣ + ∙ (λ i → (cong ∣_∣ (compPath-filler (merid x) (λ j → merid (ptSn (suc n)) (~ j ∨ i)) (~ i))) + ∙ (cong ∣_∣ (sym (compPath-filler (merid a) (λ j → merid (ptSn (suc n)) (~ j ∨ i)) (~ i))))) + + substAbove : (x : S₊ (suc n)) → transport (north≡merid a) (cong ∣_∣ (merid x) ∙ cong ∣_∣ (sym (merid a))) + ≡ cong ∣_∣ (merid x) + substAbove x i = transp (λ j → north≡merid a (i ∨ j)) i + (compPath-filler (cong ∣_∣ (merid x)) (λ j → ∣ merid a (~ j ∨ i) ∣) (~ i)) + + encode : {n : ℕ} {x : coHomK (2 + n)} → Path (coHomK (2 + n)) ∣ north ∣ x → Code n x + encode {n = n} p = transport (cong (Code n) p) ∣ (ptSn (suc n)) ∣ + + decode-encode : {n : ℕ} {x : coHomK (2 + n)} (p : Path (coHomK (2 + n)) ∣ north ∣ x) + → decode _ (encode p) ≡ p + decode-encode {n = n} = + J (λ y p → decode _ (encode p) ≡ p) + (cong (decode ∣ north ∣) (transportRefl ∣ ptSn (suc n) ∣) + ∙ cong (cong ∣_∣) (rCancel (merid (ptSn (suc n))))) + +-- We define an addition operation on Code which we can use in order to show that encode is a +-- morphism (in a very loose sense) + hLevCode : {n : ℕ} (x : coHomK (2 + n)) → isOfHLevel (3 + n) (Code n x) + hLevCode {n = n} = + trElim (λ _ → isProp→isOfHLevelSuc (3 + n) (isPropIsOfHLevel (3 + n))) + (sphereToPropElim _ + (λ _ → (isPropIsOfHLevel (3 + n))) (isOfHLevelTrunc (3 + n))) + + Code-add' : {n : ℕ} (x : _) → Code n ∣ north ∣ → Code n ∣ x ∣ → Code n ∣ x ∣ + Code-add' {n = n} north = _+ₖ_ + Code-add' {n = n} south = _+ₖ_ + Code-add' {n = n} (merid a i) = helper n a i + where + help : (n : ℕ) → (x y a : S₊ (suc n)) + → transport (λ i → Code n ∣ north ∣ → Code n ∣ merid a i ∣ → Code n ∣ merid a i ∣) + (_+ₖ_) ∣ x ∣ ∣ y ∣ + ≡ ∣ x ∣ +ₖ ∣ y ∣ + help n x y a = + (λ i → transportRefl ((∣ transportRefl x i ∣ +ₖ (∣ transportRefl y i ∣ -ₖ ∣ a ∣)) +ₖ ∣ a ∣) i) + ∙∙ cong (_+ₖ ∣ a ∣) (assocₖ _ ∣ x ∣ ∣ y ∣ (-ₖ ∣ a ∣)) + ∙∙ sym (assocₖ _ (∣ x ∣ +ₖ ∣ y ∣) (-ₖ ∣ a ∣) ∣ a ∣) + ∙∙ cong ((∣ x ∣ +ₖ ∣ y ∣) +ₖ_) (lCancelₖ _ ∣ a ∣) + ∙∙ rUnitₖ _ _ + + helper : (n : ℕ) (a : S₊ (suc n)) + → PathP (λ i → Code n ∣ north ∣ → Code n ∣ merid a i ∣ → Code n ∣ merid a i ∣) _+ₖ_ _+ₖ_ + helper n a = + toPathP (funExt + (trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelΠ (3 + n) (λ _ → isOfHLevelTrunc (3 + n))) _ _) + λ x → funExt + (trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) + λ y → help n x y a))) + + Code-add : {n : ℕ} (x : _) → Code n ∣ north ∣ → Code n x → Code n x + Code-add {n = n} = + trElim (λ x → isOfHLevelΠ (4 + n) λ _ → isOfHLevelΠ (4 + n) λ _ → isOfHLevelSuc (3 + n) (hLevCode {n = n} x)) + Code-add' + + encode-hom : {n : ℕ} {x : _} (q : 0ₖ _ ≡ 0ₖ _) (p : 0ₖ _ ≡ x) + → encode (q ∙ p) ≡ Code-add {n = n} x (encode q) (encode p) + encode-hom {n = n} q = J (λ x p → encode (q ∙ p) ≡ Code-add {n = n} x (encode q) (encode p)) + (cong encode (sym (rUnit q)) + ∙∙ sym (rUnitₖ _ (encode q)) + ∙∙ cong (encode q +ₖ_) (cong ∣_∣ (sym (transportRefl _)))) + +stabSpheres : (n : ℕ) → Iso (coHomK (suc n)) (typ (Ω (coHomK-ptd (2 + n)))) +fun (stabSpheres n) = decode _ +inv' (stabSpheres n) = encode +rightInv (stabSpheres n) p = decode-encode p +leftInv (stabSpheres n) = + trElim (λ _ → isOfHLevelPath (3 + n) (isOfHLevelTrunc (3 + n)) _ _) + λ a → cong encode (congFunct ∣_∣ (merid a) (sym (merid (ptSn (suc n))))) + ∙∙ (λ i → transport (congFunct (Code n) (cong ∣_∣ (merid a)) + (cong ∣_∣ (sym (merid (ptSn (suc n))))) i) ∣ ptSn (suc n) ∣) + ∙∙ (substComposite (λ x → x) + (cong (Code n) (cong ∣_∣ (merid a))) + (cong (Code n) (cong ∣_∣ (sym (merid (ptSn (suc n)))))) ∣ ptSn (suc n) ∣ + ∙∙ cong (transport (λ i → Code n ∣ merid (ptSn (suc n)) (~ i) ∣)) + (transportRefl (∣ (ptSn (suc n)) ∣ +ₖ ∣ a ∣) ∙ lUnitₖ (suc n) ∣ a ∣) + ∙∙ symMeridLem n (ptSn (suc n)) ∣ a ∣ + ∙∙ cong (∣ a ∣ +ₖ_) -0ₖ + ∙∙ rUnitₖ (suc n) ∣ a ∣) + +Iso-Kn-ΩKn+1 : (n : HLevel) → Iso (coHomK n) (typ (Ω (coHomK-ptd (suc n)))) +Iso-Kn-ΩKn+1 zero = invIso (compIso (congIso (truncIdempotentIso _ isGroupoidS¹)) ΩS¹Isoℤ) +Iso-Kn-ΩKn+1 (suc n) = stabSpheres n + +Kn≃ΩKn+1 : {n : ℕ} → coHomK n ≃ typ (Ω (coHomK-ptd (suc n))) +Kn≃ΩKn+1 {n = n} = isoToEquiv (Iso-Kn-ΩKn+1 n) + +-- Some properties of the Iso +Kn→ΩKn+1 : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) +Kn→ΩKn+1 n = Iso.fun (Iso-Kn-ΩKn+1 n) + +ΩKn+1→Kn : (n : ℕ) → typ (Ω (coHomK-ptd (suc n))) → coHomK n +ΩKn+1→Kn n = Iso.inv (Iso-Kn-ΩKn+1 n) + +Kn→ΩKn+10ₖ : (n : ℕ) → Kn→ΩKn+1 n (0ₖ n) ≡ refl +Kn→ΩKn+10ₖ zero = sym (rUnit refl) +Kn→ΩKn+10ₖ (suc n) i j = ∣ (rCancel (merid (ptSn (suc n))) i j) ∣ + +ΩKn+1→Kn-refl : (n : ℕ) → ΩKn+1→Kn n refl ≡ 0ₖ n +ΩKn+1→Kn-refl zero = refl +ΩKn+1→Kn-refl (suc zero) = refl +ΩKn+1→Kn-refl (suc (suc n)) = refl + +Kn≃ΩKn+1∙ : {n : ℕ} → coHomK-ptd n ≡ (Ω (coHomK-ptd (suc n))) +Kn≃ΩKn+1∙ {n = n} = ua∙ Kn≃ΩKn+1 (Kn→ΩKn+10ₖ n) + +Kn→ΩKn+1-hom : (n : ℕ) (x y : coHomK n) → Kn→ΩKn+1 n (x +[ n ]ₖ y) ≡ Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n y +Kn→ΩKn+1-hom zero x y = (λ j i → hfill (doubleComp-faces (λ i₁ → ∣ base ∣) (λ _ → ∣ base ∣) i) + (inS (∣ intLoop (x ℤ+ y) i ∣)) (~ j)) + ∙∙ (λ j i → ∣ intLoop-hom x y (~ j) i ∣) + ∙∙ (congFunct ∣_∣ (intLoop x) (intLoop y) + ∙ cong₂ _∙_ (λ j i → hfill (doubleComp-faces (λ i₁ → ∣ base ∣) (λ _ → ∣ base ∣) i) + (inS (∣ intLoop x i ∣)) j) + λ j i → hfill (doubleComp-faces (λ i₁ → ∣ base ∣) (λ _ → ∣ base ∣) i) + (inS (∣ intLoop y i ∣)) j) +Kn→ΩKn+1-hom (suc n) = σ-hom + +ΩKn+1→Kn-hom : (n : ℕ) (x y : Path (coHomK (suc n)) (0ₖ _) (0ₖ _)) + → ΩKn+1→Kn n (x ∙ y) ≡ ΩKn+1→Kn n x +[ n ]ₖ ΩKn+1→Kn n y +ΩKn+1→Kn-hom zero p q = + cong winding (congFunct (trRec isGroupoidS¹ (λ x → x)) p q) + ∙ winding-hom (cong (trRec isGroupoidS¹ (λ x → x)) p) (cong (trRec isGroupoidS¹ (λ x → x)) q) +ΩKn+1→Kn-hom (suc n) = encode-hom + +Kn→ΩKn+1-ₖ : (n : ℕ) (x : coHomK n) → Kn→ΩKn+1 n (-ₖ x) ≡ sym (Kn→ΩKn+1 n x) +Kn→ΩKn+1-ₖ n x = + lUnit _ + ∙∙ cong (_∙ Kn→ΩKn+1 n (-ₖ x)) (sym (lCancel _)) + ∙∙ sym (assoc∙ _ _ _) + ∙∙ cong (sym (Kn→ΩKn+1 n x) ∙_) help + ∙∙ sym (rUnit _) + where + help : Kn→ΩKn+1 n x ∙ Kn→ΩKn+1 n (-ₖ x) ≡ refl + help = sym (Kn→ΩKn+1-hom n x (-ₖ x)) ∙∙ cong (Kn→ΩKn+1 n) (rCancelₖ n x) ∙∙ Kn→ΩKn+10ₖ n + +isHomogeneousKn : (n : HLevel) → isHomogeneous (coHomK-ptd n) +isHomogeneousKn n = + subst isHomogeneous + (sym (ΣPathP (ua Kn≃ΩKn+1 , ua-gluePath _ (Kn→ΩKn+10ₖ n)))) + (isHomogeneousPath _ _) + +-- With the equivalence Kn≃ΩKn+1, we get that the two definitions of cohomology groups agree +open IsGroupHom + +coHom≅coHomΩ : ∀ {ℓ} (n : ℕ) (A : Type ℓ) → GroupIso (coHomGr n A) (coHomGrΩ n A) +fun (fst (coHom≅coHomΩ n A)) = map λ f a → Kn→ΩKn+1 n (f a) +inv' (fst (coHom≅coHomΩ n A)) = map λ f a → ΩKn+1→Kn n (f a) +rightInv (fst (coHom≅coHomΩ n A)) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ f → cong ∣_∣₂ (funExt λ x → rightInv (Iso-Kn-ΩKn+1 n) (f x)) +leftInv (fst (coHom≅coHomΩ n A)) = + sElim (λ _ → isOfHLevelPath 2 § _ _) + λ f → cong ∣_∣₂ (funExt λ x → leftInv (Iso-Kn-ΩKn+1 n) (f x)) +snd (coHom≅coHomΩ n A) = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 § _ _) + λ f g → cong ∣_∣₂ (funExt λ x → Kn→ΩKn+1-hom n (f x) (g x))) + +module lockedKnIso (key : Unit') where + Kn→ΩKn+1' : (n : ℕ) → coHomK n → typ (Ω (coHomK-ptd (suc n))) + Kn→ΩKn+1' n = lock key (Iso.fun (Iso-Kn-ΩKn+1 n)) + + ΩKn+1→Kn' : (n : ℕ) → typ (Ω (coHomK-ptd (suc n))) → coHomK n + ΩKn+1→Kn' n = lock key (Iso.inv (Iso-Kn-ΩKn+1 n)) + + ΩKn+1→Kn→ΩKn+1 : (n : ℕ) → (x : typ (Ω (coHomK-ptd (suc n)))) → Kn→ΩKn+1' n (ΩKn+1→Kn' n x) ≡ x + ΩKn+1→Kn→ΩKn+1 n x = pm key + where + pm : (key : Unit') → lock key (Iso.fun (Iso-Kn-ΩKn+1 n)) (lock key (Iso.inv (Iso-Kn-ΩKn+1 n)) x) ≡ x + pm unlock = Iso.rightInv (Iso-Kn-ΩKn+1 n) x + + Kn→ΩKn+1→Kn : (n : ℕ) → (x : coHomK n) → ΩKn+1→Kn' n (Kn→ΩKn+1' n x) ≡ x + Kn→ΩKn+1→Kn n x = pm key + where + pm : (key : Unit') → lock key (Iso.inv (Iso-Kn-ΩKn+1 n)) (lock key (Iso.fun (Iso-Kn-ΩKn+1 n)) x) ≡ x + pm unlock = Iso.leftInv (Iso-Kn-ΩKn+1 n) x + +-distrLemma : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n m : ℕ) (f : GroupHom (coHomGr n A) (coHomGr m B)) + (x y : coHom n A) + → fst f (x -[ n ]ₕ y) ≡ fst f x -[ m ]ₕ fst f y +-distrLemma n m f' x y = sym (-cancelRₕ m (f y) (f (x -[ n ]ₕ y))) + ∙∙ cong (λ x → x -[ m ]ₕ f y) (sym (f' .snd .pres· (x -[ n ]ₕ y) y)) + ∙∙ cong (λ x → x -[ m ]ₕ f y) ( cong f (-+cancelₕ n _ _)) + where + f = fst f' + +-- HLevel stuff for cup product +isContr-↓∙ : (n : ℕ) → isContr (coHomK-ptd (suc n) →∙ coHomK-ptd n) +fst (isContr-↓∙ zero) = (λ _ → 0) , refl +snd (isContr-↓∙ zero) (f , p) = + Σ≡Prop (λ f → isSetℤ _ _) + (funExt (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelSuc 2 isSetℤ) _ _) + (toPropElim (λ _ → isSetℤ _ _) (sym p)))) +fst (isContr-↓∙ (suc n)) = (λ _ → 0ₖ _) , refl +fst (snd (isContr-↓∙ (suc n)) f i) x = + trElim {B = λ x → 0ₖ (suc n) ≡ fst f x} + (λ _ → isOfHLevelPath (4 + n) (isOfHLevelSuc (3 + n) (isOfHLevelTrunc (3 + n))) _ _) + (sphereElim _ (λ _ → isOfHLevelTrunc (3 + n) _ _) + (sym (snd f))) x i +snd (snd (isContr-↓∙ (suc n)) f i) j = snd f (~ i ∨ j) + +isContr-↓∙' : (n : ℕ) → isContr (S₊∙ (suc n) →∙ coHomK-ptd n) +fst (isContr-↓∙' zero) = (λ _ → 0) , refl +snd (isContr-↓∙' zero) (f , p) = + Σ≡Prop (λ f → isSetℤ _ _) + (funExt (toPropElim (λ _ → isSetℤ _ _) (sym p))) +fst (isContr-↓∙' (suc n)) = (λ _ → 0ₖ _) , refl +fst (snd (isContr-↓∙' (suc n)) f i) x = + sphereElim _ {A = λ x → 0ₖ (suc n) ≡ fst f x} + (λ _ → isOfHLevelTrunc (3 + n) _ _) + (sym (snd f)) x i +snd (snd (isContr-↓∙' (suc n)) f i) j = snd f (~ i ∨ j) + +isOfHLevel→∙Kn : {A : Pointed₀} (n m : ℕ) + → isOfHLevel (suc m) (A →∙ coHomK-ptd n) → isOfHLevel (suc (suc m)) (A →∙ coHomK-ptd (suc n)) +isOfHLevel→∙Kn {A = A} n m hlev = step₃ + where + step₁ : isOfHLevel (suc m) (A →∙ Ω (coHomK-ptd (suc n))) + step₁ = + subst (isOfHLevel (suc m)) + (λ i → A →∙ ua∙ {A = Ω (coHomK-ptd (suc n))} {B = coHomK-ptd n} + (invEquiv Kn≃ΩKn+1) + (ΩKn+1→Kn-refl n) (~ i)) hlev + + step₂ : isOfHLevel (suc m) (typ (Ω (A →∙ coHomK-ptd (suc n) ∙))) + step₂ = isOfHLevelRetractFromIso (suc m) (invIso (invIso (ΩfunExtIso _ _))) step₁ + + step₃ : isOfHLevel (suc (suc m)) (A →∙ coHomK-ptd (suc n)) + step₃ = + isOfHLevelΩ→isOfHLevel m + λ f → subst (λ x → isOfHLevel (suc m) (typ (Ω x))) + (isHomogeneous→∙ (isHomogeneousKn (suc n)) f) + step₂ + +isOfHLevel↑∙ : ∀ n m → isOfHLevel (2 + n) (coHomK-ptd (suc m) →∙ coHomK-ptd (suc (n + m))) +isOfHLevel↑∙ zero m = isOfHLevel→∙Kn m 0 (isContr→isProp (isContr-↓∙ m)) +isOfHLevel↑∙ (suc n) m = isOfHLevel→∙Kn (suc (n + m)) (suc n) (isOfHLevel↑∙ n m) + +isOfHLevel↑∙' : ∀ n m → isOfHLevel (2 + n) (S₊∙ (suc m) →∙ coHomK-ptd (suc (n + m))) +isOfHLevel↑∙' zero m = isOfHLevel→∙Kn m 0 (isContr→isProp (isContr-↓∙' m)) +isOfHLevel↑∙' (suc n) m = isOfHLevel→∙Kn (suc (n + m)) (suc n) (isOfHLevel↑∙' n m) + +→∙KnPath : (A : Pointed₀) (n : ℕ) → Ω (A →∙ coHomK-ptd (suc n) ∙) ≡ (A →∙ coHomK-ptd n ∙) +→∙KnPath A n = + ua∙ (isoToEquiv (ΩfunExtIso A (coHomK-ptd (suc n)))) refl + ∙ (λ i → (A →∙ Kn≃ΩKn+1∙ {n = n} (~ i) ∙)) + +contr∙-lem : (n m : ℕ) + → isContr (coHomK-ptd (suc n) + →∙ (coHomK-ptd (suc m) + →∙ coHomK-ptd (suc (n + m)) ∙)) +fst (contr∙-lem n m) = (λ _ → (λ _ → 0ₖ _) , refl) , refl +snd (contr∙-lem n m) (f , p) = →∙Homogeneous≡ (isHomogeneous→∙ (isHomogeneousKn _)) (sym (funExt help)) + where + help : (x : _) → (f x) ≡ ((λ _ → 0ₖ _) , refl) + help = + trElim (λ _ → isOfHLevelPath (3 + n) + (subst (λ x → isOfHLevel x (coHomK-ptd (suc m) →∙ coHomK-ptd (suc (n + m)))) (λ i → suc (suc (+-comm n 1 i))) + (isOfHLevelPlus' {n = 1} (2 + n) (isOfHLevel↑∙ n m))) _ _) + (sphereElim _ (λ _ → isOfHLevel↑∙ n m _ _) p) + +isOfHLevel↑∙∙ : ∀ n m l + → isOfHLevel (2 + l) (coHomK-ptd (suc n) + →∙ (coHomK-ptd (suc m) + →∙ coHomK-ptd (suc (suc (l + n + m))) ∙)) +isOfHLevel↑∙∙ n m zero = + isOfHLevelΩ→isOfHLevel 0 λ f + → subst + isProp (cong (λ x → typ (Ω x)) + (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousKn _)) f)) + (isOfHLevelRetractFromIso 1 (ΩfunExtIso _ _) h) + where + h : isProp (coHomK-ptd (suc n) + →∙ (Ω (coHomK-ptd (suc m) + →∙ coHomK-ptd (suc (suc (n + m))) ∙))) + h = + subst isProp (λ i → coHomK-ptd (suc n) →∙ (→∙KnPath (coHomK-ptd (suc m)) (suc (n + m)) (~ i))) + (isContr→isProp (contr∙-lem n m)) +isOfHLevel↑∙∙ n m (suc l) = + isOfHLevelΩ→isOfHLevel (suc l) + λ f → + subst + (isOfHLevel (2 + l)) (cong (λ x → typ (Ω x)) + (isHomogeneous→∙ (isHomogeneous→∙ (isHomogeneousKn _)) f)) + (isOfHLevelRetractFromIso (2 + l) (ΩfunExtIso _ _) h) + where + h : isOfHLevel (2 + l) + (coHomK-ptd (suc n) + →∙ (Ω (coHomK-ptd (suc m) + →∙ coHomK-ptd (suc (suc (suc (l + n + m)))) ∙))) + h = + subst (isOfHLevel (2 + l)) + (λ i → coHomK-ptd (suc n) →∙ →∙KnPath (coHomK-ptd (suc m)) (suc (suc (l + n + m))) (~ i)) + (isOfHLevel↑∙∙ n m l) + +-- Misc. +isoType→isoCohom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) + → Iso A B + → GroupIso (coHomGr n A) (coHomGr n B) +fst (isoType→isoCohom n is) = setTruncIso (domIso is) +snd (isoType→isoCohom n is) = + makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + (λ _ _ → refl)) diff --git a/Cubical/ZCohomology/RingStructure/RingLaws.agda b/Cubical/ZCohomology/RingStructure/RingLaws.agda index 58d715584..48b9ca067 100644 --- a/Cubical/ZCohomology/RingStructure/RingLaws.agda +++ b/Cubical/ZCohomology/RingStructure/RingLaws.agda @@ -658,4 +658,9 @@ rUnit⌣ (suc n) = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ f → cong ∣_∣₂ (funExt λ x → rUnitₖ _ (f x)) +-ₕDistᵣ : ∀ {ℓ} {A : Type ℓ} (n m : ℕ) (x : coHom n A) (y : coHom m A) → (-ₕ (x ⌣ y)) ≡ x ⌣ (-ₕ y) +-ₕDistᵣ n m = + sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f g → cong ∣_∣₂ (funExt λ x → -Distᵣ n m (f x) (g x)) + -- TODO : Graded ring structure From ff786ae582fe3ae50b47f7d512b728ba6e4242c7 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 28 Sep 2021 11:54:25 +0200 Subject: [PATCH 67/89] stuff --- Cubical/ZCohomology/Gysin.agda | 1480 ++++++++++++++++++++++++++++++- Cubical/ZCohomology/Gysin2.agda | 1164 +----------------------- 2 files changed, 1460 insertions(+), 1184 deletions(-) diff --git a/Cubical/ZCohomology/Gysin.agda b/Cubical/ZCohomology/Gysin.agda index 1687d4e33..25660e4e9 100644 --- a/Cubical/ZCohomology/Gysin.agda +++ b/Cubical/ZCohomology/Gysin.agda @@ -7,10 +7,12 @@ open import Cubical.ZCohomology.GroupStructure open import Cubical.ZCohomology.Properties open import Cubical.ZCohomology.MayerVietorisUnreduced open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Wedge open import Cubical.ZCohomology.Groups.Sn open import Cubical.ZCohomology.RingStructure.CupProduct open import Cubical.ZCohomology.RingStructure.RingLaws open import Cubical.ZCohomology.RingStructure.GradedCommutativity +open import Cubical.Foundations.Path open import Cubical.Data.Sum open import Cubical.Relation.Nullary @@ -54,7 +56,7 @@ open import Cubical.HITs.S1 renaming (_·_ to _*_) open import Cubical.HITs.Truncation renaming (rec to trRec ; elim to trElim ; elim2 to trElim2) open import Cubical.HITs.SetTruncation - renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap) + renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap ; elim3 to sElim3) open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim) @@ -68,7 +70,8 @@ open import Cubical.Homotopy.Hopf open import Cubical.HITs.SetQuotients renaming (_/_ to _/'_) -SES→Iso : ∀ {ℓ ℓ'} {L R : Group ℓ-zero} +-- move to group stuff +SES→isEquiv : ∀ {ℓ ℓ'} {L R : Group ℓ-zero} → {G : Group ℓ} {H : Group ℓ'} → UnitGroup ≡ L → UnitGroup ≡ R @@ -76,7 +79,7 @@ SES→Iso : ∀ {ℓ ℓ'} {L R : Group ℓ-zero} → ((x : _) → isInKer midhom x → isInIm lhom x) → ((x : _) → isInKer rhom x → isInIm midhom x) → isEquiv (fst midhom) -SES→Iso {R = R} {G = G} {H = H} = +SES→isEquiv {R = R} {G = G} {H = H} = J (λ L _ → UnitGroup ≡ R → (lhom : GroupHom L G) (midhom : GroupHom G H) (rhom : GroupHom H R) → @@ -106,12 +109,12 @@ SES→Iso {R = R} {G = G} {H = H} = (λ s → sym (snd s) ∙ IsGroupHom.pres1 (snd lhom)) (lexact _ inker) BijectionIso.surj bijIso' x = rexact x refl +-- The pushout of the hopf invariant HopfInvariantPush : (n : ℕ) → (f : S₊ (3 +ℕ n +ℕ n) → S₊ (2 +ℕ n)) → Type _ HopfInvariantPush n f = Pushout (λ _ → tt) f - Hopfα : (n : ℕ) → (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) → coHom (2 +ℕ n) (HopfInvariantPush n (fst f)) @@ -128,6 +131,8 @@ Hopfβ : (n : ℕ) → coHom (4 +ℕ (n +ℕ n)) (HopfInvariantPush n (fst f)) Hopfβ n f = fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) ∣ ∣_∣ ∣₂ + +-- To define the Hopf invariant, we need to establish the non-trivial isos of Hⁿ(HopfInvariant). module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) where module M = MV _ _ _ (λ _ → tt) (fst f) ¬lem : (n m : ℕ) → ¬ suc (suc (m +ℕ n)) ≡ suc n @@ -148,7 +153,7 @@ module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) whe abstract help : isEquiv (fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n))) help = - SES→Iso + SES→isEquiv (fst (GroupPath _ _) (invGroupEquiv (isContr→≃Unit (isOfHLevelΣ 0 @@ -200,25 +205,23 @@ module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) whe cong (Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) (IsGroupHom.pres· (isGroupHomInv GroupEquiv1) x y) ∙ IsGroupHom.pres· (snd (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) _ _ - - -- compGroupIso (GroupEquiv→GroupIso (invGroupEquiv GroupEquiv1)) (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) -h33 : (n : ℕ) → Iso.inv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) 1 ≡ ∣ ∣_∣ ∣₂ -h33 zero = Iso.leftInv (fst (Hⁿ-Sⁿ≅ℤ (suc zero))) _ -h33 (suc n) = (λ i → Iso.inv +Hⁿ-Sⁿ≅ℤ-nice-generator : (n : ℕ) → Iso.inv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) 1 ≡ ∣ ∣_∣ ∣₂ +Hⁿ-Sⁿ≅ℤ-nice-generator zero = Iso.leftInv (fst (Hⁿ-Sⁿ≅ℤ (suc zero))) _ +Hⁿ-Sⁿ≅ℤ-nice-generator (suc n) = (λ i → Iso.inv (fst - (suspensionAx-Sn (suc n) (suc n))) (h33 n i)) + (suspensionAx-Sn (suc n) (suc n))) (Hⁿ-Sⁿ≅ℤ-nice-generator n i)) ∙ cong ∣_∣₂ (funExt λ { north → refl ; south → cong ∣_∣ₕ (merid north) ; (merid a i) j → ∣ compPath-filler (merid a) (sym (merid north)) (~ j) i ∣}) -d-β : (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) +Hopfβ↦1 : (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) → Iso.fun (fst (Hopfβ-Iso n f)) (Hopfβ n f) ≡ 1 -d-β n f = sym (cong (Iso.fun (fst (Hopfβ-Iso n f))) h) ∙ Iso.rightInv (fst (Hopfβ-Iso n f)) (pos 1) +Hopfβ↦1 n f = sym (cong (Iso.fun (fst (Hopfβ-Iso n f))) h) ∙ Iso.rightInv (fst (Hopfβ-Iso n f)) (pos 1) where h : Iso.inv (fst (Hopfβ-Iso n f)) (pos 1) ≡ Hopfβ n f h = (λ i → fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) - (h33 _ i)) + (Hⁿ-Sⁿ≅ℤ-nice-generator _ i)) ∙ cong ∣_∣₂ (funExt (λ { (inl x) → refl ; (inr x) → refl ; (push a i) → refl})) @@ -300,7 +303,7 @@ Hopfα-Iso-α : (n : ℕ) (f : _) → Iso.fun (fst (Hopfα-Iso n f)) (Hopfα n f) ≡ 1 Hopfα-Iso-α n f = - sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (h33 n)) + sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (Hⁿ-Sⁿ≅ℤ-nice-generator n)) ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) where hz : Iso.fun (h n f) (Hopfα n f) ≡ ∣ ∣_∣ ∣₂ @@ -309,6 +312,8 @@ Hopfα-Iso-α n f = _·₀ₕ_ : ∀ {ℓ} {n : ℕ} {A : Type ℓ} → ℤ → coHom n A → coHom n A _·₀ₕ_ {n = n} a = sMap λ f x → a ·₀ f x + +-- remove HopfInvariant : (n : ℕ) → (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Type _ @@ -320,9 +325,7 @@ HopfInvariant n f = (Hopfβ n f)) -⌣-α : (n : ℕ) - → (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → _ +⌣-α : (n : ℕ) → (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → _ ⌣-α n f = Hopfα n f ⌣ Hopfα n f HopfFunction : (n : ℕ) @@ -330,7 +333,7 @@ HopfFunction : (n : ℕ) → ℤ HopfFunction n f = Iso.fun (fst (Hopfβ-Iso n f)) ((subst (λ x → coHom x (HopfInvariantPush n (fst f))) - (λ i → suc (suc (suc (+-suc n n i))))) (⌣-α n f)) -- (f) + (λ i → suc (suc (suc (+-suc n n i))))) (⌣-α n f)) pres- : (e : GroupHom ℤGroup ℤGroup) → (a : ℤ) → fst e (- a) ≡ - fst e a pres- e a = +Comm (fst e (- a)) (pos 0) @@ -351,8 +354,6 @@ pres- e a = +Comm (fst e (- a)) (pos 0) ∙∙ (IsGroupHom.pres1 (snd e) ∙ sym (l _)) - - GroupHomPres· : (e : GroupHom ℤGroup ℤGroup) → (a b : ℤ) → fst e (a · b) ≡ a · fst e b GroupHomPres· e (pos zero) b = IsGroupHom.pres1 (snd e) GroupHomPres· e (pos (suc n)) b = @@ -1880,7 +1881,7 @@ snd (fst ⌣hom) = subst isEquiv (cong fst (GysinS1.ϕ∘j≡ 2)) h3 where h3 : isEquiv (GysinS1.ϕ∘j 2 .fst) h3 = - SES→Iso + SES→isEquiv (GroupPath _ _ .fst (invGroupEquiv (isContr→≃Unit isContrH³E , makeIsGroupHom λ _ _ → refl))) @@ -2150,3 +2151,1438 @@ snd (snd ⌣-pres1-CP2) = ∣ p ∣ -- -- P (inl x) = typ A -- -- P (inr x) = Hopf x -- -- P (push a i₁) = ua (v a) i₁ + + + +open import Cubical.Algebra.AbGroup + +open import Cubical.Homotopy.Loopspace + +open import Cubical.HITs.Join + +open import Cubical.Homotopy.Hopf + +open import Cubical.HITs.SetQuotients renaming (_/_ to _/'_) +-- open import Cubical.ZCohomology.Gysin + +rUnitℤ· : ∀ {ℓ} (G : Group ℓ) → (x : ℤ) → (x ℤ[ G ]· GroupStr.1g (snd G)) + ≡ GroupStr.1g (snd G) +rUnitℤ· G (pos zero) = refl +rUnitℤ· G (pos (suc n)) = + cong (GroupStr._·_ (snd G) (GroupStr.1g (snd G))) + (rUnitℤ· G (pos n)) + ∙ GroupStr.lid (snd G) (GroupStr.1g (snd G)) +rUnitℤ· G (negsuc zero) = GroupTheory.inv1g G +rUnitℤ· G (negsuc (suc n)) = + cong₂ (GroupStr._·_ (snd G)) (GroupTheory.inv1g G) (rUnitℤ· G (negsuc n)) + ∙ GroupStr.lid (snd G) (GroupStr.1g (snd G)) + +rUnitℤ·ℤ : (x : ℤ) → (x ℤ[ ℤGroup ]· 1) ≡ x +rUnitℤ·ℤ (pos zero) = refl +rUnitℤ·ℤ (pos (suc n)) = cong (pos 1 +_) (rUnitℤ·ℤ (pos n)) ∙ sym (pos+ 1 n) +rUnitℤ·ℤ (negsuc zero) = refl +rUnitℤ·ℤ (negsuc (suc n)) = cong (-1 +_) (rUnitℤ·ℤ (negsuc n)) ∙ +Comm (negsuc 0) (negsuc n) + +private + precommℤ : ∀ {ℓ} (G : Group ℓ) → (g : fst G) (y : ℤ) → (snd G GroupStr.· (y ℤ[ G ]· g)) g + ≡ (sucℤ y ℤ[ G ]· g) + precommℤ G g (pos zero) = GroupStr.lid (snd G) g + ∙ sym (GroupStr.rid (snd G) g) + precommℤ G g (pos (suc n)) = + sym (GroupStr.assoc (snd G) _ _ _) + ∙ cong ((snd G GroupStr.· g)) (precommℤ G g (pos n)) + precommℤ G g (negsuc zero) = + GroupStr.invl (snd G) g + precommℤ G g (negsuc (suc n)) = + sym (GroupStr.assoc (snd G) _ _ _) + ∙ cong ((snd G GroupStr.· GroupStr.inv (snd G) g)) (precommℤ G g (negsuc n)) + ∙ negsucLem n + where + negsucLem : (n : ℕ) → (snd G GroupStr.· GroupStr.inv (snd G) g) + (sucℤ (negsuc n) ℤ[ G ]· g) + ≡ (sucℤ (negsuc (suc n)) ℤ[ G ]· g) + negsucLem zero = (GroupStr.rid (snd G) _) + negsucLem (suc n) = refl + +distrℤ· : ∀ {ℓ} (G : Group ℓ) → (g : fst G) (x y : ℤ) + → ((x + y) ℤ[ G ]· g) + ≡ GroupStr._·_ (snd G) (x ℤ[ G ]· g) (y ℤ[ G ]· g) +distrℤ· G' g (pos zero) y = cong (_ℤ[ G' ]· g) (+Comm 0 y) + ∙ sym (GroupStr.lid (snd G') _) +distrℤ· G' g (pos (suc n)) (pos n₁) = + cong (_ℤ[ G' ]· g) (sym (pos+ (suc n) n₁)) + ∙ cong (GroupStr._·_ (snd G') g) (cong (_ℤ[ G' ]· g) (pos+ n n₁) ∙ distrℤ· G' g (pos n) (pos n₁)) + ∙ GroupStr.assoc (snd G') _ _ _ +distrℤ· G' g (pos (suc n)) (negsuc zero) = + distrℤ· G' g (pos n) 0 + ∙ ((GroupStr.rid (snd G') (pos n ℤ[ G' ]· g) ∙ sym (GroupStr.lid (snd G') (pos n ℤ[ G' ]· g))) + ∙ cong (λ x → GroupStr._·_ (snd G') x (pos n ℤ[ G' ]· g)) + (sym (GroupStr.invl (snd G') g)) ∙ sym (GroupStr.assoc (snd G') _ _ _)) + ∙ (GroupStr.assoc (snd G') _ _ _) + ∙ cong (λ x → GroupStr._·_ (snd G') x (pos n ℤ[ G' ]· g)) (GroupStr.invl (snd G') g) + ∙ GroupStr.lid (snd G') _ + ∙ sym (GroupStr.rid (snd G') _) + ∙ (cong (GroupStr._·_ (snd G') (pos n ℤ[ G' ]· g)) (sym (GroupStr.invr (snd G') g)) + ∙ GroupStr.assoc (snd G') _ _ _) + ∙ cong (λ x → GroupStr._·_ (snd G') x (GroupStr.inv (snd G') g)) + (precommℤ G' g (pos n)) +distrℤ· G' g (pos (suc n)) (negsuc (suc n₁)) = + cong (_ℤ[ G' ]· g) (predℤ+negsuc n₁ (pos (suc n))) + ∙∙ distrℤ· G' g (pos n) (negsuc n₁) + ∙∙ (cong (λ x → GroupStr._·_ (snd G') x (negsuc n₁ ℤ[ G' ]· g)) + ((sym (GroupStr.rid (snd G') (pos n ℤ[ G' ]· g)) + ∙ cong (GroupStr._·_ (snd G') (pos n ℤ[ G' ]· g)) (sym (GroupStr.invr (snd G') g))) + ∙∙ GroupStr.assoc (snd G') _ _ _ + ∙∙ cong (λ x → GroupStr._·_ (snd G') x (GroupStr.inv (snd G') g)) (precommℤ G' g (pos n) )) + ∙ sym (GroupStr.assoc (snd G') _ _ _)) +distrℤ· G' g (negsuc zero) y = + cong (_ℤ[ G' ]· g) (+Comm -1 y) ∙ lol1 y + module _ where + lol1 : (y : ℤ) → (predℤ y ℤ[ G' ]· g) ≡ (snd G' GroupStr.· GroupStr.inv (snd G') g) (y ℤ[ G' ]· g) + lol1 (pos zero) = sym (GroupStr.rid (snd G') _) + lol1 (pos (suc n)) = + sym (GroupStr.lid (snd G') ((pos n ℤ[ G' ]· g))) + ∙∙ cong (λ x → GroupStr._·_ (snd G') x (pos n ℤ[ G' ]· g)) (sym (GroupStr.invl (snd G') g)) + ∙∙ sym (GroupStr.assoc (snd G') _ _ _) + lol1 (negsuc n) = refl +distrℤ· G' g (negsuc (suc n)) y = + cong (_ℤ[ G' ]· g) (+Comm (negsuc (suc n)) y) + ∙∙ lol1 G' g 0 (y +negsuc n) + ∙∙ cong (snd G' GroupStr.· GroupStr.inv (snd G') g) + (cong (_ℤ[ G' ]· g) (+Comm y (negsuc n)) ∙ distrℤ· G' g (negsuc n) y) + ∙ (GroupStr.assoc (snd G') _ _ _) + + + +HopfInvariantPushElim : ∀ {ℓ} n → (f : _) → {P : HopfInvariantPush n f → Type ℓ} + → (isOfHLevel (suc (suc (suc (suc (n +ℕ n))))) (P (inl tt))) + → (e : P (inl tt)) + (g : (x : _) → P (inr x)) + (r : PathP (λ i → P (push north i)) e (g (f north))) + → (x : _) → P x +HopfInvariantPushElim n f {P = P} hlev e g r (inl x) = e +HopfInvariantPushElim n f {P = P} hlev e g r (inr x) = g x +HopfInvariantPushElim n f {P = P} hlev e g r (push a i₁) = help a i₁ + where + help : (a : Susp (Susp (S₊ (suc (n +ℕ n))))) → PathP (λ i → P (push a i)) e (g (f a)) + help = sphereElim _ (sphereElim _ (λ _ → isProp→isOfHLevelSuc ((suc (suc (n +ℕ n)))) (isPropIsOfHLevel _)) + (isOfHLevelPathP' (suc (suc (suc (n +ℕ n)))) + (subst (isOfHLevel (suc (suc (suc (suc (n +ℕ n)))))) + (cong P (push north)) + hlev) _ _)) r + +transportCohomIso : ∀ {ℓ} {A : Type ℓ} {n m : ℕ} + → (p : n ≡ m) + → GroupIso (coHomGr n A) (coHomGr m A) +Iso.fun (fst (transportCohomIso {A = A} p)) = subst (λ n → coHom n A) p +Iso.inv (fst (transportCohomIso {A = A} p)) = subst (λ n → coHom n A) (sym p) +Iso.rightInv (fst (transportCohomIso p)) = transportTransport⁻ _ +Iso.leftInv (fst (transportCohomIso p)) = transportTransport⁻ _ +snd (transportCohomIso {A = A} {n = n} {m = m} p) = + makeIsGroupHom (λ x y → help x y p) + where + help : (x y : coHom n A) (p : n ≡ m) → subst (λ n → coHom n A) p (x +ₕ y) + ≡ subst (λ n → coHom n A) p x +ₕ subst (λ n → coHom n A) p y + help x y = J (λ m p → subst (λ n → coHom n A) p (x +ₕ y) + ≡ subst (λ n → coHom n A) p x +ₕ subst (λ n → coHom n A) p y) + (transportRefl (x +ₕ y) ∙ cong₂ _+ₕ_ (sym (transportRefl x)) (sym (transportRefl y))) + +α-hopfMap : abs (HopfFunction zero (hopfFun , λ _ → hopfFun (snd (S₊∙ 3)))) ≡ suc zero +α-hopfMap = cong abs (cong (Iso.fun (fst (Hopfβ-Iso zero (hopfFun , refl)))) + (transportRefl (⌣-α 0 (hopfFun , refl)))) ∙ l3 (sym CP2≡CP2') GysinS1.e (Hopfα zero (hopfFun , refl)) + (l isGenerator≃ℤ-e) + (GroupIso→GroupEquiv (Hopfα-Iso 0 (hopfFun , refl)) , refl) + (snd (fst ⌣hom)) + (GroupIso→GroupEquiv (Hopfβ-Iso zero (hopfFun , refl))) + where + l : Σ-syntax (GroupIso (coHomGr 2 CP2) ℤGroup) + (λ ϕ → abs (Iso.fun (fst ϕ) GysinS1.e) ≡ 1) + → Σ-syntax (GroupEquiv (coHomGr 2 CP2) ℤGroup) + (λ ϕ → abs (fst (fst ϕ) GysinS1.e) ≡ 1) + l p = (GroupIso→GroupEquiv (fst p)) , (snd p) + +·Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (S₊∙ n →∙ A) + → (S₊∙ n →∙ A) + → (S₊∙ n →∙ A) +·Π {A = A} {n = zero} p q = (λ _ → pt A) , refl +fst (·Π {A = A} {n = suc zero} (f , p) (g , q)) base = pt A +fst (·Π {A = A} {n = suc zero} (f , p) (g , q)) (loop j) = + ((sym p ∙∙ cong f loop ∙∙ p) ∙ (sym q ∙∙ cong g loop ∙∙ q)) j +snd (·Π {A = A} {n = suc zero} (f , p) (g , q)) = refl +fst (·Π {A = A} {n = suc (suc n)} (f , p) (g , q)) north = pt A +fst (·Π {A = A} {n = suc (suc n)} (f , p) (g , q)) south = pt A +fst (·Π {A = A} {n = suc (suc n)} (f , p) (g , q)) (merid a j) = + ((sym p ∙∙ cong f (merid a ∙ sym (merid (ptSn (suc n)))) ∙∙ p) + ∙ (sym q ∙∙ cong g (merid a ∙ sym (merid (ptSn (suc n)))) ∙∙ q)) j +snd (·Π {A = A} {n = suc (suc n)} (f , p) (g , q)) = refl + +-Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (S₊∙ n →∙ A) + → (S₊∙ n →∙ A) +-Π {n = zero} f = f +fst (-Π {A = A} {n = suc zero} f) base = fst f base +fst (-Π {A = A} {n = suc zero} f) (loop j) = fst f (loop (~ j)) +snd (-Π {A = A} {n = suc zero} f) = snd f +fst (-Π {A = A} {n = suc (suc n)} f) north = fst f north +fst (-Π {A = A} {n = suc (suc n)} f) south = fst f north +fst (-Π {A = A} {n = suc (suc n)} f) (merid a j) = fst f ((merid a ∙ sym (merid (ptSn _))) (~ j)) +snd (-Π {A = A} {n = suc (suc n)} f) = snd f + +open import Cubical.Foundations.Equiv.HalfAdjoint +module _ {ℓ : Level} {A : Pointed ℓ} where +flipLoopSpace' : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → ((Ω^ (suc n)) A) ≡ (Ω^ n) (Ω A) +flipLoopSpace' {A = A} zero = refl +flipLoopSpace' {A = A} (suc n) = cong Ω (flipLoopSpace' {A = A} n) + +flipLoopSpace : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → Iso (fst ((Ω^ (suc n)) A)) (fst ((Ω^ n) (Ω A))) +flipLoopSpace {A = A} n = pathToIso (cong fst (flipLoopSpace' n)) + +Ω'∙ : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → fst ((Ω^ n) (Ω A)) → fst ((Ω^ n) (Ω A)) → fst ((Ω^ n) (Ω A)) +Ω'∙ zero p q = p ∙ q +Ω'∙ (suc n) p q = p ∙ q + +flipLoopSpacepres· : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (f g : fst ((Ω^ n) (Ω A))) + → Iso.inv (flipLoopSpace n) (Ω'∙ n f g) + ≡ (Iso.inv (flipLoopSpace n) f) ∙ (Iso.inv (flipLoopSpace n) g) +flipLoopSpacepres· zero f g = transportRefl (f ∙ g) ∙ cong₂ _∙_ (sym (transportRefl f)) (sym (transportRefl g)) +flipLoopSpacepres· {A = A} (suc n) f g i = + transp (λ j → flipLoopSpace' {A = A} n (~ i ∧ ~ j) .snd ≡ flipLoopSpace' n (~ i ∧ ~ j) .snd) i + (transp (λ j → flipLoopSpace' {A = A} n (~ i ∨ ~ j) .snd ≡ flipLoopSpace' n (~ i ∨ ~ j) .snd) (~ i) f + ∙ transp (λ j → flipLoopSpace' {A = A} n (~ i ∨ ~ j) .snd ≡ flipLoopSpace' n (~ i ∨ ~ j) .snd) (~ i) g) + +module _ {ℓ : Level} {A : Pointed ℓ} where + convertLoopFun→ : (n : ℕ) → (S₊∙ (suc n) →∙ Ω A) → (S₊∙ (suc (suc n)) →∙ A) + fst (convertLoopFun→ n f) north = pt A + fst (convertLoopFun→ n f) south = pt A + fst (convertLoopFun→ n f) (merid a i₁) = fst f a i₁ + snd (convertLoopFun→ n f) = refl + + convertLoopFun-pres∙ : + (n : ℕ) (f g : (S₊∙ (suc n) →∙ Ω A)) + → convertLoopFun→ n (·Π f g) ≡ ·Π (convertLoopFun→ n f) (convertLoopFun→ n g) + convertLoopFun-pres∙ n f g = ΣPathP ((funExt l) , refl) + where + J2 : ∀ {ℓ ℓ'} {A : Type ℓ} {x : A} + → (P : (p q : x ≡ x) → (refl ≡ p) → (refl ≡ q) → Type ℓ') + → P refl refl refl refl + → (p q : x ≡ x) (P' : _) (Q : _) → P p q P' Q + J2 P b p q = + J (λ p P' → (Q₁ : refl ≡ q) → P p q P' Q₁) + (J (λ q Q → P refl q refl Q) b) + + J4 : ∀ {ℓ ℓ'} {A : Type ℓ} {x : A} + → (P : (p q r s : x ≡ x) → (refl ≡ p) → (p ≡ q) → (refl ≡ r) → (r ≡ s) → Type ℓ') + → P refl refl refl refl refl refl refl refl + → (p q r s : x ≡ x) (P' : _) (Q : _) (R : _) (S : _) → P p q r s P' Q R S + J4 P b p q r s = + J (λ p P' → (Q₁ : p ≡ q) (R : refl ≡ r) (S₁ : r ≡ s) → P p q r s P' Q₁ R S₁) + (J (λ q Q₁ → (R : refl ≡ r) (S₁ : r ≡ s) → P refl q r s refl Q₁ R S₁) + (J (λ r R → (S₁ : r ≡ s) → P refl refl r s refl refl R S₁) + (J (λ s S₁ → P refl refl refl s refl refl refl S₁) + b))) + + looper-help : (n : ℕ) (f g : (S₊∙ (suc n) →∙ Ω A)) (a : _) + → fst (·Π f g) a ≡ fst f a ∙ fst g a + looper-help zero f g base = rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g)) + looper-help zero f g (loop i) k = + lazy-fill _ _ (sym (snd f)) (sym (snd g)) (cong (fst f) loop) (cong (fst g) loop) k i + where + lazy-fill : ∀ {ℓ} {A : Type ℓ} {x : A} (p q : x ≡ x) + (prefl : refl ≡ p) (qrefl : refl ≡ q) (fl : p ≡ p) (fr : q ≡ q) + → PathP (λ i → (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i + ≡ (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i) + ((prefl ∙∙ fl ∙∙ sym prefl) ∙ (qrefl ∙∙ fr ∙∙ sym qrefl)) + (cong₂ _∙_ fl fr) + lazy-fill {x = x} = + J2 (λ p q prefl qrefl → (fl : p ≡ p) (fr : q ≡ q) + → PathP (λ i → (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i + ≡ (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i) + ((prefl ∙∙ fl ∙∙ sym prefl) ∙ (qrefl ∙∙ fr ∙∙ sym qrefl)) + (cong₂ _∙_ fl fr)) + λ fl fr → cong₂ _∙_ (sym (rUnit fl)) (sym (rUnit fr)) + ◁ flipSquare (sym (rUnit (rUnit refl)) + ◁ (flipSquare ((λ j → cong (λ x → rUnit x j) fl + ∙ cong (λ x → lUnit x j) fr) + ▷ sym (cong₂Funct _∙_ fl fr)) + ▷ (rUnit (rUnit refl)))) + looper-help (suc n) f g north = rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g)) + looper-help (suc n) f g south = (rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) + ∙ cong₂ _∙_ (cong (fst f) (merid (ptSn _))) (cong (fst g) (merid (ptSn _))) + looper-help (suc n) f g (merid a j) i = help i j + where + help : PathP (λ i → (rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) i + ≡ ((rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) + ∙ cong₂ _∙_ (cong (fst f) (merid (ptSn _))) (cong (fst g) (merid (ptSn _)))) i) + (cong (fst (·Π f g)) (merid a)) + (cong₂ _∙_ (cong (fst f) (merid a)) (cong (fst g) (merid a))) + help = (cong₂ _∙_ (cong (sym (snd f) ∙∙_∙∙ snd f) (cong-∙ (fst f) (merid a) (sym (merid (ptSn _))))) + ((cong (sym (snd g) ∙∙_∙∙ snd g) (cong-∙ (fst g) (merid a) (sym (merid (ptSn _)))))) + ◁ lazy-help _ _ _ _ (sym (snd f)) (cong (fst f) (merid (ptSn _))) + (sym (snd g)) (cong (fst g) (merid (ptSn _))) + (cong (fst f) (merid a)) (cong (fst g) (merid a))) + where + lazy-help : ∀ {ℓ} {A : Type ℓ} {x : A} (p-north p-south q-north q-south : x ≡ x) + (prefl : refl ≡ p-north) (p-merid : p-north ≡ p-south) + (qrefl : refl ≡ q-north) (q-merid : q-north ≡ q-south) + (fl : p-north ≡ p-south) (fr : q-north ≡ q-south) + → PathP (λ i → (rUnit refl ∙ cong₂ _∙_ prefl qrefl) i + ≡ ((rUnit refl ∙ cong₂ _∙_ prefl qrefl) + ∙ cong₂ _∙_ p-merid q-merid) i) + ((prefl ∙∙ fl ∙ sym p-merid ∙∙ sym prefl) + ∙ (qrefl ∙∙ fr ∙ sym q-merid ∙∙ sym qrefl)) + (cong₂ _∙_ fl fr) + lazy-help {x = x} = + J4 (λ p-north p-south q-north q-south prefl p-merid qrefl q-merid + → (fl : p-north ≡ p-south) (fr : q-north ≡ q-south) → + PathP (λ i₁ → (rUnit refl ∙ cong₂ _∙_ prefl qrefl) i₁ ≡ + ((rUnit refl ∙ cong₂ _∙_ prefl qrefl) ∙ cong₂ _∙_ p-merid q-merid) i₁) + ((prefl ∙∙ fl ∙ sym p-merid ∙∙ sym prefl) ∙ + (qrefl ∙∙ fr ∙ sym q-merid ∙∙ sym qrefl)) + (cong₂ _∙_ fl fr)) + λ fl fr → (cong₂ _∙_ (sym (rUnit (fl ∙ refl)) ∙ sym (rUnit fl)) + (sym (rUnit (fr ∙ refl)) ∙ sym (rUnit fr)) + ◁ flipSquare ((sym (rUnit _) + ◁ flipSquare + ((λ j → cong (λ x → rUnit x j) fl + ∙ cong (λ x → lUnit x j) fr) + ▷ sym (cong₂Funct _∙_ fl fr))) + ▷ (rUnit _ ∙ rUnit _))) + l : (x : fst (S₊∙ (suc (suc n)))) → + fst (convertLoopFun→ n (·Π f g)) x ≡ + fst (·Π (convertLoopFun→ n f) (convertLoopFun→ n g)) x + l north = refl + l south = refl + l (merid a i₁) j = help j i₁ + where + help : fst (·Π f g) a ≡ cong (fst (·Π (convertLoopFun→ n f) (convertLoopFun→ n g))) (merid a) + help = looper-help n f g a + ∙ cong₂ _∙_ (sym (cong-∙ (fst (convertLoopFun→ n f)) (merid a) (sym (merid (ptSn _))) + ∙∙ cong (fst f a ∙_) (cong sym (snd f)) + ∙∙ sym (rUnit _))) + (sym (cong-∙ (fst (convertLoopFun→ n g)) (merid a) (sym (merid (ptSn _))) + ∙∙ cong (fst g a ∙_) (cong sym (snd g)) + ∙∙ sym (rUnit _))) + ∙ λ i → rUnit (cong (fst (convertLoopFun→ n f)) (merid a ∙ sym (merid (ptSn _)))) i + ∙ rUnit (cong (fst (convertLoopFun→ n g)) (merid a ∙ sym (merid (ptSn _)))) i + + myFiller : (n : ℕ) → (S₊∙ (suc (suc n)) →∙ A) → I → I → I → fst A + myFiller n f i j k = + hfill (λ k → λ { (i = i0) → doubleCompPath-filler + (sym (snd f)) + (cong (fst f) (merid (ptSn _) ∙ sym (merid (ptSn _)))) + (snd f) k j + ; (i = i1) → snd f k + ; (j = i0) → snd f k + ; (j = i1) → snd f k}) + (inS ((fst f) (rCancel' (merid (ptSn _)) i j))) + k + + convertLoopFun← : (n : ℕ) → (S₊∙ (suc (suc n)) →∙ A) → (S₊∙ (suc n) →∙ Ω A) + fst (convertLoopFun← n f) a = + sym (snd f) ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f + snd (convertLoopFun← n f) i j = myFiller n f i j i1 + + convertLoopFun : (n : ℕ) → Iso (S₊∙ (suc n) →∙ Ω A) (S₊∙ (suc (suc n)) →∙ A) + Iso.fun (convertLoopFun n) f = convertLoopFun→ n f + Iso.inv (convertLoopFun n) f = convertLoopFun← n f + Iso.rightInv (convertLoopFun n) (f , p) = + ΣPathP (funExt help , λ i j → p (~ i ∨ j)) + where + help : (x : _) → fst (convertLoopFun→ n (convertLoopFun← n (f , p))) x ≡ f x + help north = sym p + help south = sym p ∙ cong f (merid (ptSn _)) + help (merid a j) i = + hcomp (λ k → λ { (i = i1) → f (merid a j) + ; (j = i0) → p (~ i ∧ k) + ; (j = i1) → compPath-filler' (sym p) (cong f (merid (ptSn _))) k i}) + (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) + Iso.leftInv (convertLoopFun n) f = + ΣPathP (funExt help34 , help2) + where + + sillyFill₂ : (x : S₊ (suc n)) → I → I → I → fst A + sillyFill₂ x i j k = + hfill (λ k → λ { (i = i0) → convertLoopFun→ n f .fst + (compPath-filler' (merid x) (sym (merid (ptSn _))) k j) + ; (i = i1) → snd A + ; (j = i0) → fst f x (i ∨ ~ k) + ; (j = i1) → snd A}) + (inS (snd f i (~ j))) + k + + sillyFill : (x : S₊ (suc n)) → I → I → I → fst A + sillyFill x i j k = + hfill (λ k → λ { (i = i0) → doubleCompPath-filler + refl + (cong (convertLoopFun→ n f .fst) (merid x ∙ sym (merid (ptSn _)))) + refl k j + ; (i = i1) → fst f x (j ∨ ~ k) + ; (j = i0) → fst f x (i ∧ ~ k) + ; (j = i1) → snd A}) + (inS (sillyFill₂ x i j i1)) + k + + compPath-filler'' : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) + → I → I → I → A + compPath-filler'' {z = z} p q k j i = + hfill (λ k → λ { (i = i0) → p (~ j) + ; (i = i1) → q k + ; (j = i0) → q (i ∧ k) }) + (inS (p (i ∨ ~ j))) + k + + help34 : (x : _) → fst (convertLoopFun← n (convertLoopFun→ n f)) x ≡ fst f x + help34 x i j = sillyFill x i j i1 + + sideCube : Cube (λ j k → snd f j (~ k)) (λ j k → fst (convertLoopFun→ n f) (rCancel' (merid (ptSn _)) j k)) + (λ r k → convertLoopFun→ n f .fst (compPath-filler' (merid (ptSn _)) (sym (merid (ptSn _))) r k)) + (λ _ _ → pt A) + (λ r j → snd f j (~ r)) λ _ _ → pt A -- I → I → I → fst A -- r j k + sideCube r j k = + hcomp (λ i → λ {(r = i0) → snd f j (~ k ∨ ~ i) + ; (r = i1) → fst (convertLoopFun→ n f) (rCancel-filler' (merid (ptSn _)) (~ i) j k) + ; (j = i0) → convertLoopFun→ n f .fst + (compPath-filler'' (merid (ptSn _)) (sym (merid (ptSn _))) i r k) + ; (j = i1) → snd f (~ r) (~ i ∧ k) + ; (k = i0) → snd f j (~ r) + ; (k = i1) → snd f (~ r ∧ j) (~ i)}) + (hcomp (λ i → λ {(r = i0) → pt A + ; (r = i1) → snd f (~ i) k + ; (j = i0) → snd f (~ i) (k ∨ ~ r) + ; (j = i1) → snd f (~ r ∨ ~ i) k + ; (k = i0) → snd f (j ∨ ~ i) (~ r) + ; (k = i1) → pt A}) + (pt A)) + + help2 : PathP _ _ _ + help2 i j k = + hcomp (λ r → λ {(i = i0) → myFiller n (convertLoopFun→ n f) j k r + ; (i = i1) → snd f j (k ∨ ~ r) + ; (j = i0) → sillyFill (ptSn _) i k r + ; (j = i1) → snd A + ; (k = i0) → snd f j (i ∧ ~ r) + ; (k = i1) → snd A}) + (hcomp ((λ r → λ {(i = i0) → sideCube r j k + ; (i = i1) → pt A + ; (j = i0) → sillyFill₂ (ptSn _) i k r + ; (j = i1) → pt A + ; (k = i0) → snd f j (i ∨ ~ r) + ; (k = i1) → pt A})) + (snd f (i ∨ j) (~ k))) + + convertLoopFun←-pres∙ : (n : ℕ) → (f g : S₊∙ (suc (suc n)) →∙ A) + → convertLoopFun← n (·Π f g) + ≡ ·Π (convertLoopFun← n f) (convertLoopFun← n g) + convertLoopFun←-pres∙ n f g = + sym (Iso.leftInv (convertLoopFun n) (convertLoopFun← n (·Π f g))) + ∙∙ cong (convertLoopFun← n) + (Iso.rightInv (convertLoopFun n) (·Π f g) + ∙∙ cong₂ ·Π (sym (Iso.rightInv (convertLoopFun n) f)) (sym (Iso.rightInv (convertLoopFun n) g)) + ∙∙ sym (convertLoopFun-pres∙ _ (convertLoopFun← n f) (convertLoopFun← n g))) + ∙∙ (Iso.leftInv (convertLoopFun n) (·Π (convertLoopFun← n f) (convertLoopFun← n g))) + +→' : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) → (S₊∙ (suc n) →∙ A) → (fst ((Ω^ (suc n)) A)) +→' zero (f , p) = sym p ∙∙ cong f loop ∙∙ p +→' {A = A} (suc n) (f , p) = Iso.inv (flipLoopSpace _) (→' {A = Ω A} n (Iso.inv (convertLoopFun _) (f , p))) + +→'pres∙ : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (p q : S₊∙ (suc n) →∙ A) → →' n (·Π p q) ≡ →' n p ∙ →' n q +→'pres∙ zero f g = sym (rUnit _) +→'pres∙ (suc n) f g = + cong (Iso.inv (flipLoopSpace (suc n))) + (cong (→' n) (convertLoopFun←-pres∙ n f g) ∙ →'pres∙ n (convertLoopFun← n f) ((convertLoopFun← n g))) + ∙ flipLoopSpacepres· (suc n) _ _ + +←' : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) → (fst ((Ω^ (suc n)) A)) → (S₊∙ (suc n) →∙ A) +fst (←' {A = A} zero f) base = pt A +fst (←' {A = A} zero f) (loop i₁) = f i₁ +snd (←' {A = A} zero f) = refl +←' {A = A} (suc n) f = Iso.fun (convertLoopFun _) (←' n (Iso.fun (flipLoopSpace _) f)) + +cancel-r : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (x : fst ((Ω^ (suc n)) A)) + → →' n (←' n x) ≡ x +cancel-r zero x = sym (rUnit _) +cancel-r (suc n) x = + cong (Iso.inv (flipLoopSpace (suc n)) ∘ →' n) + (Iso.leftInv (convertLoopFun _) + (←' n (Iso.fun (flipLoopSpace (suc n)) x))) + ∙∙ cong (Iso.inv (flipLoopSpace (suc n))) + (cancel-r n (Iso.fun (flipLoopSpace (suc n)) x)) + ∙∙ Iso.leftInv (flipLoopSpace (suc n)) x + +cancel-l : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (x : (S₊∙ (suc n) →∙ A)) + → ←' n (→' n x) ≡ x +cancel-l zero (f , p) = ΣPathP (funExt fstp , λ i j → p (~ i ∨ j)) + where + fstp : (x : _) → _ ≡ f x + fstp base = sym p + fstp (loop i) j = doubleCompPath-filler (sym p) (cong f loop) p (~ j) i +cancel-l (suc n) f = + cong (Iso.fun (convertLoopFun n) ∘ ←' n) (Iso.rightInv (flipLoopSpace (suc n)) + (→' n (Iso.inv (convertLoopFun n) f))) + ∙∙ cong (Iso.fun (convertLoopFun n)) (cancel-l n (Iso.inv (convertLoopFun n) f)) + ∙∙ Iso.rightInv (convertLoopFun _) _ + +→∙Ω : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (S₊∙ n →∙ A) (fst ((Ω^ n) A)) +Iso.fun (→∙Ω {A = A} zero) (f , p) = f false +fst (Iso.inv (→∙Ω {A = A} zero) a) false = a +fst (Iso.inv (→∙Ω {A = A} zero) a) true = pt A +snd (Iso.inv (→∙Ω {A = A} zero) a) = refl +Iso.rightInv (→∙Ω {A = A} zero) a = refl +Iso.leftInv (→∙Ω {A = A} zero) (f , p) = + ΣPathP ((funExt (λ { false → refl ; true → sym p})) , λ i j → p (~ i ∨ j)) +Iso.fun (→∙Ω {A = A} (suc n)) = →' n +Iso.inv (→∙Ω {A = A} (suc n)) = ←' n +Iso.rightInv (→∙Ω {A = A} (suc n)) = cancel-r n +Iso.leftInv (→∙Ω {A = A} (suc n)) = cancel-l n + +→∙Ω-pres·Π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f g : S₊∙ (suc n) →∙ A) + → →' n (·Π f g) ≡ →' n f ∙ →' n g +→∙Ω-pres·Π zero f g = sym (rUnit _) +→∙Ω-pres·Π (suc n) f g = + cong (Iso.inv (flipLoopSpace (suc n))) (cong (→' n) (convertLoopFun←-pres∙ n f g) + ∙ →'pres∙ n (convertLoopFun← n f) (convertLoopFun← n g)) + ∙ (flipLoopSpacepres· (suc n) (→' n (convertLoopFun← n f)) (→' n (convertLoopFun← n g))) + +0Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → (S₊∙ n →∙ A) +fst (0Π {A = A}) _ = pt A +snd (0Π {A = A}) = refl + +rUnit·Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ·Π f 0Π ≡ f +fst (rUnit·Π {A = A} {n = zero} f i) base = snd f (~ i) +fst (rUnit·Π {A = A} {n = zero} f i) (loop j) = help i j + where + help : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) + (((sym (snd f)) ∙∙ (cong (fst f) loop) ∙∙ snd f) + ∙ (refl ∙ refl)) + (cong (fst f) loop) + help = (cong ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) ∙_) (sym (rUnit refl)) ∙ sym (rUnit _)) + ◁ λ i j → doubleCompPath-filler (sym (snd f)) (cong (fst f) loop) (snd f) (~ i) j +snd (rUnit·Π {A = A} {n = zero} f i) j = snd f (~ i ∨ j) +fst (rUnit·Π {A = A} {n = suc n} f i) north = snd f (~ i) +fst (rUnit·Π {A = A} {n = suc n} f i) south = (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i +fst (rUnit·Π {A = A} {n = suc n} f i) (merid a j) = help i j + where + help : PathP (λ i → snd f (~ i) ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) + (((sym (snd f)) ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) ∙∙ snd f) + ∙ (refl ∙ refl)) + (cong (fst f) (merid a)) + help = (cong (((sym (snd f)) ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) ∙∙ snd f) ∙_) + (sym (rUnit refl)) + ∙ sym (rUnit _)) + ◁ λ i j → hcomp (λ k → λ { (j = i0) → snd f (~ i ∧ k) + ; (j = i1) → compPath-filler' (sym (snd f)) (cong (fst f) (merid (ptSn (suc n)))) k i + ; (i = i0) → doubleCompPath-filler (sym (snd f)) + (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + (snd f) k j + ; (i = i1) → fst f (merid a j)}) + (fst f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) +snd (rUnit·Π {A = A} {n = suc n} f i) j = snd f (~ i ∨ j) + +lUnit·Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ·Π 0Π f ≡ f +fst (lUnit·Π {n = zero} f i) base = snd f (~ i) +fst (lUnit·Π {n = zero} f i) (loop j) = s i j + where + s : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) + ((refl ∙ refl) ∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) + (cong (fst f) loop) + s = (cong (_∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) (sym (rUnit refl)) ∙ sym (lUnit _)) + ◁ λ i j → doubleCompPath-filler (sym (snd f)) (cong (fst f) loop) (snd f) (~ i) j +snd (lUnit·Π {n = zero} f i) j = snd f (~ i ∨ j) +fst (lUnit·Π {n = suc n} f i) north = snd f (~ i) +fst (lUnit·Π {n = suc n} f i) south = (sym (snd f) ∙ cong (fst f) (merid (ptSn _))) i +fst (lUnit·Π {n = suc n} f i) (merid a j) = help i j + where + help : PathP (λ i → snd f (~ i) ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) + ((refl ∙ refl) ∙ ((sym (snd f)) ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) ∙∙ snd f)) + (cong (fst f) (merid a)) + help = (cong (_∙ ((sym (snd f)) ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) ∙∙ snd f)) + (sym (rUnit refl)) + ∙ sym (lUnit _)) + ◁ λ i j → hcomp (λ k → λ { (j = i0) → snd f (~ i ∧ k) + ; (j = i1) → compPath-filler' (sym (snd f)) (cong (fst f) (merid (ptSn (suc n)))) k i + ; (i = i0) → doubleCompPath-filler (sym (snd f)) + (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + (snd f) k j + ; (i = i1) → fst f (merid a j)}) + (fst f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) +snd (lUnit·Π {n = suc n} f i) j = snd f (~ i ∨ j) + +rCancel·Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ·Π f (-Π f) ≡ 0Π +fst (rCancel·Π {A = A} {n = zero} f i) base = pt A +fst (rCancel·Π {A = A} {n = zero} f i) (loop j) = rCancel (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) i j +snd (rCancel·Π {A = A} {n = zero} f i) = refl +fst (rCancel·Π {A = A} {n = suc n} f i) north = pt A +fst (rCancel·Π {A = A} {n = suc n} f i) south = pt A +fst (rCancel·Π {A = A} {n = suc n} f i) (merid a i₁) = sl i i₁ + where + pl = (sym (snd f) ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f) + sl : pl + ∙ ((sym (snd f) ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f)) ≡ refl + sl = cong (pl ∙_) (cong (sym (snd f) ∙∙_∙∙ (snd f)) + (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) + ∙∙ cong₂ _∙_ refl + (cong (cong (fst f)) (rCancel (merid (ptSn _)))) + ∙∙ sym (rUnit _))) + ∙ rCancel pl +snd (rCancel·Π {A = A} {n = suc n} f i) = refl + +lCancel·Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ·Π (-Π f) f ≡ 0Π +fst (lCancel·Π {A = A} {n = zero} f i) base = pt A +fst (lCancel·Π {A = A} {n = zero} f i) (loop j) = + rCancel (sym (snd f) ∙∙ cong (fst f) (sym loop) ∙∙ snd f) i j +fst (lCancel·Π {A = A} {n = suc n} f i) north = pt A +fst (lCancel·Π {A = A} {n = suc n} f i) south = pt A +fst (lCancel·Π {A = A} {n = suc n} f i) (merid a j) = sl i j + where + pl = (sym (snd f) ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f) + + sl : ((sym (snd f) ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f)) ∙ pl ≡ refl + sl = cong (_∙ pl) (cong (sym (snd f) ∙∙_∙∙ (snd f)) + (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) + ∙∙ cong₂ _∙_ refl (cong (cong (fst f)) (rCancel (merid (ptSn _)))) + ∙∙ sym (rUnit _))) + ∙ lCancel pl +snd (lCancel·Π {A = A} {n = zero} f i) = refl +snd (lCancel·Π {A = A} {n = suc n} f i) = refl + +assoc·Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f g h : S₊∙ (suc n) →∙ A) + → ·Π f (·Π g h) ≡ ·Π (·Π f g) h +assoc·Π {n = n} f g h = + sym (Iso.leftInv (→∙Ω (suc n)) (·Π f (·Π g h))) + ∙∙ cong (←' n) (→'pres∙ n f (·Π g h) + ∙∙ cong (→' n f ∙_) (→'pres∙ n g h) + ∙∙ assoc _ _ _ + ∙∙ cong (_∙ →' n h) (sym (→'pres∙ n f g)) + ∙∙ sym (→'pres∙ n (·Π f g) h)) + ∙∙ Iso.leftInv (→∙Ω (suc n)) (·Π (·Π f g) h) + +comm·Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f g : S₊∙ (suc (suc n)) →∙ A) + → ·Π f g ≡ ·Π g f +comm·Π {A = A} {n = n} f g = + sym (Iso.leftInv (→∙Ω (suc (suc n))) (·Π f g)) + ∙∙ cong (←' (suc n)) (→'pres∙ (suc n) f g + ∙∙ EH _ _ _ + ∙∙ sym (→'pres∙ (suc n) g f)) + ∙∙ Iso.leftInv (→∙Ω (suc (suc n))) (·Π g f) + +module _ {ℓ} {A : Pointed ℓ} {n : ℕ} where + _·π_ : ∥ S₊∙ n →∙ A ∥₂ + → ∥ S₊∙ n →∙ A ∥₂ + → ∥ S₊∙ n →∙ A ∥₂ + _·π_ = sRec2 squash₂ λ f g → ∣ ·Π f g ∣₂ + + -π : ∥ S₊∙ n →∙ A ∥₂ → ∥ S₊∙ n →∙ A ∥₂ + -π = sMap -Π + + 0π : ∥ S₊∙ n →∙ A ∥₂ + 0π = ∣ 0Π ∣₂ + + +module _ {ℓ} {A : Pointed ℓ} {n : ℕ} where + 2path : ∀ {ℓ} {A : Type ℓ} {x y : ∥ A ∥₂} → isSet (x ≡ y) + 2path = isProp→isSet (squash₂ _ _) + + rUnit·π : (f : ∥ S₊∙ (suc n) →∙ A ∥₂) → f ·π 0π ≡ f + rUnit·π = sElim (λ _ → 2path) λ f i → ∣ rUnit·Π f i ∣₂ + + lUnit·π : (f : ∥ S₊∙ (suc n) →∙ A ∥₂) → 0π ·π f ≡ f + lUnit·π = sElim (λ _ → 2path) λ f i → ∣ lUnit·Π f i ∣₂ + + rCancel·π : (f : ∥ S₊∙ (suc n) →∙ A ∥₂) → f ·π -π f ≡ 0π + rCancel·π = sElim (λ _ → 2path) λ f i → ∣ rCancel·Π f i ∣₂ + + lCancel·π : (f : ∥ S₊∙ (suc n) →∙ A ∥₂) → -π f ·π f ≡ 0π + lCancel·π = sElim (λ _ → 2path) λ f i → ∣ lCancel·Π f i ∣₂ + + assoc·π : (f g h : ∥ S₊∙ (suc n) →∙ A ∥₂) → f ·π (g ·π h) ≡ (f ·π g) ·π h + assoc·π = sElim3 (λ _ _ _ → 2path) λ f g h i → ∣ assoc·Π f g h i ∣₂ + + comm·π : (f g : ∥ S₊∙ (suc (suc n)) →∙ A ∥₂) → f ·π g ≡ g ·π f + comm·π = sElim2 (λ _ _ → 2path) λ f g i → ∣ comm·Π f g i ∣₂ + + +open import Cubical.HITs.Wedge +_∨→_ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + → (f : A →∙ C) (g : B →∙ C) + → A ⋁ B → fst C +(f ∨→ g) (inl x) = fst f x +(f ∨→ g) (inr x) = fst g x +(f ∨→ g) (push a i₁) = (snd f ∙ sym (snd g)) i₁ + + +add2 : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → (f g : S₊∙ n →∙ A) → ((S₊∙ n) ⋁ (S₊∙ n) , inl (ptSn _)) →∙ A +fst (add2 {A = A} f g) = f ∨→ g +snd (add2 {A = A} f g) = snd f + +C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Type _ +C* n f g = Pushout (λ _ → tt) (fst (add2 f g)) + +θ : ∀ {ℓ} {A : Pointed ℓ} → Susp (fst A) → (Susp (fst A) , north) ⋁ (Susp (fst A) , north) +θ north = inl north +θ south = inr north +θ {A = A} (merid a i₁) = + ((λ i → inl ((merid a ∙ sym (merid (pt A))) i)) + ∙∙ push tt + ∙∙ λ i → inr ((merid a ∙ sym (merid (pt A))) i)) i₁ + + +help3 : ∀ {ℓ} {A : Type ℓ} {x y z w : A} (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) + → p ∙∙ q ∙∙ r ≡ p ∙ q ∙ r +help3 {x = x} {y = y} {z = z} {w = w} = + J (λ y p → (q : y ≡ z) (r : z ≡ w) → + (p ∙∙ q ∙∙ r) ≡ p ∙ q ∙ r) + λ q r → lUnit (q ∙ r) + ++≡ : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → (x : _) → ·Π f g .fst x ≡ (f ∨→ g) (θ {A = S₊∙ _} x) ++≡ n f g north = sym (snd f) ++≡ n f g south = sym (snd g) ++≡ n f g (merid a i₁) j = main j i₁ + where + + help : cong (f ∨→ g) (cong (θ {A = S₊∙ _}) (merid a)) + ≡ (refl ∙∙ cong (fst f) (merid a ∙ sym (merid north)) ∙∙ snd f) + ∙ (sym (snd g) ∙∙ cong (fst g) (merid a ∙ sym (merid north)) ∙∙ refl) + help = cong-∙∙ (f ∨→ g) ((λ i → inl ((merid a ∙ sym (merid north)) i))) + (push tt) + (λ i → inr ((merid a ∙ sym (merid north)) i)) + ∙∙ help3 _ _ _ + ∙∙ cong (cong (f ∨→ g) + (λ i₂ → inl ((merid a ∙ (λ i₃ → merid north (~ i₃))) i₂)) ∙_) + (sym (assoc _ _ _)) + ∙∙ assoc _ _ _ + ∙∙ cong₂ _∙_ refl (compPath≡compPath' _ _) + + main : PathP (λ i → snd f (~ i) ≡ snd g (~ i)) (λ i₁ → ·Π f g .fst (merid a i₁)) + λ i₁ → (f ∨→ g) (θ {A = S₊∙ _} (merid a i₁)) + main = (λ i → ((λ j → snd f (~ i ∧ ~ j)) ∙∙ cong (fst f) (merid a ∙ sym (merid north)) ∙∙ snd f) + ∙ (sym (snd g) ∙∙ cong (fst g) (merid a ∙ sym (merid north)) ∙∙ λ j → snd g (~ i ∧ j))) + ▷ sym help + ++≡' : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → ·Π f g ≡ ((f ∨→ g) ∘ (θ {A = S₊∙ _}) , snd f) ++≡' n f g = ΣPathP ((funExt (+≡ n f g)) , (λ j i → snd f (i ∨ ~ j))) + +WedgeElim : (n : ℕ) → ∀ {ℓ} {P : S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) → Type ℓ} + → ((x : _) → isOfHLevel (3 +ℕ n) (P x)) + → P (inl north) + → (x : _) → P x +WedgeElim n {P = P} x s (inl x₁) = + sphereElim _ {A = P ∘ inl} + (λ _ → isOfHLevelPlus' {n = n} (3 +ℕ n) (x _)) s x₁ +WedgeElim n {P = P} x s (inr x₁) = + sphereElim _ {A = P ∘ inr} + (sphereElim _ (λ _ → isProp→isOfHLevelSuc ((suc (suc (n +ℕ n)))) + (isPropIsOfHLevel (suc (suc (suc (n +ℕ n)))))) + (subst (isOfHLevel ((3 +ℕ n) +ℕ n)) (cong P (push tt)) + (isOfHLevelPlus' {n = n} (3 +ℕ n) (x _)))) + (subst P (push tt) s) x₁ +WedgeElim n {P = P} x s (push a j) = transp (λ i → P (push tt (i ∧ j))) (~ j) s + + +H²-C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → GroupIso (coHomGr (2 +ℕ n) (C* n f g)) ℤGroup +H²-C* n f g = compGroupIso groupIso1 (Hⁿ-Sⁿ≅ℤ (suc n)) + where + help : (coHom (2 +ℕ n) (C* n f g)) → coHom (2 +ℕ n) (S₊ (2 +ℕ n)) + help = sMap λ f → f ∘ inr + + + invMapPrim : (S₊ (2 +ℕ n) → coHomK (2 +ℕ n)) → C* n f g → coHomK (2 +ℕ n) + invMapPrim h (inl x) = h (ptSn _) + invMapPrim h (inr x) = h x + invMapPrim h (push a i₁) = WedgeElim n {P = λ a → h north ≡ h (fst (add2 f g) a)} + (λ _ → isOfHLevelTrunc (4 +ℕ n) _ _) + (cong h (sym (snd f))) a i₁ + + invMap : coHom (2 +ℕ n) (S₊ (2 +ℕ n)) → (coHom (2 +ℕ n) (C* n f g)) + invMap = sMap invMapPrim + + ret1 : (x : coHom (2 +ℕ n) (S₊ (2 +ℕ n))) → help (invMap x) ≡ x + ret1 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ a → refl + + ret2 : (x : (coHom (2 +ℕ n) (C* n f g))) → invMap (help x) ≡ x + ret2 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ h → cong ∣_∣₂ (funExt λ { (inl x) → cong h ((λ i → inr ((snd f) (~ i))) + ∙ sym (push (inl north))) + ; (inr x) → refl + ; (push a i₁) → help2 h a i₁}) + where + help2 : (h : C* n f g → coHomK (2 +ℕ n)) + → (a : _) → PathP (λ i → invMapPrim (h ∘ inr) (push a i) ≡ h (push a i)) + (cong h ((λ i → inr ((snd f) (~ i))) ∙ sym (push (inl north)))) + refl + help2 h = WedgeElim n (λ _ → isOfHLevelPathP (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n) _ _) _ _) + help4 + + where + help4 : PathP (λ i → invMapPrim (h ∘ inr) (push (inl north) i) ≡ h (push (inl north) i)) + (cong h ((λ i → inr ((snd f) (~ i))) ∙ sym (push (inl north)))) + refl + help4 i j = h (hcomp (λ k → λ { (i = i1) → inr (fst f north) + ; (j = i0) → inr (snd f (~ i)) + ; (j = i1) → push (inl north) (i ∨ ~ k)}) + (inr (snd f (~ i ∧ ~ j)))) + + groupIso1 : GroupIso ((coHomGr (2 +ℕ n) (C* n f g))) (coHomGr (2 +ℕ n) (S₊ (2 +ℕ n))) + Iso.fun (fst groupIso1) = help + Iso.inv (fst groupIso1) = invMap + Iso.rightInv (fst groupIso1) = ret1 + Iso.leftInv (fst groupIso1) = ret2 + snd groupIso1 = makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f g → refl) + + +C+ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Type _ +C+ n f g = Pushout (λ _ → tt) (·Π f g .fst) + +H⁴-C∨ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → GroupIso (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C+ n f g)) + ℤGroup +H⁴-C∨ n f g = compGroupIso + (transportCohomIso (cong (3 +ℕ_) (+-suc n n))) (Hopfβ-Iso n (·Π f g)) + +lol : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) (x : ℤ) + → Iso.inv (fst (H⁴-C∨ n f g)) x + ≡ subst (λ m → coHom m (C+ n f g)) (sym (cong (3 +ℕ_) (+-suc n n))) (Iso.inv (fst (Hopfβ-Iso n (·Π f g))) x) +lol n f g = λ _ → refl + +H²-C∨ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → GroupIso (coHomGr (2 +ℕ n) (C+ n f g)) + ℤGroup +H²-C∨ n f g = Hopfα-Iso n (·Π f g) + + +H⁴-C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → GroupIso (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) + (DirProd ℤGroup ℤGroup) +H⁴-C* n f g = compGroupIso (GroupEquiv→GroupIso (invGroupEquiv fstIso)) + (compGroupIso (transportCohomIso (cong (2 +ℕ_) (+-suc n n))) + (compGroupIso (Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) _) + (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ _) (Hⁿ-Sⁿ≅ℤ _)))) + where + module Ms = MV _ _ _ (λ _ → tt) (fst (add2 f g)) + fstIso : GroupEquiv (coHomGr ((suc (suc (n +ℕ suc n)))) (S₊∙ (3 +ℕ (n +ℕ n)) ⋁ S₊∙ (3 +ℕ (n +ℕ n)))) + (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) + fst (fst fstIso) = Ms.d (suc (suc (n +ℕ suc n))) .fst + snd (fst fstIso) = + SES→isEquiv + (GroupPath _ _ .fst (compGroupEquiv (invEquiv (isContr→≃Unit ((tt , tt) , (λ _ → refl))) , makeIsGroupHom λ _ _ → refl) + (GroupEquivDirProd + (GroupIso→GroupEquiv (invGroupIso (Hⁿ-Unit≅0 _))) + (GroupIso→GroupEquiv + (invGroupIso + (Hⁿ-Sᵐ≅0 _ _ λ p → ¬lem 0 ((λ _ → north) , refl) n n (cong suc (sym (+-suc n n)) ∙ p))))))) + (GroupPath _ _ .fst + (compGroupEquiv ((invEquiv (isContr→≃Unit ((tt , tt) , (λ _ → refl))) , makeIsGroupHom λ _ _ → refl)) + ((GroupEquivDirProd + (GroupIso→GroupEquiv (invGroupIso (Hⁿ-Unit≅0 _))) + (GroupIso→GroupEquiv + (invGroupIso (Hⁿ-Sᵐ≅0 _ _ λ p → ¬lem 0 ((λ _ → north) , refl) n (suc n) (cong (2 +ℕ_) (sym (+-suc n n)) ∙ p)))))))) + (Ms.Δ (suc (suc (n +ℕ suc n)))) + (Ms.d (suc (suc (n +ℕ suc n)))) + (Ms.i (suc (suc (suc (n +ℕ suc n))))) + (Ms.Ker-d⊂Im-Δ _) + (Ms.Ker-i⊂Im-d _) + snd fstIso = Ms.d (suc (suc (n +ℕ suc n))) .snd + +Hopfie : {n : ℕ} → ∥ S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n) ∥₂ → ℤ +Hopfie = sRec isSetℤ (HopfFunction _) + +subber : (n : ℕ) (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → _ +subber n f = subst (λ x → coHom x (HopfInvariantPush n (fst f))) + (λ i → suc (suc (suc (+-suc n n i)))) + +module _ (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) where + + C+' = C+ n f g + + βₗ : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) + βₗ = Iso.inv (fst (H⁴-C* n f g)) (1 , 0) + + βᵣ : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) + βᵣ = Iso.inv (fst (H⁴-C* n f g)) (0 , 1) + + βᵣ'-fun : C* n f g → coHomK ((4 +ℕ (n +ℕ n))) + βᵣ'-fun (inl x) = 0ₖ _ + βᵣ'-fun (inr x) = 0ₖ _ + βᵣ'-fun (push (inl x) i₁) = 0ₖ _ + βᵣ'-fun (push (inr x) i₁) = Kn→ΩKn+1 _ ∣ x ∣ i₁ + βᵣ'-fun (push (push a i₂) i₁) = Kn→ΩKn+10ₖ _ (~ i₂) i₁ + + + βₗ'-fun : C* n f g → coHomK (4 +ℕ (n +ℕ n)) + βₗ'-fun (inl x) = 0ₖ _ + βₗ'-fun (inr x) = 0ₖ _ + βₗ'-fun (push (inl x) i₁) = Kn→ΩKn+1 _ ∣ x ∣ i₁ + βₗ'-fun (push (inr x) i₁) = 0ₖ _ + βₗ'-fun (push (push a i₂) i₁) = Kn→ΩKn+10ₖ _ i₂ i₁ + + βₗ''-fun : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) + βₗ''-fun = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) ∣ βₗ'-fun ∣₂ + + βᵣ''-fun : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) + βᵣ''-fun = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) ∣ βᵣ'-fun ∣₂ + + α∨ : coHom (2 +ℕ n) (C* n f g) + α∨ = Iso.inv (fst (H²-C* n f g)) 1 + + private + pp : (a : _) → 0ₖ (suc (suc n)) ≡ ∣ (f ∨→ g) a ∣ + pp = WedgeElim _ (λ _ → isOfHLevelTrunc (4 +ℕ n) _ _) + (cong ∣_∣ₕ (sym (snd f))) + + pp-inr : pp (inr north) ≡ cong ∣_∣ₕ (sym (snd g)) + pp-inr = (λ j → transport (λ i → 0ₖ (2 +ℕ n) ≡ ∣ compPath-filler' (snd f) (sym (snd g)) (~ j) i ∣ₕ) + λ i → ∣ snd f (~ i ∨ j) ∣ₕ) + ∙ λ j → transp (λ i → 0ₖ (2 +ℕ n) ≡ ∣ snd g (~ i ∧ ~ j) ∣ₕ) j λ i → ∣ snd g (~ i ∨ ~ j) ∣ₕ + + α∨'-fun : C* n f g → coHomK (2 +ℕ n) + α∨'-fun (inl x) = 0ₖ _ + α∨'-fun (inr x) = ∣ x ∣ + α∨'-fun (push a i₁) = pp a i₁ + + α∨' : coHom (2 +ℕ n) (C* n f g) + α∨' = ∣ α∨'-fun ∣₂ + + + β+ : coHom ((2 +ℕ n) +' (2 +ℕ n)) C+' + β+ = Iso.inv (fst (H⁴-C∨ n f g)) 1 + + q : C+' → C* n f g + q (inl x) = inl x + q (inr x) = inr x + q (push a i₁) = (push (θ a) ∙ λ i → inr (+≡ n f g a (~ i))) i₁ + + jₗ : HopfInvariantPush n (fst f) → C* n f g + jₗ (inl x) = inl x + jₗ (inr x) = inr x + jₗ (push a i₁) = push (inl a) i₁ + + jᵣ : HopfInvariantPush n (fst g) → C* n f g + jᵣ (inl x) = inl x + jᵣ (inr x) = inr x + jᵣ (push a i₁) = push (inr a) i₁ + +α-∨→1 : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → Iso.fun (fst (H²-C* n f g)) (α∨' n f g) ≡ 1 +α-∨→1 n f g = sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (Hⁿ-Sⁿ≅ℤ-nice-generator n)) + ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) + +α-∨-lem : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → α∨ n f g ≡ α∨' n f g +α-∨-lem n f g = sym (Iso.leftInv (fst (H²-C* n f g)) _) + ∙∙ cong (Iso.inv (fst (H²-C* n f g))) help + ∙∙ Iso.leftInv (fst (H²-C* n f g)) _ + where + help : Iso.fun (fst (H²-C* n f g)) (α∨ n f g) ≡ Iso.fun (fst (H²-C* n f g)) (α∨' n f g) + help = (Iso.rightInv (fst (H²-C* n f g)) (pos 1)) ∙ sym (α-∨→1 n f g) + +q-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (q n f g) (α∨ n f g) ≡ Hopfα n (·Π f g) +q-α n f g = (λ i → coHomFun _ (q n f g) (α-∨-lem n f g i)) + ∙ sym (Iso.leftInv is _) + ∙∙ cong (Iso.inv is) help + ∙∙ Iso.leftInv is _ + where + is = fst (Hopfα-Iso n (·Π f g)) + + help : Iso.fun is (coHomFun _ (q n f g) (α∨' n f g)) + ≡ Iso.fun is (Hopfα n (·Π f g)) + help = sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (Hⁿ-Sⁿ≅ℤ-nice-generator n)) + ∙∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) + ∙∙ sym (Hopfα-Iso-α n (·Π f g)) + + +βₗ≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → βₗ n f g ≡ βₗ''-fun n f g +βₗ≡ n f g = cong (FF ∘ subber2) + (cong (Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) + (S₊∙ (suc (suc (suc (n +ℕ n))))) + (((suc (suc (n +ℕ n))))))))) help + ∙ help2) + ∙ funExt⁻ FF∘subber ∣ wedgeGen ∣₂ + ∙ cong subber3 (sym βₗ'-fun≡) + where + FF = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (n +ℕ suc n))) .fst + FF' = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (suc (n +ℕ n)))) .fst + + help : Iso.inv (fst (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (1 , 0) + ≡ (∣ ∣_∣ ∣₂ , 0ₕ _) + help = ΣPathP ((Hⁿ-Sⁿ≅ℤ-nice-generator _) , IsGroupHom.pres1 (snd (invGroupIso (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) + + subber3 = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + + subber2 = (subst (λ m → coHom m (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) + (sym (cong (2 +ℕ_) (+-suc n n)))) + + FF∘subber : FF ∘ subber2 ≡ subber3 ∘ FF' + FF∘subber = + funExt (λ a → (λ j → transp (λ i → coHom ((suc ∘ suc ∘ suc) (+-suc n n (~ i ∧ j))) (C* n f g)) (~ j) + (MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) ((suc (suc (+-suc n n j)))) .fst + (transp (λ i → coHom (2 +ℕ (+-suc n n (j ∨ ~ i))) + (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) j + a)))) + + wedgeGen : (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) → + coHomK (suc (suc (suc (n +ℕ n))))) + wedgeGen (inl x) = ∣ x ∣ + wedgeGen (inr x) = 0ₖ _ + wedgeGen (push a i₁) = 0ₖ _ + + βₗ'-fun≡ : ∣ βₗ'-fun n f g ∣₂ ≡ FF' ∣ wedgeGen ∣₂ + βₗ'-fun≡ = cong ∣_∣₂ (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push (inl x) i₁) → refl + ; (push (inr x) i₁) j → Kn→ΩKn+10ₖ _ (~ j) i₁ + ; (push (push a i₂) i₁) j → Kn→ΩKn+10ₖ _ (~ j ∧ i₂) i₁}) + + help2 : Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) (((suc (suc (n +ℕ n)))))))) + (∣ ∣_∣ ∣₂ , 0ₕ _) + ≡ ∣ wedgeGen ∣₂ + help2 = cong ∣_∣₂ (funExt λ { (inl x) → rUnitₖ (suc (suc (suc (n +ℕ n)))) ∣ x ∣ + ; (inr x) → refl + ; (push a i₁) → refl}) + +βᵣ≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → βᵣ n f g ≡ βᵣ''-fun n f g +βᵣ≡ n f g = + cong (FF ∘ subber2) + (cong (Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) + (S₊∙ (suc (suc (suc (n +ℕ n))))) + (((suc (suc (n +ℕ n))))))))) + help + ∙ help2) + ∙ funExt⁻ FF∘subber ∣ wedgeGen ∣₂ + ∙ cong (subber3) (sym βᵣ'-fun≡) + where + FF = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (n +ℕ suc n))) .fst + FF' = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (suc (n +ℕ n)))) .fst + + help : Iso.inv (fst (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (0 , 1) + ≡ (0ₕ _ , ∣ ∣_∣ ∣₂) + help = ΣPathP (IsGroupHom.pres1 (snd (invGroupIso (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) , (Hⁿ-Sⁿ≅ℤ-nice-generator _)) + + subber3 = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + + subber2 = (subst (λ m → coHom m (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) + (sym (cong (2 +ℕ_) (+-suc n n)))) + + FF∘subber : FF ∘ subber2 ≡ subber3 ∘ FF' + FF∘subber = + funExt (λ a → (λ j → transp (λ i → coHom ((suc ∘ suc ∘ suc) (+-suc n n (~ i ∧ j))) (C* n f g)) (~ j) + (MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) ((suc (suc (+-suc n n j)))) .fst + (transp (λ i → coHom (2 +ℕ (+-suc n n (j ∨ ~ i))) + (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) j + a)))) + + wedgeGen : (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) → + coHomK (suc (suc (suc (n +ℕ n))))) + wedgeGen (inl x) = 0ₖ _ + wedgeGen (inr x) = ∣ x ∣ + wedgeGen (push a i₁) = 0ₖ _ + + + help2 : Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) (((suc (suc (n +ℕ n)))))))) + (0ₕ _ , ∣ ∣_∣ ∣₂) + ≡ ∣ wedgeGen ∣₂ + help2 = cong ∣_∣₂ (funExt λ { (inl x) → refl + ; (inr x) → lUnitₖ (suc (suc (suc (n +ℕ n)))) ∣ x ∣ + ; (push a i₁) → refl}) + + βᵣ'-fun≡ : ∣ βᵣ'-fun n f g ∣₂ ≡ FF' ∣ wedgeGen ∣₂ + βᵣ'-fun≡ = cong ∣_∣₂ (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push (inl x) i₁) j → Kn→ΩKn+10ₖ _ (~ j) i₁ + ; (push (inr x) i₁) → refl + ; (push (push a i₂) i₁) j → Kn→ΩKn+10ₖ _ (~ j ∧ ~ i₂) i₁}) + +q-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (q n f g) (βₗ n f g) + ≡ β+ n f g +q-βₗ n f g = cong (coHomFun _ (q n f g)) (βₗ≡ n f g) + ∙ transportLem + ∙ cong (subst (λ m → coHom m (C+' n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + (sym (Iso.leftInv (fst (Hopfβ-Iso n (·Π f g))) (Hopfβ n (·Π f g))) + ∙ cong (Iso.inv ((fst (Hopfβ-Iso n (·Π f g))))) (Hopfβ↦1 n (·Π f g))) + where + transportLem : coHomFun (suc (suc (suc (n +ℕ suc n)))) (q n f g) + (βₗ''-fun n f g) + ≡ subst (λ m → coHom m (C+' n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + (Hopfβ n (·Π f g)) + transportLem = + natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (q n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∙ cong (subst (λ m → coHom m (C+' n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + (cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → loll a i₁})) + where + open import Cubical.Foundations.Path + loll : (x : fst (S₊∙ (3 +ℕ (n +ℕ n)))) → + PathP + (λ i₁ → + βₗ'-fun n f g ((push (θ x) ∙ λ i → inr (+≡ n f g x (~ i))) i₁) ≡ + MV.d-pre Unit (fst (S₊∙ (2 +ℕ n))) (fst (S₊∙ (3 +ℕ n +ℕ n))) + (λ _ → tt) (fst (·Π f g)) (3 +ℕ n +ℕ n) ∣_∣ (push x i₁)) + (λ _ → βₗ'-fun n f g (q n f g (inl tt))) + (λ _ → βₗ'-fun n f g (q n f g (inr (·Π f g .fst x)))) + loll x = + flipSquare (cong-∙ (βₗ'-fun n f g) (push (θ x)) (λ i → inr (+≡ n f g x (~ i))) + ∙∙ sym (rUnit _) + ∙∙ lem₁ x) + + where + lem₁ : (x : _) → cong (βₗ'-fun n f g) (push (θ {A = S₊∙ _} x)) ≡ Kn→ΩKn+1 _ ∣ x ∣ + lem₁ north = refl + lem₁ south = sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)) + lem₁ (merid a j) k = lala k j + where + lala : PathP (λ k → Kn→ΩKn+1 _ ∣ north ∣ₕ ≡ (sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north))) k) + (λ j → cong (βₗ'-fun n f g) (push (θ {A = S₊∙ _} (merid a j)))) + (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) + lala = (cong-∙∙ (cong (βₗ'-fun n f g) ∘ push) + ((λ i₁ → inl ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁))) + (push tt) + ((λ i₁ → inr ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁))) + ∙ sym (compPath≡compPath' _ _) + ∙ (λ _ → (λ j → Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∣ (merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) j ∣ₕ) + ∙ Kn→ΩKn+10ₖ _) + ∙ cong (_∙ Kn→ΩKn+10ₖ _) (cong-∙ ((Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ)) + (merid a) (sym (merid north))) + ∙ sym (assoc _ _ _) + ∙ cong (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a) ∙_) + (sym (symDistr (sym ((Kn→ΩKn+10ₖ _))) (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)))))) + ◁ λ i j → compPath-filler (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) + (sym (sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) + (cong ∣_∣ₕ (merid north)))) + (~ i) j + +q-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (q n f g) (βᵣ n f g) + ≡ β+ n f g +q-βᵣ n f g = cong (coHomFun _ (q n f g)) (βᵣ≡ n f g) + ∙ transportLem + ∙ cong (subst (λ m → coHom m (C+' n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + (sym (Iso.leftInv (fst (Hopfβ-Iso n (·Π f g))) (Hopfβ n (·Π f g))) + ∙ cong (Iso.inv ((fst (Hopfβ-Iso n (·Π f g))))) (Hopfβ↦1 n (·Π f g))) + where + transportLem : coHomFun (suc (suc (suc (n +ℕ suc n)))) (q n f g) + (βᵣ''-fun n f g) + ≡ subst (λ m → coHom m (C+' n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + (Hopfβ n (·Π f g)) + transportLem = + natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (q n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∙ cong (subst (λ m → coHom m (C+' n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + (cong ∣_∣₂ (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a i₁) → loll a i₁})) + where + open import Cubical.Foundations.Path + loll : (x : fst (S₊∙ (3 +ℕ (n +ℕ n)))) → + PathP + (λ i₁ → + βᵣ'-fun n f g ((push (θ x) ∙ λ i → inr (+≡ n f g x (~ i))) i₁) ≡ + MV.d-pre Unit (fst (S₊∙ (2 +ℕ n))) (fst (S₊∙ (3 +ℕ n +ℕ n))) + (λ _ → tt) (fst (·Π f g)) (3 +ℕ n +ℕ n) ∣_∣ (push x i₁)) + (λ _ → βᵣ'-fun n f g (q n f g (inl tt))) + (λ _ → βᵣ'-fun n f g (q n f g (inr (·Π f g .fst x)))) + loll x = flipSquare (cong-∙ (βᵣ'-fun n f g) (push (θ x)) (λ i → inr (+≡ n f g x (~ i))) + ∙∙ sym (rUnit _) + ∙∙ lem₁ x) + where + lem₁ : (x : _) → cong (βᵣ'-fun n f g) (push (θ {A = S₊∙ _} x)) ≡ Kn→ΩKn+1 _ ∣ x ∣ + lem₁ north = sym (Kn→ΩKn+10ₖ _) + lem₁ south = cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)) + lem₁ (merid a j) k = lala k j -- lala k j + where + lala : PathP (λ k → (Kn→ΩKn+10ₖ _) (~ k) ≡ (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north))) k) + (λ j → cong (βᵣ'-fun n f g) (push (θ {A = S₊∙ _} (merid a j)))) + (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) + lala = (cong-∙∙ (cong (βᵣ'-fun n f g) ∘ push) + (λ i₁ → inl ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁)) + (push tt) + (λ i₁ → inr ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁)) + ∙∙ (cong (sym (Kn→ΩKn+10ₖ _) ∙_) (cong-∙ (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a) (sym (merid (ptSn _))))) + ∙∙ sym (help3 _ _ _)) + ◁ symP (doubleCompPath-filler + (sym (Kn→ΩKn+10ₖ _)) + (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) + (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (sym (merid north)))) +open import Cubical.Foundations.Path +jₗ-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (jₗ n f g) (α∨ n f g) + ≡ Hopfα n f +jₗ-α n f g = + cong (coHomFun _ (jₗ n f g)) (α-∨-lem n f g) + ∙ cong ∣_∣₂ (funExt (HopfInvariantPushElim n (fst f) + (isOfHLevelPath ((suc (suc (suc (suc (n +ℕ n)))))) + (isOfHLevelPlus' {n = n} (4 +ℕ n) (isOfHLevelTrunc (4 +ℕ n))) _ _) + refl + (λ _ → refl) + λ j → refl)) + +jₗ-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (jₗ n f g) (βₗ n f g) + ≡ subst (λ m → coHom m (HopfInvariantPush n (fst f))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + (Hopfβ n f) +jₗ-βₗ n f g = + cong (coHomFun _ (jₗ n f g)) (βₗ≡ n f g) + ∙∙ natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (jₗ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + (cong ∣_∣₂ + (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a i₁) → refl})) + +jₗ-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (jₗ n f g) (βᵣ n f g) + ≡ 0ₕ _ +jₗ-βᵣ n f g = + cong (coHomFun _ (jₗ n f g)) (βᵣ≡ n f g) + ∙∙ natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (jₗ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + cool + where + cool : coHomFun (suc (suc (suc (suc (n +ℕ n))))) (jₗ n f g) + ∣ βᵣ'-fun n f g ∣₂ ≡ 0ₕ _ + cool = cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → refl}) + +jᵣ-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (jᵣ n f g) (α∨ n f g) + ≡ Hopfα n g +jᵣ-α n f g = cong (coHomFun _ (jᵣ n f g)) (α-∨-lem n f g) + ∙ cong ∣_∣₂ (funExt (HopfInvariantPushElim n (fst g) + (isOfHLevelPath ((suc (suc (suc (suc (n +ℕ n)))))) + (isOfHLevelPlus' {n = n} (4 +ℕ n) (isOfHLevelTrunc (4 +ℕ n))) _ _) + refl + (λ _ → refl) + (flipSquare (pp-inr n f g)))) + +jᵣ-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (jᵣ n f g) (βₗ n f g) ≡ 0ₕ _ +jᵣ-βₗ n f g = + cong (coHomFun _ (jᵣ n f g)) (βₗ≡ n f g) + ∙∙ natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (jᵣ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + cool + where + cool : coHomFun (suc (suc (suc (suc (n +ℕ n))))) (jᵣ n f g) + ∣ βₗ'-fun n f g ∣₂ ≡ 0ₕ _ + cool = cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → refl}) + + +jᵣ-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (jᵣ n f g) (βᵣ n f g) + ≡ subst (λ m → coHom m (HopfInvariantPush n (fst g))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + (Hopfβ n g) +jᵣ-βᵣ n f g = + cong (coHomFun _ (jᵣ n f g)) (βᵣ≡ n f g) + ∙∙ natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (jᵣ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst g))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + (cong ∣_∣₂ + (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a i₁) → refl})) + +gen₂ℤ×ℤ : gen₂-by (DirProd ℤGroup ℤGroup) (1 , 0) (0 , 1) +fst (gen₂ℤ×ℤ (x , y)) = x , y +snd (gen₂ℤ×ℤ (x , y)) = + ΣPathP ((cong₂ _+_ ((·Comm 1 x) ∙ cong fst (sym (distrLem 1 0 x))) ((·Comm 0 y) ∙ cong fst (sym (distrLem 0 1 y)))) + , +Comm y 0 ∙ cong₂ _+_ (·Comm 0 x ∙ cong snd (sym (distrLem 1 0 x))) (·Comm 1 y ∙ cong snd (sym (distrLem 0 1 y)))) + where + ℤ×ℤ = DirProd ℤGroup ℤGroup + _+''_ = GroupStr._·_ (snd ℤ×ℤ) + + ll3 : (x : ℤ) → - x ≡ 0 - x + ll3 (pos zero) = refl + ll3 (pos (suc zero)) = refl + ll3 (pos (suc (suc n))) = + cong predℤ (ll3 (pos (suc n))) + ll3 (negsuc zero) = refl + ll3 (negsuc (suc n)) = cong sucℤ (ll3 (negsuc n)) + + distrLem : (x y : ℤ) (z : ℤ) + → Path (ℤ × ℤ) (z ℤ[ ℤ×ℤ ]· (x , y)) (z · x , z · y) + distrLem x y (pos zero) = refl + distrLem x y (pos (suc n)) = + (cong ((x , y) +''_) (distrLem x y (pos n))) + distrLem x y (negsuc zero) = ΣPathP (sym (ll3 x) , sym (ll3 y)) + distrLem x y (negsuc (suc n)) = + cong₂ _+''_ (ΣPathP (sym (ll3 x) , sym (ll3 y))) + (distrLem x y (negsuc n)) + +genH²ⁿC* : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → gen₂-by (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) + (βₗ n f g) + (βᵣ n f g) +genH²ⁿC* n f g = + Iso-pres-gen₂ (DirProd ℤGroup ℤGroup) + (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) + (1 , 0) + (0 , 1) + gen₂ℤ×ℤ + (invGroupIso (H⁴-C* n f g)) + +private + + H⁴C* : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Group _ + H⁴C* n f g = coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) + + X : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → ℤ + X n f g = (genH²ⁿC* n f g) (α∨ n f g ⌣ α∨ n f g) .fst .fst + Y : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → ℤ + Y n f g = (genH²ⁿC* n f g) (α∨ n f g ⌣ α∨ n f g) .fst .snd + + α∨≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → α∨ n f g ⌣ α∨ n f g ≡ ((X n f g) ℤ[ H⁴C* n f g ]· βₗ n f g) + +ₕ ((Y n f g) ℤ[ H⁴C* n f g ]· βᵣ n f g) + α∨≡ n f g = (genH²ⁿC* n f g) (α∨ n f g ⌣ α∨ n f g) .snd + + +eq₁ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → (Hopfα n (·Π f g)) ⌣ (Hopfα n (·Π f g)) + ≡ ((X n f g + Y n f g) ℤ[ coHomGr _ _ ]· (β+ n f g)) +eq₁ n f g = + cong (λ x → x ⌣ x) (sym (q-α n f g) ∙ cong (coHomFun (suc (suc n)) (q n f g)) (α-∨-lem n f g)) + ∙ cong ((coHomFun _) (q _ f g)) (cong (λ x → x ⌣ x) (sym (α-∨-lem n f g)) + ∙ α∨≡ n f g) + ∙ IsGroupHom.pres· (coHomMorph _ (q n f g) .snd) _ _ + ∙ cong₂ _+ₕ_ (hompres· (coHomMorph _ (q n f g)) (βₗ n f g) (X n f g) + ∙ cong (λ z → (X n f g) ℤ[ coHomGr _ _ ]· z) + (q-βₗ n f g)) + (hompres· (coHomMorph _ (q n f g)) (βᵣ n f g) (Y n f g) + ∙ cong (λ z → (Y n f g) ℤ[ coHomGr _ _ ]· z) + (q-βᵣ n f g)) + ∙ sym (distrℤ· (coHomGr _ _) (β+ n f g) (X n f g) (Y n f g)) + +coHomFun⌣ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) + → (n : ℕ) (x y : coHom n B) + → coHomFun _ f (x ⌣ y) ≡ coHomFun _ f x ⌣ coHomFun _ f y +coHomFun⌣ f n = sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl + +eq₂ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → Hopfα n f ⌣ Hopfα n f + ≡ (X n f g ℤ[ coHomGr _ _ ]· subst (λ m → coHom m (HopfInvariantPush n (fst f))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + (Hopfβ n f)) +eq₂ n f g = + cong (λ x → x ⌣ x) (sym (jₗ-α n f g)) + ∙∙ sym (coHomFun⌣ (jₗ n f g) _ _ _) + ∙∙ cong (coHomFun _ (jₗ n f g)) (α∨≡ n f g) + ∙∙ IsGroupHom.pres· (snd (coHomMorph _ (jₗ n f g))) _ _ + ∙∙ cong₂ _+ₕ_ (hompres· (coHomMorph _ (jₗ n f g)) (βₗ n f g) (X n f g)) + (hompres· (coHomMorph _ (jₗ n f g)) (βᵣ n f g) (Y n f g) + ∙ cong (λ z → (Y n f g) ℤ[ coHomGr _ _ ]· z) + (jₗ-βᵣ n f g)) + ∙∙ cong₂ _+ₕ_ refl (rUnitℤ· (coHomGr _ _) (Y n f g)) + ∙∙ (rUnitₕ _ _ + ∙ cong (X n f g ℤ[ coHomGr _ _ ]·_) (jₗ-βₗ n f g)) + +eq₃ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → Hopfα n g ⌣ Hopfα n g + ≡ (Y n f g ℤ[ coHomGr _ _ ]· subst (λ m → coHom m (HopfInvariantPush n (fst f))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + (Hopfβ n g)) +eq₃ n f g = + cong (λ x → x ⌣ x) (sym (jᵣ-α n f g)) + ∙∙ sym (coHomFun⌣ (jᵣ n f g) _ _ _) + ∙∙ cong (coHomFun _ (jᵣ n f g)) (α∨≡ n f g) + ∙∙ IsGroupHom.pres· (snd (coHomMorph _ (jᵣ n f g))) _ _ + ∙∙ cong₂ _+ₕ_ (hompres· (coHomMorph _ (jᵣ n f g)) (βₗ n f g) (X n f g) + ∙ cong (λ z → (X n f g) ℤ[ coHomGr _ _ ]· z) + (jᵣ-βₗ n f g)) + (hompres· (coHomMorph _ (jᵣ n f g)) (βᵣ n f g) (Y n f g)) + ∙∙ cong₂ _+ₕ_ (rUnitℤ· (coHomGr _ _) (X n f g)) refl + ∙∙ ((lUnitₕ _ _) ∙ cong (Y n f g ℤ[ coHomGr _ _ ]·_) (jᵣ-βᵣ n f g)) + +eq₂-eq₂ : (n : ℕ) (f g : S₊∙ (suc (suc (suc (n +ℕ n)))) →∙ S₊∙ (suc (suc n))) + → HopfFunction n (·Π f g) ≡ HopfFunction n f + HopfFunction n g +eq₂-eq₂ n f g = + mainL + ∙ sym (cong₂ _+_ (help n f g) (helpg n f g)) + where + transpLem : ∀ {ℓ} {G : ℕ → Group ℓ} + → (n m : ℕ) + → (x : ℤ) + → (p : m ≡ n) + → (g : fst (G n)) + → subst (fst ∘ G) p (x ℤ[ G m ]· subst (fst ∘ G) (sym p) g) ≡ (x ℤ[ G n ]· g) + transpLem {G = G} n m x = + J (λ n p → (g : fst (G n)) → subst (fst ∘ G) p (x ℤ[ G m ]· subst (fst ∘ G) (sym p) g) + ≡ (x ℤ[ G n ]· g)) + λ g → transportRefl _ ∙ cong (x ℤ[ G m ]·_) (transportRefl _) + + mainL : HopfFunction n (·Π f g) ≡ X n f g + Y n f g + mainL = cong (Iso.fun (fst (Hopfβ-Iso n (·Π f g)))) + (cong (subst (λ x → coHom x (HopfInvariantPush n (fst (·Π f g)))) + λ i₁ → suc (suc (suc (+-suc n n i₁)))) + (eq₁ n f g)) + ∙∙ cong (Iso.fun (fst (Hopfβ-Iso n (·Π f g)))) + (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst (·Π f g)))} _ _ + (X n f g + Y n f g) (λ i₁ → suc (suc (suc (+-suc n n i₁)))) + (Iso.inv (fst (Hopfβ-Iso n (·Π f g))) 1)) + ∙∙ hompres· (_ , snd (Hopfβ-Iso n (·Π f g))) (Iso.inv (fst (Hopfβ-Iso n (·Π f g))) 1) + (X n f g + Y n f g) + ∙∙ cong ((X n f g + Y n f g) ℤ[ ℤ , ℤGroup .snd ]·_) + (Iso.rightInv ((fst (Hopfβ-Iso n (·Π f g)))) 1) + ∙∙ rUnitℤ·ℤ (X n f g + Y n f g) + + + help : (n : ℕ) (f g : _) → HopfFunction n f ≡ X n f g + help n f g = cong (Iso.fun (fst (Hopfβ-Iso n f))) + (cong (subst (λ x → coHom x (HopfInvariantPush n (fst f))) + (λ i₁ → suc (suc (suc (+-suc n n i₁))))) (eq₂ n f g)) + ∙ cong (Iso.fun (fst (Hopfβ-Iso n f))) + (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst f))} _ _ + (X n f g) + ((cong (suc ∘ suc ∘ suc) (+-suc n n))) + (Hopfβ n f)) + ∙ hompres· (_ , snd (Hopfβ-Iso n f)) (Hopfβ n f) (X n f g) + ∙ cong (X n f g ℤ[ ℤ , ℤGroup .snd ]·_) (Hopfβ↦1 n f) + ∙ rUnitℤ·ℤ (X n f g) + + helpg : (n : ℕ) (f g : _) → HopfFunction n g ≡ Y n f g + helpg n f g = + cong (Iso.fun (fst (Hopfβ-Iso n g))) + (cong (subst (λ x → coHom x (HopfInvariantPush n (fst g))) + (λ i₁ → suc (suc (suc (+-suc n n i₁))))) + (eq₃ n f g)) + ∙∙ cong (Iso.fun (fst (Hopfβ-Iso n g))) + (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst g))} _ _ + (Y n f g) + ((cong (suc ∘ suc ∘ suc) (+-suc n n))) + (Hopfβ n g)) + ∙∙ hompres· (_ , snd (Hopfβ-Iso n g)) (Hopfβ n g) (Y n f g) + ∙∙ cong (Y n f g ℤ[ ℤ , ℤGroup .snd ]·_) (Hopfβ↦1 n g) + ∙∙ rUnitℤ·ℤ (Y n f g) diff --git a/Cubical/ZCohomology/Gysin2.agda b/Cubical/ZCohomology/Gysin2.agda index b986144f4..48293e586 100644 --- a/Cubical/ZCohomology/Gysin2.agda +++ b/Cubical/ZCohomology/Gysin2.agda @@ -54,1168 +54,8 @@ open import Cubical.HITs.Sn open import Cubical.HITs.Susp open import Cubical.HITs.S1 renaming (_·_ to _*_) open import Cubical.HITs.Truncation - renaming (rec to trRec ; elim to trElim ; elim2 to trElim2) + renaming (rec to trRec ; elim to trElim ; elim2 to trElim2 ; map to trMap) open import Cubical.HITs.SetTruncation - renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap) + renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; elim3 to sElim3 ; map to sMap) open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim) - -open import Cubical.Algebra.AbGroup - -open import Cubical.Homotopy.Loopspace - -open import Cubical.HITs.Join - -open import Cubical.Homotopy.Hopf - -open import Cubical.HITs.SetQuotients renaming (_/_ to _/'_) -open import Cubical.ZCohomology.Gysin - - -rUnitℤ· : ∀ {ℓ} (G : Group ℓ) → (x : ℤ) → (x ℤ[ G ]· GroupStr.1g (snd G)) - ≡ GroupStr.1g (snd G) -rUnitℤ· G (pos zero) = refl -rUnitℤ· G (pos (suc n)) = - cong (GroupStr._·_ (snd G) (GroupStr.1g (snd G))) - (rUnitℤ· G (pos n)) - ∙ GroupStr.lid (snd G) (GroupStr.1g (snd G)) -rUnitℤ· G (negsuc zero) = GroupTheory.inv1g G -rUnitℤ· G (negsuc (suc n)) = - cong₂ (GroupStr._·_ (snd G)) (GroupTheory.inv1g G) (rUnitℤ· G (negsuc n)) - ∙ GroupStr.lid (snd G) (GroupStr.1g (snd G)) - -rUnitℤ·ℤ : (x : ℤ) → (x ℤ[ ℤGroup ]· 1) ≡ x -rUnitℤ·ℤ (pos zero) = refl -rUnitℤ·ℤ (pos (suc n)) = cong (pos 1 +_) (rUnitℤ·ℤ (pos n)) ∙ sym (pos+ 1 n) -rUnitℤ·ℤ (negsuc zero) = refl -rUnitℤ·ℤ (negsuc (suc n)) = cong (-1 +_) (rUnitℤ·ℤ (negsuc n)) ∙ +Comm (negsuc 0) (negsuc n) - -private - precommℤ : ∀ {ℓ} (G : Group ℓ) → (g : fst G) (y : ℤ) → (snd G GroupStr.· (y ℤ[ G ]· g)) g - ≡ (sucℤ y ℤ[ G ]· g) - precommℤ G g (pos zero) = GroupStr.lid (snd G) g - ∙ sym (GroupStr.rid (snd G) g) - precommℤ G g (pos (suc n)) = - sym (GroupStr.assoc (snd G) _ _ _) - ∙ cong ((snd G GroupStr.· g)) (precommℤ G g (pos n)) - precommℤ G g (negsuc zero) = - GroupStr.invl (snd G) g - precommℤ G g (negsuc (suc n)) = - sym (GroupStr.assoc (snd G) _ _ _) - ∙ cong ((snd G GroupStr.· GroupStr.inv (snd G) g)) (precommℤ G g (negsuc n)) - ∙ negsucLem n - where - negsucLem : (n : ℕ) → (snd G GroupStr.· GroupStr.inv (snd G) g) - (sucℤ (negsuc n) ℤ[ G ]· g) - ≡ (sucℤ (negsuc (suc n)) ℤ[ G ]· g) - negsucLem zero = (GroupStr.rid (snd G) _) - negsucLem (suc n) = refl - -distrℤ· : ∀ {ℓ} (G : Group ℓ) → (g : fst G) (x y : ℤ) - → ((x + y) ℤ[ G ]· g) - ≡ GroupStr._·_ (snd G) (x ℤ[ G ]· g) (y ℤ[ G ]· g) -distrℤ· G' g (pos zero) y = cong (_ℤ[ G' ]· g) (+Comm 0 y) - ∙ sym (GroupStr.lid (snd G') _) -distrℤ· G' g (pos (suc n)) (pos n₁) = - cong (_ℤ[ G' ]· g) (sym (pos+ (suc n) n₁)) - ∙ cong (GroupStr._·_ (snd G') g) (cong (_ℤ[ G' ]· g) (pos+ n n₁) ∙ distrℤ· G' g (pos n) (pos n₁)) - ∙ GroupStr.assoc (snd G') _ _ _ -distrℤ· G' g (pos (suc n)) (negsuc zero) = - distrℤ· G' g (pos n) 0 - ∙ ((GroupStr.rid (snd G') (pos n ℤ[ G' ]· g) ∙ sym (GroupStr.lid (snd G') (pos n ℤ[ G' ]· g))) - ∙ cong (λ x → GroupStr._·_ (snd G') x (pos n ℤ[ G' ]· g)) - (sym (GroupStr.invl (snd G') g)) ∙ sym (GroupStr.assoc (snd G') _ _ _)) - ∙ (GroupStr.assoc (snd G') _ _ _) - ∙ cong (λ x → GroupStr._·_ (snd G') x (pos n ℤ[ G' ]· g)) (GroupStr.invl (snd G') g) - ∙ GroupStr.lid (snd G') _ - ∙ sym (GroupStr.rid (snd G') _) - ∙ (cong (GroupStr._·_ (snd G') (pos n ℤ[ G' ]· g)) (sym (GroupStr.invr (snd G') g)) - ∙ GroupStr.assoc (snd G') _ _ _) - ∙ cong (λ x → GroupStr._·_ (snd G') x (GroupStr.inv (snd G') g)) - (precommℤ G' g (pos n)) -distrℤ· G' g (pos (suc n)) (negsuc (suc n₁)) = - cong (_ℤ[ G' ]· g) (predℤ+negsuc n₁ (pos (suc n))) - ∙∙ distrℤ· G' g (pos n) (negsuc n₁) - ∙∙ (cong (λ x → GroupStr._·_ (snd G') x (negsuc n₁ ℤ[ G' ]· g)) - ((sym (GroupStr.rid (snd G') (pos n ℤ[ G' ]· g)) - ∙ cong (GroupStr._·_ (snd G') (pos n ℤ[ G' ]· g)) (sym (GroupStr.invr (snd G') g))) - ∙∙ GroupStr.assoc (snd G') _ _ _ - ∙∙ cong (λ x → GroupStr._·_ (snd G') x (GroupStr.inv (snd G') g)) (precommℤ G' g (pos n) )) - ∙ sym (GroupStr.assoc (snd G') _ _ _)) -distrℤ· G' g (negsuc zero) y = - cong (_ℤ[ G' ]· g) (+Comm -1 y) ∙ lol1 y - module _ where - lol1 : (y : ℤ) → (predℤ y ℤ[ G' ]· g) ≡ (snd G' GroupStr.· GroupStr.inv (snd G') g) (y ℤ[ G' ]· g) - lol1 (pos zero) = sym (GroupStr.rid (snd G') _) - lol1 (pos (suc n)) = - sym (GroupStr.lid (snd G') ((pos n ℤ[ G' ]· g))) - ∙∙ cong (λ x → GroupStr._·_ (snd G') x (pos n ℤ[ G' ]· g)) (sym (GroupStr.invl (snd G') g)) - ∙∙ sym (GroupStr.assoc (snd G') _ _ _) - lol1 (negsuc n) = refl -distrℤ· G' g (negsuc (suc n)) y = - cong (_ℤ[ G' ]· g) (+Comm (negsuc (suc n)) y) - ∙∙ lol1 G' g 0 (y +negsuc n) - ∙∙ cong (snd G' GroupStr.· GroupStr.inv (snd G') g) - (cong (_ℤ[ G' ]· g) (+Comm y (negsuc n)) ∙ distrℤ· G' g (negsuc n) y) - ∙ (GroupStr.assoc (snd G') _ _ _) - - - -HopfInvariantPushElim : ∀ {ℓ} n → (f : _) → {P : HopfInvariantPush n f → Type ℓ} - → (isOfHLevel (suc (suc (suc (suc (n +ℕ n))))) (P (inl tt))) - → (e : P (inl tt)) - (g : (x : _) → P (inr x)) - (r : PathP (λ i → P (push north i)) e (g (f north))) - → (x : _) → P x -HopfInvariantPushElim n f {P = P} hlev e g r (inl x) = e -HopfInvariantPushElim n f {P = P} hlev e g r (inr x) = g x -HopfInvariantPushElim n f {P = P} hlev e g r (push a i₁) = help a i₁ - where - help : (a : Susp (Susp (S₊ (suc (n +ℕ n))))) → PathP (λ i → P (push a i)) e (g (f a)) - help = sphereElim _ (sphereElim _ (λ _ → isProp→isOfHLevelSuc ((suc (suc (n +ℕ n)))) (isPropIsOfHLevel _)) - (isOfHLevelPathP' (suc (suc (suc (n +ℕ n)))) - (subst (isOfHLevel (suc (suc (suc (suc (n +ℕ n)))))) - (cong P (push north)) - hlev) _ _)) r - -transportCohomIso : ∀ {ℓ} {A : Type ℓ} {n m : ℕ} - → (p : n ≡ m) - → GroupIso (coHomGr n A) (coHomGr m A) -Iso.fun (fst (transportCohomIso {A = A} p)) = subst (λ n → coHom n A) p -Iso.inv (fst (transportCohomIso {A = A} p)) = subst (λ n → coHom n A) (sym p) -Iso.rightInv (fst (transportCohomIso p)) = transportTransport⁻ _ -Iso.leftInv (fst (transportCohomIso p)) = transportTransport⁻ _ -snd (transportCohomIso {A = A} {n = n} {m = m} p) = - makeIsGroupHom (λ x y → help x y p) - where - help : (x y : coHom n A) (p : n ≡ m) → subst (λ n → coHom n A) p (x +ₕ y) - ≡ subst (λ n → coHom n A) p x +ₕ subst (λ n → coHom n A) p y - help x y = J (λ m p → subst (λ n → coHom n A) p (x +ₕ y) - ≡ subst (λ n → coHom n A) p x +ₕ subst (λ n → coHom n A) p y) - (transportRefl (x +ₕ y) ∙ cong₂ _+ₕ_ (sym (transportRefl x)) (sym (transportRefl y))) - -α-hopfMap : abs (HopfFunction zero (hopfFun , λ _ → hopfFun (snd (S₊∙ 3)))) ≡ suc zero -α-hopfMap = cong abs (cong (Iso.fun (fst (Hopfβ-Iso zero (hopfFun , refl)))) - (transportRefl (⌣-α 0 (hopfFun , refl)))) ∙ l3 (sym CP2≡CP2') GysinS1.e (Hopfα zero (hopfFun , refl)) - (l isGenerator≃ℤ-e) - (GroupIso→GroupEquiv (Hopfα-Iso 0 (hopfFun , refl)) , refl) - (snd (fst ⌣hom)) - (GroupIso→GroupEquiv (Hopfβ-Iso zero (hopfFun , refl))) - where - l : Σ-syntax (GroupIso (coHomGr 2 CP2) ℤGroup) - (λ ϕ → abs (Iso.fun (fst ϕ) GysinS1.e) ≡ 1) - → Σ-syntax (GroupEquiv (coHomGr 2 CP2) ℤGroup) - (λ ϕ → abs (fst (fst ϕ) GysinS1.e) ≡ 1) - l p = (GroupIso→GroupEquiv (fst p)) , (snd p) - -·Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (S₊∙ n →∙ A) - → (S₊∙ n →∙ A) - → (S₊∙ n →∙ A) -·Π {A = A} {n = zero} p q = (λ _ → pt A) , refl -fst (·Π {A = A} {n = suc zero} (f , p) (g , q)) base = pt A -fst (·Π {A = A} {n = suc zero} (f , p) (g , q)) (loop j) = - ((sym p ∙∙ cong f loop ∙∙ p) ∙ (sym q ∙∙ cong g loop ∙∙ q)) j -snd (·Π {A = A} {n = suc zero} (f , p) (g , q)) = refl -fst (·Π {A = A} {n = suc (suc n)} (f , p) (g , q)) north = pt A -fst (·Π {A = A} {n = suc (suc n)} (f , p) (g , q)) south = pt A -fst (·Π {A = A} {n = suc (suc n)} (f , p) (g , q)) (merid a j) = - ((sym p ∙∙ cong f (merid a ∙ sym (merid (ptSn (suc n)))) ∙∙ p) - ∙ (sym q ∙∙ cong g (merid a ∙ sym (merid (ptSn (suc n)))) ∙∙ q)) j -snd (·Π {A = A} {n = suc (suc n)} (f , p) (g , q)) = refl - -flipLoopSpace : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → Iso (fst ((Ω^ (suc n)) A)) (fst ((Ω^ n) (Ω A))) -flipLoopSpace n = {!!} - -module _ {ℓ : Level} {A : Pointed ℓ} where - convertLoopFun→ : (n : ℕ) → (S₊∙ (suc n) →∙ Ω A) → (S₊∙ (suc (suc n)) →∙ A) - fst (convertLoopFun→ n f) north = pt A - fst (convertLoopFun→ n f) south = pt A - fst (convertLoopFun→ n f) (merid a i₁) = fst f a i₁ - snd (convertLoopFun→ n f) = refl - - myFiller : (n : ℕ) → (S₊∙ (suc (suc n)) →∙ A) → I → I → I → fst A - myFiller n f i j k = - hfill (λ k → λ { (i = i0) → doubleCompPath-filler (sym (snd f)) (cong (fst f) (merid (ptSn _) ∙ sym (merid (ptSn _)))) (snd f) k j - ; (i = i1) → snd f k - ; (j = i0) → snd f k - ; (j = i1) → snd f k}) - (inS ((fst f) (rCancel (merid (ptSn _)) i j))) - k - - convertLoopFun← : (n : ℕ) → (S₊∙ (suc (suc n)) →∙ A) → (S₊∙ (suc n) →∙ Ω A) - fst (convertLoopFun← n f) a = sym (snd f) ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f - snd (convertLoopFun← n f) i j = myFiller n f i j i1 - - convertLoopFun : (n : ℕ) → Iso (S₊∙ (suc n) →∙ Ω A) (S₊∙ (suc (suc n)) →∙ A) - Iso.fun (convertLoopFun n) f = convertLoopFun→ n f - Iso.inv (convertLoopFun n) f = convertLoopFun← n f - Iso.rightInv (convertLoopFun n) (f , p) = - ΣPathP (funExt help , λ i j → p (~ i ∨ j)) - where - help : (x : _) → fst (convertLoopFun→ n (convertLoopFun← n (f , p))) x ≡ f x - help north = sym p - help south = sym p ∙ cong f (merid (ptSn _)) - help (merid a j) i = - hcomp (λ k → λ { (i = i1) → f (merid a j) - ; (j = i0) → p (~ i ∧ k) - ; (j = i1) → compPath-filler' (sym p) (cong f (merid (ptSn _))) k i}) - (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) - Iso.leftInv (convertLoopFun n) f = - ΣPathP (funExt help34 , help2) - where - - sillyFill₂ : (x : S₊ (suc n)) → I → I → I → fst A - sillyFill₂ x i j k = - hfill (λ k → λ { (i = i0) → convertLoopFun→ n f .fst (compPath-filler' (merid x) (sym (merid (ptSn _))) k j) - ; (i = i1) → snd A - ; (j = i0) → fst f x (i ∨ ~ k) - ; (j = i1) → snd A}) - (inS (snd f i (~ j))) - k - - sillyFill : (x : S₊ (suc n)) → I → I → I → fst A - sillyFill x i j k = - hfill (λ k → λ { (i = i0) → doubleCompPath-filler refl (cong (convertLoopFun→ n f .fst) (merid x ∙ sym (merid (ptSn _)))) refl k j - ; (i = i1) → fst f x (j ∨ ~ k) - ; (j = i0) → fst f x (i ∧ ~ k) - ; (j = i1) → snd A}) - (inS (sillyFill₂ x i j i1)) - k - - - - help34 : (x : _) → fst (convertLoopFun← n (convertLoopFun→ n f)) x ≡ fst f x - help34 x i j = sillyFill x i j i1 - - help2 : PathP _ _ _ - help2 i j k = - hcomp (λ r → λ {(i = i0) → myFiller n (convertLoopFun→ n f) j k r - ; (i = i1) → snd f j (k ∨ ~ r) - ; (j = i0) → sillyFill (ptSn _) i k r - ; (j = i1) → snd A - ; (k = i0) → snd f j (i ∧ ~ r) - ; (k = i1) → snd A}) - (hcomp ((λ r → λ {(i = i0) → {!!} - ; (i = i1) → {!!} - ; (j = i0) → {!!} - ; (j = i1) → {!!} - ; (k = i0) → {!!} - ; (k = i1) → {!!}})) - {!!}) - - {- - i = i0 ⊢ snd (convertLoopFun← n (Iso.fun (convertLoopFun n) f)) j k -i = i1 ⊢ snd f j k -j = i0 ⊢ funExt help i (snd (S₊∙ (suc n))) k -j = i1 ⊢ snd (Ω A) k -k = i0 ⊢ snd A -k = i1 ⊢ snd A - -} - -→' : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) → (S₊∙ (suc n) →∙ A) → (fst ((Ω^ (suc n)) A)) -→' zero (f , p) = sym p ∙∙ cong f loop ∙∙ p -→' {A = A} (suc n) (f , p) = Iso.inv (flipLoopSpace _) (→' {A = Ω A} n (Iso.inv (convertLoopFun _) (f , p))) - -←' : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) → (fst ((Ω^ (suc n)) A)) → (S₊∙ (suc n) →∙ A) -fst (←' {A = A} zero f) base = pt A -fst (←' {A = A} zero f) (loop i₁) = f i₁ -snd (←' {A = A} zero f) = refl -←' {A = A} (suc n) f = Iso.fun (convertLoopFun _) (←' n (Iso.fun (flipLoopSpace _) f)) - -cancel-r : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → (x : fst ((Ω^ (suc n)) A)) - → →' n (←' n x) ≡ x -cancel-r zero x = sym (rUnit _) -cancel-r (suc n) x = - cong (Iso.inv (flipLoopSpace (suc n)) ∘ →' n) - (Iso.leftInv (convertLoopFun _) - (←' n (Iso.fun (flipLoopSpace (suc n)) x))) - ∙∙ cong (Iso.inv (flipLoopSpace (suc n))) - (cancel-r n (Iso.fun (flipLoopSpace (suc n)) x)) - ∙∙ Iso.leftInv (flipLoopSpace (suc n)) x - -cancel-l : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → (x : (S₊∙ (suc n) →∙ A)) - → ←' n (→' n x) ≡ x -cancel-l zero (f , p) = ΣPathP (funExt fstp , λ i j → p (~ i ∨ j)) - where - fstp : (x : _) → _ ≡ f x - fstp base = sym p - fstp (loop i) j = doubleCompPath-filler (sym p) (cong f loop) p (~ j) i -cancel-l (suc n) f = - cong (Iso.fun (convertLoopFun n) ∘ ←' n) (Iso.rightInv (flipLoopSpace (suc n)) - (→' n (Iso.inv (convertLoopFun n) f))) - ∙∙ cong (Iso.fun (convertLoopFun n)) (cancel-l n (Iso.inv (convertLoopFun n) f)) - ∙∙ Iso.rightInv (convertLoopFun _) _ - -→∙Ω : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → Iso (S₊∙ n →∙ A) (fst ((Ω^ n) A)) -Iso.fun (→∙Ω {A = A} zero) (f , p) = f false -fst (Iso.inv (→∙Ω {A = A} zero) a) false = a -fst (Iso.inv (→∙Ω {A = A} zero) a) true = pt A -snd (Iso.inv (→∙Ω {A = A} zero) a) = refl -Iso.rightInv (→∙Ω {A = A} zero) a = refl -Iso.leftInv (→∙Ω {A = A} zero) (f , p) = ΣPathP ((funExt (λ { false → refl ; true → sym p})) , λ i j → p (~ i ∨ j)) -→∙Ω {A = A} (suc n) = {!!} - --- 0Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → (S₊∙ n →∙ A) --- fst (0Π {A = A}) _ = pt A --- snd (0Π {A = A}) = refl - --- comm·Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} --- → (f g : S₊∙ (suc (suc n)) →∙ A) --- → ·Π f g ≡ ·Π g f --- fst (comm·Π {A = A} {n = n} f g i) north = pt A --- fst (comm·Π {A = A} {n = n} f g i) south = pt A --- fst (comm·Π {A = A} {n = zero} f g i) (merid base j) = {!!} --- fst (comm·Π {A = A} {n = zero} f g i) (merid (loop k) j) = {!!} --- fst (comm·Π {A = A} {n = suc n} f g i) (merid a j) = {!!} --- snd (comm·Π {A = A} {n = n} f g i) = refl - --- rUnit·Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} --- → (f : S₊∙ (suc n) →∙ A) --- → ·Π f 0Π ≡ f --- fst (rUnit·Π {A = A} {n = zero} f i) base = snd f (~ i) --- fst (rUnit·Π {A = A} {n = zero} f i) (loop j) = help i j --- where --- help : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) --- (((sym (snd f)) ∙∙ (cong (fst f) loop) ∙∙ snd f) --- ∙ (refl ∙ refl)) --- (cong (fst f) loop) --- help = (cong ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) ∙_) (sym (rUnit refl)) ∙ sym (rUnit _)) --- ◁ λ i j → doubleCompPath-filler (sym (snd f)) (cong (fst f) loop) (snd f) (~ i) j --- snd (rUnit·Π {A = A} {n = zero} f i) j = snd f (~ i ∨ j) --- fst (rUnit·Π {A = A} {n = suc n} f i) north = snd f (~ i) --- fst (rUnit·Π {A = A} {n = suc n} f i) south = (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i --- fst (rUnit·Π {A = A} {n = suc n} f i) (merid a j) = help i j --- where --- help : PathP (λ i → snd f (~ i) ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) --- (((sym (snd f)) ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) ∙∙ snd f) --- ∙ (refl ∙ refl)) --- (cong (fst f) (merid a)) --- help = (cong (((sym (snd f)) ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) ∙∙ snd f) ∙_) --- (sym (rUnit refl)) --- ∙ sym (rUnit _)) --- ◁ λ i j → hcomp (λ k → λ { (j = i0) → snd f (~ i ∧ k) --- ; (j = i1) → compPath-filler' (sym (snd f)) (cong (fst f) (merid (ptSn (suc n)))) k i --- ; (i = i0) → doubleCompPath-filler (sym (snd f)) --- (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) --- (snd f) k j --- ; (i = i1) → fst f (merid a j)}) --- (fst f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) --- snd (rUnit·Π {A = A} {n = suc n} f i) j = snd f (~ i ∨ j) - --- -- _·π_ : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} --- -- → ∥ S₊∙ n →∙ A ∥₂ --- -- → ∥ S₊∙ n →∙ A ∥₂ --- -- → ∥ S₊∙ n →∙ A ∥₂ --- -- _·π_ = sRec2 squash₂ λ f g → ∣ ·Π f g ∣₂ - --- -- open import Cubical.HITs.Wedge --- -- _∨→_ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} --- -- → (f : A →∙ C) (g : B →∙ C) --- -- → A ⋁ B → fst C --- -- (f ∨→ g) (inl x) = fst f x --- -- (f ∨→ g) (inr x) = fst g x --- -- (f ∨→ g) (push a i₁) = (snd f ∙ sym (snd g)) i₁ - - --- -- add2 : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → (f g : S₊∙ n →∙ A) → ((S₊∙ n) ⋁ (S₊∙ n) , inl (ptSn _)) →∙ A --- -- fst (add2 {A = A} f g) = f ∨→ g --- -- snd (add2 {A = A} f g) = snd f - --- -- C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Type _ --- -- C* n f g = Pushout (λ _ → tt) (fst (add2 f g)) - --- -- θ : ∀ {ℓ} {A : Pointed ℓ} → Susp (fst A) → (Susp (fst A) , north) ⋁ (Susp (fst A) , north) --- -- θ north = inl north --- -- θ south = inr north --- -- θ {A = A} (merid a i₁) = --- -- ((λ i → inl ((merid a ∙ sym (merid (pt A))) i)) --- -- ∙∙ push tt --- -- ∙∙ λ i → inr ((merid a ∙ sym (merid (pt A))) i)) i₁ - - --- -- help3 : ∀ {ℓ} {A : Type ℓ} {x y z w : A} (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) --- -- → p ∙∙ q ∙∙ r ≡ p ∙ q ∙ r --- -- help3 {x = x} {y = y} {z = z} {w = w} = --- -- J (λ y p → (q : y ≡ z) (r : z ≡ w) → --- -- (p ∙∙ q ∙∙ r) ≡ p ∙ q ∙ r) --- -- λ q r → lUnit (q ∙ r) - --- -- +≡ : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → (x : _) → ·Π f g .fst x ≡ (f ∨→ g) (θ {A = S₊∙ _} x) --- -- +≡ n f g north = sym (snd f) --- -- +≡ n f g south = sym (snd g) --- -- +≡ n f g (merid a i₁) j = main j i₁ --- -- where - --- -- help : cong (f ∨→ g) (cong (θ {A = S₊∙ _}) (merid a)) --- -- ≡ (refl ∙∙ cong (fst f) (merid a ∙ sym (merid north)) ∙∙ snd f) --- -- ∙ (sym (snd g) ∙∙ cong (fst g) (merid a ∙ sym (merid north)) ∙∙ refl) --- -- help = cong-∙∙ (f ∨→ g) ((λ i → inl ((merid a ∙ sym (merid north)) i))) --- -- (push tt) --- -- (λ i → inr ((merid a ∙ sym (merid north)) i)) --- -- ∙∙ help3 _ _ _ --- -- ∙∙ cong (cong (f ∨→ g) --- -- (λ i₂ → inl ((merid a ∙ (λ i₃ → merid north (~ i₃))) i₂)) ∙_) --- -- (sym (assoc _ _ _)) --- -- ∙∙ assoc _ _ _ --- -- ∙∙ cong₂ _∙_ refl (compPath≡compPath' _ _) - --- -- main : PathP (λ i → snd f (~ i) ≡ snd g (~ i)) (λ i₁ → ·Π f g .fst (merid a i₁)) --- -- λ i₁ → (f ∨→ g) (θ {A = S₊∙ _} (merid a i₁)) --- -- main = (λ i → ((λ j → snd f (~ i ∧ ~ j)) ∙∙ cong (fst f) (merid a ∙ sym (merid north)) ∙∙ snd f) --- -- ∙ (sym (snd g) ∙∙ cong (fst g) (merid a ∙ sym (merid north)) ∙∙ λ j → snd g (~ i ∧ j))) --- -- ▷ sym help - --- -- +≡' : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → ·Π f g ≡ ((f ∨→ g) ∘ (θ {A = S₊∙ _}) , snd f) --- -- +≡' n f g = ΣPathP ((funExt (+≡ n f g)) , (λ j i → snd f (i ∨ ~ j))) - --- -- WedgeElim : (n : ℕ) → ∀ {ℓ} {P : S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) → Type ℓ} --- -- → ((x : _) → isOfHLevel (3 +ℕ n) (P x)) --- -- → P (inl north) --- -- → (x : _) → P x --- -- WedgeElim n {P = P} x s (inl x₁) = --- -- sphereElim _ {A = P ∘ inl} --- -- (λ _ → isOfHLevelPlus' {n = n} (3 +ℕ n) (x _)) s x₁ --- -- WedgeElim n {P = P} x s (inr x₁) = --- -- sphereElim _ {A = P ∘ inr} --- -- (sphereElim _ (λ _ → isProp→isOfHLevelSuc ((suc (suc (n +ℕ n)))) --- -- (isPropIsOfHLevel (suc (suc (suc (n +ℕ n)))))) --- -- (subst (isOfHLevel ((3 +ℕ n) +ℕ n)) (cong P (push tt)) --- -- (isOfHLevelPlus' {n = n} (3 +ℕ n) (x _)))) --- -- (subst P (push tt) s) x₁ --- -- WedgeElim n {P = P} x s (push a j) = transp (λ i → P (push tt (i ∧ j))) (~ j) s - - --- -- H²-C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → GroupIso (coHomGr (2 +ℕ n) (C* n f g)) ℤGroup --- -- H²-C* n f g = compGroupIso groupIso1 (Hⁿ-Sⁿ≅ℤ (suc n)) --- -- where --- -- help : (coHom (2 +ℕ n) (C* n f g)) → coHom (2 +ℕ n) (S₊ (2 +ℕ n)) --- -- help = sMap λ f → f ∘ inr - - --- -- invMapPrim : (S₊ (2 +ℕ n) → coHomK (2 +ℕ n)) → C* n f g → coHomK (2 +ℕ n) --- -- invMapPrim h (inl x) = h (ptSn _) --- -- invMapPrim h (inr x) = h x --- -- invMapPrim h (push a i₁) = WedgeElim n {P = λ a → h north ≡ h (fst (add2 f g) a)} --- -- (λ _ → isOfHLevelTrunc (4 +ℕ n) _ _) --- -- (cong h (sym (snd f))) a i₁ - --- -- invMap : coHom (2 +ℕ n) (S₊ (2 +ℕ n)) → (coHom (2 +ℕ n) (C* n f g)) --- -- invMap = sMap invMapPrim - --- -- ret1 : (x : coHom (2 +ℕ n) (S₊ (2 +ℕ n))) → help (invMap x) ≡ x --- -- ret1 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) --- -- λ a → refl - --- -- ret2 : (x : (coHom (2 +ℕ n) (C* n f g))) → invMap (help x) ≡ x --- -- ret2 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) --- -- λ h → cong ∣_∣₂ (funExt λ { (inl x) → cong h ((λ i → inr ((snd f) (~ i))) --- -- ∙ sym (push (inl north))) --- -- ; (inr x) → refl --- -- ; (push a i₁) → help2 h a i₁}) --- -- where --- -- help2 : (h : C* n f g → coHomK (2 +ℕ n)) --- -- → (a : _) → PathP (λ i → invMapPrim (h ∘ inr) (push a i) ≡ h (push a i)) --- -- (cong h ((λ i → inr ((snd f) (~ i))) ∙ sym (push (inl north)))) --- -- refl --- -- help2 h = WedgeElim n (λ _ → isOfHLevelPathP (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n) _ _) _ _) --- -- help4 - --- -- where --- -- help4 : PathP (λ i → invMapPrim (h ∘ inr) (push (inl north) i) ≡ h (push (inl north) i)) --- -- (cong h ((λ i → inr ((snd f) (~ i))) ∙ sym (push (inl north)))) --- -- refl --- -- help4 i j = h (hcomp (λ k → λ { (i = i1) → inr (fst f north) --- -- ; (j = i0) → inr (snd f (~ i)) --- -- ; (j = i1) → push (inl north) (i ∨ ~ k)}) --- -- (inr (snd f (~ i ∧ ~ j)))) - --- -- groupIso1 : GroupIso ((coHomGr (2 +ℕ n) (C* n f g))) (coHomGr (2 +ℕ n) (S₊ (2 +ℕ n))) --- -- Iso.fun (fst groupIso1) = help --- -- Iso.inv (fst groupIso1) = invMap --- -- Iso.rightInv (fst groupIso1) = ret1 --- -- Iso.leftInv (fst groupIso1) = ret2 --- -- snd groupIso1 = makeIsGroupHom --- -- (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) --- -- λ f g → refl) - - --- -- C+ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Type _ --- -- C+ n f g = Pushout (λ _ → tt) (·Π f g .fst) - --- -- H⁴-C∨ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → GroupIso (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C+ n f g)) --- -- ℤGroup --- -- H⁴-C∨ n f g = compGroupIso --- -- (transportCohomIso (cong (3 +ℕ_) (+-suc n n))) (Hopfβ-Iso n (·Π f g)) - --- -- lol : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) (x : ℤ) --- -- → Iso.inv (fst (H⁴-C∨ n f g)) x --- -- ≡ subst (λ m → coHom m (C+ n f g)) (sym (cong (3 +ℕ_) (+-suc n n))) (Iso.inv (fst (Hopfβ-Iso n (·Π f g))) x) --- -- lol n f g = λ _ → refl - --- -- H²-C∨ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → GroupIso (coHomGr (2 +ℕ n) (C+ n f g)) --- -- ℤGroup --- -- H²-C∨ n f g = Hopfα-Iso n (·Π f g) - - --- -- H⁴-C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → GroupIso (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) --- -- (DirProd ℤGroup ℤGroup) --- -- H⁴-C* n f g = compGroupIso (GroupEquiv→GroupIso (invGroupEquiv fstIso)) --- -- (compGroupIso (transportCohomIso (cong (2 +ℕ_) (+-suc n n))) --- -- (compGroupIso (Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) _) --- -- (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ _) (Hⁿ-Sⁿ≅ℤ _)))) --- -- where --- -- module Ms = MV _ _ _ (λ _ → tt) (fst (add2 f g)) --- -- fstIso : GroupEquiv (coHomGr ((suc (suc (n +ℕ suc n)))) (S₊∙ (3 +ℕ (n +ℕ n)) ⋁ S₊∙ (3 +ℕ (n +ℕ n)))) --- -- (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) --- -- fst (fst fstIso) = Ms.d (suc (suc (n +ℕ suc n))) .fst --- -- snd (fst fstIso) = --- -- SES→Iso --- -- (GroupPath _ _ .fst (compGroupEquiv (invEquiv (isContr→≃Unit ((tt , tt) , (λ _ → refl))) , makeIsGroupHom λ _ _ → refl) --- -- (GroupEquivDirProd --- -- (GroupIso→GroupEquiv (invGroupIso (Hⁿ-Unit≅0 _))) --- -- (GroupIso→GroupEquiv --- -- (invGroupIso --- -- (Hⁿ-Sᵐ≅0 _ _ λ p → ¬lem 0 ((λ _ → north) , refl) n n (cong suc (sym (+-suc n n)) ∙ p))))))) --- -- (GroupPath _ _ .fst --- -- (compGroupEquiv ((invEquiv (isContr→≃Unit ((tt , tt) , (λ _ → refl))) , makeIsGroupHom λ _ _ → refl)) --- -- ((GroupEquivDirProd --- -- (GroupIso→GroupEquiv (invGroupIso (Hⁿ-Unit≅0 _))) --- -- (GroupIso→GroupEquiv --- -- (invGroupIso (Hⁿ-Sᵐ≅0 _ _ λ p → ¬lem 0 ((λ _ → north) , refl) n (suc n) (cong (2 +ℕ_) (sym (+-suc n n)) ∙ p)))))))) --- -- (Ms.Δ (suc (suc (n +ℕ suc n)))) --- -- (Ms.d (suc (suc (n +ℕ suc n)))) --- -- (Ms.i (suc (suc (suc (n +ℕ suc n))))) --- -- (Ms.Ker-d⊂Im-Δ _) --- -- (Ms.Ker-i⊂Im-d _) --- -- snd fstIso = Ms.d (suc (suc (n +ℕ suc n))) .snd - --- -- Hopfie : {n : ℕ} → ∥ S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n) ∥₂ → ℤ --- -- Hopfie = sRec isSetℤ (HopfFunction _) - --- -- subber : (n : ℕ) (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → _ --- -- subber n f = subst (λ x → coHom x (HopfInvariantPush n (fst f))) --- -- (λ i → suc (suc (suc (+-suc n n i)))) - --- -- module _ (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) where - --- -- C+' = C+ n f g - --- -- βₗ : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) --- -- βₗ = Iso.inv (fst (H⁴-C* n f g)) (1 , 0) - --- -- βᵣ : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) --- -- βᵣ = Iso.inv (fst (H⁴-C* n f g)) (0 , 1) - --- -- βᵣ'-fun : C* n f g → coHomK ((4 +ℕ (n +ℕ n))) --- -- βᵣ'-fun (inl x) = 0ₖ _ --- -- βᵣ'-fun (inr x) = 0ₖ _ --- -- βᵣ'-fun (push (inl x) i₁) = 0ₖ _ --- -- βᵣ'-fun (push (inr x) i₁) = Kn→ΩKn+1 _ ∣ x ∣ i₁ --- -- βᵣ'-fun (push (push a i₂) i₁) = Kn→ΩKn+10ₖ _ (~ i₂) i₁ - - --- -- βₗ'-fun : C* n f g → coHomK (4 +ℕ (n +ℕ n)) --- -- βₗ'-fun (inl x) = 0ₖ _ --- -- βₗ'-fun (inr x) = 0ₖ _ --- -- βₗ'-fun (push (inl x) i₁) = Kn→ΩKn+1 _ ∣ x ∣ i₁ --- -- βₗ'-fun (push (inr x) i₁) = 0ₖ _ --- -- βₗ'-fun (push (push a i₂) i₁) = Kn→ΩKn+10ₖ _ i₂ i₁ - --- -- βₗ''-fun : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) --- -- βₗ''-fun = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) ∣ βₗ'-fun ∣₂ - --- -- βᵣ''-fun : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) --- -- βᵣ''-fun = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) ∣ βᵣ'-fun ∣₂ - --- -- α∨ : coHom (2 +ℕ n) (C* n f g) --- -- α∨ = Iso.inv (fst (H²-C* n f g)) 1 - --- -- private --- -- pp : (a : _) → 0ₖ (suc (suc n)) ≡ ∣ (f ∨→ g) a ∣ --- -- pp = WedgeElim _ (λ _ → isOfHLevelTrunc (4 +ℕ n) _ _) --- -- (cong ∣_∣ₕ (sym (snd f))) - --- -- pp-inr : pp (inr north) ≡ cong ∣_∣ₕ (sym (snd g)) --- -- pp-inr = (λ j → transport (λ i → 0ₖ (2 +ℕ n) ≡ ∣ compPath-filler' (snd f) (sym (snd g)) (~ j) i ∣ₕ) --- -- λ i → ∣ snd f (~ i ∨ j) ∣ₕ) --- -- ∙ λ j → transp (λ i → 0ₖ (2 +ℕ n) ≡ ∣ snd g (~ i ∧ ~ j) ∣ₕ) j λ i → ∣ snd g (~ i ∨ ~ j) ∣ₕ - --- -- α∨'-fun : C* n f g → coHomK (2 +ℕ n) --- -- α∨'-fun (inl x) = 0ₖ _ --- -- α∨'-fun (inr x) = ∣ x ∣ --- -- α∨'-fun (push a i₁) = pp a i₁ - --- -- α∨' : coHom (2 +ℕ n) (C* n f g) --- -- α∨' = ∣ α∨'-fun ∣₂ - - --- -- β+ : coHom ((2 +ℕ n) +' (2 +ℕ n)) C+' --- -- β+ = Iso.inv (fst (H⁴-C∨ n f g)) 1 - --- -- q : C+' → C* n f g --- -- q (inl x) = inl x --- -- q (inr x) = inr x --- -- q (push a i₁) = (push (θ a) ∙ λ i → inr (+≡ n f g a (~ i))) i₁ - --- -- jₗ : HopfInvariantPush n (fst f) → C* n f g --- -- jₗ (inl x) = inl x --- -- jₗ (inr x) = inr x --- -- jₗ (push a i₁) = push (inl a) i₁ - --- -- jᵣ : HopfInvariantPush n (fst g) → C* n f g --- -- jᵣ (inl x) = inl x --- -- jᵣ (inr x) = inr x --- -- jᵣ (push a i₁) = push (inr a) i₁ - --- -- α-∨→1 : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → Iso.fun (fst (H²-C* n f g)) (α∨' n f g) ≡ 1 --- -- α-∨→1 n f g = sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (h33 n)) --- -- ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) - --- -- α-∨-lem : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → α∨ n f g ≡ α∨' n f g --- -- α-∨-lem n f g = sym (Iso.leftInv (fst (H²-C* n f g)) _) --- -- ∙∙ cong (Iso.inv (fst (H²-C* n f g))) help --- -- ∙∙ Iso.leftInv (fst (H²-C* n f g)) _ --- -- where --- -- help : Iso.fun (fst (H²-C* n f g)) (α∨ n f g) ≡ Iso.fun (fst (H²-C* n f g)) (α∨' n f g) --- -- help = (Iso.rightInv (fst (H²-C* n f g)) (pos 1)) ∙ sym (α-∨→1 n f g) - --- -- q-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → coHomFun _ (q n f g) (α∨ n f g) ≡ Hopfα n (·Π f g) --- -- q-α n f g = (λ i → coHomFun _ (q n f g) (α-∨-lem n f g i)) --- -- ∙ sym (Iso.leftInv is _) --- -- ∙∙ cong (Iso.inv is) help --- -- ∙∙ Iso.leftInv is _ --- -- where --- -- is = fst (Hopfα-Iso n (·Π f g)) - --- -- help : Iso.fun is (coHomFun _ (q n f g) (α∨' n f g)) --- -- ≡ Iso.fun is (Hopfα n (·Π f g)) --- -- help = sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (h33 n)) --- -- ∙∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) --- -- ∙∙ sym (Hopfα-Iso-α n (·Π f g)) - - --- -- βₗ≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → βₗ n f g ≡ βₗ''-fun n f g --- -- βₗ≡ n f g = cong (FF ∘ subber2) --- -- (cong (Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) --- -- (S₊∙ (suc (suc (suc (n +ℕ n))))) --- -- (((suc (suc (n +ℕ n))))))))) help --- -- ∙ help2) --- -- ∙ funExt⁻ FF∘subber ∣ wedgeGen ∣₂ --- -- ∙ cong subber3 (sym βₗ'-fun≡) --- -- where --- -- FF = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (n +ℕ suc n))) .fst --- -- FF' = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (suc (n +ℕ n)))) .fst - --- -- help : Iso.inv (fst (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (1 , 0) --- -- ≡ (∣ ∣_∣ ∣₂ , 0ₕ _) --- -- help = ΣPathP ((h33 _) , IsGroupHom.pres1 (snd (invGroupIso (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) - --- -- subber3 = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - --- -- subber2 = (subst (λ m → coHom m (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) --- -- (sym (cong (2 +ℕ_) (+-suc n n)))) - --- -- FF∘subber : FF ∘ subber2 ≡ subber3 ∘ FF' --- -- FF∘subber = --- -- funExt (λ a → (λ j → transp (λ i → coHom ((suc ∘ suc ∘ suc) (+-suc n n (~ i ∧ j))) (C* n f g)) (~ j) --- -- (MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) ((suc (suc (+-suc n n j)))) .fst --- -- (transp (λ i → coHom (2 +ℕ (+-suc n n (j ∨ ~ i))) --- -- (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) j --- -- a)))) - --- -- wedgeGen : (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) → --- -- coHomK (suc (suc (suc (n +ℕ n))))) --- -- wedgeGen (inl x) = ∣ x ∣ --- -- wedgeGen (inr x) = 0ₖ _ --- -- wedgeGen (push a i₁) = 0ₖ _ - --- -- βₗ'-fun≡ : ∣ βₗ'-fun n f g ∣₂ ≡ FF' ∣ wedgeGen ∣₂ --- -- βₗ'-fun≡ = cong ∣_∣₂ (funExt λ { (inl x) → refl --- -- ; (inr x) → refl --- -- ; (push (inl x) i₁) → refl --- -- ; (push (inr x) i₁) j → Kn→ΩKn+10ₖ _ (~ j) i₁ --- -- ; (push (push a i₂) i₁) j → Kn→ΩKn+10ₖ _ (~ j ∧ i₂) i₁}) - --- -- help2 : Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) (((suc (suc (n +ℕ n)))))))) --- -- (∣ ∣_∣ ∣₂ , 0ₕ _) --- -- ≡ ∣ wedgeGen ∣₂ --- -- help2 = cong ∣_∣₂ (funExt λ { (inl x) → rUnitₖ (suc (suc (suc (n +ℕ n)))) ∣ x ∣ --- -- ; (inr x) → refl --- -- ; (push a i₁) → refl}) - --- -- βᵣ≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → βᵣ n f g ≡ βᵣ''-fun n f g --- -- βᵣ≡ n f g = --- -- cong (FF ∘ subber2) --- -- (cong (Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) --- -- (S₊∙ (suc (suc (suc (n +ℕ n))))) --- -- (((suc (suc (n +ℕ n))))))))) --- -- help --- -- ∙ help2) --- -- ∙ funExt⁻ FF∘subber ∣ wedgeGen ∣₂ --- -- ∙ cong (subber3) (sym βᵣ'-fun≡) --- -- where --- -- FF = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (n +ℕ suc n))) .fst --- -- FF' = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (suc (n +ℕ n)))) .fst - --- -- help : Iso.inv (fst (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (0 , 1) --- -- ≡ (0ₕ _ , ∣ ∣_∣ ∣₂) --- -- help = ΣPathP (IsGroupHom.pres1 (snd (invGroupIso (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) , (h33 _)) - --- -- subber3 = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - --- -- subber2 = (subst (λ m → coHom m (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) --- -- (sym (cong (2 +ℕ_) (+-suc n n)))) - --- -- FF∘subber : FF ∘ subber2 ≡ subber3 ∘ FF' --- -- FF∘subber = --- -- funExt (λ a → (λ j → transp (λ i → coHom ((suc ∘ suc ∘ suc) (+-suc n n (~ i ∧ j))) (C* n f g)) (~ j) --- -- (MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) ((suc (suc (+-suc n n j)))) .fst --- -- (transp (λ i → coHom (2 +ℕ (+-suc n n (j ∨ ~ i))) --- -- (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) j --- -- a)))) - --- -- wedgeGen : (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) → --- -- coHomK (suc (suc (suc (n +ℕ n))))) --- -- wedgeGen (inl x) = 0ₖ _ --- -- wedgeGen (inr x) = ∣ x ∣ --- -- wedgeGen (push a i₁) = 0ₖ _ - - --- -- help2 : Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) (((suc (suc (n +ℕ n)))))))) --- -- (0ₕ _ , ∣ ∣_∣ ∣₂) --- -- ≡ ∣ wedgeGen ∣₂ --- -- help2 = cong ∣_∣₂ (funExt λ { (inl x) → refl --- -- ; (inr x) → lUnitₖ (suc (suc (suc (n +ℕ n)))) ∣ x ∣ --- -- ; (push a i₁) → refl}) - --- -- βᵣ'-fun≡ : ∣ βᵣ'-fun n f g ∣₂ ≡ FF' ∣ wedgeGen ∣₂ --- -- βᵣ'-fun≡ = cong ∣_∣₂ (funExt λ { (inl x) → refl --- -- ; (inr x) → refl --- -- ; (push (inl x) i₁) j → Kn→ΩKn+10ₖ _ (~ j) i₁ --- -- ; (push (inr x) i₁) → refl --- -- ; (push (push a i₂) i₁) j → Kn→ΩKn+10ₖ _ (~ j ∧ ~ i₂) i₁}) - --- -- q-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → coHomFun _ (q n f g) (βₗ n f g) --- -- ≡ β+ n f g --- -- q-βₗ n f g = cong (coHomFun _ (q n f g)) (βₗ≡ n f g) --- -- ∙ transportLem --- -- ∙ cong (subst (λ m → coHom m (C+' n f g)) --- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) --- -- (sym (Iso.leftInv (fst (Hopfβ-Iso n (·Π f g))) (Hopfβ n (·Π f g))) --- -- ∙ cong (Iso.inv ((fst (Hopfβ-Iso n (·Π f g))))) (d-β n (·Π f g))) --- -- where --- -- transportLem : coHomFun (suc (suc (suc (n +ℕ suc n)))) (q n f g) --- -- (βₗ''-fun n f g) --- -- ≡ subst (λ m → coHom m (C+' n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- -- (Hopfβ n (·Π f g)) --- -- transportLem = --- -- natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (q n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- -- ∙ cong (subst (λ m → coHom m (C+' n f g)) --- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) --- -- (cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → loll a i₁})) --- -- where --- -- open import Cubical.Foundations.Path --- -- loll : (x : fst (S₊∙ (3 +ℕ (n +ℕ n)))) → --- -- PathP --- -- (λ i₁ → --- -- βₗ'-fun n f g ((push (θ x) ∙ λ i → inr (+≡ n f g x (~ i))) i₁) ≡ --- -- MV.d-pre Unit (fst (S₊∙ (2 +ℕ n))) (fst (S₊∙ (3 +ℕ n +ℕ n))) --- -- (λ _ → tt) (fst (·Π f g)) (3 +ℕ n +ℕ n) ∣_∣ (push x i₁)) --- -- (λ _ → βₗ'-fun n f g (q n f g (inl tt))) --- -- (λ _ → βₗ'-fun n f g (q n f g (inr (·Π f g .fst x)))) --- -- loll x = --- -- flipSquare (cong-∙ (βₗ'-fun n f g) (push (θ x)) (λ i → inr (+≡ n f g x (~ i))) --- -- ∙∙ sym (rUnit _) --- -- ∙∙ lem₁ x) - --- -- where --- -- lem₁ : (x : _) → cong (βₗ'-fun n f g) (push (θ {A = S₊∙ _} x)) ≡ Kn→ΩKn+1 _ ∣ x ∣ --- -- lem₁ north = refl --- -- lem₁ south = sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)) --- -- lem₁ (merid a j) k = lala k j --- -- where --- -- lala : PathP (λ k → Kn→ΩKn+1 _ ∣ north ∣ₕ ≡ (sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north))) k) --- -- (λ j → cong (βₗ'-fun n f g) (push (θ {A = S₊∙ _} (merid a j)))) --- -- (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) --- -- lala = (cong-∙∙ (cong (βₗ'-fun n f g) ∘ push) --- -- ((λ i₁ → inl ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁))) --- -- (push tt) --- -- ((λ i₁ → inr ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁))) --- -- ∙ sym (compPath≡compPath' _ _) --- -- ∙ (λ _ → (λ j → Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∣ (merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) j ∣ₕ) --- -- ∙ Kn→ΩKn+10ₖ _) --- -- ∙ cong (_∙ Kn→ΩKn+10ₖ _) (cong-∙ ((Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ)) --- -- (merid a) (sym (merid north))) --- -- ∙ sym (assoc _ _ _) --- -- ∙ cong (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a) ∙_) --- -- (sym (symDistr (sym ((Kn→ΩKn+10ₖ _))) (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)))))) --- -- ◁ λ i j → compPath-filler (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) --- -- (sym (sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) --- -- (cong ∣_∣ₕ (merid north)))) --- -- (~ i) j - --- -- q-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → coHomFun _ (q n f g) (βᵣ n f g) --- -- ≡ β+ n f g --- -- q-βᵣ n f g = cong (coHomFun _ (q n f g)) (βᵣ≡ n f g) --- -- ∙ transportLem --- -- ∙ cong (subst (λ m → coHom m (C+' n f g)) --- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) --- -- (sym (Iso.leftInv (fst (Hopfβ-Iso n (·Π f g))) (Hopfβ n (·Π f g))) --- -- ∙ cong (Iso.inv ((fst (Hopfβ-Iso n (·Π f g))))) (d-β n (·Π f g))) --- -- where --- -- transportLem : coHomFun (suc (suc (suc (n +ℕ suc n)))) (q n f g) --- -- (βᵣ''-fun n f g) --- -- ≡ subst (λ m → coHom m (C+' n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- -- (Hopfβ n (·Π f g)) --- -- transportLem = --- -- natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (q n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- -- ∙ cong (subst (λ m → coHom m (C+' n f g)) --- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) --- -- (cong ∣_∣₂ (funExt λ { (inl x) → refl --- -- ; (inr x) → refl --- -- ; (push a i₁) → loll a i₁})) --- -- where --- -- open import Cubical.Foundations.Path --- -- loll : (x : fst (S₊∙ (3 +ℕ (n +ℕ n)))) → --- -- PathP --- -- (λ i₁ → --- -- βᵣ'-fun n f g ((push (θ x) ∙ λ i → inr (+≡ n f g x (~ i))) i₁) ≡ --- -- MV.d-pre Unit (fst (S₊∙ (2 +ℕ n))) (fst (S₊∙ (3 +ℕ n +ℕ n))) --- -- (λ _ → tt) (fst (·Π f g)) (3 +ℕ n +ℕ n) ∣_∣ (push x i₁)) --- -- (λ _ → βᵣ'-fun n f g (q n f g (inl tt))) --- -- (λ _ → βᵣ'-fun n f g (q n f g (inr (·Π f g .fst x)))) --- -- loll x = flipSquare (cong-∙ (βᵣ'-fun n f g) (push (θ x)) (λ i → inr (+≡ n f g x (~ i))) --- -- ∙∙ sym (rUnit _) --- -- ∙∙ lem₁ x) --- -- where --- -- lem₁ : (x : _) → cong (βᵣ'-fun n f g) (push (θ {A = S₊∙ _} x)) ≡ Kn→ΩKn+1 _ ∣ x ∣ --- -- lem₁ north = sym (Kn→ΩKn+10ₖ _) --- -- lem₁ south = cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)) --- -- lem₁ (merid a j) k = lala k j -- lala k j --- -- where --- -- lala : PathP (λ k → (Kn→ΩKn+10ₖ _) (~ k) ≡ (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north))) k) --- -- (λ j → cong (βᵣ'-fun n f g) (push (θ {A = S₊∙ _} (merid a j)))) --- -- (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) --- -- lala = (cong-∙∙ (cong (βᵣ'-fun n f g) ∘ push) --- -- (λ i₁ → inl ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁)) --- -- (push tt) --- -- (λ i₁ → inr ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁)) --- -- ∙∙ (cong (sym (Kn→ΩKn+10ₖ _) ∙_) (cong-∙ (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a) (sym (merid (ptSn _))))) --- -- ∙∙ sym (help3 _ _ _)) --- -- ◁ symP (doubleCompPath-filler --- -- (sym (Kn→ΩKn+10ₖ _)) --- -- (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) --- -- (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (sym (merid north)))) --- -- open import Cubical.Foundations.Path --- -- jₗ-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → coHomFun _ (jₗ n f g) (α∨ n f g) --- -- ≡ Hopfα n f --- -- jₗ-α n f g = --- -- cong (coHomFun _ (jₗ n f g)) (α-∨-lem n f g) --- -- ∙ cong ∣_∣₂ (funExt (HopfInvariantPushElim n (fst f) --- -- (isOfHLevelPath ((suc (suc (suc (suc (n +ℕ n)))))) --- -- (isOfHLevelPlus' {n = n} (4 +ℕ n) (isOfHLevelTrunc (4 +ℕ n))) _ _) --- -- refl --- -- (λ _ → refl) --- -- λ j → refl)) - --- -- jₗ-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → coHomFun _ (jₗ n f g) (βₗ n f g) --- -- ≡ subst (λ m → coHom m (HopfInvariantPush n (fst f))) --- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- -- (Hopfβ n f) --- -- jₗ-βₗ n f g = --- -- cong (coHomFun _ (jₗ n f g)) (βₗ≡ n f g) --- -- ∙∙ natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (jₗ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- -- ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) --- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) --- -- (cong ∣_∣₂ --- -- (funExt λ { (inl x) → refl --- -- ; (inr x) → refl --- -- ; (push a i₁) → refl})) - --- -- jₗ-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → coHomFun _ (jₗ n f g) (βᵣ n f g) --- -- ≡ 0ₕ _ --- -- jₗ-βᵣ n f g = --- -- cong (coHomFun _ (jₗ n f g)) (βᵣ≡ n f g) --- -- ∙∙ natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (jₗ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- -- ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) --- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) --- -- cool --- -- where --- -- cool : coHomFun (suc (suc (suc (suc (n +ℕ n))))) (jₗ n f g) --- -- ∣ βᵣ'-fun n f g ∣₂ ≡ 0ₕ _ --- -- cool = cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → refl}) - --- -- jᵣ-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → coHomFun _ (jᵣ n f g) (α∨ n f g) --- -- ≡ Hopfα n g --- -- jᵣ-α n f g = cong (coHomFun _ (jᵣ n f g)) (α-∨-lem n f g) --- -- ∙ cong ∣_∣₂ (funExt (HopfInvariantPushElim n (fst g) --- -- (isOfHLevelPath ((suc (suc (suc (suc (n +ℕ n)))))) --- -- (isOfHLevelPlus' {n = n} (4 +ℕ n) (isOfHLevelTrunc (4 +ℕ n))) _ _) --- -- refl --- -- (λ _ → refl) --- -- (flipSquare (pp-inr n f g)))) - --- -- jᵣ-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → coHomFun _ (jᵣ n f g) (βₗ n f g) ≡ 0ₕ _ --- -- jᵣ-βₗ n f g = --- -- cong (coHomFun _ (jᵣ n f g)) (βₗ≡ n f g) --- -- ∙∙ natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (jᵣ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- -- ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) --- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) --- -- cool --- -- where --- -- cool : coHomFun (suc (suc (suc (suc (n +ℕ n))))) (jᵣ n f g) --- -- ∣ βₗ'-fun n f g ∣₂ ≡ 0ₕ _ --- -- cool = cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → refl}) - - --- -- jᵣ-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → coHomFun _ (jᵣ n f g) (βᵣ n f g) --- -- ≡ subst (λ m → coHom m (HopfInvariantPush n (fst g))) --- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- -- (Hopfβ n g) --- -- jᵣ-βᵣ n f g = --- -- cong (coHomFun _ (jᵣ n f g)) (βᵣ≡ n f g) --- -- ∙∙ natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (jᵣ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- -- ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst g))) --- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) --- -- (cong ∣_∣₂ --- -- (funExt λ { (inl x) → refl --- -- ; (inr x) → refl --- -- ; (push a i₁) → refl})) - --- -- gen₂ℤ×ℤ : gen₂-by (DirProd ℤGroup ℤGroup) (1 , 0) (0 , 1) --- -- fst (gen₂ℤ×ℤ (x , y)) = x , y --- -- snd (gen₂ℤ×ℤ (x , y)) = --- -- ΣPathP ((cong₂ _+_ ((·Comm 1 x) ∙ cong fst (sym (distrLem 1 0 x))) ((·Comm 0 y) ∙ cong fst (sym (distrLem 0 1 y)))) --- -- , +Comm y 0 ∙ cong₂ _+_ (·Comm 0 x ∙ cong snd (sym (distrLem 1 0 x))) (·Comm 1 y ∙ cong snd (sym (distrLem 0 1 y)))) --- -- where --- -- ℤ×ℤ = DirProd ℤGroup ℤGroup --- -- _+''_ = GroupStr._·_ (snd ℤ×ℤ) - --- -- ll3 : (x : ℤ) → - x ≡ 0 - x --- -- ll3 (pos zero) = refl --- -- ll3 (pos (suc zero)) = refl --- -- ll3 (pos (suc (suc n))) = --- -- cong predℤ (ll3 (pos (suc n))) --- -- ll3 (negsuc zero) = refl --- -- ll3 (negsuc (suc n)) = cong sucℤ (ll3 (negsuc n)) - --- -- distrLem : (x y : ℤ) (z : ℤ) --- -- → Path (ℤ × ℤ) (z ℤ[ ℤ×ℤ ]· (x , y)) (z · x , z · y) --- -- distrLem x y (pos zero) = refl --- -- distrLem x y (pos (suc n)) = --- -- (cong ((x , y) +''_) (distrLem x y (pos n))) --- -- distrLem x y (negsuc zero) = ΣPathP (sym (ll3 x) , sym (ll3 y)) --- -- distrLem x y (negsuc (suc n)) = --- -- cong₂ _+''_ (ΣPathP (sym (ll3 x) , sym (ll3 y))) --- -- (distrLem x y (negsuc n)) - --- -- genH²ⁿC* : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → gen₂-by (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) --- -- (βₗ n f g) --- -- (βᵣ n f g) --- -- genH²ⁿC* n f g = --- -- Iso-pres-gen₂ (DirProd ℤGroup ℤGroup) --- -- (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) --- -- (1 , 0) --- -- (0 , 1) --- -- gen₂ℤ×ℤ --- -- (invGroupIso (H⁴-C* n f g)) - --- -- private - --- -- H⁴C* : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Group _ --- -- H⁴C* n f g = coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) - --- -- X : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → ℤ --- -- X n f g = (genH²ⁿC* n f g) (α∨ n f g ⌣ α∨ n f g) .fst .fst --- -- Y : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → ℤ --- -- Y n f g = (genH²ⁿC* n f g) (α∨ n f g ⌣ α∨ n f g) .fst .snd - --- -- α∨≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → α∨ n f g ⌣ α∨ n f g ≡ ((X n f g) ℤ[ H⁴C* n f g ]· βₗ n f g) --- -- +ₕ ((Y n f g) ℤ[ H⁴C* n f g ]· βᵣ n f g) --- -- α∨≡ n f g = (genH²ⁿC* n f g) (α∨ n f g ⌣ α∨ n f g) .snd - - --- -- eq₁ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → (Hopfα n (·Π f g)) ⌣ (Hopfα n (·Π f g)) --- -- ≡ ((X n f g + Y n f g) ℤ[ coHomGr _ _ ]· (β+ n f g)) --- -- eq₁ n f g = --- -- cong (λ x → x ⌣ x) (sym (q-α n f g) ∙ cong (coHomFun (suc (suc n)) (q n f g)) (α-∨-lem n f g)) --- -- ∙ cong ((coHomFun _) (q _ f g)) (cong (λ x → x ⌣ x) (sym (α-∨-lem n f g)) --- -- ∙ α∨≡ n f g) --- -- ∙ IsGroupHom.pres· (coHomMorph _ (q n f g) .snd) _ _ --- -- ∙ cong₂ _+ₕ_ (hompres· (coHomMorph _ (q n f g)) (βₗ n f g) (X n f g) --- -- ∙ cong (λ z → (X n f g) ℤ[ coHomGr _ _ ]· z) --- -- (q-βₗ n f g)) --- -- (hompres· (coHomMorph _ (q n f g)) (βᵣ n f g) (Y n f g) --- -- ∙ cong (λ z → (Y n f g) ℤ[ coHomGr _ _ ]· z) --- -- (q-βᵣ n f g)) --- -- ∙ sym (distrℤ· (coHomGr _ _) (β+ n f g) (X n f g) (Y n f g)) - --- -- coHomFun⌣ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) --- -- → (n : ℕ) (x y : coHom n B) --- -- → coHomFun _ f (x ⌣ y) ≡ coHomFun _ f x ⌣ coHomFun _ f y --- -- coHomFun⌣ f n = sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl - --- -- eq₂ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → Hopfα n f ⌣ Hopfα n f --- -- ≡ (X n f g ℤ[ coHomGr _ _ ]· subst (λ m → coHom m (HopfInvariantPush n (fst f))) --- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- -- (Hopfβ n f)) --- -- eq₂ n f g = --- -- cong (λ x → x ⌣ x) (sym (jₗ-α n f g)) --- -- ∙∙ sym (coHomFun⌣ (jₗ n f g) _ _ _) --- -- ∙∙ cong (coHomFun _ (jₗ n f g)) (α∨≡ n f g) --- -- ∙∙ IsGroupHom.pres· (snd (coHomMorph _ (jₗ n f g))) _ _ --- -- ∙∙ cong₂ _+ₕ_ (hompres· (coHomMorph _ (jₗ n f g)) (βₗ n f g) (X n f g)) --- -- (hompres· (coHomMorph _ (jₗ n f g)) (βᵣ n f g) (Y n f g) --- -- ∙ cong (λ z → (Y n f g) ℤ[ coHomGr _ _ ]· z) --- -- (jₗ-βᵣ n f g)) --- -- ∙∙ cong₂ _+ₕ_ refl (rUnitℤ· (coHomGr _ _) (Y n f g)) --- -- ∙∙ (rUnitₕ _ _ --- -- ∙ cong (X n f g ℤ[ coHomGr _ _ ]·_) (jₗ-βₗ n f g)) - --- -- eq₃ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- -- → Hopfα n g ⌣ Hopfα n g --- -- ≡ (Y n f g ℤ[ coHomGr _ _ ]· subst (λ m → coHom m (HopfInvariantPush n (fst f))) --- -- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- -- (Hopfβ n g)) --- -- eq₃ n f g = --- -- cong (λ x → x ⌣ x) (sym (jᵣ-α n f g)) --- -- ∙∙ sym (coHomFun⌣ (jᵣ n f g) _ _ _) --- -- ∙∙ cong (coHomFun _ (jᵣ n f g)) (α∨≡ n f g) --- -- ∙∙ IsGroupHom.pres· (snd (coHomMorph _ (jᵣ n f g))) _ _ --- -- ∙∙ cong₂ _+ₕ_ (hompres· (coHomMorph _ (jᵣ n f g)) (βₗ n f g) (X n f g) --- -- ∙ cong (λ z → (X n f g) ℤ[ coHomGr _ _ ]· z) --- -- (jᵣ-βₗ n f g)) --- -- (hompres· (coHomMorph _ (jᵣ n f g)) (βᵣ n f g) (Y n f g)) --- -- ∙∙ cong₂ _+ₕ_ (rUnitℤ· (coHomGr _ _) (X n f g)) refl --- -- ∙∙ ((lUnitₕ _ _) ∙ cong (Y n f g ℤ[ coHomGr _ _ ]·_) (jᵣ-βᵣ n f g)) - --- -- eq₂-eq₂ : (n : ℕ) (f g : S₊∙ (suc (suc (suc (n +ℕ n)))) →∙ S₊∙ (suc (suc n))) --- -- → HopfFunction n (·Π f g) ≡ HopfFunction n f + HopfFunction n g --- -- eq₂-eq₂ n f g = --- -- mainL --- -- ∙ sym (cong₂ _+_ (help n f g) (helpg n f g)) --- -- where --- -- transpLem : ∀ {ℓ} {G : ℕ → Group ℓ} --- -- → (n m : ℕ) --- -- → (x : ℤ) --- -- → (p : m ≡ n) --- -- → (g : fst (G n)) --- -- → subst (fst ∘ G) p (x ℤ[ G m ]· subst (fst ∘ G) (sym p) g) ≡ (x ℤ[ G n ]· g) --- -- transpLem {G = G} n m x = --- -- J (λ n p → (g : fst (G n)) → subst (fst ∘ G) p (x ℤ[ G m ]· subst (fst ∘ G) (sym p) g) --- -- ≡ (x ℤ[ G n ]· g)) --- -- λ g → transportRefl _ ∙ cong (x ℤ[ G m ]·_) (transportRefl _) - --- -- mainL : HopfFunction n (·Π f g) ≡ X n f g + Y n f g --- -- mainL = cong (Iso.fun (fst (Hopfβ-Iso n (·Π f g)))) --- -- (cong (subst (λ x → coHom x (HopfInvariantPush n (fst (·Π f g)))) --- -- λ i₁ → suc (suc (suc (+-suc n n i₁)))) --- -- (eq₁ n f g)) --- -- ∙∙ cong (Iso.fun (fst (Hopfβ-Iso n (·Π f g)))) --- -- (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst (·Π f g)))} _ _ --- -- (X n f g + Y n f g) (λ i₁ → suc (suc (suc (+-suc n n i₁)))) --- -- (Iso.inv (fst (Hopfβ-Iso n (·Π f g))) 1)) --- -- ∙∙ hompres· (_ , snd (Hopfβ-Iso n (·Π f g))) (Iso.inv (fst (Hopfβ-Iso n (·Π f g))) 1) --- -- (X n f g + Y n f g) --- -- ∙∙ cong ((X n f g + Y n f g) ℤ[ ℤ , ℤGroup .snd ]·_) --- -- (Iso.rightInv ((fst (Hopfβ-Iso n (·Π f g)))) 1) --- -- ∙∙ rUnitℤ·ℤ (X n f g + Y n f g) - - --- -- help : (n : ℕ) (f g : _) → HopfFunction n f ≡ X n f g --- -- help n f g = cong (Iso.fun (fst (Hopfβ-Iso n f))) --- -- (cong (subst (λ x → coHom x (HopfInvariantPush n (fst f))) --- -- (λ i₁ → suc (suc (suc (+-suc n n i₁))))) (eq₂ n f g)) --- -- ∙ cong (Iso.fun (fst (Hopfβ-Iso n f))) --- -- (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst f))} _ _ --- -- (X n f g) --- -- ((cong (suc ∘ suc ∘ suc) (+-suc n n))) --- -- (Hopfβ n f)) --- -- ∙ hompres· (_ , snd (Hopfβ-Iso n f)) (Hopfβ n f) (X n f g) --- -- ∙ cong (X n f g ℤ[ ℤ , ℤGroup .snd ]·_) (d-β n f) --- -- ∙ rUnitℤ·ℤ (X n f g) - --- -- helpg : (n : ℕ) (f g : _) → HopfFunction n g ≡ Y n f g --- -- helpg n f g = --- -- cong (Iso.fun (fst (Hopfβ-Iso n g))) --- -- (cong (subst (λ x → coHom x (HopfInvariantPush n (fst g))) --- -- (λ i₁ → suc (suc (suc (+-suc n n i₁))))) --- -- (eq₃ n f g)) --- -- ∙∙ cong (Iso.fun (fst (Hopfβ-Iso n g))) --- -- (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst g))} _ _ --- -- (Y n f g) --- -- ((cong (suc ∘ suc ∘ suc) (+-suc n n))) --- -- (Hopfβ n g)) --- -- ∙∙ hompres· (_ , snd (Hopfβ-Iso n g)) (Hopfβ n g) (Y n f g) --- -- ∙∙ cong (Y n f g ℤ[ ℤ , ℤGroup .snd ]·_) (d-β n g) --- -- ∙∙ rUnitℤ·ℤ (Y n f g) - - --- -- {- --- -- q-α : coHomFun _ q α∨ ≡ Hopfα n (·Π f g) --- -- q-α with n --- -- ... | zero = ? --- -- ... | (suc n) = ? --- -- where --- -- helpMe! : Iso.fun (fst (Hopfα-Iso _ (·Π f g))) (Hopfα n (·Π f g)) --- -- ≡ Iso.fun (fst (Hopfα-Iso _ (·Π f g))) (coHomFun _ q α∨) --- -- helpMe! = {!!} - --- -- q-βₗ : coHomFun _ q βₗ ≡ {!(Hopfβ n ?)!} --- -- q-βₗ = {!!} - --- -- q-βᵣ : coHomFun _ q βᵣ ≡ {!!} --- -- q-βᵣ = {!!} -} - --- -- -- helpi : {!(f : ? →∙ ?) (HopfInvariantPush n (fst f) → ?)!} --- -- -- helpi = {!!} - --- -- -- lal : {!!} --- -- -- lal = {!!} - --- -- -- isHomHopf : ∀ n → (x y : ∥ S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n) ∥₂) --- -- -- → Hopfie (x ·π y) ≡ Hopfie x + Hopfie y --- -- -- isHomHopf n = --- -- -- sElim2 (λ _ _ → isOfHLevelPath 2 isSetℤ _ _) --- -- -- λ f g → (λ i → Iso.fun (fst (Hopfβ-Iso n (·Π f g))) (subber n (·Π f g) (⌣-α n (·Π f g)))) --- -- -- ∙ cong (Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) --- -- -- {!!} --- -- -- ∙∙ IsGroupHom.pres· (((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) .snd) --- -- -- (invEq (fst (GroupEquiv1 n f)) (subber n f (⌣-α n f))) --- -- -- (invEq (fst (GroupEquiv1 n g)) (subber n g (⌣-α n g))) --- -- -- ∙∙ λ i → Iso.fun (fst (Hopfβ-Iso n f)) (subber n f (⌣-α n f)) --- -- -- + Iso.fun (fst (Hopfβ-Iso n g)) (subber n g (⌣-α n g)) --- -- -- where --- -- -- cool : (f g : _) → Iso.fun (fst (Hopfβ-Iso n (·Π f g))) (subber n (·Π f g) (⌣-α n f)) --- -- -- ≡ Iso.fun (fst (Hopfβ-Iso n f)) (subber n f (⌣-α n f)) --- -- -- cool f g = {!!} From 7dee57bdaec3fd10c6408c909d912431054a8bc0 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Thu, 30 Sep 2021 10:48:17 +0200 Subject: [PATCH 68/89] stuff --- Cubical/Algebra/Group/GroupPath.agda | 199 +++++++-------- Cubical/HITs/Hopf.agda | 41 +++ Cubical/ZCohomology/Gysin.agda | 358 +++++++++------------------ 3 files changed, 260 insertions(+), 338 deletions(-) diff --git a/Cubical/Algebra/Group/GroupPath.agda b/Cubical/Algebra/Group/GroupPath.agda index 60d5900cc..32c53d72c 100644 --- a/Cubical/Algebra/Group/GroupPath.agda +++ b/Cubical/Algebra/Group/GroupPath.agda @@ -195,101 +195,104 @@ snd (Iso-pres-gen₂ G H g₁ g₂ genG is h) = (hompres· (_ , snd is) g₁ (fst (fst (genG (inv (fst is) h))))) (hompres· (_ , snd is) g₂ (snd (fst (genG (inv (fst is) h)))))) --- module _ {ℓ : Level} {G' : Group ℓ} where --- private --- G = fst G' - --- _+G_ = GroupStr._·_ (snd G') --- -G = GroupStr.inv (snd G') --- 1G = GroupStr.1g (snd G') - --- _ℤ·_ : ℤ → G → G --- pos zero ℤ· g = 1G --- pos (suc n) ℤ· g = g +G (pos n ℤ· g) --- negsuc zero ℤ· g = -G g --- negsuc (suc n) ℤ· g = -G g +G (negsuc n ℤ· g) - --- rUnitℤ : (g : G) → 1 ℤ· g ≡ g --- rUnitℤ = {!!} - --- copyList : {!length!} --- copyList = {!!} --- open import Cubical.Data.Empty renaming (rec to ⊥-rec) --- module _ {ℓ : Level} (G' : Group ℓ) where --- private --- G = fst G' - --- _+G_ = GroupStr._·_ (snd G') --- -G = GroupStr.inv (snd G') --- 1G = GroupStr.1g (snd G') --- sum : List (ℤ × G) → G --- sum [] = 1G --- sum (x ∷ x₁) = (_ℤ·_ {G' = G'} (fst x) (snd x)) +G sum x₁ - --- sum'help : (l1 : List ℤ) (l2 : List G) (n m : ℕ) --- → (n ≡ m) --- → length l1 ≡ n --- → length l2 ≡ m --- → G --- sum'help [] [] n m p q r = 1G --- sum'help [] (x ∷ l2) n m p q r = --- ⊥-rec (snotz (r ∙∙ sym p ∙∙ sym q)) --- sum'help (x ∷ l1) [] n m p q r = --- ⊥-rec (snotz (q ∙∙ p ∙∙ sym r)) --- sum'help (x ∷ l1) (x₁ ∷ l2) n m p q r = --- _ℤ·_ {G' = G'} x x₁ +G sum'help l1 l2 _ _ (cong predℕ p) (cong predℕ q) (cong predℕ r) - --- sum' : (l1 : List ℤ) (l2 : List G) → length l1 ≡ length l2 --- → G --- sum' l1 l2 p = sum'help l1 l2 _ _ p refl refl - --- isFinGen : Type _ --- isFinGen = --- Σ[ l ∈ List G ] --- ((g : G) → Σ[ l2 ∈ List ℤ ] --- Σ[ p ∈ (length l2 ≡ length l) ] --- sum' l2 l p ≡ g) - --- isFinGenℤ : isFinGen ℤGroup --- isFinGenℤ = (1 ∷ []) --- , λ g → (g ∷ []) , refl , help g --- where --- help : (g : ℤ) → sum' ℤGroup (g ∷ []) (pos 1 ∷ []) (λ _ → 1) ≡ g --- help (pos zero) = refl --- help (pos (suc n)) = cong (pos 1 +ℤ_) (help (pos n)) ∙ +Comm (pos 1) (pos n) --- help (negsuc zero) = refl --- help (negsuc (suc n)) = cong (negsuc 0 +ℤ_) (help (negsuc n)) ∙ +Comm (negsuc 0) (negsuc n) - --- open import Cubical.Algebra.Group.DirProd --- isFinGenℤ×ℤ : isFinGen (DirProd ℤGroup ℤGroup) --- fst isFinGenℤ×ℤ = (1 , 0) ∷ (0 , 1) ∷ [] --- fst (snd isFinGenℤ×ℤ (x , y)) = x ∷ y ∷ [] --- fst (snd (snd isFinGenℤ×ℤ (x , y))) = refl --- snd (snd (snd isFinGenℤ×ℤ (x , y))) = --- ΣPathP ((λ i → fst ((distrLem 1 0 x) i) +ℤ fst (distrLem 0 1 y i)) --- ∙ (λ i → (·Comm x 1 i) +ℤ (·Comm y 0 i)) --- , (λ i → snd ((distrLem 1 0 x) i) +ℤ snd (distrLem 0 1 y i)) --- ∙ (λ i → (·Comm x 0 i) +ℤ (·Comm y 1 i)) --- ∙ sym (+Comm y 0)) --- where --- ll : (x : ℤ) → - x ≡ 0 -ℤ x --- ll (pos zero) = refl --- ll (pos (suc zero)) = refl --- ll (pos (suc (suc n))) = --- cong predℤ (ll (pos (suc n))) --- ll (negsuc zero) = refl --- ll (negsuc (suc n)) = cong sucℤ (ll (negsuc n)) - --- ℤ×ℤ = DirProd ℤGroup ℤGroup --- _+''_ = GroupStr._·_ (snd ℤ×ℤ) --- distrLem : (x y : ℤ) (z : ℤ) --- → Path (ℤ × ℤ) (_ℤ·_ {G' = DirProd ℤGroup ℤGroup} z (x , y)) (z * x , z * y) --- distrLem x y (pos zero) = refl --- distrLem x y (pos (suc n)) = cong ((x , y) +''_) (distrLem x y (pos n)) --- distrLem x y (negsuc zero) = ΣPathP (sym (ll x) , sym (ll y)) --- distrLem x y (negsuc (suc n)) = --- cong₂ _+''_ (ΣPathP (sym (ll x) , sym (ll y))) (distrLem x y (negsuc n)) - --- maini : ∀ {ℓ} {G : Type ℓ} --- → {!GroupIso!} --- maini = {!!} +rUnitℤ· : ∀ {ℓ} (G : Group ℓ) → (x : ℤ) → (x ℤ[ G ]· GroupStr.1g (snd G)) + ≡ GroupStr.1g (snd G) +rUnitℤ· G (pos zero) = refl +rUnitℤ· G (pos (suc n)) = + cong (GroupStr._·_ (snd G) (GroupStr.1g (snd G))) + (rUnitℤ· G (pos n)) + ∙ GroupStr.lid (snd G) (GroupStr.1g (snd G)) +rUnitℤ· G (negsuc zero) = GroupTheory.inv1g G +rUnitℤ· G (negsuc (suc n)) = + cong₂ (GroupStr._·_ (snd G)) (GroupTheory.inv1g G) (rUnitℤ· G (negsuc n)) + ∙ GroupStr.lid (snd G) (GroupStr.1g (snd G)) + +rUnitℤ·ℤ : (x : ℤ) → (x ℤ[ ℤGroup ]· 1) ≡ x +rUnitℤ·ℤ (pos zero) = refl +rUnitℤ·ℤ (pos (suc n)) = cong (pos 1 +ℤ_) (rUnitℤ·ℤ (pos n)) ∙ sym (pos+ 1 n) +rUnitℤ·ℤ (negsuc zero) = refl +rUnitℤ·ℤ (negsuc (suc n)) = cong (-1 +ℤ_) (rUnitℤ·ℤ (negsuc n)) ∙ +Comm (negsuc 0) (negsuc n) + +private + precommℤ : ∀ {ℓ} (G : Group ℓ) → (g : fst G) (y : ℤ) → (snd G GroupStr.· (y ℤ[ G ]· g)) g + ≡ (sucℤ y ℤ[ G ]· g) + precommℤ G g (pos zero) = GroupStr.lid (snd G) g + ∙ sym (GroupStr.rid (snd G) g) + precommℤ G g (pos (suc n)) = + sym (GroupStr.assoc (snd G) _ _ _) + ∙ cong ((snd G GroupStr.· g)) (precommℤ G g (pos n)) + precommℤ G g (negsuc zero) = + GroupStr.invl (snd G) g + precommℤ G g (negsuc (suc n)) = + sym (GroupStr.assoc (snd G) _ _ _) + ∙ cong ((snd G GroupStr.· GroupStr.inv (snd G) g)) (precommℤ G g (negsuc n)) + ∙ negsucLem n + where + negsucLem : (n : ℕ) → (snd G GroupStr.· GroupStr.inv (snd G) g) + (sucℤ (negsuc n) ℤ[ G ]· g) + ≡ (sucℤ (negsuc (suc n)) ℤ[ G ]· g) + negsucLem zero = (GroupStr.rid (snd G) _) + negsucLem (suc n) = refl + +distrℤ· : ∀ {ℓ} (G : Group ℓ) → (g : fst G) (x y : ℤ) + → ((x +ℤ y) ℤ[ G ]· g) + ≡ GroupStr._·_ (snd G) (x ℤ[ G ]· g) (y ℤ[ G ]· g) +distrℤ· G' g (pos zero) y = cong (_ℤ[ G' ]· g) (+Comm 0 y) + ∙ sym (GroupStr.lid (snd G') _) +distrℤ· G' g (pos (suc n)) (pos n₁) = + cong (_ℤ[ G' ]· g) (sym (pos+ (suc n) n₁)) + ∙ cong (GroupStr._·_ (snd G') g) (cong (_ℤ[ G' ]· g) (pos+ n n₁) ∙ distrℤ· G' g (pos n) (pos n₁)) + ∙ GroupStr.assoc (snd G') _ _ _ +distrℤ· G' g (pos (suc n)) (negsuc zero) = + distrℤ· G' g (pos n) 0 + ∙ ((GroupStr.rid (snd G') (pos n ℤ[ G' ]· g) ∙ sym (GroupStr.lid (snd G') (pos n ℤ[ G' ]· g))) + ∙ cong (λ x → GroupStr._·_ (snd G') x (pos n ℤ[ G' ]· g)) + (sym (GroupStr.invl (snd G') g)) ∙ sym (GroupStr.assoc (snd G') _ _ _)) + ∙ (GroupStr.assoc (snd G') _ _ _) + ∙ cong (λ x → GroupStr._·_ (snd G') x (pos n ℤ[ G' ]· g)) (GroupStr.invl (snd G') g) + ∙ GroupStr.lid (snd G') _ + ∙ sym (GroupStr.rid (snd G') _) + ∙ (cong (GroupStr._·_ (snd G') (pos n ℤ[ G' ]· g)) (sym (GroupStr.invr (snd G') g)) + ∙ GroupStr.assoc (snd G') _ _ _) + ∙ cong (λ x → GroupStr._·_ (snd G') x (GroupStr.inv (snd G') g)) + (precommℤ G' g (pos n)) +distrℤ· G' g (pos (suc n)) (negsuc (suc n₁)) = + cong (_ℤ[ G' ]· g) (predℤ+negsuc n₁ (pos (suc n))) + ∙∙ distrℤ· G' g (pos n) (negsuc n₁) + ∙∙ (cong (λ x → GroupStr._·_ (snd G') x (negsuc n₁ ℤ[ G' ]· g)) + ((sym (GroupStr.rid (snd G') (pos n ℤ[ G' ]· g)) + ∙ cong (GroupStr._·_ (snd G') (pos n ℤ[ G' ]· g)) (sym (GroupStr.invr (snd G') g))) + ∙∙ GroupStr.assoc (snd G') _ _ _ + ∙∙ cong (λ x → GroupStr._·_ (snd G') x (GroupStr.inv (snd G') g)) (precommℤ G' g (pos n) )) + ∙ sym (GroupStr.assoc (snd G') _ _ _)) +distrℤ· G' g (negsuc zero) y = + cong (_ℤ[ G' ]· g) (+Comm -1 y) ∙ lol1 y + module _ where + lol1 : (y : ℤ) → (predℤ y ℤ[ G' ]· g) ≡ (snd G' GroupStr.· GroupStr.inv (snd G') g) (y ℤ[ G' ]· g) + lol1 (pos zero) = sym (GroupStr.rid (snd G') _) + lol1 (pos (suc n)) = + sym (GroupStr.lid (snd G') ((pos n ℤ[ G' ]· g))) + ∙∙ cong (λ x → GroupStr._·_ (snd G') x (pos n ℤ[ G' ]· g)) (sym (GroupStr.invl (snd G') g)) + ∙∙ sym (GroupStr.assoc (snd G') _ _ _) + lol1 (negsuc n) = refl +distrℤ· G' g (negsuc (suc n)) y = + cong (_ℤ[ G' ]· g) (+Comm (negsuc (suc n)) y) + ∙∙ lol1 G' g 0 (y +negsuc n) + ∙∙ cong (snd G' GroupStr.· GroupStr.inv (snd G') g) + (cong (_ℤ[ G' ]· g) (+Comm y (negsuc n)) ∙ distrℤ· G' g (negsuc n) y) + ∙ (GroupStr.assoc (snd G') _ _ _) + +minus≡0- : (x : ℤ) → - x ≡ (0 -ℤ x) +minus≡0- x = sym (GroupStr.lid (snd ℤGroup) (- x)) + +GroupHomℤ→ℤpres- : (e : GroupHom ℤGroup ℤGroup) → (a : ℤ) → fst e (- a) ≡ - fst e a +GroupHomℤ→ℤpres- e a = cong (fst e) (minus≡0- a) ∙∙ IsGroupHom.presinv (snd e) a ∙∙ sym (minus≡0- _) + +ℤ·≡· : (a b : ℤ) → a * b ≡ (a ℤ[ ℤGroup ]· b) +ℤ·≡· (pos zero) b = refl +ℤ·≡· (pos (suc n)) b = cong (b +ℤ_) (ℤ·≡· (pos n) b) +ℤ·≡· (negsuc zero) b = minus≡0- b +ℤ·≡· (negsuc (suc n)) b = cong₂ _+ℤ_ (minus≡0- b) (ℤ·≡· (negsuc n) b) + +GroupHomℤ→ℤPres· : (e : GroupHom ℤGroup ℤGroup) → (a b : ℤ) → fst e (a * b) ≡ a * fst e b +GroupHomℤ→ℤPres· e a b = cong (fst e) (ℤ·≡· a b) ∙∙ hompres· e b a ∙∙ sym (ℤ·≡· a (fst e b)) diff --git a/Cubical/HITs/Hopf.agda b/Cubical/HITs/Hopf.agda index 0b72621fb..7c26352e6 100644 --- a/Cubical/HITs/Hopf.agda +++ b/Cubical/HITs/Hopf.agda @@ -338,3 +338,44 @@ leftInv IsoS³TotalHopf x = (JoinS¹S¹→TotalHopf→JoinS¹S¹ (S³→joinS¹S¹ (inv IsoS³S3 x))) ∙∙ cong (fun IsoS³S3) (S³→joinS¹S¹→S³ (inv IsoS³S3 x)) ∙∙ Iso.rightInv IsoS³S3 x + +-- kebab : TotalHopf → TotalHopf → TotalHopf +-- kebab (north , b) (c , d) = c , {!b!} +-- kebab (south , b) (c , d) = {!!} +-- kebab (merid a i , b) (c , d) = {!!} + +-- S¹→Space : (x : S¹) → Type +-- S¹→Space base = TotalHopf +-- S¹→Space (loop i) = {!!} + +-- TotalHopf→TotalHopf : TotalHopf → TotalHopf +-- TotalHopf→TotalHopf (north , p) = north , invLooper p +-- TotalHopf→TotalHopf (south , p) = south , p +-- TotalHopf→TotalHopf (merid a i , p) = {!transport (λ i → Glue S¹ (Border a i) → TotalHopf) (λ x → north , invLooper x)!} +-- where +-- kebab2 : PathP (λ i → Glue S¹ (Border a i) → TotalHopf) (λ x → north , invLooper x) λ p → south , p +-- kebab2 = toPathP (funExt λ x → transportRefl _ ∙ ΣPathP (refl , {!p!}) ∙ {!!}) + + +-- open import Cubical.Data.Sum + +-- code : ∀ {ℓ} {A : Type ℓ} {B : Type ℓ} (x : A) → A ⊎ B → Type _ +-- code {A = A} {B = B} x (inl y) = x ≡ y +-- code {A = A} {B = B} y (inr x) = Path (A ⊎ B) (inl y) (inr x) + +-- decode' : ∀ {ℓ} {A : Type ℓ} {B : Type ℓ} → (y : A ⊎ B) (x : A) → code x y → inl x ≡ y +-- decode' (inl x₁) x p = cong inl p +-- decode' (inr x₁) x p = p + +-- encode' : ∀ {ℓ} {A : Type ℓ} {B : Type ℓ} (x : A) (y : A ⊎ B) → inl x ≡ y → code x y +-- encode' x y = J (λ y _ → code x y) refl + +-- encode-decode : ∀ {ℓ} {A : Type ℓ} {B : Type ℓ} (x : A) (y : A) → (s : code {A = A} {B = B} x (inl y)) → encode' {A = A} {B = B} x (inl y) (decode' (inl y) x s) ≡ s +-- encode-decode {A = A} {B = B} x y = J (λ y s → encode' {A = A} {B = B} x (inl y) (λ i → inl (s i)) ≡ s) (transportRefl refl) + +-- en + +-- lal : S₊ 1 → TotalHopf → TotalHopf +-- lal a (north , p) = {!merid a ?!} , {!!} +-- lal a (south , p) = {!!} +-- lal a (merid a₁ i , p) = {!!} diff --git a/Cubical/ZCohomology/Gysin.agda b/Cubical/ZCohomology/Gysin.agda index 25660e4e9..8043a7867 100644 --- a/Cubical/ZCohomology/Gysin.agda +++ b/Cubical/ZCohomology/Gysin.agda @@ -70,6 +70,20 @@ open import Cubical.Homotopy.Hopf open import Cubical.HITs.SetQuotients renaming (_/_ to _/'_) +-- move to Sigma +contrFstΣ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} + → (e : isContr A) + → Iso (Σ A B) (B (fst e)) +Iso.fun (contrFstΣ {B = B} e) (a , b) = subst B (sym (snd e a)) b +Iso.inv (contrFstΣ {B = B} e) b = (fst e) , b +Iso.rightInv (contrFstΣ {B = B} e) b = cong (λ x → subst B x b) h3 ∙ transportRefl b + where + h3 : sym (snd e (fst e)) ≡ refl + h3 = isProp→isSet (isContr→isProp e) _ _ _ _ +Iso.leftInv (contrFstΣ {B = B} e) b = + ΣPathP ((snd e (fst b)) + , (λ j → transp (λ i → B (snd e (fst b) (~ i ∨ j))) j (snd b))) + -- move to group stuff SES→isEquiv : ∀ {ℓ ℓ'} {L R : Group ℓ-zero} → {G : Group ℓ} {H : Group ℓ'} @@ -309,10 +323,32 @@ Hopfα-Iso-α n f = hz : Iso.fun (h n f) (Hopfα n f) ≡ ∣ ∣_∣ ∣₂ hz = refl -_·₀ₕ_ : ∀ {ℓ} {n : ℕ} {A : Type ℓ} → ℤ → coHom n A → coHom n A -_·₀ₕ_ {n = n} a = sMap λ f x → a ·₀ f x + +⌣-α : (n : ℕ) → (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → _ +⌣-α n f = Hopfα n f ⌣ Hopfα n f + +HopfFunction : (n : ℕ) + → (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → ℤ +HopfFunction n f = Iso.fun (fst (Hopfβ-Iso n f)) + ((subst (λ x → coHom x (HopfInvariantPush n (fst f))) + (λ i → suc (suc (suc (+-suc n n i))))) (⌣-α n f)) +open import Cubical.Algebra.AbGroup + +open import Cubical.Homotopy.Loopspace + +open import Cubical.HITs.Join + +open import Cubical.Homotopy.Hopf + +open import Cubical.HITs.SetQuotients renaming (_/_ to _/'_) +-- open import Cubical.ZCohomology.Gysin + +_·₀ₕ_ : ∀ {ℓ} {n : ℕ} {A : Type ℓ} → ℤ → coHom n A → coHom n A +_·₀ₕ_ {n = n} a b = a ℤ[ coHomGr n _ ]· b + -- remove HopfInvariant : (n : ℕ) → (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) @@ -324,79 +360,15 @@ HopfInvariant n f = (λ i → suc (suc (suc (+-suc n n (~ i))))) (Hopfβ n f)) - -⌣-α : (n : ℕ) → (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → _ -⌣-α n f = Hopfα n f ⌣ Hopfα n f - -HopfFunction : (n : ℕ) - → (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → ℤ -HopfFunction n f = Iso.fun (fst (Hopfβ-Iso n f)) - ((subst (λ x → coHom x (HopfInvariantPush n (fst f))) - (λ i → suc (suc (suc (+-suc n n i))))) (⌣-α n f)) - -pres- : (e : GroupHom ℤGroup ℤGroup) → (a : ℤ) → fst e (- a) ≡ - fst e a -pres- e a = +Comm (fst e (- a)) (pos 0) - ∙ cong (_+ (fst e (- a))) (sym (l (fst e a)) ∙ +Comm (fst e a) (- fst e a)) - ∙∙ sym (+Assoc _ _ _) - ∙∙ cong (- (fst e a) +_) (pp ∙ l (fst e a)) - where - l : (a : ℤ) → a + (- a) ≡ 0 - l (pos zero) = refl - l (pos (suc zero)) = refl - l (pos (suc (suc n))) = predℤ+negsuc n (pos (suc (suc n))) ∙ l (pos (suc n)) - l (negsuc zero) = refl - l (negsuc (suc n)) = (cong sucℤ (sucℤ+pos n (negsuc (suc n)))) ∙ l (negsuc n) - - pp : fst e a + fst e (- a) ≡ fst e a + (- fst e a) - pp = sym (IsGroupHom.pres· (snd e) a (- a)) - ∙∙ cong (fst e) (l a) - ∙∙ (IsGroupHom.pres1 (snd e) - ∙ sym (l _)) - -GroupHomPres· : (e : GroupHom ℤGroup ℤGroup) → (a b : ℤ) → fst e (a · b) ≡ a · fst e b -GroupHomPres· e (pos zero) b = IsGroupHom.pres1 (snd e) -GroupHomPres· e (pos (suc n)) b = - IsGroupHom.pres· (snd e) b (pos n · b) ∙ cong (fst e b +_) (GroupHomPres· e (pos n) b) -GroupHomPres· e (negsuc zero) b = pres- e b -GroupHomPres· e (negsuc (suc n)) b = - IsGroupHom.pres· (snd e) (- b) (negsuc n · b) - ∙ cong₂ _+_ (pres- e b) (GroupHomPres· e (negsuc n) b) - -lUnit·₀ₕ : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → (x : coHom n A) → (pos zero) ·₀ₕ x ≡ 0ₕ _ -lUnit·₀ₕ n = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl - -minus≡0- : (x : ℤ) → - x ≡ (0 - x) -minus≡0- x = sym (GroupStr.lid (snd ℤGroup) (- x)) - -GroupHomPres·₀ : ∀ {ℓ} {A : Type ℓ} (n : ℕ) +GroupHomℤ→ℤPres·₀ : ∀ {ℓ} {A : Type ℓ} (n : ℕ) (e : GroupHom ℤGroup (coHomGr n A)) → (a b : ℤ) → fst e (a · b) - ≡ (a ·₀ₕ fst e b) -GroupHomPres·₀ n e (pos zero) b = IsGroupHom.pres1 (snd e) ∙ sym (lUnit·₀ₕ n (fst e b)) -GroupHomPres·₀ {A = A} n e (pos (suc a)) b = - IsGroupHom.pres· (snd e) b (pos a · b) - ∙ (λ i → fst e b +ₕ GroupHomPres·₀ n e (pos a) b i) - ∙ s (fst e b) - where - s : (x : coHom n A) → x +ₕ (pos a ·₀ₕ x) ≡ pos (suc a) ·₀ₕ x - s = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl -GroupHomPres·₀ n e (negsuc zero) b = - (cong (fst e) (sym (GroupStr.lid (snd ℤGroup) (- b))) - ∙ IsGroupHom.presinv (snd e) b) -GroupHomPres·₀ {A = A} n e (negsuc (suc a)) b = - IsGroupHom.pres· (snd e) (- b) (negsuc a · b) - ∙∙ cong₂ _+ₕ_ ((cong (fst e) (sym (GroupStr.lid (snd ℤGroup) (- b))) - ∙ IsGroupHom.presinv (snd e) b)) - (GroupHomPres·₀ n e (negsuc a) b) - ∙∙ helper (fst e b) - where - helper : (x : coHom n A) → (-ₕ x) +ₕ (negsuc a ·₀ₕ x) ≡ (negsuc (suc a) ·₀ₕ x) - helper = - sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - λ f → cong ∣_∣₂ (funExt λ x → commₖ n (-ₖ (f x)) (negsuc a ·₀ f x)) + ≡ (a ·₀ₕ fst e b) +GroupHomℤ→ℤPres·₀ {A = A} n e a b = cong (fst e) (ℤ·≡· a b) ∙ hompres· {H = coHomGr n A} (_ , snd e) b a + +-- goal: remove genH : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → (e : GroupIso (coHomGr n A) ℤGroup) @@ -405,57 +377,72 @@ genH : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → Σ[ a ∈ ℤ ] a ·₀ₕ x ≡ y) genH {A = A} n e = (Iso.inv (fst e) 1) , λ y → (Iso.fun (fst e) y) - , (sym (GroupHomPres·₀ n (_ , snd (invGroupIso e)) (Iso.fun (fst e) y) (pos 1)) + , (sym (GroupHomℤ→ℤPres·₀ n (_ , snd (invGroupIso e)) (Iso.fun (fst e) y) (pos 1)) ∙∙ cong (Iso.inv (fst e)) (·Comm (Iso.fun (fst e) y) (pos 1)) ∙∙ Iso.leftInv (fst e) y) +-- keep +genH' : ∀ {ℓ} {A : Type ℓ} (n : ℕ) + → (e : GroupIso (coHomGr n A) + ℤGroup) + → Σ[ x ∈ coHom n A ] + gen₁-by (coHomGr n A) x +genH' {A = A} n e = (Iso.inv (fst e) 1) + , λ h → (Iso.fun (fst e) h) , + (sym (Iso.leftInv (fst e) h) + ∙∙ sym (cong (Iso.inv (fst e)) (·Comm (Iso.fun (fst e) h) (pos 1))) + ∙∙ (cong (Iso.inv (fst e)) (ℤ·≡· _ _) + ∙ (hompres· (_ , snd (invGroupIso e)) (pos 1) (Iso.fun (fst e) h)))) + +-- Move to intGroup +private + l2 : (n m : ℕ) → Σ[ a ∈ ℕ ] (negsuc n · pos (suc m)) ≡ negsuc a + l2 n zero = n , ·Comm (negsuc n) (pos 1) + l2 n (suc m) = h3 _ _ .fst , + (·Comm (negsuc n) (pos (suc (suc m))) + ∙∙ cong (negsuc n +_) (·Comm (pos (suc m)) (negsuc n) ∙ (l2 n m .snd)) + ∙∙ h3 _ _ .snd) + where + h3 : (x y : ℕ) → Σ[ a ∈ ℕ ] negsuc x + negsuc y ≡ negsuc a + h3 zero zero = 1 , refl + h3 zero (suc y) = (suc (suc y)) , +Comm (negsuc zero) (negsuc (suc y)) + h3 (suc x) zero = (suc (suc x)) , refl + h3 (suc x) (suc y) = (h3 (suc (suc x)) y .fst) , (predℤ+negsuc y (negsuc (suc x)) ∙ snd ((h3 (suc (suc x))) y)) + + zz : (n x : ℕ) → Σ[ a ∈ ℕ ] ((pos (suc x)) · pos (suc (suc n)) ≡ pos (suc (suc a))) + zz n zero = n , refl + zz n (suc x) = h3 _ _ (zz n x) + where + h3 : (x : ℤ) (n : ℕ) → Σ[ a ∈ ℕ ] (x ≡ pos (suc (suc a))) + → Σ[ a ∈ ℕ ] pos n + x ≡ pos (suc (suc a)) + h3 x n (a , p) = n +ℕ a + , cong (pos n +_) p ∙ cong sucℤ (sucℤ+pos a (pos n)) + ∙ sucℤ+pos a (pos (suc n)) ∙ (sym (pos+ (suc (suc n)) a)) -l2 : (n m : ℕ) → Σ[ a ∈ ℕ ] (negsuc n · pos (suc m)) ≡ negsuc a -l2 n zero = n , ·Comm (negsuc n) (pos 1) -l2 n (suc m) = h3 _ _ .fst , - (·Comm (negsuc n) (pos (suc (suc m))) - ∙∙ cong (negsuc n +_) (·Comm (pos (suc m)) (negsuc n) ∙ (l2 n m .snd)) - ∙∙ h3 _ _ .snd) - where - h3 : (x y : ℕ) → Σ[ a ∈ ℕ ] negsuc x + negsuc y ≡ negsuc a - h3 zero zero = 1 , refl - h3 zero (suc y) = (suc (suc y)) , +Comm (negsuc zero) (negsuc (suc y)) - h3 (suc x) zero = (suc (suc x)) , refl - h3 (suc x) (suc y) = (h3 (suc (suc x)) y .fst) , (predℤ+negsuc y (negsuc (suc x)) ∙ snd ((h3 (suc (suc x))) y)) - -zz : (n x : ℕ) → Σ[ a ∈ ℕ ] ((pos (suc x)) · pos (suc (suc n)) ≡ pos (suc (suc a))) -zz n zero = n , refl -zz n (suc x) = h3 _ _ (zz n x) - where - h3 : (x : ℤ) (n : ℕ) → Σ[ a ∈ ℕ ] (x ≡ pos (suc (suc a))) - → Σ[ a ∈ ℕ ] pos n + x ≡ pos (suc (suc a)) - h3 x n (a , p) = n +ℕ a - , cong (pos n +_) p ∙ cong sucℤ (sucℤ+pos a (pos n)) - ∙ sucℤ+pos a (pos (suc n)) ∙ (sym (pos+ (suc (suc n)) a)) - -lem22 : (n : ℕ) (x : ℤ) → ¬ pos 1 ≡ x · pos (suc (suc n)) -lem22 n (pos zero) p = snotz (injPos p) -lem22 n (pos (suc n₁)) p = snotz (injPos (sym (cong predℤ (snd (zz n n₁))) ∙ sym (cong predℤ p))) -lem22 n (negsuc n₁) p = posNotnegsuc _ _ (p ∙ l2 _ _ .snd) + lem22 : (n : ℕ) (x : ℤ) → ¬ pos 1 ≡ x · pos (suc (suc n)) + lem22 n (pos zero) p = snotz (injPos p) + lem22 n (pos (suc n₁)) p = snotz (injPos (sym (cong predℤ (snd (zz n n₁))) ∙ sym (cong predℤ p))) + lem22 n (negsuc n₁) p = posNotnegsuc _ _ (p ∙ l2 _ _ .snd) -grr : (e : GroupEquiv ℤGroup ℤGroup) (x : ℤ) → (fst (fst e) 1) ≡ x → abs (fst (fst e) 1) ≡ 1 -grr e (pos zero) p = +-- Move to intGroup +GroupEquivℤ-pres1 : (e : GroupEquiv ℤGroup ℤGroup) (x : ℤ) → (fst (fst e) 1) ≡ x → abs (fst (fst e) 1) ≡ 1 +GroupEquivℤ-pres1 e (pos zero) p = ⊥-rec (snotz (injPos (sym (retEq (fst e) 1) ∙∙ (cong (fst (fst (invGroupEquiv e))) p) ∙∙ IsGroupHom.pres1 (snd (invGroupEquiv e))))) -grr e (pos (suc zero)) p = cong abs p -grr e (pos (suc (suc n))) p = ⊥-rec (lem22 _ _ (h3 ∙ ·Comm (pos (suc (suc n))) (invEq (fst e) 1))) +GroupEquivℤ-pres1 e (pos (suc zero)) p = cong abs p +GroupEquivℤ-pres1 e (pos (suc (suc n))) p = ⊥-rec (lem22 _ _ (h3 ∙ ·Comm (pos (suc (suc n))) (invEq (fst e) 1))) where h3 : pos 1 ≡ _ h3 = sym (retEq (fst e) 1) ∙∙ (cong (fst (fst (invGroupEquiv e))) (p ∙ ·Comm 1 (pos (suc (suc n))))) - ∙∙ GroupHomPres· (_ , snd (invGroupEquiv e)) (pos (suc (suc n))) 1 -grr e (negsuc zero) p = cong abs p -grr e (negsuc (suc n)) p = ⊥-rec (lem22 _ _ l33) + ∙∙ GroupHomℤ→ℤPres· (_ , snd (invGroupEquiv e)) (pos (suc (suc n))) 1 +GroupEquivℤ-pres1 e (negsuc zero) p = cong abs p +GroupEquivℤ-pres1 e (negsuc (suc n)) p = ⊥-rec (lem22 _ _ l33) where h3 : fst (fst e) (negsuc zero) ≡ pos (suc (suc n)) - h3 = pres- (_ , snd e) (pos 1) ∙ cong -_ p + h3 = GroupHomℤ→ℤpres- (_ , snd e) (pos 1) ∙ cong -_ p compGroup : GroupEquiv ℤGroup ℤGroup fst compGroup = isoToEquiv (iso -_ -_ -Involutive -Involutive) @@ -471,9 +458,10 @@ grr e (negsuc (suc n)) p = ⊥-rec (lem22 _ _ l33) l33 = sym (retEq (fst compHom) (pos 1)) ∙∙ cong (invEq (fst compHom)) l32 ∙∙ (cong (invEq (fst compHom)) (·Comm (pos 1) (pos (suc (suc n)))) - ∙ GroupHomPres· (_ , (snd (invGroupEquiv compHom))) (pos (suc (suc n))) (pos 1) + ∙ GroupHomℤ→ℤPres· (_ , (snd (invGroupEquiv compHom))) (pos (suc (suc n))) (pos 1) ∙ ·Comm (pos (suc (suc n))) (invEq (fst compHom) (pos 1))) +-- move to Int abs→⊎ : (x : ℤ) (n : ℕ) → abs x ≡ n → (x ≡ pos n) ⊎ (x ≡ - pos n) abs→⊎ x n = J (λ n _ → (x ≡ pos n) ⊎ (x ≡ - pos n)) (help x) where @@ -487,6 +475,7 @@ abs→⊎ x n = J (λ n _ → (x ≡ pos n) ⊎ (x ≡ - pos n)) (help x) ⊎→abs x zero (inr x₁) = cong abs x₁ ⊎→abs x (suc n) (inr x₁) = cong abs x₁ +-- Move to group JGroupEquiv : ∀ {ℓ ℓ'} {G : Group ℓ} (P : (H : Group ℓ) → GroupEquiv G H → Type ℓ') → P G idGroupEquiv → ∀ {H} e → P H e @@ -502,6 +491,8 @@ JGroupEquiv {G = G} P p {H} e = (transportRefl (fst (fst e) (transportRefl x i)) i)))) ∙ retEq (fst e) x)) +-- really needed? +{- sesIsoPresGen : ∀ (G : Group ℓ-zero) → (iso : GroupEquiv ℤGroup G) @@ -525,20 +516,22 @@ sesIsoPresGen G = JGroupEquiv (λ G iso → where autoPres1 : (e : GroupEquiv ℤGroup ℤGroup) → abs (fst (fst e) 1) ≡ 1 - autoPres1 e = grr e _ refl - + autoPres1 e = GroupEquivℤ-pres1 e _ refl +-} +-- move to intGroup allisoPresGen : ∀ {ℓ} (G : Group ℓ) (ϕ : GroupEquiv G ℤGroup) (x : fst G) → (fst (fst ϕ) x ≡ 1) ⊎ (fst (fst ϕ) x ≡ - 1) - → (ψ : GroupEquiv G ℤGroup) → (fst (fst ψ) x ≡ 1) ⊎ (fst (fst ψ) x ≡ - 1) + → (ψ : GroupEquiv G ℤGroup) + → (fst (fst ψ) x ≡ 1) ⊎ (fst (fst ψ) x ≡ - 1) allisoPresGen G (ϕeq , ϕhom) x (inl r) (ψeq , ψhom) = abs→⊎ _ _ (cong abs (cong (fst ψeq) (sym (retEq ϕeq x) ∙ cong (invEq ϕeq) r)) - ∙ grr (compGroupEquiv (invGroupEquiv (ϕeq , ϕhom)) (ψeq , ψhom)) _ refl) + ∙ GroupEquivℤ-pres1 (compGroupEquiv (invGroupEquiv (ϕeq , ϕhom)) (ψeq , ψhom)) _ refl) allisoPresGen G (ϕeq , ϕhom) x (inr r) (ψeq , ψhom) = abs→⊎ _ _ (cong abs (cong (fst ψeq) (sym (retEq ϕeq x) ∙ cong (invEq ϕeq) r)) ∙ cong abs (IsGroupHom.presinv (snd (compGroupEquiv (invGroupEquiv (ϕeq , ϕhom)) (ψeq , ψhom))) 1) - ∙ absLem _ (grr (compGroupEquiv (invGroupEquiv (ϕeq , ϕhom)) (ψeq , ψhom)) _ refl)) + ∙ absLem _ (GroupEquivℤ-pres1 (compGroupEquiv (invGroupEquiv (ϕeq , ϕhom)) (ψeq , ψhom)) _ refl)) where absLem : ∀ x → abs x ≡ 1 → abs (pos 0 - x) ≡ 1 absLem (pos zero) p = ⊥-rec (snotz (sym p)) @@ -547,14 +540,13 @@ allisoPresGen G (ϕeq , ϕhom) x (inr r) (ψeq , ψhom) = absLem (negsuc zero) p = refl absLem (negsuc (suc n)) p = ⊥-rec (snotz (cong predℕ p)) - -l3 : ∀ {ℓ} {G H : Type ℓ} → (G ≡ H) → (g₁ : coHom 2 G) (h₁ : coHom 2 H) +⌣equiv→pres1 : ∀ {ℓ} {G H : Type ℓ} → (G ≡ H) → (g₁ : coHom 2 G) (h₁ : coHom 2 H) → (fstEq : (Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] abs (fst (fst ϕ) g₁) ≡ 1)) → (sndEq : ((Σ[ ϕ ∈ GroupEquiv (coHomGr 2 H) ℤGroup ] abs (fst (fst ϕ) h₁) ≡ 1))) → isEquiv {A = coHom 2 G} {B = coHom 4 G} (_⌣ g₁) → (3rdEq : GroupEquiv (coHomGr 4 H) ℤGroup) → abs (fst (fst 3rdEq) (h₁ ⌣ h₁)) ≡ 1 -l3 {G = G} = J (λ H _ → (g₁ : coHom 2 G) (h₁ : coHom 2 H) +⌣equiv→pres1 {G = G} = J (λ H _ → (g₁ : coHom 2 G) (h₁ : coHom 2 H) → (fstEq : (Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] abs (fst (fst ϕ) g₁) ≡ 1)) → (sndEq : ((Σ[ ϕ ∈ GroupEquiv (coHomGr 2 H) ℤGroup ] abs (fst (fst ϕ) h₁) ≡ 1))) → isEquiv {A = coHom 2 G} {B = coHom 4 G} (_⌣ g₁) @@ -610,7 +602,7 @@ l3 {G = G} = J (λ H _ → (g₁ : coHom 2 G) (h₁ : coHom 2 H) sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ f → cong ∣_∣₂ (funExt λ x → -ₖ^2 (f x)) - theEq : coHom 2 G ≃ coHom 4 G -- (λ x → -ₕ (x ⌣ g)) + theEq : coHom 2 G ≃ coHom 4 G theEq = compEquiv (_ , ⌣eq) (isoToEquiv (-ₕeq 4)) main3 : (h ≡ g) ⊎ (h ≡ (-ₕ g)) → isEquiv {A = coHom 2 G} (_⌣ h) @@ -624,6 +616,7 @@ l3 {G = G} = J (λ H _ → (g₁ : coHom 2 G) (h₁ : coHom 2 H) snd main4 = makeIsGroupHom λ g1 g2 → rightDistr-⌣ _ _ g1 g2 h +-- move to S1 (not there already -- check Sn?) characFunSpaceS¹ : ∀ {ℓ} {A : Type ℓ} → Iso (S₊ 1 → A) (Σ[ x ∈ A ] x ≡ x) Iso.fun characFunSpaceS¹ f = f base , cong f loop @@ -653,19 +646,6 @@ fib n x = (λ f → ∥ (Σ[ g ∈ (S₊ 3 → coHomK (3 +ℕ n)) ] ((a : S₊ 1) (b : S₊ 3) → f a ≡ g b)) ∥₂ , squash₂) x .fst -contrFstΣ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} - → (e : isContr A) - → Iso (Σ A B) (B (fst e)) -Iso.fun (contrFstΣ {B = B} e) (a , b) = subst B (sym (snd e a)) b -Iso.inv (contrFstΣ {B = B} e) b = (fst e) , b -Iso.rightInv (contrFstΣ {B = B} e) b = cong (λ x → subst B x b) h3 ∙ transportRefl b - where - h3 : sym (snd e (fst e)) ≡ refl - h3 = isProp→isSet (isContr→isProp e) _ _ _ _ -Iso.leftInv (contrFstΣ {B = B} e) b = - ΣPathP ((snd e (fst b)) - , (λ j → transp (λ i → B (snd e (fst b) (~ i ∨ j))) j (snd b))) - Iso1 : (n : ℕ) → Iso (coHom (3 +ℕ n) (join S¹ (S₊ 3))) ∥ Σ (coHomS¹-ish n) (fib n) ∥₂ Iso1 n = compIso (setTruncIso characFunSpace) Iso2 where @@ -1872,8 +1852,8 @@ snd e' = makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ f g → refl) -grrg : GroupIso (coHomGr 2 CP2) ℤGroup -grrg = compGroupIso e' (Hⁿ-Sⁿ≅ℤ 1) +GroupEquivℤ-pres1g : GroupIso (coHomGr 2 CP2) ℤGroup +GroupEquivℤ-pres1g = compGroupIso e' (Hⁿ-Sⁿ≅ℤ 1) ⌣hom : GroupEquiv (coHomGr 2 CP2) (coHomGr 4 CP2) fst (fst ⌣hom) = GysinS1.⌣-hom 2 .fst @@ -2059,7 +2039,7 @@ ll = ≡CP2 _ _ ∣ funExt (λ x → funExt⁻ (cong fst (ll32 x)) south) ∣ isGenerator≃ℤ-e : isGenerator≃ℤ (coHomGr 2 CP2) GysinS1.e isGenerator≃ℤ-e = subst (isGenerator≃ℤ (coHomGr 2 CP2)) (sym ll) - (grrg , refl) + (GroupEquivℤ-pres1g , refl) hopfFun : S₊ 3 → S₊ 2 hopfFun x = fibS1.TotalSpaceHopf'→TotalSpace @@ -2083,15 +2063,15 @@ fst (snd ⌣-pres1-CP2) = isGenerator≃ℤ-e snd (snd ⌣-pres1-CP2) = ∣ p ∣ where p : Σ _ _ - fst p = compGroupIso (GroupEquiv→GroupIso (invGroupEquiv ⌣hom)) grrg + fst p = compGroupIso (GroupEquiv→GroupIso (invGroupEquiv ⌣hom)) GroupEquivℤ-pres1g snd p = {!!} - ∙ sesIsoPresGen _ (invGroupEquiv (GroupIso→GroupEquiv grrg)) _ - (compGroupEquiv (invGroupEquiv (GroupIso→GroupEquiv grrg)) (invGroupEquiv (invGroupEquiv ⌣hom))) + ∙ sesIsoPresGen _ (invGroupEquiv (GroupIso→GroupEquiv GroupEquivℤ-pres1g)) _ + (compGroupEquiv (invGroupEquiv (GroupIso→GroupEquiv GroupEquivℤ-pres1g)) (invGroupEquiv (invGroupEquiv ⌣hom))) GysinS1.e {!!} ⌣hom {- GysinS1.e , (isGenerator≃ℤ-e - , ∣ compGroupIso (GroupEquiv→GroupIso (invGroupEquiv ⌣hom)) grrg + , ∣ compGroupIso (GroupEquiv→GroupIso (invGroupEquiv ⌣hom)) GroupEquivℤ-pres1g , {!!} ∣) -} @@ -2152,108 +2132,6 @@ snd (snd ⌣-pres1-CP2) = ∣ p ∣ -- -- P (inr x) = Hopf x -- -- P (push a i₁) = ua (v a) i₁ - - -open import Cubical.Algebra.AbGroup - -open import Cubical.Homotopy.Loopspace - -open import Cubical.HITs.Join - -open import Cubical.Homotopy.Hopf - -open import Cubical.HITs.SetQuotients renaming (_/_ to _/'_) --- open import Cubical.ZCohomology.Gysin - -rUnitℤ· : ∀ {ℓ} (G : Group ℓ) → (x : ℤ) → (x ℤ[ G ]· GroupStr.1g (snd G)) - ≡ GroupStr.1g (snd G) -rUnitℤ· G (pos zero) = refl -rUnitℤ· G (pos (suc n)) = - cong (GroupStr._·_ (snd G) (GroupStr.1g (snd G))) - (rUnitℤ· G (pos n)) - ∙ GroupStr.lid (snd G) (GroupStr.1g (snd G)) -rUnitℤ· G (negsuc zero) = GroupTheory.inv1g G -rUnitℤ· G (negsuc (suc n)) = - cong₂ (GroupStr._·_ (snd G)) (GroupTheory.inv1g G) (rUnitℤ· G (negsuc n)) - ∙ GroupStr.lid (snd G) (GroupStr.1g (snd G)) - -rUnitℤ·ℤ : (x : ℤ) → (x ℤ[ ℤGroup ]· 1) ≡ x -rUnitℤ·ℤ (pos zero) = refl -rUnitℤ·ℤ (pos (suc n)) = cong (pos 1 +_) (rUnitℤ·ℤ (pos n)) ∙ sym (pos+ 1 n) -rUnitℤ·ℤ (negsuc zero) = refl -rUnitℤ·ℤ (negsuc (suc n)) = cong (-1 +_) (rUnitℤ·ℤ (negsuc n)) ∙ +Comm (negsuc 0) (negsuc n) - -private - precommℤ : ∀ {ℓ} (G : Group ℓ) → (g : fst G) (y : ℤ) → (snd G GroupStr.· (y ℤ[ G ]· g)) g - ≡ (sucℤ y ℤ[ G ]· g) - precommℤ G g (pos zero) = GroupStr.lid (snd G) g - ∙ sym (GroupStr.rid (snd G) g) - precommℤ G g (pos (suc n)) = - sym (GroupStr.assoc (snd G) _ _ _) - ∙ cong ((snd G GroupStr.· g)) (precommℤ G g (pos n)) - precommℤ G g (negsuc zero) = - GroupStr.invl (snd G) g - precommℤ G g (negsuc (suc n)) = - sym (GroupStr.assoc (snd G) _ _ _) - ∙ cong ((snd G GroupStr.· GroupStr.inv (snd G) g)) (precommℤ G g (negsuc n)) - ∙ negsucLem n - where - negsucLem : (n : ℕ) → (snd G GroupStr.· GroupStr.inv (snd G) g) - (sucℤ (negsuc n) ℤ[ G ]· g) - ≡ (sucℤ (negsuc (suc n)) ℤ[ G ]· g) - negsucLem zero = (GroupStr.rid (snd G) _) - negsucLem (suc n) = refl - -distrℤ· : ∀ {ℓ} (G : Group ℓ) → (g : fst G) (x y : ℤ) - → ((x + y) ℤ[ G ]· g) - ≡ GroupStr._·_ (snd G) (x ℤ[ G ]· g) (y ℤ[ G ]· g) -distrℤ· G' g (pos zero) y = cong (_ℤ[ G' ]· g) (+Comm 0 y) - ∙ sym (GroupStr.lid (snd G') _) -distrℤ· G' g (pos (suc n)) (pos n₁) = - cong (_ℤ[ G' ]· g) (sym (pos+ (suc n) n₁)) - ∙ cong (GroupStr._·_ (snd G') g) (cong (_ℤ[ G' ]· g) (pos+ n n₁) ∙ distrℤ· G' g (pos n) (pos n₁)) - ∙ GroupStr.assoc (snd G') _ _ _ -distrℤ· G' g (pos (suc n)) (negsuc zero) = - distrℤ· G' g (pos n) 0 - ∙ ((GroupStr.rid (snd G') (pos n ℤ[ G' ]· g) ∙ sym (GroupStr.lid (snd G') (pos n ℤ[ G' ]· g))) - ∙ cong (λ x → GroupStr._·_ (snd G') x (pos n ℤ[ G' ]· g)) - (sym (GroupStr.invl (snd G') g)) ∙ sym (GroupStr.assoc (snd G') _ _ _)) - ∙ (GroupStr.assoc (snd G') _ _ _) - ∙ cong (λ x → GroupStr._·_ (snd G') x (pos n ℤ[ G' ]· g)) (GroupStr.invl (snd G') g) - ∙ GroupStr.lid (snd G') _ - ∙ sym (GroupStr.rid (snd G') _) - ∙ (cong (GroupStr._·_ (snd G') (pos n ℤ[ G' ]· g)) (sym (GroupStr.invr (snd G') g)) - ∙ GroupStr.assoc (snd G') _ _ _) - ∙ cong (λ x → GroupStr._·_ (snd G') x (GroupStr.inv (snd G') g)) - (precommℤ G' g (pos n)) -distrℤ· G' g (pos (suc n)) (negsuc (suc n₁)) = - cong (_ℤ[ G' ]· g) (predℤ+negsuc n₁ (pos (suc n))) - ∙∙ distrℤ· G' g (pos n) (negsuc n₁) - ∙∙ (cong (λ x → GroupStr._·_ (snd G') x (negsuc n₁ ℤ[ G' ]· g)) - ((sym (GroupStr.rid (snd G') (pos n ℤ[ G' ]· g)) - ∙ cong (GroupStr._·_ (snd G') (pos n ℤ[ G' ]· g)) (sym (GroupStr.invr (snd G') g))) - ∙∙ GroupStr.assoc (snd G') _ _ _ - ∙∙ cong (λ x → GroupStr._·_ (snd G') x (GroupStr.inv (snd G') g)) (precommℤ G' g (pos n) )) - ∙ sym (GroupStr.assoc (snd G') _ _ _)) -distrℤ· G' g (negsuc zero) y = - cong (_ℤ[ G' ]· g) (+Comm -1 y) ∙ lol1 y - module _ where - lol1 : (y : ℤ) → (predℤ y ℤ[ G' ]· g) ≡ (snd G' GroupStr.· GroupStr.inv (snd G') g) (y ℤ[ G' ]· g) - lol1 (pos zero) = sym (GroupStr.rid (snd G') _) - lol1 (pos (suc n)) = - sym (GroupStr.lid (snd G') ((pos n ℤ[ G' ]· g))) - ∙∙ cong (λ x → GroupStr._·_ (snd G') x (pos n ℤ[ G' ]· g)) (sym (GroupStr.invl (snd G') g)) - ∙∙ sym (GroupStr.assoc (snd G') _ _ _) - lol1 (negsuc n) = refl -distrℤ· G' g (negsuc (suc n)) y = - cong (_ℤ[ G' ]· g) (+Comm (negsuc (suc n)) y) - ∙∙ lol1 G' g 0 (y +negsuc n) - ∙∙ cong (snd G' GroupStr.· GroupStr.inv (snd G') g) - (cong (_ℤ[ G' ]· g) (+Comm y (negsuc n)) ∙ distrℤ· G' g (negsuc n) y) - ∙ (GroupStr.assoc (snd G') _ _ _) - - - HopfInvariantPushElim : ∀ {ℓ} n → (f : _) → {P : HopfInvariantPush n f → Type ℓ} → (isOfHLevel (suc (suc (suc (suc (n +ℕ n))))) (P (inl tt))) → (e : P (inl tt)) @@ -2289,7 +2167,7 @@ snd (transportCohomIso {A = A} {n = n} {m = m} p) = α-hopfMap : abs (HopfFunction zero (hopfFun , λ _ → hopfFun (snd (S₊∙ 3)))) ≡ suc zero α-hopfMap = cong abs (cong (Iso.fun (fst (Hopfβ-Iso zero (hopfFun , refl)))) - (transportRefl (⌣-α 0 (hopfFun , refl)))) ∙ l3 (sym CP2≡CP2') GysinS1.e (Hopfα zero (hopfFun , refl)) + (transportRefl (⌣-α 0 (hopfFun , refl)))) ∙ ⌣equiv→pres1 (sym CP2≡CP2') GysinS1.e (Hopfα zero (hopfFun , refl)) (l isGenerator≃ℤ-e) (GroupIso→GroupEquiv (Hopfα-Iso 0 (hopfFun , refl)) , refl) (snd (fst ⌣hom)) From 0ca5eda765c4ce4146d24cc7688e9a1360b141bf Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 4 Oct 2021 12:04:34 +0200 Subject: [PATCH 69/89] stuff --- Cubical/Algebra/Group/GroupPath.agda | 183 +------------ Cubical/Algebra/Group/ZModule.agda | 305 +++++++++++++++++++++ Cubical/Data/Int/Properties.agda | 14 + Cubical/Data/Sigma/Properties.agda | 19 +- Cubical/ZCohomology/Gysin.agda | 396 ++++++++------------------- 5 files changed, 463 insertions(+), 454 deletions(-) create mode 100644 Cubical/Algebra/Group/ZModule.agda diff --git a/Cubical/Algebra/Group/GroupPath.agda b/Cubical/Algebra/Group/GroupPath.agda index 32c53d72c..2ba25af6f 100644 --- a/Cubical/Algebra/Group/GroupPath.agda +++ b/Cubical/Algebra/Group/GroupPath.agda @@ -128,171 +128,18 @@ uaCompGroupEquiv f g = caracGroup≡ _ _ ( ≡⟨ sym (cong-∙ ⟨_⟩ (uaGroup f) (uaGroup g)) ⟩ cong ⟨_⟩ (uaGroup f ∙ uaGroup g) ∎) -open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) -open import Cubical.Data.Int renaming (_·_ to _*_ ; _+_ to _+ℤ_ ; _-_ to _-ℤ_) -open import Cubical.Data.Nat -open import Cubical.Data.List - -_ℤ[_]·_ : ∀ {ℓ} → ℤ → (G : Group ℓ) → fst G → fst G -pos zero ℤ[ G' ]· g = GroupStr.1g (snd G') -pos (suc n) ℤ[ G' ]· g = GroupStr._·_ (snd G') g (pos n ℤ[ G' ]· g) -negsuc zero ℤ[ G' ]· g = GroupStr.inv (snd G') g -negsuc (suc n) ℤ[ G' ]· g = - GroupStr._·_ (snd G') (GroupStr.inv (snd G') g) (negsuc n ℤ[ G' ]· g) - -gen₁-by : ∀ {ℓ} (G : Group ℓ) - → (g : fst G) - → Type _ -gen₁-by G g = (h : fst G) - → Σ[ a ∈ ℤ ] h ≡ (a ℤ[ G ]· g) - -gen₂-by : ∀ {ℓ} (G : Group ℓ) - → (g₁ g₂ : fst G) - → Type _ -gen₂-by G g₁ g₂ = - (h : fst G) - → Σ[ a ∈ ℤ × ℤ ] h - ≡ GroupStr._·_ (snd G) - ((fst a) ℤ[ G ]· g₁) - ((snd a) ℤ[ G ]· g₂) - -hompres· : ∀ {ℓ ℓ'} {G : Group ℓ} {H : Group ℓ'} - → (ϕ : GroupHom G H) - → (x : fst G) (z : ℤ) - → fst ϕ (z ℤ[ G ]· x) ≡ (z ℤ[ H ]· fst ϕ x) -hompres· (ϕ , ϕhom) x (pos zero) = IsGroupHom.pres1 ϕhom -hompres· {H = H} (ϕ , ϕhom) x (pos (suc n)) = - IsGroupHom.pres· ϕhom _ _ - ∙ cong (GroupStr._·_ (snd H) (ϕ x)) (hompres· (ϕ , ϕhom) x (pos n)) -hompres· (ϕ , ϕhom) x (negsuc zero) = IsGroupHom.presinv ϕhom _ -hompres· {H = H} (ϕ , ϕhom) x (negsuc (suc n)) = - IsGroupHom.pres· ϕhom _ _ - ∙ cong₂ (GroupStr._·_ (snd H)) (IsGroupHom.presinv ϕhom x) - ((hompres· (ϕ , ϕhom) x (negsuc n))) - -Iso-pres-gen₁ : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') - → (g : fst G) - → gen₁-by G g - → (e : GroupIso G H) - → gen₁-by H (Iso.fun (fst e) g) -Iso-pres-gen₁ G H g genG is h = - (fst (genG (Iso.inv (fst is) h))) - , (sym (Iso.rightInv (fst is) h) - ∙∙ cong (Iso.fun (fst is)) (snd (genG (Iso.inv (fst is) h))) - ∙∙ (hompres· (_ , snd is) g (fst (genG (Iso.inv (fst is) h))))) - -Iso-pres-gen₂ : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') - → (g₁ g₂ : fst G) - → gen₂-by G g₁ g₂ - → (e : GroupIso G H) - → gen₂-by H (Iso.fun (fst e) g₁) (Iso.fun (fst e) g₂) -fst (Iso-pres-gen₂ G H g₁ g₂ genG is h) = genG (Iso.inv (fst is) h) .fst -snd (Iso-pres-gen₂ G H g₁ g₂ genG is h) = - sym (Iso.rightInv (fst is) h) - ∙∙ cong (fun (fst is)) (snd (genG (Iso.inv (fst is) h))) - ∙∙ (IsGroupHom.pres· (snd is) _ _ - ∙ cong₂ (GroupStr._·_ (snd H)) - (hompres· (_ , snd is) g₁ (fst (fst (genG (inv (fst is) h))))) - (hompres· (_ , snd is) g₂ (snd (fst (genG (inv (fst is) h)))))) - -rUnitℤ· : ∀ {ℓ} (G : Group ℓ) → (x : ℤ) → (x ℤ[ G ]· GroupStr.1g (snd G)) - ≡ GroupStr.1g (snd G) -rUnitℤ· G (pos zero) = refl -rUnitℤ· G (pos (suc n)) = - cong (GroupStr._·_ (snd G) (GroupStr.1g (snd G))) - (rUnitℤ· G (pos n)) - ∙ GroupStr.lid (snd G) (GroupStr.1g (snd G)) -rUnitℤ· G (negsuc zero) = GroupTheory.inv1g G -rUnitℤ· G (negsuc (suc n)) = - cong₂ (GroupStr._·_ (snd G)) (GroupTheory.inv1g G) (rUnitℤ· G (negsuc n)) - ∙ GroupStr.lid (snd G) (GroupStr.1g (snd G)) - -rUnitℤ·ℤ : (x : ℤ) → (x ℤ[ ℤGroup ]· 1) ≡ x -rUnitℤ·ℤ (pos zero) = refl -rUnitℤ·ℤ (pos (suc n)) = cong (pos 1 +ℤ_) (rUnitℤ·ℤ (pos n)) ∙ sym (pos+ 1 n) -rUnitℤ·ℤ (negsuc zero) = refl -rUnitℤ·ℤ (negsuc (suc n)) = cong (-1 +ℤ_) (rUnitℤ·ℤ (negsuc n)) ∙ +Comm (negsuc 0) (negsuc n) - -private - precommℤ : ∀ {ℓ} (G : Group ℓ) → (g : fst G) (y : ℤ) → (snd G GroupStr.· (y ℤ[ G ]· g)) g - ≡ (sucℤ y ℤ[ G ]· g) - precommℤ G g (pos zero) = GroupStr.lid (snd G) g - ∙ sym (GroupStr.rid (snd G) g) - precommℤ G g (pos (suc n)) = - sym (GroupStr.assoc (snd G) _ _ _) - ∙ cong ((snd G GroupStr.· g)) (precommℤ G g (pos n)) - precommℤ G g (negsuc zero) = - GroupStr.invl (snd G) g - precommℤ G g (negsuc (suc n)) = - sym (GroupStr.assoc (snd G) _ _ _) - ∙ cong ((snd G GroupStr.· GroupStr.inv (snd G) g)) (precommℤ G g (negsuc n)) - ∙ negsucLem n - where - negsucLem : (n : ℕ) → (snd G GroupStr.· GroupStr.inv (snd G) g) - (sucℤ (negsuc n) ℤ[ G ]· g) - ≡ (sucℤ (negsuc (suc n)) ℤ[ G ]· g) - negsucLem zero = (GroupStr.rid (snd G) _) - negsucLem (suc n) = refl - -distrℤ· : ∀ {ℓ} (G : Group ℓ) → (g : fst G) (x y : ℤ) - → ((x +ℤ y) ℤ[ G ]· g) - ≡ GroupStr._·_ (snd G) (x ℤ[ G ]· g) (y ℤ[ G ]· g) -distrℤ· G' g (pos zero) y = cong (_ℤ[ G' ]· g) (+Comm 0 y) - ∙ sym (GroupStr.lid (snd G') _) -distrℤ· G' g (pos (suc n)) (pos n₁) = - cong (_ℤ[ G' ]· g) (sym (pos+ (suc n) n₁)) - ∙ cong (GroupStr._·_ (snd G') g) (cong (_ℤ[ G' ]· g) (pos+ n n₁) ∙ distrℤ· G' g (pos n) (pos n₁)) - ∙ GroupStr.assoc (snd G') _ _ _ -distrℤ· G' g (pos (suc n)) (negsuc zero) = - distrℤ· G' g (pos n) 0 - ∙ ((GroupStr.rid (snd G') (pos n ℤ[ G' ]· g) ∙ sym (GroupStr.lid (snd G') (pos n ℤ[ G' ]· g))) - ∙ cong (λ x → GroupStr._·_ (snd G') x (pos n ℤ[ G' ]· g)) - (sym (GroupStr.invl (snd G') g)) ∙ sym (GroupStr.assoc (snd G') _ _ _)) - ∙ (GroupStr.assoc (snd G') _ _ _) - ∙ cong (λ x → GroupStr._·_ (snd G') x (pos n ℤ[ G' ]· g)) (GroupStr.invl (snd G') g) - ∙ GroupStr.lid (snd G') _ - ∙ sym (GroupStr.rid (snd G') _) - ∙ (cong (GroupStr._·_ (snd G') (pos n ℤ[ G' ]· g)) (sym (GroupStr.invr (snd G') g)) - ∙ GroupStr.assoc (snd G') _ _ _) - ∙ cong (λ x → GroupStr._·_ (snd G') x (GroupStr.inv (snd G') g)) - (precommℤ G' g (pos n)) -distrℤ· G' g (pos (suc n)) (negsuc (suc n₁)) = - cong (_ℤ[ G' ]· g) (predℤ+negsuc n₁ (pos (suc n))) - ∙∙ distrℤ· G' g (pos n) (negsuc n₁) - ∙∙ (cong (λ x → GroupStr._·_ (snd G') x (negsuc n₁ ℤ[ G' ]· g)) - ((sym (GroupStr.rid (snd G') (pos n ℤ[ G' ]· g)) - ∙ cong (GroupStr._·_ (snd G') (pos n ℤ[ G' ]· g)) (sym (GroupStr.invr (snd G') g))) - ∙∙ GroupStr.assoc (snd G') _ _ _ - ∙∙ cong (λ x → GroupStr._·_ (snd G') x (GroupStr.inv (snd G') g)) (precommℤ G' g (pos n) )) - ∙ sym (GroupStr.assoc (snd G') _ _ _)) -distrℤ· G' g (negsuc zero) y = - cong (_ℤ[ G' ]· g) (+Comm -1 y) ∙ lol1 y - module _ where - lol1 : (y : ℤ) → (predℤ y ℤ[ G' ]· g) ≡ (snd G' GroupStr.· GroupStr.inv (snd G') g) (y ℤ[ G' ]· g) - lol1 (pos zero) = sym (GroupStr.rid (snd G') _) - lol1 (pos (suc n)) = - sym (GroupStr.lid (snd G') ((pos n ℤ[ G' ]· g))) - ∙∙ cong (λ x → GroupStr._·_ (snd G') x (pos n ℤ[ G' ]· g)) (sym (GroupStr.invl (snd G') g)) - ∙∙ sym (GroupStr.assoc (snd G') _ _ _) - lol1 (negsuc n) = refl -distrℤ· G' g (negsuc (suc n)) y = - cong (_ℤ[ G' ]· g) (+Comm (negsuc (suc n)) y) - ∙∙ lol1 G' g 0 (y +negsuc n) - ∙∙ cong (snd G' GroupStr.· GroupStr.inv (snd G') g) - (cong (_ℤ[ G' ]· g) (+Comm y (negsuc n)) ∙ distrℤ· G' g (negsuc n) y) - ∙ (GroupStr.assoc (snd G') _ _ _) - -minus≡0- : (x : ℤ) → - x ≡ (0 -ℤ x) -minus≡0- x = sym (GroupStr.lid (snd ℤGroup) (- x)) - -GroupHomℤ→ℤpres- : (e : GroupHom ℤGroup ℤGroup) → (a : ℤ) → fst e (- a) ≡ - fst e a -GroupHomℤ→ℤpres- e a = cong (fst e) (minus≡0- a) ∙∙ IsGroupHom.presinv (snd e) a ∙∙ sym (minus≡0- _) - -ℤ·≡· : (a b : ℤ) → a * b ≡ (a ℤ[ ℤGroup ]· b) -ℤ·≡· (pos zero) b = refl -ℤ·≡· (pos (suc n)) b = cong (b +ℤ_) (ℤ·≡· (pos n) b) -ℤ·≡· (negsuc zero) b = minus≡0- b -ℤ·≡· (negsuc (suc n)) b = cong₂ _+ℤ_ (minus≡0- b) (ℤ·≡· (negsuc n) b) - -GroupHomℤ→ℤPres· : (e : GroupHom ℤGroup ℤGroup) → (a b : ℤ) → fst e (a * b) ≡ a * fst e b -GroupHomℤ→ℤPres· e a b = cong (fst e) (ℤ·≡· a b) ∙∙ hompres· e b a ∙∙ sym (ℤ·≡· a (fst e b)) +-- J-rule for GroupEquivs +JGroupEquiv : {G : Group ℓ} (P : (H : Group ℓ) → GroupEquiv G H → Type ℓ') + → P G idGroupEquiv + → ∀ {H} e → P H e +JGroupEquiv {G = G} P p {H} e = + transport (λ i → P (GroupPath G H .fst e i) + (transp (λ j → GroupEquiv G (GroupPath G H .fst e (i ∨ ~ j))) i e)) + (subst (P G) (sym l) p) + where + l : (transp (λ j → GroupEquiv G (GroupPath G H .fst e (~ j))) i0 e) ≡ idGroupEquiv + l = Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (Σ≡Prop (λ _ → isPropIsEquiv _) + (funExt λ x → (λ i → fst (fst (fst e .snd .equiv-proof + (transportRefl (fst (fst e) (transportRefl x i)) i)))) + ∙ retEq (fst e) x)) diff --git a/Cubical/Algebra/Group/ZModule.agda b/Cubical/Algebra/Group/ZModule.agda new file mode 100644 index 000000000..a0e6aa647 --- /dev/null +++ b/Cubical/Algebra/Group/ZModule.agda @@ -0,0 +1,305 @@ +-- Left ℤ-multiplication on groups and some of its properties +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Algebra.Group.ZModule where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Sigma +open import Cubical.Data.Int renaming (_·_ to _*_ ; _+_ to _+ℤ_ ; _-_ to _-ℤ_) +open import Cubical.Data.Nat hiding (_·_) renaming (_+_ to _+ℕ_) +open import Cubical.Data.Empty renaming (rec to ⊥-rec) +open import Cubical.Data.Sum + +open import Cubical.Displayed.Base +open import Cubical.Displayed.Auto +open import Cubical.Displayed.Record +open import Cubical.Displayed.Universe + +open import Cubical.Algebra.Group.Base +open import Cubical.Algebra.Group.Properties +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) + +open import Cubical.Relation.Nullary + +private + variable + ℓ ℓ' : Level + +open Iso +open GroupStr +open IsGroupHom + +private + minus≡0- : (x : ℤ) → - x ≡ (0 -ℤ x) + minus≡0- x = sym (lid (snd ℤGroup) (- x)) + +_ℤ[_]·_ : ℤ → (G : Group ℓ) → fst G → fst G +pos zero ℤ[ G' ]· g = 1g (snd G') +pos (suc n) ℤ[ G' ]· g = _·_ (snd G') g (pos n ℤ[ G' ]· g) +negsuc zero ℤ[ G' ]· g = inv (snd G') g +negsuc (suc n) ℤ[ G' ]· g = + _·_ (snd G') (inv (snd G') g) (negsuc n ℤ[ G' ]· g) + +homPresℤ· : {G : Group ℓ} {H : Group ℓ'} + → (ϕ : GroupHom G H) + → (x : fst G) (z : ℤ) + → fst ϕ (z ℤ[ G ]· x) ≡ (z ℤ[ H ]· fst ϕ x) +homPresℤ· (ϕ , ϕhom) x (pos zero) = pres1 ϕhom +homPresℤ· {H = H} (ϕ , ϕhom) x (pos (suc n)) = + pres· ϕhom _ _ + ∙ cong (_·_ (snd H) (ϕ x)) (homPresℤ· (ϕ , ϕhom) x (pos n)) +homPresℤ· (ϕ , ϕhom) x (negsuc zero) = presinv ϕhom _ +homPresℤ· {H = H} (ϕ , ϕhom) x (negsuc (suc n)) = + pres· ϕhom _ _ + ∙ cong₂ (_·_ (snd H)) (presinv ϕhom x) + ((homPresℤ· (ϕ , ϕhom) x (negsuc n))) + +rUnitℤ· : (G : Group ℓ) (x : ℤ) → (x ℤ[ G ]· 1g (snd G)) ≡ 1g (snd G) +rUnitℤ· G (pos zero) = refl +rUnitℤ· G (pos (suc n)) = + cong (_·_ (snd G) (1g (snd G))) + (rUnitℤ· G (pos n)) + ∙ lid (snd G) (1g (snd G)) +rUnitℤ· G (negsuc zero) = GroupTheory.inv1g G +rUnitℤ· G (negsuc (suc n)) = + cong₂ (_·_ (snd G)) (GroupTheory.inv1g G) (rUnitℤ· G (negsuc n)) + ∙ lid (snd G) (1g (snd G)) + +rUnitℤ·ℤ : (x : ℤ) → (x ℤ[ ℤGroup ]· 1) ≡ x +rUnitℤ·ℤ (pos zero) = refl +rUnitℤ·ℤ (pos (suc n)) = cong (pos 1 +ℤ_) (rUnitℤ·ℤ (pos n)) ∙ sym (pos+ 1 n) +rUnitℤ·ℤ (negsuc zero) = refl +rUnitℤ·ℤ (negsuc (suc n)) = cong (-1 +ℤ_) (rUnitℤ·ℤ (negsuc n)) + ∙ +Comm (negsuc 0) (negsuc n) + +private + precommℤ : (G : Group ℓ) (g : fst G) (y : ℤ) + → (snd G · (y ℤ[ G ]· g)) g ≡ (sucℤ y ℤ[ G ]· g) + precommℤ G g (pos zero) = lid (snd G) g ∙ sym (rid (snd G) g) + precommℤ G g (pos (suc n)) = + sym (assoc (snd G) _ _ _) + ∙ cong ((snd G · g)) (precommℤ G g (pos n)) + precommℤ G g (negsuc zero) = invl (snd G) g + precommℤ G g (negsuc (suc n)) = + sym (assoc (snd G) _ _ _) + ∙ cong ((snd G · inv (snd G) g)) (precommℤ G g (negsuc n)) + ∙ negsucLem n + where + negsucLem : (n : ℕ) → (snd G · inv (snd G) g) + (sucℤ (negsuc n) ℤ[ G ]· g) + ≡ (sucℤ (negsuc (suc n)) ℤ[ G ]· g) + negsucLem zero = (rid (snd G) _) + negsucLem (suc n) = refl + +module _ (G : Group ℓ) (g : fst G) where + private + lem : (y : ℤ) → (predℤ y ℤ[ G ]· g) ≡ (snd G · inv (snd G) g) (y ℤ[ G ]· g) + lem (pos zero) = sym (rid (snd G) _) + lem (pos (suc n)) = + sym (lid (snd G) ((pos n ℤ[ G ]· g))) + ∙∙ cong (λ x → _·_ (snd G) x (pos n ℤ[ G ]· g)) (sym (invl (snd G) g)) + ∙∙ sym (assoc (snd G) _ _ _) + lem (negsuc n) = refl + + distrℤ· : (x y : ℤ) → ((x +ℤ y) ℤ[ G ]· g) + ≡ _·_ (snd G) (x ℤ[ G ]· g) (y ℤ[ G ]· g) + distrℤ· (pos zero) y = cong (_ℤ[ G ]· g) (+Comm 0 y) + ∙ sym (lid (snd G) _) + distrℤ· (pos (suc n)) (pos n₁) = + cong (_ℤ[ G ]· g) (sym (pos+ (suc n) n₁)) + ∙ cong (_·_ (snd G) g) (cong (_ℤ[ G ]· g) (pos+ n n₁) ∙ distrℤ· (pos n) (pos n₁)) + ∙ assoc (snd G) _ _ _ + distrℤ· (pos (suc n)) (negsuc zero) = + distrℤ· (pos n) 0 + ∙ ((rid (snd G) (pos n ℤ[ G ]· g) ∙ sym (lid (snd G) (pos n ℤ[ G ]· g))) + ∙ cong (λ x → _·_ (snd G) x (pos n ℤ[ G ]· g)) + (sym (invl (snd G) g)) ∙ sym (assoc (snd G) _ _ _)) + ∙ (assoc (snd G) _ _ _) + ∙ cong (λ x → _·_ (snd G) x (pos n ℤ[ G ]· g)) (invl (snd G) g) + ∙ lid (snd G) _ + ∙ sym (rid (snd G) _) + ∙ (cong (_·_ (snd G) (pos n ℤ[ G ]· g)) (sym (invr (snd G) g)) + ∙ assoc (snd G) _ _ _) + ∙ cong (λ x → _·_ (snd G) x (inv (snd G) g)) + (precommℤ G g (pos n)) + distrℤ· (pos (suc n)) (negsuc (suc n₁)) = + cong (_ℤ[ G ]· g) (predℤ+negsuc n₁ (pos (suc n))) + ∙∙ distrℤ· (pos n) (negsuc n₁) + ∙∙ (cong (λ x → _·_ (snd G) x (negsuc n₁ ℤ[ G ]· g)) + ((sym (rid (snd G) (pos n ℤ[ G ]· g)) + ∙ cong (_·_ (snd G) (pos n ℤ[ G ]· g)) (sym (invr (snd G) g))) + ∙∙ assoc (snd G) _ _ _ + ∙∙ cong (λ x → _·_ (snd G) x (inv (snd G) g)) (precommℤ G g (pos n) )) + ∙ sym (assoc (snd G) _ _ _)) + distrℤ· (negsuc zero) y = + cong (_ℤ[ G ]· g) (+Comm -1 y) ∙ lem y + distrℤ· (negsuc (suc n)) y = + cong (_ℤ[ G ]· g) (+Comm (negsuc (suc n)) y) + ∙∙ lem (y +negsuc n) + ∙∙ cong (snd G · inv (snd G) g) + (cong (_ℤ[ G ]· g) (+Comm y (negsuc n)) ∙ distrℤ· (negsuc n) y) + ∙ (assoc (snd G) _ _ _) + +GroupHomℤ→ℤpres- : (e : GroupHom ℤGroup ℤGroup) (a : ℤ) + → fst e (- a) ≡ - fst e a +GroupHomℤ→ℤpres- e a = cong (fst e) (minus≡0- a) ∙∙ presinv (snd e) a ∙∙ sym (minus≡0- _) + +ℤ·≡· : (a b : ℤ) → a * b ≡ (a ℤ[ ℤGroup ]· b) +ℤ·≡· (pos zero) b = refl +ℤ·≡· (pos (suc n)) b = cong (b +ℤ_) (ℤ·≡· (pos n) b) +ℤ·≡· (negsuc zero) b = minus≡0- b +ℤ·≡· (negsuc (suc n)) b = cong₂ _+ℤ_ (minus≡0- b) (ℤ·≡· (negsuc n) b) + +GroupHomℤ→ℤPres· : (e : GroupHom ℤGroup ℤGroup) (a b : ℤ) → fst e (a * b) ≡ a * fst e b +GroupHomℤ→ℤPres· e a b = cong (fst e) (ℤ·≡· a b) ∙∙ homPresℤ· e b a ∙∙ sym (ℤ·≡· a (fst e b)) + +-- Generators in terms of ℤ-multiplication +-- Todo : generalise +gen₁-by : (G : Group ℓ) → (g : fst G) → Type _ +gen₁-by G g = (h : fst G) + → Σ[ a ∈ ℤ ] h ≡ (a ℤ[ G ]· g) + +gen₂-by : ∀ {ℓ} (G : Group ℓ) → (g₁ g₂ : fst G) → Type _ +gen₂-by G g₁ g₂ = + (h : fst G) → Σ[ a ∈ ℤ × ℤ ] h ≡ _·_ (snd G) ((fst a) ℤ[ G ]· g₁) ((snd a) ℤ[ G ]· g₂) + +Iso-pres-gen₁ : ∀ {ℓ ℓ'} (G : Group ℓ) (H : Group ℓ') (g : fst G) + → gen₁-by G g → (e : GroupIso G H) + → gen₁-by H (Iso.fun (fst e) g) +Iso-pres-gen₁ G H g genG is h = + (fst (genG (Iso.inv (fst is) h))) + , (sym (Iso.rightInv (fst is) h) + ∙∙ cong (Iso.fun (fst is)) (snd (genG (Iso.inv (fst is) h))) + ∙∙ (homPresℤ· (_ , snd is) g (fst (genG (Iso.inv (fst is) h))))) + +Iso-pres-gen₂ : (G : Group ℓ) (H : Group ℓ') (g₁ g₂ : fst G) + → gen₂-by G g₁ g₂ → (e : GroupIso G H) + → gen₂-by H (Iso.fun (fst e) g₁) (Iso.fun (fst e) g₂) +fst (Iso-pres-gen₂ G H g₁ g₂ genG is h) = genG (Iso.inv (fst is) h) .fst +snd (Iso-pres-gen₂ G H g₁ g₂ genG is h) = + sym (Iso.rightInv (fst is) h) + ∙∙ cong (fun (fst is)) (snd (genG (Iso.inv (fst is) h))) + ∙∙ (pres· (snd is) _ _ + ∙ cong₂ (_·_ (snd H)) + (homPresℤ· (_ , snd is) g₁ (fst (fst (genG (inv (fst is) h))))) + (homPresℤ· (_ , snd is) g₂ (snd (fst (genG (inv (fst is) h)))))) + + +private + intLem₁ : (n m : ℕ) → Σ[ a ∈ ℕ ] (negsuc n * pos (suc m)) ≡ negsuc a + intLem₁ n zero = n , ·Comm (negsuc n) (pos 1) + intLem₁ n (suc m) = lem _ _ .fst , + (·Comm (negsuc n) (pos (suc (suc m))) + ∙∙ cong (negsuc n +ℤ_) (·Comm (pos (suc m)) (negsuc n) ∙ (intLem₁ n m .snd)) + ∙∙ lem _ _ .snd) + where + lem : (x y : ℕ) → Σ[ a ∈ ℕ ] negsuc x +ℤ negsuc y ≡ negsuc a + lem zero zero = 1 , refl + lem zero (suc y) = (suc (suc y)) , +Comm (negsuc zero) (negsuc (suc y)) + lem (suc x) zero = (suc (suc x)) , refl + lem (suc x) (suc y) = (lem (suc (suc x)) y .fst) , (predℤ+negsuc y (negsuc (suc x)) ∙ snd ((lem (suc (suc x))) y)) + + intLem₂ : (n x : ℕ) → Σ[ a ∈ ℕ ] ((pos (suc x)) * pos (suc (suc n)) ≡ pos (suc (suc a))) + intLem₂ n zero = n , refl + intLem₂ n (suc x) = h3 _ _ (intLem₂ n x) + where + h3 : (x : ℤ) (n : ℕ) → Σ[ a ∈ ℕ ] (x ≡ pos (suc (suc a))) + → Σ[ a ∈ ℕ ] pos n +ℤ x ≡ pos (suc (suc a)) + h3 x n (a , p) = n +ℕ a + , cong (pos n +ℤ_) p ∙ cong sucℤ (sucℤ+pos a (pos n)) + ∙ sucℤ+pos a (pos (suc n)) ∙ (sym (pos+ (suc (suc n)) a)) + + ¬1=x·suc-suc : (n : ℕ) (x : ℤ) → ¬ pos 1 ≡ x * pos (suc (suc n)) + ¬1=x·suc-suc n (pos zero) p = snotz (injPos p) + ¬1=x·suc-suc n (pos (suc n₁)) p = + snotz (injPos (sym (cong predℤ (snd (intLem₂ n n₁))) ∙ sym (cong predℤ p))) + ¬1=x·suc-suc n (negsuc n₁) p = posNotnegsuc _ _ (p ∙ intLem₁ _ _ .snd) + +GroupEquivℤ-pres1 : (e : GroupEquiv ℤGroup ℤGroup) (x : ℤ) + → (fst (fst e) 1) ≡ x → abs (fst (fst e) 1) ≡ 1 +GroupEquivℤ-pres1 e (pos zero) p = + ⊥-rec (snotz (injPos (sym (retEq (fst e) 1) + ∙∙ (cong (fst (fst (invGroupEquiv e))) p) + ∙∙ IsGroupHom.pres1 (snd (invGroupEquiv e))))) +GroupEquivℤ-pres1 e (pos (suc zero)) p = cong abs p +GroupEquivℤ-pres1 e (pos (suc (suc n))) p = + ⊥-rec (¬1=x·suc-suc _ _ (h3 ∙ ·Comm (pos (suc (suc n))) (invEq (fst e) 1))) + where + h3 : pos 1 ≡ _ + h3 = sym (retEq (fst e) 1) + ∙∙ (cong (fst (fst (invGroupEquiv e))) (p ∙ ·Comm 1 (pos (suc (suc n))))) + ∙∙ GroupHomℤ→ℤPres· (_ , snd (invGroupEquiv e)) (pos (suc (suc n))) 1 +GroupEquivℤ-pres1 e (negsuc zero) p = cong abs p +GroupEquivℤ-pres1 e (negsuc (suc n)) p = ⊥-rec (¬1=x·suc-suc _ _ l33) + where + h3 : fst (fst e) (negsuc zero) ≡ pos (suc (suc n)) + h3 = GroupHomℤ→ℤpres- (_ , snd e) (pos 1) ∙ cong -_ p + + compGroup : GroupEquiv ℤGroup ℤGroup + fst compGroup = isoToEquiv (iso -_ -_ -Involutive -Involutive) + snd compGroup = makeIsGroupHom -Dist+ + + compHom : GroupEquiv ℤGroup ℤGroup + compHom = compGroupEquiv compGroup e + + l32 : fst (fst compHom) (pos 1) ≡ pos (suc (suc n)) + l32 = h3 + + l33 : pos 1 ≡ invEq (fst compHom) (pos 1) * pos (suc (suc n)) + l33 = sym (retEq (fst compHom) (pos 1)) + ∙∙ cong (invEq (fst compHom)) l32 + ∙∙ (cong (invEq (fst compHom)) (·Comm (pos 1) (pos (suc (suc n)))) + ∙ GroupHomℤ→ℤPres· (_ , (snd (invGroupEquiv compHom))) (pos (suc (suc n))) (pos 1) + ∙ ·Comm (pos (suc (suc n))) (invEq (fst compHom) (pos 1))) + +-- really needed? +{- +sesIsoPresGen : + ∀ (G : Group ℓ-zero) + → (iso : GroupEquiv ℤGroup G) + → (H : Group ℓ-zero) + → (iso2 : GroupEquiv ℤGroup H) + → (e : fst G) + → invEq (fst iso) e ≡ 1 + → (hom : GroupEquiv G H) + → abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1 +sesIsoPresGen G = JGroupEquiv (λ G iso → + (H : Group ℓ-zero) + (iso2 : GroupEquiv ℤGroup H) + → (e : fst G) + → invEq (fst iso) e ≡ 1 + → (hom : GroupEquiv G H) + → abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1) + λ H → JGroupEquiv (λ H iso2 → (e : ℤ) → e ≡ pos 1 → + (hom : GroupEquiv ℤGroup H) → + abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1) + λ e p hom → cong (abs ∘ (fst (fst hom))) p ∙ autoPres1 hom + where + autoPres1 : (e : GroupEquiv ℤGroup ℤGroup) + → abs (fst (fst e) 1) ≡ 1 + autoPres1 e = GroupEquivℤ-pres1 e _ refl +-} + +allisoPresGen : ∀ {ℓ} (G : Group ℓ) (ϕ : GroupEquiv G ℤGroup) (x : fst G) + → (fst (fst ϕ) x ≡ 1) ⊎ (fst (fst ϕ) x ≡ - 1) + → (ψ : GroupEquiv G ℤGroup) + → (fst (fst ψ) x ≡ 1) ⊎ (fst (fst ψ) x ≡ - 1) +allisoPresGen G (ϕeq , ϕhom) x (inl r) (ψeq , ψhom) = + abs→⊎ _ _ (cong abs (cong (fst ψeq) (sym (retEq ϕeq x) ∙ cong (invEq ϕeq) r)) + ∙ GroupEquivℤ-pres1 (compGroupEquiv (invGroupEquiv (ϕeq , ϕhom)) (ψeq , ψhom)) _ refl) +allisoPresGen G (ϕeq , ϕhom) x (inr r) (ψeq , ψhom) = + abs→⊎ _ _ + (cong abs (cong (fst ψeq) (sym (retEq ϕeq x) ∙ cong (invEq ϕeq) r)) + ∙ cong abs (IsGroupHom.presinv (snd (compGroupEquiv (invGroupEquiv (ϕeq , ϕhom)) (ψeq , ψhom))) 1) + ∙ absLem _ (GroupEquivℤ-pres1 (compGroupEquiv (invGroupEquiv (ϕeq , ϕhom)) (ψeq , ψhom)) _ refl)) + where + absLem : ∀ x → abs x ≡ 1 → abs (pos 0 -ℤ x) ≡ 1 + absLem (pos zero) p = ⊥-rec (snotz (sym p)) + absLem (pos (suc zero)) p = refl + absLem (pos (suc (suc n))) p = ⊥-rec (snotz (cong predℕ p)) + absLem (negsuc zero) p = refl + absLem (negsuc (suc n)) p = ⊥-rec (snotz (cong predℕ p)) diff --git a/Cubical/Data/Int/Properties.agda b/Cubical/Data/Int/Properties.agda index 14e2c120f..f1342736d 100644 --- a/Cubical/Data/Int/Properties.agda +++ b/Cubical/Data/Int/Properties.agda @@ -456,3 +456,17 @@ private cong ((- (b · c)) +_) (·Assoc (negsuc n) b c) ∙∙ cong (_+ ((negsuc n · b) · c)) (-DistL· b c) ∙∙ sym (·DistL+ (- b) (negsuc n · b) c) + +-- Absolute values +abs→⊎ : (x : ℤ) (n : ℕ) → abs x ≡ n → (x ≡ pos n) ⊎ (x ≡ - pos n) +abs→⊎ x n = J (λ n _ → (x ≡ pos n) ⊎ (x ≡ - pos n)) (help x) + where + help : (x : ℤ) → (x ≡ pos (abs x)) ⊎ (x ≡ - pos (abs x)) + help (pos n) = inl refl + help (negsuc n) = inr refl + +⊎→abs : (x : ℤ) (n : ℕ) → (x ≡ pos n) ⊎ (x ≡ - pos n) → abs x ≡ n +⊎→abs (pos n₁) n (inl x₁) = cong abs x₁ +⊎→abs (negsuc n₁) n (inl x₁) = cong abs x₁ +⊎→abs x zero (inr x₁) = cong abs x₁ +⊎→abs x (suc n) (inr x₁) = cong abs x₁ diff --git a/Cubical/Data/Sigma/Properties.agda b/Cubical/Data/Sigma/Properties.agda index 0679e8a88..607a86415 100644 --- a/Cubical/Data/Sigma/Properties.agda +++ b/Cubical/Data/Sigma/Properties.agda @@ -263,16 +263,17 @@ PathΣ→ΣPathTransport a b = Iso.inv (IsoΣPathTransportPathΣ a b) ΣPathTransport≡PathΣ : (a b : Σ A B) → ΣPathTransport a b ≡ (a ≡ b) ΣPathTransport≡PathΣ a b = ua (ΣPathTransport≃PathΣ a b) +Σ-contractFstIso : (c : isContr A) → Iso (Σ A B) (B (c .fst)) +fun (Σ-contractFstIso {B = B} c) p = subst B (sym (c .snd (fst p))) (snd p) +inv (Σ-contractFstIso {B = B} c) b = _ , b +rightInv (Σ-contractFstIso {B = B} c) b = + cong (λ p → subst B p b) (isProp→isSet (isContr→isProp c) _ _ _ _) ∙ transportRefl _ +fst (leftInv (Σ-contractFstIso {B = B} c) p j) = c .snd (fst p) j +snd (leftInv (Σ-contractFstIso {B = B} c) p j) = + transp (λ i → B (c .snd (fst p) (~ i ∨ j))) j (snd p) + Σ-contractFst : (c : isContr A) → Σ A B ≃ B (c .fst) -Σ-contractFst {B = B} c = isoToEquiv isom - where - isom : Iso _ _ - isom .fun (a , b) = subst B (sym (c .snd a)) b - isom .inv b = (c .fst , b) - isom .rightInv b = - cong (λ p → subst B p b) (isProp→isSet (isContr→isProp c) _ _ _ _) ∙ transportRefl _ - isom .leftInv (a , b) = - ΣPathTransport≃PathΣ _ _ .fst (c .snd a , transportTransport⁻ (cong B (c .snd a)) _) +Σ-contractFst {B = B} c = isoToEquiv (Σ-contractFstIso c) -- a special case of the above module _ (A : Unit → Type ℓ) where diff --git a/Cubical/ZCohomology/Gysin.agda b/Cubical/ZCohomology/Gysin.agda index 8043a7867..ac1cf61c4 100644 --- a/Cubical/ZCohomology/Gysin.agda +++ b/Cubical/ZCohomology/Gysin.agda @@ -45,6 +45,7 @@ open import Cubical.Data.Unit open import Cubical.Data.Bool open import Cubical.Algebra.Group renaming (ℤ to ℤGroup ; Unit to UnitGroup) hiding (Bool) +open import Cubical.Algebra.Group.ZModule open import Cubical.HITs.Pushout.Flattening open import Cubical.Homotopy.Connected @@ -70,19 +71,15 @@ open import Cubical.Homotopy.Hopf open import Cubical.HITs.SetQuotients renaming (_/_ to _/'_) --- move to Sigma -contrFstΣ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} - → (e : isContr A) - → Iso (Σ A B) (B (fst e)) -Iso.fun (contrFstΣ {B = B} e) (a , b) = subst B (sym (snd e a)) b -Iso.inv (contrFstΣ {B = B} e) b = (fst e) , b -Iso.rightInv (contrFstΣ {B = B} e) b = cong (λ x → subst B x b) h3 ∙ transportRefl b - where - h3 : sym (snd e (fst e)) ≡ refl - h3 = isProp→isSet (isContr→isProp e) _ _ _ _ -Iso.leftInv (contrFstΣ {B = B} e) b = - ΣPathP ((snd e (fst b)) - , (λ j → transp (λ i → B (snd e (fst b) (~ i ∨ j))) j (snd b))) +isContr→≡UnitGroup : {G : Group ℓ-zero} → isContr (fst G) → UnitGroup ≡ G +isContr→≡UnitGroup c = + fst (GroupPath _ _) + (invGroupEquiv ((isContr→≃Unit c) + , (makeIsGroupHom (λ _ _ → refl)))) + +GroupIsoUnitGroup→isContr : {G : Group ℓ-zero} → GroupIso UnitGroup G → isContr (fst G) +GroupIsoUnitGroup→isContr is = + isOfHLevelRetractFromIso 0 (invIso (fst is)) isContrUnit -- move to group stuff SES→isEquiv : ∀ {ℓ ℓ'} {L R : Group ℓ-zero} @@ -120,7 +117,8 @@ SES→isEquiv {R = R} {G = G} {H = H} = BijectionIso.fun bijIso' = midhom BijectionIso.inj bijIso' x inker = pRec (GroupStr.is-set (snd G) _ _) - (λ s → sym (snd s) ∙ IsGroupHom.pres1 (snd lhom)) (lexact _ inker) + (λ s → sym (snd s) ∙ IsGroupHom.pres1 (snd lhom)) + (lexact _ inker) BijectionIso.surj bijIso' x = rexact x refl -- The pushout of the hopf invariant @@ -145,7 +143,6 @@ Hopfβ : (n : ℕ) → coHom (4 +ℕ (n +ℕ n)) (HopfInvariantPush n (fst f)) Hopfβ n f = fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) ∣ ∣_∣ ∣₂ - -- To define the Hopf invariant, we need to establish the non-trivial isos of Hⁿ(HopfInvariant). module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) where module M = MV _ _ _ (λ _ → tt) (fst f) @@ -157,126 +154,128 @@ module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) whe ¬lem n (suc m) (cong (suc ∘ suc ) (sym (+-suc m n)) ∙ (cong predℕ p)) - GroupEquiv1 : + SphereHopfCohomIso : GroupEquiv - (coHomGr (3 +ℕ n +ℕ n) (fst (S₊∙ (3 +ℕ n +ℕ n)))) - ((coHomGr (suc (3 +ℕ n +ℕ n)) (Pushout (λ _ → tt) (fst f)))) - fst (fst GroupEquiv1) = fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) - snd (fst GroupEquiv1) = help + (coHomGr (3 +ℕ n +ℕ n) (S₊ (3 +ℕ n +ℕ n))) + ((coHomGr (suc (3 +ℕ n +ℕ n)) (HopfInvariantPush _ (fst f)))) + fst (fst SphereHopfCohomIso) = fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) + snd (fst SphereHopfCohomIso) = help where abstract help : isEquiv (fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n))) help = SES→isEquiv - (fst (GroupPath _ _) - (invGroupEquiv (isContr→≃Unit - (isOfHLevelΣ 0 - (isOfHLevelRetractFromIso 0 (fst (Hⁿ-Unit≅0 _)) isContrUnit) - λ _ → isOfHLevelRetractFromIso 0 - (fst (Hⁿ-Sᵐ≅0 _ _ (¬lem n n))) isContrUnit) - , makeIsGroupHom λ _ _ → refl))) - (fst (GroupPath _ _) - (invGroupEquiv - (isContr→≃Unit - (isContrΣ - (isOfHLevelRetractFromIso 0 (fst (Hⁿ-Unit≅0 _)) - isContrUnit) - λ _ → isOfHLevelRetractFromIso 0 - (fst (Hⁿ-Sᵐ≅0 _ _ (¬lem n (suc n)))) isContrUnit) - , makeIsGroupHom λ _ _ → refl))) + (isContr→≡UnitGroup + (isContrΣ (GroupIsoUnitGroup→isContr (invGroupIso (Hⁿ-Unit≅0 _))) + λ _ → GroupIsoUnitGroup→isContr + (invGroupIso (Hⁿ-Sᵐ≅0 _ _ (¬lem n n))))) + (isContr→≡UnitGroup + (isContrΣ (GroupIsoUnitGroup→isContr + (invGroupIso (Hⁿ-Unit≅0 _))) + λ _ → GroupIsoUnitGroup→isContr + (invGroupIso (Hⁿ-Sᵐ≅0 _ _ (¬lem n (suc n)))))) (M.Δ (3 +ℕ n +ℕ n)) (M.d (3 +ℕ n +ℕ n)) (M.i (4 +ℕ n +ℕ n)) (M.Ker-d⊂Im-Δ _) (M.Ker-i⊂Im-d _) - snd GroupEquiv1 = s + snd SphereHopfCohomIso = s where + -- Currently, type checking is very slow without the abstract flag. + -- TODO : Remove abstract abstract - s : IsGroupHom (snd (coHomGr (3 +ℕ n +ℕ n) (fst (S₊∙ (3 +ℕ n +ℕ n))))) + s : IsGroupHom (snd (coHomGr (3 +ℕ n +ℕ n) (S₊ (3 +ℕ n +ℕ n)))) (fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n))) (snd (coHomGr (suc (3 +ℕ n +ℕ n)) (Pushout (λ _ → tt) (fst f)))) s = snd (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) - Hopfβ-Iso : GroupIso ((coHomGr (suc (3 +ℕ n +ℕ n)) (Pushout (λ _ → tt) (fst f)))) ℤGroup - Iso.fun (fst Hopfβ-Iso) x = Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (invEq (fst GroupEquiv1) x) - Iso.inv (fst Hopfβ-Iso) f = (fst (fst GroupEquiv1)) (Iso.inv (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) f) + Hopfβ-Iso : GroupIso (coHomGr (suc (3 +ℕ n +ℕ n)) (HopfInvariantPush _ (fst f))) + ℤGroup + Iso.fun (fst Hopfβ-Iso) x = + Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) + (invEq (fst SphereHopfCohomIso) x) + Iso.inv (fst Hopfβ-Iso) f = + (fst (fst SphereHopfCohomIso)) + (Iso.inv (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) f) Iso.rightInv (fst Hopfβ-Iso) f = - cong (Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) ) - (retEq (fst GroupEquiv1) (Iso.inv (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) f)) - ∙ Iso.rightInv (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) f + cong (Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) + (retEq (fst SphereHopfCohomIso) (Iso.inv (fst (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) f)) + ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) f Iso.leftInv (fst Hopfβ-Iso) x = - cong (fst (fst GroupEquiv1)) - (Iso.leftInv (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) - (invEq (fst GroupEquiv1) x)) - ∙ secEq (fst GroupEquiv1) x + cong (fst (fst SphereHopfCohomIso)) + (Iso.leftInv (fst (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) + (invEq (fst SphereHopfCohomIso) x)) + ∙ secEq (fst SphereHopfCohomIso) x snd Hopfβ-Iso = grHom where abstract grHom : IsGroupHom (coHomGr (suc (3 +ℕ n +ℕ n)) (Pushout (λ _ → tt) (fst f)) .snd) - (λ x → Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (invEq (fst GroupEquiv1) x)) + (λ x → Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) + (invEq (fst SphereHopfCohomIso) x)) (ℤGroup .snd) grHom = makeIsGroupHom λ x y → cong (Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) - (IsGroupHom.pres· (isGroupHomInv GroupEquiv1) x y) + (IsGroupHom.pres· (isGroupHomInv SphereHopfCohomIso) x y) ∙ IsGroupHom.pres· (snd (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) _ _ Hⁿ-Sⁿ≅ℤ-nice-generator : (n : ℕ) → Iso.inv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) 1 ≡ ∣ ∣_∣ ∣₂ Hⁿ-Sⁿ≅ℤ-nice-generator zero = Iso.leftInv (fst (Hⁿ-Sⁿ≅ℤ (suc zero))) _ -Hⁿ-Sⁿ≅ℤ-nice-generator (suc n) = (λ i → Iso.inv - (fst - (suspensionAx-Sn (suc n) (suc n))) (Hⁿ-Sⁿ≅ℤ-nice-generator n i)) +Hⁿ-Sⁿ≅ℤ-nice-generator (suc n) = + (λ i → Iso.inv (fst (suspensionAx-Sn (suc n) (suc n))) (Hⁿ-Sⁿ≅ℤ-nice-generator n i)) ∙ cong ∣_∣₂ (funExt λ { north → refl ; south → cong ∣_∣ₕ (merid north) - ; (merid a i) j → ∣ compPath-filler (merid a) (sym (merid north)) (~ j) i ∣}) + ; (merid a i) j → ∣ compPath-filler (merid a) (sym (merid north)) (~ j) i ∣ₕ}) Hopfβ↦1 : (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) → Iso.fun (fst (Hopfβ-Iso n f)) (Hopfβ n f) ≡ 1 -Hopfβ↦1 n f = sym (cong (Iso.fun (fst (Hopfβ-Iso n f))) h) ∙ Iso.rightInv (fst (Hopfβ-Iso n f)) (pos 1) +Hopfβ↦1 n f = sym (cong (Iso.fun (fst (Hopfβ-Iso n f))) lem) ∙ Iso.rightInv (fst (Hopfβ-Iso n f)) (pos 1) where - h : Iso.inv (fst (Hopfβ-Iso n f)) (pos 1) ≡ Hopfβ n f - h = (λ i → fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) + lem : Iso.inv (fst (Hopfβ-Iso n f)) (pos 1) ≡ Hopfβ n f + lem = (λ i → fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) (Hⁿ-Sⁿ≅ℤ-nice-generator _ i)) ∙ cong ∣_∣₂ (funExt (λ { (inl x) → refl ; (inr x) → refl ; (push a i) → refl})) module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) where - 2+n = 2 +ℕ n - H = HopfInvariantPush n (fst f) + private + 2+n = 2 +ℕ n + H = HopfInvariantPush n (fst f) - ff : coHom 2+n H → coHom 2+n (S₊ (suc (suc n))) - ff = sMap (_∘ inr) + H→Sphere : coHom 2+n H → coHom 2+n (S₊ (suc (suc n))) + H→Sphere = sMap (_∘ inr) - grHom : IsGroupHom (snd (coHomGr 2+n H)) ff (snd (coHomGr 2+n (S₊ (suc (suc n))))) - grHom = - makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) - λ f g → refl) + grHom : IsGroupHom (snd (coHomGr 2+n H)) H→Sphere (snd (coHomGr 2+n (S₊ (suc (suc n))))) + grHom = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f g → refl) - h' : (g : (S₊ (suc (suc n)) → coHomK 2+n)) → H → coHomK (2 +ℕ n) - h' g (inl x) = 0ₖ _ - h' g (inr x) = g x -ₖ g north - h' g (push a i) = h23 a i - where - h23 : (a : S₊ (suc (suc (suc (n +ℕ n))))) → 0ₖ (suc (suc n)) ≡ g (fst f a) -ₖ g north - h23 = sphereElim _ (λ x → isOfHLevelPlus' {n = n} (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n) _ _)) - (sym (rCancelₖ _ (g north)) ∙ cong (λ x → g x -ₖ g north) (sym (snd f))) + h' : (g : (S₊ (suc (suc n)) → coHomK 2+n)) → H → coHomK (2 +ℕ n) + h' g (inl x) = 0ₖ _ + h' g (inr x) = g x -ₖ g north + h' g (push a i) = h23 a i + where + h23 : (a : S₊ (suc (suc (suc (n +ℕ n))))) → 0ₖ (suc (suc n)) ≡ g (fst f a) -ₖ g north + h23 = sphereElim _ (λ x → isOfHLevelPlus' {n = n} (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n) _ _)) + (sym (rCancelₖ _ (g north)) ∙ cong (λ x → g x -ₖ g north) (sym (snd f))) - ff⁻ : coHom 2+n (S₊ (suc (suc n))) → coHom 2+n H - ff⁻ = sMap h' + Sphere→H : coHom 2+n (S₊ (suc (suc n))) → coHom 2+n H + Sphere→H = sMap h' - h23 : (x : _) → ∥ x ≡ 0ₖ (suc (suc n)) ∥ - h23 = coHomK-elim _ (λ _ → isProp→isOfHLevelSuc (suc n) squash) ∣ refl ∣ + conCohom2+n : (x : _) → ∥ x ≡ 0ₖ (suc (suc n)) ∥ + conCohom2+n = coHomK-elim _ (λ _ → isProp→isOfHLevelSuc (suc n) squash) ∣ refl ∣ - h : Iso (coHom (2 +ℕ n) (HopfInvariantPush n (fst f))) (coHom (2 +ℕ n) ((S₊ (suc (suc n))))) - Iso.fun h = ff - Iso.inv h = ff⁻ - Iso.rightInv h = + HIPSphereCohomIso : Iso (coHom (2 +ℕ n) (HopfInvariantPush n (fst f))) + (coHom (2 +ℕ n) ((S₊ (suc (suc n))))) + Iso.fun HIPSphereCohomIso = H→Sphere + Iso.inv HIPSphereCohomIso = Sphere→H + Iso.rightInv HIPSphereCohomIso = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ g → pRec (squash₂ _ _) (λ p → cong ∣_∣₂ (funExt λ x → cong (g x +ₖ_) (cong (-ₖ_) p) ∙ rUnitₖ _ (g x))) - (h23 (g north)) - Iso.leftInv h = + (conCohom2+n (g north)) + Iso.leftInv HIPSphereCohomIso = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ g → pRec (squash₂ _ _) (pRec (isPropΠ (λ _ → squash₂ _ _)) @@ -296,8 +295,8 @@ module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) whe (isOfHLevelTrunc (suc (suc (suc (suc n)))))) _ _) _ _) r a i})) (l g gtt gn)) - (h23 (g (inr north)))) - (h23 (g (inl tt))) + (conCohom2+n (g (inr north)))) + (conCohom2+n (g (inl tt))) where l : (g : HopfInvariantPush n (fst f) → coHomK (suc (suc n))) → (gtt : g (inl tt) ≡ 0ₖ (suc (suc n))) @@ -306,11 +305,10 @@ module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) whe (sym gtt) ((λ i → g (inr (fst f north)) -ₖ gn i) ∙ rUnitₖ _ (g (inr (fst f north))))) l g gtt gn = isConnectedPathP _ (isConnectedPath _ (isConnectedKn _) _ _) _ _ .fst - Hopfα-Iso : GroupIso (coHomGr (2 +ℕ n) (HopfInvariantPush n (fst f))) ℤGroup Hopfα-Iso = compGroupIso - (h , grHom) + (HIPSphereCohomIso , grHom) (Hⁿ-Sⁿ≅ℤ (suc n)) Hopfα-Iso-α : (n : ℕ) (f : _) @@ -320,21 +318,19 @@ Hopfα-Iso-α n f = sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (Hⁿ-Sⁿ≅ℤ-nice-generator n)) ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) where - hz : Iso.fun (h n f) (Hopfα n f) ≡ ∣ ∣_∣ ∣₂ + hz : Iso.fun (HIPSphereCohomIso n f) (Hopfα n f) ≡ ∣ ∣_∣ ∣₂ hz = refl - ⌣-α : (n : ℕ) → (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → _ ⌣-α n f = Hopfα n f ⌣ Hopfα n f -HopfFunction : (n : ℕ) +HopfInvariant : (n : ℕ) → (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → ℤ -HopfFunction n f = Iso.fun (fst (Hopfβ-Iso n f)) +HopfInvariant n f = Iso.fun (fst (Hopfβ-Iso n f)) ((subst (λ x → coHom x (HopfInvariantPush n (fst f))) (λ i → suc (suc (suc (+-suc n n i))))) (⌣-α n f)) - open import Cubical.Algebra.AbGroup open import Cubical.Homotopy.Loopspace @@ -346,27 +342,18 @@ open import Cubical.Homotopy.Hopf open import Cubical.HITs.SetQuotients renaming (_/_ to _/'_) -- open import Cubical.ZCohomology.Gysin -_·₀ₕ_ : ∀ {ℓ} {n : ℕ} {A : Type ℓ} → ℤ → coHom n A → coHom n A -_·₀ₕ_ {n = n} a b = a ℤ[ coHomGr n _ ]· b - --- remove -HopfInvariant : (n : ℕ) - → (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → Type _ -HopfInvariant n f = - Σ[ x ∈ ℤ ] - Hopfα n f ⌣ Hopfα n f - ≡ (x ·₀ₕ subst (λ x → coHom x (HopfInvariantPush n (fst f))) - (λ i → suc (suc (suc (+-suc n n (~ i))))) - (Hopfβ n f)) - -GroupHomℤ→ℤPres·₀ : ∀ {ℓ} {A : Type ℓ} (n : ℕ) - (e : GroupHom ℤGroup (coHomGr n A)) - → (a b : ℤ) - → fst e (a · b) - ≡ (a ·₀ₕ fst e b) -GroupHomℤ→ℤPres·₀ {A = A} n e a b = cong (fst e) (ℤ·≡· a b) ∙ hompres· {H = coHomGr n A} (_ , snd e) b a - +private + _·₀ₕ_ : ∀ {ℓ} {n : ℕ} {A : Type ℓ} → ℤ → coHom n A → coHom n A + _·₀ₕ_ {n = n} a b = a ℤ[ coHomGr n _ ]· b + + GroupHomℤ→ℤPres·₀ : ∀ {ℓ} {A : Type ℓ} (n : ℕ) + (e : GroupHom ℤGroup (coHomGr n A)) + → (a b : ℤ) + → fst e (a · b) + ≡ (a ·₀ₕ fst e b) + GroupHomℤ→ℤPres·₀ {A = A} n e a b = + cong (fst e) (ℤ·≡· a b) + ∙ homPresℤ· {H = coHomGr n A} (_ , snd e) b a -- goal: remove genH : ∀ {ℓ} {A : Type ℓ} (n : ℕ) @@ -391,154 +378,9 @@ genH' {A = A} n e = (Iso.inv (fst e) 1) (sym (Iso.leftInv (fst e) h) ∙∙ sym (cong (Iso.inv (fst e)) (·Comm (Iso.fun (fst e) h) (pos 1))) ∙∙ (cong (Iso.inv (fst e)) (ℤ·≡· _ _) - ∙ (hompres· (_ , snd (invGroupIso e)) (pos 1) (Iso.fun (fst e) h)))) + ∙ (homPresℤ· (_ , snd (invGroupIso e)) (pos 1) (Iso.fun (fst e) h)))) -- Move to intGroup -private - l2 : (n m : ℕ) → Σ[ a ∈ ℕ ] (negsuc n · pos (suc m)) ≡ negsuc a - l2 n zero = n , ·Comm (negsuc n) (pos 1) - l2 n (suc m) = h3 _ _ .fst , - (·Comm (negsuc n) (pos (suc (suc m))) - ∙∙ cong (negsuc n +_) (·Comm (pos (suc m)) (negsuc n) ∙ (l2 n m .snd)) - ∙∙ h3 _ _ .snd) - where - h3 : (x y : ℕ) → Σ[ a ∈ ℕ ] negsuc x + negsuc y ≡ negsuc a - h3 zero zero = 1 , refl - h3 zero (suc y) = (suc (suc y)) , +Comm (negsuc zero) (negsuc (suc y)) - h3 (suc x) zero = (suc (suc x)) , refl - h3 (suc x) (suc y) = (h3 (suc (suc x)) y .fst) , (predℤ+negsuc y (negsuc (suc x)) ∙ snd ((h3 (suc (suc x))) y)) - - zz : (n x : ℕ) → Σ[ a ∈ ℕ ] ((pos (suc x)) · pos (suc (suc n)) ≡ pos (suc (suc a))) - zz n zero = n , refl - zz n (suc x) = h3 _ _ (zz n x) - where - h3 : (x : ℤ) (n : ℕ) → Σ[ a ∈ ℕ ] (x ≡ pos (suc (suc a))) - → Σ[ a ∈ ℕ ] pos n + x ≡ pos (suc (suc a)) - h3 x n (a , p) = n +ℕ a - , cong (pos n +_) p ∙ cong sucℤ (sucℤ+pos a (pos n)) - ∙ sucℤ+pos a (pos (suc n)) ∙ (sym (pos+ (suc (suc n)) a)) - - lem22 : (n : ℕ) (x : ℤ) → ¬ pos 1 ≡ x · pos (suc (suc n)) - lem22 n (pos zero) p = snotz (injPos p) - lem22 n (pos (suc n₁)) p = snotz (injPos (sym (cong predℤ (snd (zz n n₁))) ∙ sym (cong predℤ p))) - lem22 n (negsuc n₁) p = posNotnegsuc _ _ (p ∙ l2 _ _ .snd) - - --- Move to intGroup -GroupEquivℤ-pres1 : (e : GroupEquiv ℤGroup ℤGroup) (x : ℤ) → (fst (fst e) 1) ≡ x → abs (fst (fst e) 1) ≡ 1 -GroupEquivℤ-pres1 e (pos zero) p = - ⊥-rec (snotz (injPos (sym (retEq (fst e) 1) - ∙∙ (cong (fst (fst (invGroupEquiv e))) p) - ∙∙ IsGroupHom.pres1 (snd (invGroupEquiv e))))) -GroupEquivℤ-pres1 e (pos (suc zero)) p = cong abs p -GroupEquivℤ-pres1 e (pos (suc (suc n))) p = ⊥-rec (lem22 _ _ (h3 ∙ ·Comm (pos (suc (suc n))) (invEq (fst e) 1))) - where - - h3 : pos 1 ≡ _ - h3 = sym (retEq (fst e) 1) - ∙∙ (cong (fst (fst (invGroupEquiv e))) (p ∙ ·Comm 1 (pos (suc (suc n))))) - ∙∙ GroupHomℤ→ℤPres· (_ , snd (invGroupEquiv e)) (pos (suc (suc n))) 1 -GroupEquivℤ-pres1 e (negsuc zero) p = cong abs p -GroupEquivℤ-pres1 e (negsuc (suc n)) p = ⊥-rec (lem22 _ _ l33) - where - h3 : fst (fst e) (negsuc zero) ≡ pos (suc (suc n)) - h3 = GroupHomℤ→ℤpres- (_ , snd e) (pos 1) ∙ cong -_ p - - compGroup : GroupEquiv ℤGroup ℤGroup - fst compGroup = isoToEquiv (iso -_ -_ -Involutive -Involutive) - snd compGroup = makeIsGroupHom -Dist+ - - compHom : GroupEquiv ℤGroup ℤGroup - compHom = compGroupEquiv compGroup e - - l32 : fst (fst compHom) (pos 1) ≡ pos (suc (suc n)) - l32 = h3 - - l33 : pos 1 ≡ invEq (fst compHom) (pos 1) · pos (suc (suc n)) - l33 = sym (retEq (fst compHom) (pos 1)) - ∙∙ cong (invEq (fst compHom)) l32 - ∙∙ (cong (invEq (fst compHom)) (·Comm (pos 1) (pos (suc (suc n)))) - ∙ GroupHomℤ→ℤPres· (_ , (snd (invGroupEquiv compHom))) (pos (suc (suc n))) (pos 1) - ∙ ·Comm (pos (suc (suc n))) (invEq (fst compHom) (pos 1))) - --- move to Int -abs→⊎ : (x : ℤ) (n : ℕ) → abs x ≡ n → (x ≡ pos n) ⊎ (x ≡ - pos n) -abs→⊎ x n = J (λ n _ → (x ≡ pos n) ⊎ (x ≡ - pos n)) (help x) - where - help : (x : ℤ) → (x ≡ pos (abs x)) ⊎ (x ≡ - pos (abs x)) - help (pos n) = inl refl - help (negsuc n) = inr refl - -⊎→abs : (x : ℤ) (n : ℕ) → (x ≡ pos n) ⊎ (x ≡ - pos n) → abs x ≡ n -⊎→abs (pos n₁) n (inl x₁) = cong abs x₁ -⊎→abs (negsuc n₁) n (inl x₁) = cong abs x₁ -⊎→abs x zero (inr x₁) = cong abs x₁ -⊎→abs x (suc n) (inr x₁) = cong abs x₁ - --- Move to group -JGroupEquiv : ∀ {ℓ ℓ'} {G : Group ℓ} (P : (H : Group ℓ) → GroupEquiv G H → Type ℓ') - → P G idGroupEquiv - → ∀ {H} e → P H e -JGroupEquiv {G = G} P p {H} e = - transport (λ i → P (GroupPath G H .fst e i) - (transp (λ j → GroupEquiv G (GroupPath G H .fst e (i ∨ ~ j))) i e)) - (subst (P G) (sym l) p) - where - l : (transp (λ j → GroupEquiv G (GroupPath G H .fst e (~ j))) i0 e) ≡ idGroupEquiv - l = Σ≡Prop (λ _ → isPropIsGroupHom _ _) - (Σ≡Prop (λ _ → isPropIsEquiv _) - (funExt λ x → (λ i → fst (fst (fst e .snd .equiv-proof - (transportRefl (fst (fst e) (transportRefl x i)) i)))) - ∙ retEq (fst e) x)) - --- really needed? -{- -sesIsoPresGen : - ∀ (G : Group ℓ-zero) - → (iso : GroupEquiv ℤGroup G) - → (H : Group ℓ-zero) - → (iso2 : GroupEquiv ℤGroup H) - → (e : fst G) - → invEq (fst iso) e ≡ 1 - → (hom : GroupEquiv G H) - → abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1 -sesIsoPresGen G = JGroupEquiv (λ G iso → - (H : Group ℓ-zero) - (iso2 : GroupEquiv ℤGroup H) - → (e : fst G) - → invEq (fst iso) e ≡ 1 - → (hom : GroupEquiv G H) - → abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1) - λ H → JGroupEquiv (λ H iso2 → (e : ℤ) → e ≡ pos 1 → - (hom : GroupEquiv ℤGroup H) → - abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1) - λ e p hom → cong (abs ∘ (fst (fst hom))) p ∙ autoPres1 hom - where - autoPres1 : (e : GroupEquiv ℤGroup ℤGroup) - → abs (fst (fst e) 1) ≡ 1 - autoPres1 e = GroupEquivℤ-pres1 e _ refl --} - --- move to intGroup -allisoPresGen : ∀ {ℓ} (G : Group ℓ) (ϕ : GroupEquiv G ℤGroup) (x : fst G) - → (fst (fst ϕ) x ≡ 1) ⊎ (fst (fst ϕ) x ≡ - 1) - → (ψ : GroupEquiv G ℤGroup) - → (fst (fst ψ) x ≡ 1) ⊎ (fst (fst ψ) x ≡ - 1) -allisoPresGen G (ϕeq , ϕhom) x (inl r) (ψeq , ψhom) = - abs→⊎ _ _ (cong abs (cong (fst ψeq) (sym (retEq ϕeq x) ∙ cong (invEq ϕeq) r)) - ∙ GroupEquivℤ-pres1 (compGroupEquiv (invGroupEquiv (ϕeq , ϕhom)) (ψeq , ψhom)) _ refl) -allisoPresGen G (ϕeq , ϕhom) x (inr r) (ψeq , ψhom) = - abs→⊎ _ _ - (cong abs (cong (fst ψeq) (sym (retEq ϕeq x) ∙ cong (invEq ϕeq) r)) - ∙ cong abs (IsGroupHom.presinv (snd (compGroupEquiv (invGroupEquiv (ϕeq , ϕhom)) (ψeq , ψhom))) 1) - ∙ absLem _ (GroupEquivℤ-pres1 (compGroupEquiv (invGroupEquiv (ϕeq , ϕhom)) (ψeq , ψhom)) _ refl)) - where - absLem : ∀ x → abs x ≡ 1 → abs (pos 0 - x) ≡ 1 - absLem (pos zero) p = ⊥-rec (snotz (sym p)) - absLem (pos (suc zero)) p = refl - absLem (pos (suc (suc n))) p = ⊥-rec (snotz (cong predℕ p)) - absLem (negsuc zero) p = refl - absLem (negsuc (suc n)) p = ⊥-rec (snotz (cong predℕ p)) ⌣equiv→pres1 : ∀ {ℓ} {G H : Type ℓ} → (G ≡ H) → (g₁ : coHom 2 G) (h₁ : coHom 2 H) → (fstEq : (Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] abs (fst (fst ϕ) g₁) ≡ 1)) @@ -678,7 +520,7 @@ isContrcoHomS¹-ish n = isOfHLevelRetractFromIso 0 (mapCompIso characFunSpaceS¹ (isConnectedPath _ (isConnectedKn (2 +ℕ n)) (0ₖ _) (0ₖ _)) ∣ refl ∣ ∣ p ∣)))))) Iso2' : (n : ℕ) → Iso (∥ Σ (coHomS¹-ish n) (fib n) ∥₂) (fib n ∣ (λ _ → 0ₖ _) ∣) -Iso2' n = compIso (setTruncIso (contrFstΣ centerChange)) +Iso2' n = compIso (setTruncIso (Σ-contractFstIso centerChange)) (setTruncIdempotentIso squash₂) where centerChange : isContr (coHomS¹-ish n) @@ -694,7 +536,7 @@ ll5 n p = (setTruncIso (compIso (Σ-cong-iso-snd λ f → (compIso characFunSpaceS¹ (invIso (Σ-cong-iso-fst (iso funExt⁻ funExt (λ _ → refl) λ _ → refl))))) - (compIso ΣΣ (contrFstΣ (isContrSingl _))))) + (compIso ΣΣ (Σ-contractFstIso (isContrSingl _))))) (compIso (setTruncIso (iso funExt⁻ funExt (λ _ → refl) λ _ → refl)) (compIso (setTruncIso (codomainIso (congIso (invIso (Iso-Kn-ΩKn+1 _))))) ((compIso (invIso (fst (coHom≅coHomΩ _ (S₊ _)))) @@ -2165,7 +2007,7 @@ snd (transportCohomIso {A = A} {n = n} {m = m} p) = ≡ subst (λ n → coHom n A) p x +ₕ subst (λ n → coHom n A) p y) (transportRefl (x +ₕ y) ∙ cong₂ _+ₕ_ (sym (transportRefl x)) (sym (transportRefl y))) -α-hopfMap : abs (HopfFunction zero (hopfFun , λ _ → hopfFun (snd (S₊∙ 3)))) ≡ suc zero +α-hopfMap : abs (HopfInvariant zero (hopfFun , λ _ → hopfFun (snd (S₊∙ 3)))) ≡ suc zero α-hopfMap = cong abs (cong (Iso.fun (fst (Hopfβ-Iso zero (hopfFun , refl)))) (transportRefl (⌣-α 0 (hopfFun , refl)))) ∙ ⌣equiv→pres1 (sym CP2≡CP2') GysinS1.e (Hopfα zero (hopfFun , refl)) (l isGenerator≃ℤ-e) @@ -2886,7 +2728,7 @@ H⁴-C* n f g = compGroupIso (GroupEquiv→GroupIso (invGroupEquiv fstIso)) snd fstIso = Ms.d (suc (suc (n +ℕ suc n))) .snd Hopfie : {n : ℕ} → ∥ S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n) ∥₂ → ℤ -Hopfie = sRec isSetℤ (HopfFunction _) +Hopfie = sRec isSetℤ (HopfInvariant _) subber : (n : ℕ) (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → _ subber n f = subst (λ x → coHom x (HopfInvariantPush n (fst f))) @@ -3356,10 +3198,10 @@ eq₁ n f g = ∙ cong ((coHomFun _) (q _ f g)) (cong (λ x → x ⌣ x) (sym (α-∨-lem n f g)) ∙ α∨≡ n f g) ∙ IsGroupHom.pres· (coHomMorph _ (q n f g) .snd) _ _ - ∙ cong₂ _+ₕ_ (hompres· (coHomMorph _ (q n f g)) (βₗ n f g) (X n f g) + ∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (q n f g)) (βₗ n f g) (X n f g) ∙ cong (λ z → (X n f g) ℤ[ coHomGr _ _ ]· z) (q-βₗ n f g)) - (hompres· (coHomMorph _ (q n f g)) (βᵣ n f g) (Y n f g) + (homPresℤ· (coHomMorph _ (q n f g)) (βᵣ n f g) (Y n f g) ∙ cong (λ z → (Y n f g) ℤ[ coHomGr _ _ ]· z) (q-βᵣ n f g)) ∙ sym (distrℤ· (coHomGr _ _) (β+ n f g) (X n f g) (Y n f g)) @@ -3379,8 +3221,8 @@ eq₂ n f g = ∙∙ sym (coHomFun⌣ (jₗ n f g) _ _ _) ∙∙ cong (coHomFun _ (jₗ n f g)) (α∨≡ n f g) ∙∙ IsGroupHom.pres· (snd (coHomMorph _ (jₗ n f g))) _ _ - ∙∙ cong₂ _+ₕ_ (hompres· (coHomMorph _ (jₗ n f g)) (βₗ n f g) (X n f g)) - (hompres· (coHomMorph _ (jₗ n f g)) (βᵣ n f g) (Y n f g) + ∙∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (jₗ n f g)) (βₗ n f g) (X n f g)) + (homPresℤ· (coHomMorph _ (jₗ n f g)) (βᵣ n f g) (Y n f g) ∙ cong (λ z → (Y n f g) ℤ[ coHomGr _ _ ]· z) (jₗ-βᵣ n f g)) ∙∙ cong₂ _+ₕ_ refl (rUnitℤ· (coHomGr _ _) (Y n f g)) @@ -3397,15 +3239,15 @@ eq₃ n f g = ∙∙ sym (coHomFun⌣ (jᵣ n f g) _ _ _) ∙∙ cong (coHomFun _ (jᵣ n f g)) (α∨≡ n f g) ∙∙ IsGroupHom.pres· (snd (coHomMorph _ (jᵣ n f g))) _ _ - ∙∙ cong₂ _+ₕ_ (hompres· (coHomMorph _ (jᵣ n f g)) (βₗ n f g) (X n f g) + ∙∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (jᵣ n f g)) (βₗ n f g) (X n f g) ∙ cong (λ z → (X n f g) ℤ[ coHomGr _ _ ]· z) (jᵣ-βₗ n f g)) - (hompres· (coHomMorph _ (jᵣ n f g)) (βᵣ n f g) (Y n f g)) + (homPresℤ· (coHomMorph _ (jᵣ n f g)) (βᵣ n f g) (Y n f g)) ∙∙ cong₂ _+ₕ_ (rUnitℤ· (coHomGr _ _) (X n f g)) refl ∙∙ ((lUnitₕ _ _) ∙ cong (Y n f g ℤ[ coHomGr _ _ ]·_) (jᵣ-βᵣ n f g)) eq₂-eq₂ : (n : ℕ) (f g : S₊∙ (suc (suc (suc (n +ℕ n)))) →∙ S₊∙ (suc (suc n))) - → HopfFunction n (·Π f g) ≡ HopfFunction n f + HopfFunction n g + → HopfInvariant n (·Π f g) ≡ HopfInvariant n f + HopfInvariant n g eq₂-eq₂ n f g = mainL ∙ sym (cong₂ _+_ (help n f g) (helpg n f g)) @@ -3421,7 +3263,7 @@ eq₂-eq₂ n f g = ≡ (x ℤ[ G n ]· g)) λ g → transportRefl _ ∙ cong (x ℤ[ G m ]·_) (transportRefl _) - mainL : HopfFunction n (·Π f g) ≡ X n f g + Y n f g + mainL : HopfInvariant n (·Π f g) ≡ X n f g + Y n f g mainL = cong (Iso.fun (fst (Hopfβ-Iso n (·Π f g)))) (cong (subst (λ x → coHom x (HopfInvariantPush n (fst (·Π f g)))) λ i₁ → suc (suc (suc (+-suc n n i₁)))) @@ -3430,14 +3272,14 @@ eq₂-eq₂ n f g = (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst (·Π f g)))} _ _ (X n f g + Y n f g) (λ i₁ → suc (suc (suc (+-suc n n i₁)))) (Iso.inv (fst (Hopfβ-Iso n (·Π f g))) 1)) - ∙∙ hompres· (_ , snd (Hopfβ-Iso n (·Π f g))) (Iso.inv (fst (Hopfβ-Iso n (·Π f g))) 1) + ∙∙ homPresℤ· (_ , snd (Hopfβ-Iso n (·Π f g))) (Iso.inv (fst (Hopfβ-Iso n (·Π f g))) 1) (X n f g + Y n f g) ∙∙ cong ((X n f g + Y n f g) ℤ[ ℤ , ℤGroup .snd ]·_) (Iso.rightInv ((fst (Hopfβ-Iso n (·Π f g)))) 1) ∙∙ rUnitℤ·ℤ (X n f g + Y n f g) - help : (n : ℕ) (f g : _) → HopfFunction n f ≡ X n f g + help : (n : ℕ) (f g : _) → HopfInvariant n f ≡ X n f g help n f g = cong (Iso.fun (fst (Hopfβ-Iso n f))) (cong (subst (λ x → coHom x (HopfInvariantPush n (fst f))) (λ i₁ → suc (suc (suc (+-suc n n i₁))))) (eq₂ n f g)) @@ -3446,11 +3288,11 @@ eq₂-eq₂ n f g = (X n f g) ((cong (suc ∘ suc ∘ suc) (+-suc n n))) (Hopfβ n f)) - ∙ hompres· (_ , snd (Hopfβ-Iso n f)) (Hopfβ n f) (X n f g) + ∙ homPresℤ· (_ , snd (Hopfβ-Iso n f)) (Hopfβ n f) (X n f g) ∙ cong (X n f g ℤ[ ℤ , ℤGroup .snd ]·_) (Hopfβ↦1 n f) ∙ rUnitℤ·ℤ (X n f g) - helpg : (n : ℕ) (f g : _) → HopfFunction n g ≡ Y n f g + helpg : (n : ℕ) (f g : _) → HopfInvariant n g ≡ Y n f g helpg n f g = cong (Iso.fun (fst (Hopfβ-Iso n g))) (cong (subst (λ x → coHom x (HopfInvariantPush n (fst g))) @@ -3461,6 +3303,6 @@ eq₂-eq₂ n f g = (Y n f g) ((cong (suc ∘ suc ∘ suc) (+-suc n n))) (Hopfβ n g)) - ∙∙ hompres· (_ , snd (Hopfβ-Iso n g)) (Hopfβ n g) (Y n f g) + ∙∙ homPresℤ· (_ , snd (Hopfβ-Iso n g)) (Hopfβ n g) (Y n f g) ∙∙ cong (Y n f g ℤ[ ℤ , ℤGroup .snd ]·_) (Hopfβ↦1 n g) ∙∙ rUnitℤ·ℤ (Y n f g) From 2a298247ee5775c5acb3d7de188e60f6b18e418e Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Fri, 8 Oct 2021 16:17:13 +0200 Subject: [PATCH 70/89] st uff --- Cubical/Foundations/Isomorphism.agda | 13 +- Cubical/HITs/Join/Properties.agda | 48 + Cubical/HITs/SetTruncation/Properties.agda | 3 + Cubical/HITs/Sn/Properties.agda | 23 + Cubical/Homotopy/HomotopyGroup.agda | 651 +++++ Cubical/Homotopy/HopfInvariant/Base.agda | 334 +++ Cubical/Homotopy/HopfInvariant/HopfMap.agda | 1447 ++++++++++ Cubical/Homotopy/Loopspace.agda | 4 - Cubical/ZCohomology/Gysin.agda | 2824 ++----------------- 9 files changed, 2819 insertions(+), 2528 deletions(-) create mode 100644 Cubical/Homotopy/HomotopyGroup.agda create mode 100644 Cubical/Homotopy/HopfInvariant/Base.agda create mode 100644 Cubical/Homotopy/HopfInvariant/HopfMap.agda diff --git a/Cubical/Foundations/Isomorphism.agda b/Cubical/Foundations/Isomorphism.agda index 1746086b1..a6a9d874f 100644 --- a/Cubical/Foundations/Isomorphism.agda +++ b/Cubical/Foundations/Isomorphism.agda @@ -171,10 +171,15 @@ A ∎Iso = idIso {A = A} infixr 0 _Iso⟨_⟩_ infix 1 _∎Iso +codomainIsoDep : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : A → Type ℓ'} {C : A → Type ℓ''} + → ((a : A) → Iso (B a) (C a)) + → Iso ((a : A) → B a) ((a : A) → C a) +Iso.fun (codomainIsoDep is) f a = Iso.fun (is a) (f a) +Iso.inv (codomainIsoDep is) f a = Iso.inv (is a) (f a) +Iso.rightInv (codomainIsoDep is) f = funExt λ a → Iso.rightInv (is a) (f a) +Iso.leftInv (codomainIsoDep is) f = funExt λ a → Iso.leftInv (is a) (f a) + codomainIso : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} → Iso B C → Iso (A → B) (A → C) -Iso.fun (codomainIso is) f a = Iso.fun is (f a) -Iso.inv (codomainIso is) f a = Iso.inv is (f a) -Iso.rightInv (codomainIso is) f = funExt λ a → Iso.rightInv is (f a) -Iso.leftInv (codomainIso is) f = funExt λ a → Iso.leftInv is (f a) +codomainIso z = codomainIsoDep λ _ → z diff --git a/Cubical/HITs/Join/Properties.agda b/Cubical/HITs/Join/Properties.agda index fe0efa926..24a6fe649 100644 --- a/Cubical/HITs/Join/Properties.agda +++ b/Cubical/HITs/Join/Properties.agda @@ -18,6 +18,7 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Function open import Cubical.Data.Sigma renaming (fst to proj₁; snd to proj₂) @@ -28,6 +29,19 @@ private variable ℓ ℓ' : Level +-- Characterisation of function type join A B → C +IsoFunSpaceJoin : ∀ {ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + → Iso (join A B → C) + (Σ[ f ∈ (A → C) ] Σ[ g ∈ (B → C) ] + ((a : A) (b : B) → f a ≡ g b)) +Iso.fun IsoFunSpaceJoin f = (f ∘ inl) , ((f ∘ inr) , (λ a b → cong f (push a b))) +Iso.inv IsoFunSpaceJoin (f , g , p) (inl x) = f x +Iso.inv IsoFunSpaceJoin (f , g , p) (inr x) = g x +Iso.inv IsoFunSpaceJoin (f , g , p) (push a b i) = p a b i +Iso.rightInv IsoFunSpaceJoin (f , g , p) = refl +Iso.leftInv IsoFunSpaceJoin f = + funExt λ { (inl x) → refl ; (inr x) → refl ; (push a b i) → refl} + -- Alternative definition of the join using a pushout joinPushout : (A : Type ℓ) → (B : Type ℓ') → Type (ℓ-max ℓ ℓ') joinPushout A B = Pushout {A = A × B} proj₁ proj₂ @@ -417,3 +431,37 @@ joinAssocDirect {A = A} {B} {C} = ; (l = i1) → push (push a b i) c j }) (push (push a b i) c j)) + +-- commutativity +join-commFun : ∀ {ℓ'} {A : Type ℓ} {B : Type ℓ'} → join A B → join B A +join-commFun (inl x) = inr x +join-commFun (inr x) = inl x +join-commFun (push a b i) = push b a (~ i) + +join-commFun² : ∀ {ℓ'} {A : Type ℓ} {B : Type ℓ'} (x : join A B) + → join-commFun (join-commFun x) ≡ x +join-commFun² (inl x) = refl +join-commFun² (inr x) = refl +join-commFun² (push a b i) = refl + +join-comm : ∀ {ℓ'} {A : Type ℓ} {B : Type ℓ'} → Iso (join A B) (join B A) +Iso.fun join-comm = join-commFun +Iso.inv join-comm = join-commFun +Iso.rightInv join-comm = join-commFun² +Iso.leftInv join-comm = join-commFun² + +-- Applying Isos to joins (more efficient than transports) +Iso→joinIso : ∀ {ℓ' ℓ'' ℓ'''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} + → Iso A C → Iso B D → Iso (join A B) (join C D) +Iso.fun (Iso→joinIso is1 is2) (inl x) = inl (Iso.fun is1 x) +Iso.fun (Iso→joinIso is1 is2) (inr x) = inr (Iso.fun is2 x) +Iso.fun (Iso→joinIso is1 is2) (push a b i) = push (Iso.fun is1 a) (Iso.fun is2 b) i +Iso.inv (Iso→joinIso is1 is2) (inl x) = inl (Iso.inv is1 x) +Iso.inv (Iso→joinIso is1 is2) (inr x) = inr (Iso.inv is2 x) +Iso.inv (Iso→joinIso is1 is2) (push a b i) = push (Iso.inv is1 a) (Iso.inv is2 b) i +Iso.rightInv (Iso→joinIso is1 is2) (inl x) i = inl (Iso.rightInv is1 x i) +Iso.rightInv (Iso→joinIso is1 is2) (inr x) i = inr (Iso.rightInv is2 x i) +Iso.rightInv (Iso→joinIso is1 is2) (push a b j) i = push (Iso.rightInv is1 a i) (Iso.rightInv is2 b i) j +Iso.leftInv (Iso→joinIso is1 is2) (inl x) i = inl (Iso.leftInv is1 x i) +Iso.leftInv (Iso→joinIso is1 is2) (inr x) i = inr (Iso.leftInv is2 x i) +Iso.leftInv (Iso→joinIso is1 is2) (push a b i) j = push (Iso.leftInv is1 a j) (Iso.leftInv is2 b j) i diff --git a/Cubical/HITs/SetTruncation/Properties.agda b/Cubical/HITs/SetTruncation/Properties.agda index e48526721..7066a1aa4 100644 --- a/Cubical/HITs/SetTruncation/Properties.agda +++ b/Cubical/HITs/SetTruncation/Properties.agda @@ -26,6 +26,9 @@ private ℓ ℓ' ℓ'' : Level A B C D : Type ℓ +isSetPathImplicit : {x y : ∥ A ∥₂} → isSet (x ≡ y) +isSetPathImplicit = isOfHLevelPath 2 squash₂ _ _ + rec : isSet B → (A → B) → ∥ A ∥₂ → B rec Bset f ∣ x ∣₂ = f x rec Bset f (squash₂ x y p q i j) = diff --git a/Cubical/HITs/Sn/Properties.agda b/Cubical/HITs/Sn/Properties.agda index 4dc17ee2e..6af47f72f 100644 --- a/Cubical/HITs/Sn/Properties.agda +++ b/Cubical/HITs/Sn/Properties.agda @@ -19,11 +19,17 @@ open import Cubical.HITs.Susp open import Cubical.HITs.Truncation open import Cubical.Homotopy.Loopspace open import Cubical.Homotopy.Connected +open import Cubical.HITs.Join +open import Cubical.Data.Bool private variable ℓ : Level +IsoSucSphereSusp : (n : ℕ) → Iso (S₊ (suc n)) (Susp (S₊ n)) +IsoSucSphereSusp zero = S¹IsoSuspBool +IsoSucSphereSusp (suc n) = idIso + -- Elimination principles for spheres sphereElim : (n : ℕ) {A : (S₊ (suc n)) → Type ℓ} → ((x : S₊ (suc n)) → isOfHLevel (suc n) (A x)) → A (ptSn (suc n)) @@ -308,3 +314,20 @@ isConnectedPathSⁿ n x y = (pathIdTruncSⁿretract n x y) ((isContr→isProp (sphereConnected (suc n)) ∣ x ∣ ∣ y ∣) , isProp→isSet (isContr→isProp (sphereConnected (suc n))) _ _ _) + + +-- Equivalence Sⁿ*Sᵐ≃Sⁿ⁺ᵐ⁺¹ +IsoSphereJoin : (n m : ℕ) → Iso (join (S₊ n) (S₊ m)) (S₊ (suc (n + m))) +IsoSphereJoin zero m = + compIso join-comm + (compIso (invIso Susp-iso-joinBool) + (invIso (IsoSucSphereSusp m))) +IsoSphereJoin (suc n) m = + compIso (Iso→joinIso + (subst (λ x → Iso (S₊ (suc x)) (join (S₊ n) Bool)) (+-comm n 0) (invIso (IsoSphereJoin n 0))) + idIso) + (compIso (equivToIso joinAssocDirect) + (compIso (Iso→joinIso idIso (compIso join-comm (compIso (invIso Susp-iso-joinBool) (invIso (IsoSucSphereSusp m))))) + (compIso + (IsoSphereJoin n (suc m)) + (pathToIso λ i → S₊ (suc (+-suc n m i)))))) diff --git a/Cubical/Homotopy/HomotopyGroup.agda b/Cubical/Homotopy/HomotopyGroup.agda new file mode 100644 index 000000000..ba0cb6b8c --- /dev/null +++ b/Cubical/Homotopy/HomotopyGroup.agda @@ -0,0 +1,651 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +{- +This file contains +1. The definition of πₙ as a truncated loop space +2. The definition of πₙ as a truncated function space (Sⁿ →∙ A) +3. A structure preserving equivalence Ωⁿ A ≃ (Sⁿ →∙ A) +4. A proof that the two constructions of homotopy groups are isomorphic +-} +module Cubical.Homotopy.HomotopyGroup where + +open import Cubical.Homotopy.Loopspace + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Path +open import Cubical.Foundations.Transport + +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim + ; elim2 to sElim2 ; elim3 to sElim3 ; map to sMap) +open import Cubical.HITs.Sn +open import Cubical.Data.Bool +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 +open import Cubical.Data.Sigma +open import Cubical.Data.Nat + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid + +open Iso +open IsGroup +open IsSemigroup +open IsMonoid +open GroupStr + + +{- homotopy Group -} +π : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Type ℓ +π n A = ∥ typ ((Ω^ n) A) ∥₂ + +π' : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Type ℓ +π' n A = ∥ S₊∙ n →∙ A ∥₂ + + +{- π as a group -} +1π : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π n A +1π zero {A = A} = ∣ pt A ∣₂ +1π (suc n) = ∣ refl ∣₂ + +·π : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π (suc n) A → π (suc n) A → π (suc n) A +·π n = sRec2 squash₂ λ p q → ∣ p ∙ q ∣₂ + +-π : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π (suc n) A → π (suc n) A +-π n = sMap sym + +π-rUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π (suc n) A) + → (·π n x (1π (suc n))) ≡ x +π-rUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ rUnit p (~ i) ∣₂ + +π-lUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π (suc n) A) + → (·π n (1π (suc n)) x) ≡ x +π-lUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ lUnit p (~ i) ∣₂ + +π-rCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π (suc n) A) + → (·π n x (-π n x)) ≡ 1π (suc n) +π-rCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ rCancel p i ∣₂ + +π-lCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π (suc n) A) + → (·π n (-π n x) x) ≡ 1π (suc n) +π-lCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ lCancel p i ∣₂ + +π-assoc : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y z : π (suc n) A) + → ·π n x (·π n y z) ≡ ·π n (·π n x y) z +π-assoc n = sElim3 (λ _ _ _ → isSetPathImplicit) λ p q r i → ∣ ∙assoc p q r i ∣₂ + +π-comm : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y : π (suc (suc n)) A) + → ·π (suc n) x y ≡ ·π (suc n) y x +π-comm n = sElim2 (λ _ _ → isSetPathImplicit) λ p q i → ∣ EH n p q i ∣₂ + +πGr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Group ℓ +fst (πGr n A) = π (suc n) A +1g (snd (πGr n A)) = 1π (suc n) +GroupStr._·_ (snd (πGr n A)) = ·π n +inv (snd (πGr n A)) = -π n +is-set (isSemigroup (isMonoid (isGroup (snd (πGr n A))))) = squash₂ +assoc (isSemigroup (isMonoid (isGroup (snd (πGr n A))))) = π-assoc n +identity (isMonoid (isGroup (snd (πGr n A)))) x = (π-rUnit n x) , (π-lUnit n x) +inverse (isGroup (snd (πGr n A))) x = (π-rCancel n x) , (π-lCancel n x) + +-- We now define the group structure on π'. +-- We first define the corresponding structure on the untruncated +-- (S₊∙ n →∙ A). We first give multiplication and inversion and use this +-- to give a structure preserving equivalence (S₊∙ n →∙ A) ≃ Ωⁿ A. + +∙Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (S₊∙ n →∙ A) + → (S₊∙ n →∙ A) + → (S₊∙ n →∙ A) +∙Π {A = A} {n = zero} p q = (λ _ → pt A) , refl +fst (∙Π {A = A} {n = suc zero} (f , p) (g , q)) base = pt A +fst (∙Π {A = A} {n = suc zero} (f , p) (g , q)) (loop j) = + ((sym p ∙∙ cong f loop ∙∙ p) ∙ (sym q ∙∙ cong g loop ∙∙ q)) j +snd (∙Π {A = A} {n = suc zero} (f , p) (g , q)) = refl +fst (∙Π {A = A} {n = suc (suc n)} (f , p) (g , q)) north = pt A +fst (∙Π {A = A} {n = suc (suc n)} (f , p) (g , q)) south = pt A +fst (∙Π {A = A} {n = suc (suc n)} (f , p) (g , q)) (merid a j) = + ((sym p ∙∙ cong f (merid a ∙ sym (merid (ptSn (suc n)))) ∙∙ p) + ∙ (sym q ∙∙ cong g (merid a ∙ sym (merid (ptSn (suc n)))) ∙∙ q)) j +snd (∙Π {A = A} {n = suc (suc n)} (f , p) (g , q)) = refl + +-Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (S₊∙ n →∙ A) + → (S₊∙ n →∙ A) +-Π {n = zero} f = f +fst (-Π {A = A} {n = suc zero} f) base = fst f base +fst (-Π {A = A} {n = suc zero} f) (loop j) = fst f (loop (~ j)) +snd (-Π {A = A} {n = suc zero} f) = snd f +fst (-Π {A = A} {n = suc (suc n)} f) north = fst f north +fst (-Π {A = A} {n = suc (suc n)} f) south = fst f north +fst (-Π {A = A} {n = suc (suc n)} f) (merid a j) = + fst f ((merid a ∙ sym (merid (ptSn _))) (~ j)) +snd (-Π {A = A} {n = suc (suc n)} f) = snd f + +-- +module _ {ℓ : Level} {A : Pointed ℓ} where +flipΩPath : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → ((Ω^ (suc n)) A) ≡ (Ω^ n) (Ω A) +flipΩPath {A = A} zero = refl +flipΩPath {A = A} (suc n) = cong Ω (flipΩPath {A = A} n) + +flipΩIso : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → Iso (fst ((Ω^ (suc n)) A)) (fst ((Ω^ n) (Ω A))) +flipΩIso {A = A} n = pathToIso (cong fst (flipΩPath n)) + +flipΩIsopres· : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (f g : fst ((Ω^ (suc n)) (Ω A))) + → Iso.inv (flipΩIso (suc n)) (f ∙ g) + ≡ (Iso.inv (flipΩIso (suc n)) f) ∙ (Iso.inv (flipΩIso (suc n)) g) +flipΩIsopres· {A = A} n f g i = + transp (λ j → flipΩPath {A = A} n (~ i ∧ ~ j) .snd ≡ flipΩPath n (~ i ∧ ~ j) .snd) i + (transp (λ j → flipΩPath {A = A} n (~ i ∨ ~ j) .snd + ≡ flipΩPath n (~ i ∨ ~ j) .snd) (~ i) f + ∙ transp (λ j → flipΩPath {A = A} n (~ i ∨ ~ j) .snd + ≡ flipΩPath n (~ i ∨ ~ j) .snd) (~ i) g) + +module _ {ℓ : Level} {A : Pointed ℓ} where + suspension→ : (n : ℕ) → (S₊∙ (suc n) →∙ Ω A) → (S₊∙ (suc (suc n)) →∙ A) + fst (suspension→ n f) north = pt A + fst (suspension→ n f) south = pt A + fst (suspension→ n f) (merid a i₁) = fst f a i₁ + snd (suspension→ n f) = refl + + suspension-pres∙ : + (n : ℕ) (f g : (S₊∙ (suc n) →∙ Ω A)) + → suspension→ n (∙Π f g) ≡ ∙Π (suspension→ n f) (suspension→ n g) + suspension-pres∙ n f g = ΣPathP ((funExt main) , refl) + where + J2 : ∀ {ℓ ℓ'} {A : Type ℓ} {x : A} + → (P : (p q : x ≡ x) → (refl ≡ p) → (refl ≡ q) → Type ℓ') + → P refl refl refl refl + → (p q : x ≡ x) (P' : _) (Q : _) → P p q P' Q + J2 P b p q = + J (λ p P' → (Q₁ : refl ≡ q) → P p q P' Q₁) + (J (λ q Q → P refl q refl Q) b) + + J4 : ∀ {ℓ ℓ'} {A : Type ℓ} {x : A} + → (P : (p q r s : x ≡ x) → (refl ≡ p) → (p ≡ q) → (refl ≡ r) → (r ≡ s) → Type ℓ') + → P refl refl refl refl refl refl refl refl + → (p q r s : x ≡ x) (P' : _) (Q : _) (R : _) (S : _) → P p q r s P' Q R S + J4 P b p q r s = + J (λ p P' → (Q₁ : p ≡ q) (R : refl ≡ r) (S₁ : r ≡ s) → P p q r s P' Q₁ R S₁) + (J (λ q Q₁ → (R : refl ≡ r) (S₁ : r ≡ s) → P refl q r s refl Q₁ R S₁) + (J (λ r R → (S₁ : r ≡ s) → P refl refl r s refl refl R S₁) + (J (λ s S₁ → P refl refl refl s refl refl refl S₁) + b))) + + looper-help : (n : ℕ) (f g : (S₊∙ (suc n) →∙ Ω A)) (a : _) + → fst (∙Π f g) a ≡ fst f a ∙ fst g a + looper-help zero f g base = rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g)) + looper-help zero f g (loop i) k = + lazy-fill _ _ (sym (snd f)) (sym (snd g)) (cong (fst f) loop) (cong (fst g) loop) k i + where + lazy-fill : ∀ {ℓ} {A : Type ℓ} {x : A} (p q : x ≡ x) + (prefl : refl ≡ p) (qrefl : refl ≡ q) (fl : p ≡ p) (fr : q ≡ q) + → PathP (λ i → (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i + ≡ (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i) + ((prefl ∙∙ fl ∙∙ sym prefl) ∙ (qrefl ∙∙ fr ∙∙ sym qrefl)) + (cong₂ _∙_ fl fr) + lazy-fill {x = x} = + J2 (λ p q prefl qrefl → (fl : p ≡ p) (fr : q ≡ q) + → PathP (λ i → (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i + ≡ (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i) + ((prefl ∙∙ fl ∙∙ sym prefl) ∙ (qrefl ∙∙ fr ∙∙ sym qrefl)) + (cong₂ _∙_ fl fr)) + λ fl fr → cong₂ _∙_ (sym (rUnit fl)) (sym (rUnit fr)) + ◁ flipSquare (sym (rUnit (rUnit refl)) + ◁ (flipSquare ((λ j → cong (λ x → rUnit x j) fl + ∙ cong (λ x → lUnit x j) fr) + ▷ sym (cong₂Funct _∙_ fl fr)) + ▷ (rUnit (rUnit refl)))) + looper-help (suc n) f g north = rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g)) + looper-help (suc n) f g south = (rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) + ∙ cong₂ _∙_ (cong (fst f) (merid (ptSn _))) + (cong (fst g) (merid (ptSn _))) + looper-help (suc n) f g (merid a j) i = help i j + where + help : PathP (λ i → (rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) i + ≡ ((rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) + ∙ cong₂ _∙_ (cong (fst f) (merid (ptSn _))) + (cong (fst g) (merid (ptSn _)))) i) + (cong (fst (∙Π f g)) (merid a)) + (cong₂ _∙_ (cong (fst f) (merid a)) (cong (fst g) (merid a))) + help = (cong₂ _∙_ (cong (sym (snd f) ∙∙_∙∙ snd f) + (cong-∙ (fst f) (merid a) (sym (merid (ptSn _))))) + ((cong (sym (snd g) ∙∙_∙∙ snd g) + (cong-∙ (fst g) (merid a) (sym (merid (ptSn _)))))) + ◁ lazy-help _ _ _ _ (sym (snd f)) (cong (fst f) (merid (ptSn _))) + (sym (snd g)) (cong (fst g) (merid (ptSn _))) + (cong (fst f) (merid a)) (cong (fst g) (merid a))) + where + lazy-help : ∀ {ℓ} {A : Type ℓ} {x : A} (p-north p-south q-north q-south : x ≡ x) + (prefl : refl ≡ p-north) (p-merid : p-north ≡ p-south) + (qrefl : refl ≡ q-north) (q-merid : q-north ≡ q-south) + (fl : p-north ≡ p-south) (fr : q-north ≡ q-south) + → PathP (λ i → (rUnit refl ∙ cong₂ _∙_ prefl qrefl) i + ≡ ((rUnit refl ∙ cong₂ _∙_ prefl qrefl) + ∙ cong₂ _∙_ p-merid q-merid) i) + ((prefl ∙∙ fl ∙ sym p-merid ∙∙ sym prefl) + ∙ (qrefl ∙∙ fr ∙ sym q-merid ∙∙ sym qrefl)) + (cong₂ _∙_ fl fr) + lazy-help {x = x} = + J4 (λ p-north p-south q-north q-south prefl p-merid qrefl q-merid + → (fl : p-north ≡ p-south) (fr : q-north ≡ q-south) → + PathP (λ i₁ → (rUnit refl ∙ cong₂ _∙_ prefl qrefl) i₁ ≡ + ((rUnit refl ∙ cong₂ _∙_ prefl qrefl) ∙ cong₂ _∙_ p-merid q-merid) i₁) + ((prefl ∙∙ fl ∙ sym p-merid ∙∙ sym prefl) ∙ + (qrefl ∙∙ fr ∙ sym q-merid ∙∙ sym qrefl)) + (cong₂ _∙_ fl fr)) + λ fl fr → (cong₂ _∙_ (sym (rUnit (fl ∙ refl)) ∙ sym (rUnit fl)) + (sym (rUnit (fr ∙ refl)) ∙ sym (rUnit fr)) + ◁ flipSquare ((sym (rUnit _) + ◁ flipSquare + ((λ j → cong (λ x → rUnit x j) fl + ∙ cong (λ x → lUnit x j) fr) + ▷ sym (cong₂Funct _∙_ fl fr))) + ▷ (rUnit _ ∙ rUnit _))) + main : (x : fst (S₊∙ (suc (suc n)))) → + fst (suspension→ n (∙Π f g)) x ≡ + fst (∙Π (suspension→ n f) (suspension→ n g)) x + main north = refl + main south = refl + main (merid a i₁) j = help j i₁ + where + help : fst (∙Π f g) a ≡ cong (fst (∙Π (suspension→ n f) (suspension→ n g))) (merid a) + help = looper-help n f g a + ∙ cong₂ _∙_ (sym (cong-∙ (fst (suspension→ n f)) (merid a) (sym (merid (ptSn _))) + ∙∙ cong (fst f a ∙_) (cong sym (snd f)) + ∙∙ sym (rUnit _))) + (sym (cong-∙ (fst (suspension→ n g)) (merid a) (sym (merid (ptSn _))) + ∙∙ cong (fst g a ∙_) (cong sym (snd g)) + ∙∙ sym (rUnit _))) + ∙ λ i → rUnit (cong (fst (suspension→ n f)) (merid a ∙ sym (merid (ptSn _)))) i + ∙ rUnit (cong (fst (suspension→ n g)) (merid a ∙ sym (merid (ptSn _)))) i + + suspension←filler : (n : ℕ) → (S₊∙ (suc (suc n)) →∙ A) → I → I → I → fst A + suspension←filler n f i j k = + hfill (λ k → λ { (i = i0) → doubleCompPath-filler + (sym (snd f)) + (cong (fst f) (merid (ptSn _) ∙ sym (merid (ptSn _)))) + (snd f) k j + ; (i = i1) → snd f k + ; (j = i0) → snd f k + ; (j = i1) → snd f k}) + (inS ((fst f) (rCancel' (merid (ptSn _)) i j))) + k + + suspension← : (n : ℕ) → (S₊∙ (suc (suc n)) →∙ A) → (S₊∙ (suc n) →∙ Ω A) + fst (suspension← n f) a = + sym (snd f) ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f + snd (suspension← n f) i j = suspension←filler n f i j i1 + + suspensionIso : (n : ℕ) → Iso (S₊∙ (suc n) →∙ Ω A) (S₊∙ (suc (suc n)) →∙ A) + Iso.fun (suspensionIso n) f = suspension→ n f + Iso.inv (suspensionIso n) f = suspension← n f + Iso.rightInv (suspensionIso n) (f , p) = + ΣPathP (funExt help , λ i j → p (~ i ∨ j)) + where + help : (x : _) → fst (suspension→ n (suspension← n (f , p))) x ≡ f x + help north = sym p + help south = sym p ∙ cong f (merid (ptSn _)) + help (merid a j) i = + hcomp (λ k → λ { (i = i1) → f (merid a j) + ; (j = i0) → p (~ i ∧ k) + ; (j = i1) → compPath-filler' (sym p) (cong f (merid (ptSn _))) k i}) + (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) + Iso.leftInv (suspensionIso n) f = + ΣPathP (funExt lem , PathP-filler) + where + + side₁ : (x : S₊ (suc n)) → I → I → I → fst A + side₁ x i j k = + hfill (λ k → λ { (i = i0) → suspension→ n f .fst + (compPath-filler' (merid x) (sym (merid (ptSn _))) k j) + ; (i = i1) → snd A + ; (j = i0) → fst f x (i ∨ ~ k) + ; (j = i1) → snd A}) + (inS (snd f i (~ j))) + k + + side₂ : (x : S₊ (suc n)) → I → I → I → fst A + side₂ x i j k = + hfill (λ k → λ { (i = i0) → doubleCompPath-filler + refl + (cong (suspension→ n f .fst) + (merid x ∙ sym (merid (ptSn _)))) + refl k j + ; (i = i1) → fst f x (j ∨ ~ k) + ; (j = i0) → fst f x (i ∧ ~ k) + ; (j = i1) → snd A}) + (inS (side₁ x i j i1)) + k + + compPath-filler'' : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) + → I → I → I → A + compPath-filler'' {z = z} p q k j i = + hfill (λ k → λ { (i = i0) → p (~ j) + ; (i = i1) → q k + ; (j = i0) → q (i ∧ k) }) + (inS (p (i ∨ ~ j))) + k + + lem : (x : _) → fst (suspension← n (suspension→ n f)) x ≡ fst f x + lem x i j = side₂ x i j i1 + + sideCube : Cube (λ j k → snd f j (~ k)) + (λ j k → fst (suspension→ n f) (rCancel' (merid (ptSn _)) j k)) + (λ r k → suspension→ n f .fst + (compPath-filler' (merid (ptSn _)) (sym (merid (ptSn _))) r k)) + (λ _ _ → pt A) + (λ r j → snd f j (~ r)) λ _ _ → pt A + sideCube r j k = + hcomp (λ i → λ {(r = i0) → snd f j (~ k ∨ ~ i) + ; (r = i1) → fst (suspension→ n f) (rCancel-filler' (merid (ptSn _)) (~ i) j k) + ; (j = i0) → suspension→ n f .fst + (compPath-filler'' (merid (ptSn _)) (sym (merid (ptSn _))) i r k) + ; (j = i1) → snd f (~ r) (~ i ∧ k) + ; (k = i0) → snd f j (~ r) + ; (k = i1) → snd f (~ r ∧ j) (~ i)}) + (hcomp (λ i → λ {(r = i0) → pt A + ; (r = i1) → snd f (~ i) k + ; (j = i0) → snd f (~ i) (k ∨ ~ r) + ; (j = i1) → snd f (~ r ∨ ~ i) k + ; (k = i0) → snd f (j ∨ ~ i) (~ r) + ; (k = i1) → pt A}) + (pt A)) + + PathP-filler : PathP _ _ _ + PathP-filler i j k = + hcomp (λ r → λ {(i = i0) → suspension←filler n (suspension→ n f) j k r + ; (i = i1) → snd f j (k ∨ ~ r) + ; (j = i0) → side₂ (ptSn _) i k r + ; (j = i1) → snd A + ; (k = i0) → snd f j (i ∧ ~ r) + ; (k = i1) → snd A}) + (hcomp ((λ r → λ {(i = i0) → sideCube r j k + ; (i = i1) → pt A + ; (j = i0) → side₁ (ptSn _) i k r + ; (j = i1) → pt A + ; (k = i0) → snd f j (i ∨ ~ r) + ; (k = i1) → pt A})) + (snd f (i ∨ j) (~ k))) + + suspension←-pres∙ : (n : ℕ) → (f g : S₊∙ (suc (suc n)) →∙ A) + → suspension← n (∙Π f g) + ≡ ∙Π (suspension← n f) (suspension← n g) + suspension←-pres∙ n f g = + sym (Iso.leftInv (suspensionIso n) (suspension← n (∙Π f g))) + ∙∙ cong (suspension← n) + (Iso.rightInv (suspensionIso n) (∙Π f g) + ∙∙ cong₂ ∙Π (sym (Iso.rightInv (suspensionIso n) f)) + (sym (Iso.rightInv (suspensionIso n) g)) + ∙∙ sym (suspension-pres∙ _ (suspension← n f) (suspension← n g))) + ∙∙ (Iso.leftInv (suspensionIso n) (∙Π (suspension← n f) (suspension← n g))) + +SphereMap→Ω : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (S₊∙ (suc n) →∙ A) → (fst ((Ω^ (suc n)) A)) +SphereMap→Ω zero (f , p) = sym p ∙∙ cong f loop ∙∙ p +SphereMap→Ω {A = A} (suc n) (f , p) = + Iso.inv (flipΩIso _) (SphereMap→Ω {A = Ω A} n (Iso.inv (suspensionIso _) (f , p))) + +SphereMap→Ωpres∙ : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (p q : S₊∙ (suc n) →∙ A) + → SphereMap→Ω n (∙Π p q) ≡ SphereMap→Ω n p ∙ SphereMap→Ω n q +SphereMap→Ωpres∙ zero f g = sym (rUnit _) +SphereMap→Ωpres∙ (suc n) f g = + cong (Iso.inv (flipΩIso (suc n))) + (cong (SphereMap→Ω n) (suspension←-pres∙ n f g) + ∙ SphereMap→Ωpres∙ n (suspension← n f) ((suspension← n g))) + ∙ flipΩIsopres· n _ _ + +Ω→SphereMap : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) → (fst ((Ω^ (suc n)) A)) → (S₊∙ (suc n) →∙ A) +fst (Ω→SphereMap {A = A} zero f) base = pt A +fst (Ω→SphereMap {A = A} zero f) (loop i₁) = f i₁ +snd (Ω→SphereMap {A = A} zero f) = refl +Ω→SphereMap {A = A} (suc n) f = Iso.fun (suspensionIso _) (Ω→SphereMap n (Iso.fun (flipΩIso _) f)) + +Ω→SphereMap→Ω : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (x : fst ((Ω^ (suc n)) A)) + → SphereMap→Ω n (Ω→SphereMap n x) ≡ x +Ω→SphereMap→Ω zero x = sym (rUnit _) +Ω→SphereMap→Ω (suc n) x = + cong (Iso.inv (flipΩIso (suc n)) ∘ SphereMap→Ω n) + (Iso.leftInv (suspensionIso _) + (Ω→SphereMap n (Iso.fun (flipΩIso (suc n)) x))) + ∙∙ cong (Iso.inv (flipΩIso (suc n))) + (Ω→SphereMap→Ω n (Iso.fun (flipΩIso (suc n)) x)) + ∙∙ Iso.leftInv (flipΩIso (suc n)) x + +SphereMap→Ω→SphereMap : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (x : (S₊∙ (suc n) →∙ A)) + → Ω→SphereMap n (SphereMap→Ω n x) ≡ x +SphereMap→Ω→SphereMap zero (f , p) = ΣPathP (funExt fstp , λ i j → p (~ i ∨ j)) + where + fstp : (x : _) → _ ≡ f x + fstp base = sym p + fstp (loop i) j = doubleCompPath-filler (sym p) (cong f loop) p (~ j) i +SphereMap→Ω→SphereMap (suc n) f = + cong (Iso.fun (suspensionIso n) ∘ Ω→SphereMap n) (Iso.rightInv (flipΩIso (suc n)) + (SphereMap→Ω n (Iso.inv (suspensionIso n) f))) + ∙∙ cong (Iso.fun (suspensionIso n)) (SphereMap→Ω→SphereMap n (Iso.inv (suspensionIso n) f)) + ∙∙ Iso.rightInv (suspensionIso _) _ + + +-- We finally arrive at the main result +IsoSphereMapΩ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (S₊∙ n →∙ A) (fst ((Ω^ n) A)) +Iso.fun (IsoSphereMapΩ {A = A} zero) (f , p) = f false +fst (Iso.inv (IsoSphereMapΩ {A = A} zero) a) false = a +fst (Iso.inv (IsoSphereMapΩ {A = A} zero) a) true = pt A +snd (Iso.inv (IsoSphereMapΩ {A = A} zero) a) = refl +Iso.rightInv (IsoSphereMapΩ {A = A} zero) a = refl +Iso.leftInv (IsoSphereMapΩ {A = A} zero) (f , p) = + ΣPathP ((funExt (λ { false → refl ; true → sym p})) , λ i j → p (~ i ∨ j)) +Iso.fun (IsoSphereMapΩ {A = A} (suc n)) = SphereMap→Ω n +Iso.inv (IsoSphereMapΩ {A = A} (suc n)) = Ω→SphereMap n +Iso.rightInv (IsoSphereMapΩ {A = A} (suc n)) = Ω→SphereMap→Ω n +Iso.leftInv (IsoSphereMapΩ {A = A} (suc n)) = SphereMap→Ω→SphereMap n + +-- The iso is structure preserving +IsoSphereMapΩ-pres∙Π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f g : S₊∙ (suc n) →∙ A) + → SphereMap→Ω n (∙Π f g) ≡ SphereMap→Ω n f ∙ SphereMap→Ω n g +IsoSphereMapΩ-pres∙Π zero f g = sym (rUnit _) +IsoSphereMapΩ-pres∙Π (suc n) f g = + cong (Iso.inv (flipΩIso (suc n))) + (cong (SphereMap→Ω n) (suspension←-pres∙ n f g) + ∙ SphereMap→Ωpres∙ n (suspension← n f) (suspension← n g)) + ∙ (flipΩIsopres· n (SphereMap→Ω n (suspension← n f)) (SphereMap→Ω n (suspension← n g))) + +-- It is useful to define the ``Group Structure'' on (S₊∙ n →∙ A) before doing it on π' +-- These will be the equivalents of the usual groupoid laws on Ω A. +1Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → (S₊∙ n →∙ A) +fst (1Π {A = A}) _ = pt A +snd (1Π {A = A}) = refl + +∙Π-rUnit : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ∙Π f 1Π ≡ f +fst (∙Π-rUnit {A = A} {n = zero} f i) base = snd f (~ i) +fst (∙Π-rUnit {A = A} {n = zero} f i) (loop j) = help i j + where + help : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) + (((sym (snd f)) ∙∙ (cong (fst f) loop) ∙∙ snd f) + ∙ (refl ∙ refl)) + (cong (fst f) loop) + help = (cong ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) ∙_) (sym (rUnit refl)) ∙ sym (rUnit _)) + ◁ λ i j → doubleCompPath-filler (sym (snd f)) (cong (fst f) loop) (snd f) (~ i) j +snd (∙Π-rUnit {A = A} {n = zero} f i) j = snd f (~ i ∨ j) +fst (∙Π-rUnit {A = A} {n = suc n} f i) north = snd f (~ i) +fst (∙Π-rUnit {A = A} {n = suc n} f i) south = (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i +fst (∙Π-rUnit {A = A} {n = suc n} f i) (merid a j) = help i j + where + help : PathP (λ i → snd f (~ i) ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) + (((sym (snd f)) ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) ∙∙ snd f) + ∙ (refl ∙ refl)) + (cong (fst f) (merid a)) + help = (cong (((sym (snd f)) ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) ∙∙ snd f) ∙_) + (sym (rUnit refl)) + ∙ sym (rUnit _)) + ◁ λ i j → hcomp (λ k → + λ { (j = i0) → snd f (~ i ∧ k) + ; (j = i1) → compPath-filler' (sym (snd f)) (cong (fst f) (merid (ptSn (suc n)))) k i + ; (i = i0) → doubleCompPath-filler (sym (snd f)) + (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + (snd f) k j + ; (i = i1) → fst f (merid a j)}) + (fst f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) +snd (∙Π-rUnit {A = A} {n = suc n} f i) j = snd f (~ i ∨ j) + +∙Π-lUnit : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ∙Π 1Π f ≡ f +fst (∙Π-lUnit {n = zero} f i) base = snd f (~ i) +fst (∙Π-lUnit {n = zero} f i) (loop j) = s i j + where + s : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) + ((refl ∙ refl) ∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) + (cong (fst f) loop) + s = (cong (_∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) (sym (rUnit refl)) ∙ sym (lUnit _)) + ◁ λ i j → doubleCompPath-filler (sym (snd f)) (cong (fst f) loop) (snd f) (~ i) j +snd (∙Π-lUnit {n = zero} f i) j = snd f (~ i ∨ j) +fst (∙Π-lUnit {n = suc n} f i) north = snd f (~ i) +fst (∙Π-lUnit {n = suc n} f i) south = (sym (snd f) ∙ cong (fst f) (merid (ptSn _))) i +fst (∙Π-lUnit {n = suc n} f i) (merid a j) = help i j + where + help : PathP (λ i → snd f (~ i) ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) + ((refl ∙ refl) ∙ ((sym (snd f)) ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) ∙∙ snd f)) + (cong (fst f) (merid a)) + help = + (cong (_∙ ((sym (snd f)) ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) ∙∙ snd f)) + (sym (rUnit refl)) + ∙ sym (lUnit _)) + ◁ λ i j → hcomp (λ k → λ { (j = i0) → snd f (~ i ∧ k) + ; (j = i1) → compPath-filler' (sym (snd f)) (cong (fst f) (merid (ptSn (suc n)))) k i + ; (i = i0) → doubleCompPath-filler (sym (snd f)) + (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + (snd f) k j + ; (i = i1) → fst f (merid a j)}) + (fst f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) +snd (∙Π-lUnit {n = suc n} f i) j = snd f (~ i ∨ j) + +∙Π-rCancel : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ∙Π f (-Π f) ≡ 1Π +fst (∙Π-rCancel {A = A} {n = zero} f i) base = pt A +fst (∙Π-rCancel {A = A} {n = zero} f i) (loop j) = rCancel (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) i j +snd (∙Π-rCancel {A = A} {n = zero} f i) = refl +fst (∙Π-rCancel {A = A} {n = suc n} f i) north = pt A +fst (∙Π-rCancel {A = A} {n = suc n} f i) south = pt A +fst (∙Π-rCancel {A = A} {n = suc n} f i) (merid a i₁) = sl i i₁ + where + pl = (sym (snd f) ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f) + sl : pl + ∙ ((sym (snd f) ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f)) ≡ refl + sl = cong (pl ∙_) (cong (sym (snd f) ∙∙_∙∙ (snd f)) + (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) + ∙∙ cong₂ _∙_ refl + (cong (cong (fst f)) (rCancel (merid (ptSn _)))) + ∙∙ sym (rUnit _))) + ∙ rCancel pl +snd (∙Π-rCancel {A = A} {n = suc n} f i) = refl + +∙Π-lCancel : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ∙Π (-Π f) f ≡ 1Π +fst (∙Π-lCancel {A = A} {n = zero} f i) base = pt A +fst (∙Π-lCancel {A = A} {n = zero} f i) (loop j) = + rCancel (sym (snd f) ∙∙ cong (fst f) (sym loop) ∙∙ snd f) i j +fst (∙Π-lCancel {A = A} {n = suc n} f i) north = pt A +fst (∙Π-lCancel {A = A} {n = suc n} f i) south = pt A +fst (∙Π-lCancel {A = A} {n = suc n} f i) (merid a j) = sl i j + where + pl = (sym (snd f) ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f) + + sl : ((sym (snd f) ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f)) ∙ pl ≡ refl + sl = cong (_∙ pl) (cong (sym (snd f) ∙∙_∙∙ (snd f)) + (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) + ∙∙ cong₂ _∙_ refl (cong (cong (fst f)) (rCancel (merid (ptSn _)))) + ∙∙ sym (rUnit _))) + ∙ lCancel pl +snd (∙Π-lCancel {A = A} {n = zero} f i) = refl +snd (∙Π-lCancel {A = A} {n = suc n} f i) = refl + +∙Π-assoc : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f g h : S₊∙ (suc n) →∙ A) + → ∙Π f (∙Π g h) ≡ ∙Π (∙Π f g) h +∙Π-assoc {n = n} f g h = + sym (Iso.leftInv (IsoSphereMapΩ (suc n)) (∙Π f (∙Π g h))) + ∙∙ cong (Ω→SphereMap n) (SphereMap→Ωpres∙ n f (∙Π g h) + ∙∙ cong (SphereMap→Ω n f ∙_) (SphereMap→Ωpres∙ n g h) + ∙∙ ∙assoc (SphereMap→Ω n f) (SphereMap→Ω n g) (SphereMap→Ω n h) + ∙∙ cong (_∙ SphereMap→Ω n h) (sym (SphereMap→Ωpres∙ n f g)) + ∙∙ sym (SphereMap→Ωpres∙ n (∙Π f g) h)) + ∙∙ Iso.leftInv (IsoSphereMapΩ (suc n)) (∙Π (∙Π f g) h) + +∙Π-comm : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f g : S₊∙ (suc (suc n)) →∙ A) + → ∙Π f g ≡ ∙Π g f +∙Π-comm {A = A} {n = n} f g = + sym (Iso.leftInv (IsoSphereMapΩ (suc (suc n))) (∙Π f g)) + ∙∙ cong (Ω→SphereMap (suc n)) (SphereMap→Ωpres∙ (suc n) f g + ∙∙ EH _ _ _ + ∙∙ sym (SphereMap→Ωpres∙ (suc n) g f)) + ∙∙ Iso.leftInv (IsoSphereMapΩ (suc (suc n))) (∙Π g f) + + +{- π'' as a group -} +1π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' n A +1π' n {A = A} = ∣ 1Π ∣₂ + +·π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' (suc n) A → π' (suc n) A → π' (suc n) A +·π' n = sRec2 squash₂ λ p q → ∣ ∙Π p q ∣₂ + +-π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' (suc n) A → π' (suc n) A +-π' n = sMap -Π + +π'-rUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) + → (·π' n x (1π' (suc n))) ≡ x +π'-rUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-rUnit p i ∣₂ + +π'-lUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) + → (·π' n (1π' (suc n)) x) ≡ x +π'-lUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-lUnit p i ∣₂ + +π'-rCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) + → (·π' n x (-π' n x)) ≡ 1π' (suc n) +π'-rCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-rCancel p i ∣₂ + +π'-lCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) + → (·π' n (-π' n x) x) ≡ 1π' (suc n) +π'-lCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-lCancel p i ∣₂ + +π'-assoc : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y z : π' (suc n) A) + → ·π' n x (·π' n y z) ≡ ·π' n (·π' n x y) z +π'-assoc n = sElim3 (λ _ _ _ → isSetPathImplicit) λ p q r i → ∣ ∙Π-assoc p q r i ∣₂ + +π'-comm : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y : π' (suc (suc n)) A) + → ·π' (suc n) x y ≡ ·π' (suc n) y x +π'-comm n = sElim2 (λ _ _ → isSetPathImplicit) λ p q i → ∣ ∙Π-comm p q i ∣₂ + +π'Gr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Group ℓ +fst (π'Gr n A) = π' (suc n) A +1g (snd (π'Gr n A)) = 1π' (suc n) +GroupStr._·_ (snd (π'Gr n A)) = ·π' n +inv (snd (π'Gr n A)) = -π' n +is-set (isSemigroup (isMonoid (isGroup (snd (π'Gr n A))))) = squash₂ +assoc (isSemigroup (isMonoid (isGroup (snd (π'Gr n A))))) = π'-assoc n +identity (isMonoid (isGroup (snd (π'Gr n A)))) x = (π'-rUnit n x) , (π'-lUnit n x) +inverse (isGroup (snd (π'Gr n A))) x = (π'-rCancel n x) , (π'-lCancel n x) + +-- Finally, the Iso +π'Gr≅πGr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → GroupIso (π'Gr n A) (πGr n A) +fst (π'Gr≅πGr n A) = setTruncIso (IsoSphereMapΩ (suc n)) +snd (π'Gr≅πGr n A) = + makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) + λ p q i → ∣ IsoSphereMapΩ-pres∙Π n p q i ∣₂) diff --git a/Cubical/Homotopy/HopfInvariant/Base.agda b/Cubical/Homotopy/HopfInvariant/Base.agda new file mode 100644 index 000000000..cdfb55efe --- /dev/null +++ b/Cubical/Homotopy/HopfInvariant/Base.agda @@ -0,0 +1,334 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Homotopy.HopfInvariant.Base where + +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Groups.Connected +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.MayerVietorisUnreduced +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Wedge +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.RingStructure.CupProduct +open import Cubical.ZCohomology.RingStructure.RingLaws +open import Cubical.ZCohomology.RingStructure.GradedCommutativity +open import Cubical.Relation.Nullary +open import Cubical.ZCohomology.Gysin + +open import Cubical.Functions.Embedding + +open import Cubical.Foundations.Path +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Pointed.Homogeneous + +open import Cubical.Homotopy.HomotopyGroup + +open import Cubical.Foundations.Univalence + +open import Cubical.Relation.Nullary + +open import Cubical.Data.Sum +open import Cubical.Data.Fin +open import Cubical.Data.Empty renaming (rec to ⊥-rec) +open import Cubical.Data.Sigma +open import Cubical.Data.Int hiding (_+'_) +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_ ; _·_ to _·ℕ_) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Unit +open import Cubical.Data.Bool +open import Cubical.Algebra.Group + renaming (ℤ to ℤGroup ; Unit to UnitGroup) hiding (Bool) +open import Cubical.Algebra.Group.ZModule + +open import Cubical.HITs.Pushout.Flattening +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.EilenbergSteenrod +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 renaming (_·_ to _*_) +open import Cubical.HITs.Truncation + renaming (rec to trRec ; elim to trElim ; elim2 to trElim2) +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap ; elim3 to sElim3) +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; elim to pElim) + +open import Cubical.Algebra.AbGroup + +open import Cubical.Homotopy.Loopspace + +open import Cubical.HITs.Join + +open import Cubical.Homotopy.Hopf + +open import Cubical.HITs.SetQuotients renaming (_/_ to _/'_) + +isContr→≡UnitGroup : {G : Group ℓ-zero} → isContr (fst G) → UnitGroup ≡ G +isContr→≡UnitGroup c = + fst (GroupPath _ _) + (invGroupEquiv ((isContr→≃Unit c) + , (makeIsGroupHom (λ _ _ → refl)))) + +GroupIsoUnitGroup→isContr : {G : Group ℓ-zero} → GroupIso UnitGroup G → isContr (fst G) +GroupIsoUnitGroup→isContr is = + isOfHLevelRetractFromIso 0 (invIso (fst is)) isContrUnit + +-- move to group stuff +SES→isEquiv : ∀ {ℓ ℓ'} {L R : Group ℓ-zero} + → {G : Group ℓ} {H : Group ℓ'} + → UnitGroup ≡ L + → UnitGroup ≡ R + → (lhom : GroupHom L G) (midhom : GroupHom G H) (rhom : GroupHom H R) + → ((x : _) → isInKer midhom x → isInIm lhom x) + → ((x : _) → isInKer rhom x → isInIm midhom x) + → isEquiv (fst midhom) +SES→isEquiv {R = R} {G = G} {H = H} = + J (λ L _ → UnitGroup ≡ R → + (lhom : GroupHom L G) (midhom : GroupHom G H) + (rhom : GroupHom H R) → + ((x : fst G) → isInKer midhom x → isInIm lhom x) → + ((x : fst H) → isInKer rhom x → isInIm midhom x) → + isEquiv (fst midhom)) + ((J (λ R _ → (lhom : GroupHom UnitGroup G) (midhom : GroupHom G H) + (rhom : GroupHom H R) → + ((x : fst G) → isInKer midhom x → isInIm lhom x) → + ((x : _) → isInKer rhom x → isInIm midhom x) → + isEquiv (fst midhom)) + main)) + where + main : (lhom : GroupHom UnitGroup G) (midhom : GroupHom G H) + (rhom : GroupHom H UnitGroup) → + ((x : fst G) → isInKer midhom x → isInIm lhom x) → + ((x : fst H) → isInKer rhom x → isInIm midhom x) → + isEquiv (fst midhom) + main lhom midhom rhom lexact rexact = + BijectionIsoToGroupEquiv {G = G} {H = H} + bijIso' .fst .snd + where + bijIso' : BijectionIso G H + BijectionIso.fun bijIso' = midhom + BijectionIso.inj bijIso' x inker = + pRec (GroupStr.is-set (snd G) _ _) + (λ s → sym (snd s) ∙ IsGroupHom.pres1 (snd lhom)) + (lexact _ inker) + BijectionIso.surj bijIso' x = rexact x refl + +-- The pushout of the hopf invariant +HopfInvariantPush : (n : ℕ) + → (f : S₊ (3 +ℕ n +ℕ n) → S₊ (2 +ℕ n)) + → Type _ +HopfInvariantPush n f = Pushout (λ _ → tt) f + +Hopfα : (n : ℕ) + → (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) + → coHom (2 +ℕ n) (HopfInvariantPush n (fst f)) +Hopfα n f = ∣ (λ { (inl x) → 0ₖ _ + ; (inr x) → ∣ x ∣ + ; (push a i) → help a (~ i)}) ∣₂ + where + help : (a : S₊ (3 +ℕ n +ℕ n)) → ∣ fst f a ∣ ≡ 0ₖ (2 +ℕ n) + help = sphereElim _ (λ _ → isOfHLevelPlus' {n = n} (3 +ℕ n) + (isOfHLevelPath' (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n)) _ _)) (cong ∣_∣ₕ (snd f)) + +Hopfβ : (n : ℕ) + → (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) + → coHom (4 +ℕ (n +ℕ n)) (HopfInvariantPush n (fst f)) +Hopfβ n f = fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) ∣ ∣_∣ ∣₂ + +-- To define the Hopf invariant, we need to establish the non-trivial isos of Hⁿ(HopfInvariant). +module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) where + module M = MV _ _ _ (λ _ → tt) (fst f) + ¬lemHopf : (n m : ℕ) → ¬ suc (suc (m +ℕ n)) ≡ suc n + ¬lemHopf zero zero p = snotz (cong predℕ p) + ¬lemHopf (suc n) zero p = ¬lemHopf n zero (cong predℕ p) + ¬lemHopf zero (suc m) p = snotz (cong predℕ p) + ¬lemHopf (suc n) (suc m) p = + ¬lemHopf n (suc m) (cong (suc ∘ suc ) + (sym (+-suc m n)) ∙ (cong predℕ p)) + + SphereHopfCohomIso : + GroupEquiv + (coHomGr (3 +ℕ n +ℕ n) (S₊ (3 +ℕ n +ℕ n))) + ((coHomGr (suc (3 +ℕ n +ℕ n)) (HopfInvariantPush _ (fst f)))) + fst (fst SphereHopfCohomIso) = fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) + snd (fst SphereHopfCohomIso) = help + where + abstract + help : isEquiv (fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n))) + help = + SES→isEquiv + (isContr→≡UnitGroup + (isContrΣ (GroupIsoUnitGroup→isContr (invGroupIso (Hⁿ-Unit≅0 _))) + λ _ → GroupIsoUnitGroup→isContr + (invGroupIso (Hⁿ-Sᵐ≅0 _ _ (¬lemHopf n n))))) + (isContr→≡UnitGroup + (isContrΣ (GroupIsoUnitGroup→isContr + (invGroupIso (Hⁿ-Unit≅0 _))) + λ _ → GroupIsoUnitGroup→isContr + (invGroupIso (Hⁿ-Sᵐ≅0 _ _ (¬lemHopf n (suc n)))))) + (M.Δ (3 +ℕ n +ℕ n)) + (M.d (3 +ℕ n +ℕ n)) + (M.i (4 +ℕ n +ℕ n)) + (M.Ker-d⊂Im-Δ _) + (M.Ker-i⊂Im-d _) + snd SphereHopfCohomIso = d-inv + where + -- Currently, type checking is very slow without the abstract flag. + -- TODO : Remove abstract + abstract + d-inv : IsGroupHom (snd (coHomGr (3 +ℕ n +ℕ n) (S₊ (3 +ℕ n +ℕ n)))) + (fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n))) + (snd (coHomGr (suc (3 +ℕ n +ℕ n)) (Pushout (λ _ → tt) (fst f)))) + d-inv = snd (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) + + Hopfβ-Iso : GroupIso (coHomGr (suc (3 +ℕ n +ℕ n)) (HopfInvariantPush _ (fst f))) + ℤGroup + Iso.fun (fst Hopfβ-Iso) x = + Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) + (invEq (fst SphereHopfCohomIso) x) + Iso.inv (fst Hopfβ-Iso) f = + (fst (fst SphereHopfCohomIso)) + (Iso.inv (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) f) + Iso.rightInv (fst Hopfβ-Iso) f = + cong (Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) + (retEq (fst SphereHopfCohomIso) (Iso.inv (fst (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) f)) + ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) f + Iso.leftInv (fst Hopfβ-Iso) x = + cong (fst (fst SphereHopfCohomIso)) + (Iso.leftInv (fst (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) + (invEq (fst SphereHopfCohomIso) x)) + ∙ secEq (fst SphereHopfCohomIso) x + snd Hopfβ-Iso = grHom + where + abstract + grHom : IsGroupHom (coHomGr (suc (3 +ℕ n +ℕ n)) (Pushout (λ _ → tt) (fst f)) .snd) + (λ x → Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) + (invEq (fst SphereHopfCohomIso) x)) + (ℤGroup .snd) + grHom = makeIsGroupHom λ x y → + cong (Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) + (IsGroupHom.pres· (isGroupHomInv SphereHopfCohomIso) x y) + ∙ IsGroupHom.pres· (snd (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) _ _ + +Hⁿ-Sⁿ≅ℤ-nice-generator : (n : ℕ) → Iso.inv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) 1 ≡ ∣ ∣_∣ ∣₂ +Hⁿ-Sⁿ≅ℤ-nice-generator zero = Iso.leftInv (fst (Hⁿ-Sⁿ≅ℤ (suc zero))) _ +Hⁿ-Sⁿ≅ℤ-nice-generator (suc n) = + (λ i → Iso.inv (fst (suspensionAx-Sn (suc n) (suc n))) (Hⁿ-Sⁿ≅ℤ-nice-generator n i)) + ∙ cong ∣_∣₂ (funExt λ { north → refl + ; south → cong ∣_∣ₕ (merid north) + ; (merid a i) j → ∣ compPath-filler (merid a) (sym (merid north)) (~ j) i ∣ₕ}) + +Hopfβ↦1 : (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) + → Iso.fun (fst (Hopfβ-Iso n f)) (Hopfβ n f) ≡ 1 +Hopfβ↦1 n f = sym (cong (Iso.fun (fst (Hopfβ-Iso n f))) lem) ∙ Iso.rightInv (fst (Hopfβ-Iso n f)) (pos 1) + where + lem : Iso.inv (fst (Hopfβ-Iso n f)) (pos 1) ≡ Hopfβ n f + lem = (λ i → fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) + (Hⁿ-Sⁿ≅ℤ-nice-generator _ i)) + ∙ cong ∣_∣₂ (funExt (λ { (inl x) → refl + ; (inr x) → refl + ; (push a i) → refl})) + +module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) where + private + 2+n = 2 +ℕ n + H = HopfInvariantPush n (fst f) + + H→Sphere : coHom 2+n H → coHom 2+n (S₊ (suc (suc n))) + H→Sphere = sMap (_∘ inr) + + grHom : IsGroupHom (snd (coHomGr 2+n H)) H→Sphere (snd (coHomGr 2+n (S₊ (suc (suc n))))) + grHom = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f g → refl) + + preSphere→H : (g : (S₊ (suc (suc n)) → coHomK 2+n)) → H → coHomK (2 +ℕ n) + preSphere→H g (inl x) = 0ₖ _ + preSphere→H g (inr x) = g x -ₖ g north + preSphere→H g (push a i) = h23 a i + where + h23 : (a : S₊ (suc (suc (suc (n +ℕ n))))) → 0ₖ (suc (suc n)) ≡ g (fst f a) -ₖ g north + h23 = sphereElim _ (λ x → isOfHLevelPlus' {n = n} (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n) _ _)) + (sym (rCancelₖ _ (g north)) ∙ cong (λ x → g x -ₖ g north) (sym (snd f))) + + Sphere→H : coHom 2+n (S₊ (suc (suc n))) → coHom 2+n H + Sphere→H = sMap preSphere→H + + conCohom2+n : (x : _) → ∥ x ≡ 0ₖ (suc (suc n)) ∥ + conCohom2+n = coHomK-elim _ (λ _ → isProp→isOfHLevelSuc (suc n) squash) ∣ refl ∣ + + HIPSphereCohomIso : Iso (coHom (2 +ℕ n) (HopfInvariantPush n (fst f))) + (coHom (2 +ℕ n) ((S₊ (suc (suc n))))) + Iso.fun HIPSphereCohomIso = H→Sphere + Iso.inv HIPSphereCohomIso = Sphere→H + Iso.rightInv HIPSphereCohomIso = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ g → pRec (squash₂ _ _) + (λ p → cong ∣_∣₂ (funExt λ x → cong (g x +ₖ_) (cong (-ₖ_) p) ∙ rUnitₖ _ (g x))) + (conCohom2+n (g north)) + Iso.leftInv HIPSphereCohomIso = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ g → pRec (squash₂ _ _) + (pRec (isPropΠ (λ _ → squash₂ _ _)) + (λ gn gtt → + trRec (isProp→isOfHLevelSuc n (squash₂ _ _)) + (λ r → cong ∣_∣₂ (funExt λ { + (inl x) → sym gtt + ; (inr x) → (λ i → g (inr x) -ₖ gn i) ∙ rUnitₖ _ (g (inr x)) + ; (push a i) + → sphereElim _ + {A = λ a → PathP (λ i → preSphere→H (λ x → g (inr x)) (push a i) ≡ g (push a i)) + (sym gtt) + ((λ i → g (inr (fst f a)) -ₖ gn i) ∙ rUnitₖ _ (g (inr (fst f a))))} + (λ _ → isOfHLevelPathP' (suc (suc (suc (n +ℕ n)))) + (isOfHLevelPath (suc (suc (suc (suc (n +ℕ n))))) + (isOfHLevelPlus' {n = n} (suc (suc (suc (suc n)))) + (isOfHLevelTrunc (suc (suc (suc (suc n)))))) _ _) _ _) + r a i})) + (push-helper g gtt gn)) + (conCohom2+n (g (inr north)))) + (conCohom2+n (g (inl tt))) + where + push-helper : (g : HopfInvariantPush n (fst f) → coHomK (suc (suc n))) + → (gtt : g (inl tt) ≡ 0ₖ (suc (suc n))) + → (gn : g (inr north) ≡ 0ₖ (suc (suc n))) + → hLevelTrunc (suc n) (PathP (λ i → preSphere→H (λ x → g (inr x)) (push north i) ≡ g (push north i)) + (sym gtt) ((λ i → g (inr (fst f north)) -ₖ gn i) ∙ rUnitₖ _ (g (inr (fst f north))))) + push-helper g gtt gn = isConnectedPathP _ (isConnectedPath _ (isConnectedKn _) _ _) _ _ .fst + + Hopfα-Iso : GroupIso (coHomGr (2 +ℕ n) (HopfInvariantPush n (fst f))) ℤGroup + Hopfα-Iso = + compGroupIso + (HIPSphereCohomIso , grHom) + (Hⁿ-Sⁿ≅ℤ (suc n)) + +Hopfα-Iso-α : (n : ℕ) (f : _) + → Iso.fun (fst (Hopfα-Iso n f)) (Hopfα n f) + ≡ 1 +Hopfα-Iso-α n f = + sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (Hⁿ-Sⁿ≅ℤ-nice-generator n)) + ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) + where + hz : Iso.fun (HIPSphereCohomIso n f) (Hopfα n f) ≡ ∣ ∣_∣ ∣₂ + hz = refl + +⌣-α : (n : ℕ) → (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → _ +⌣-α n f = Hopfα n f ⌣ Hopfα n f + +HopfInvariant : (n : ℕ) + → (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → ℤ +HopfInvariant n f = Iso.fun (fst (Hopfβ-Iso n f)) + ((subst (λ x → coHom x (HopfInvariantPush n (fst f))) + (λ i → suc (suc (suc (+-suc n n i))))) (⌣-α n f)) diff --git a/Cubical/Homotopy/HopfInvariant/HopfMap.agda b/Cubical/Homotopy/HopfInvariant/HopfMap.agda new file mode 100644 index 000000000..06eb05722 --- /dev/null +++ b/Cubical/Homotopy/HopfInvariant/HopfMap.agda @@ -0,0 +1,1447 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +{- +This file contains a proof that the generator of Π₃S² has +hopf invariant ±1. +-} +module Cubical.Homotopy.HopfInvariant.HopfMap where + +open import Cubical.Homotopy.HopfInvariant.Base +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Groups.Connected +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.MayerVietorisUnreduced +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Wedge +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.RingStructure.CupProduct +open import Cubical.ZCohomology.RingStructure.RingLaws +open import Cubical.ZCohomology.RingStructure.GradedCommutativity +open import Cubical.Relation.Nullary +open import Cubical.ZCohomology.Gysin + +open import Cubical.Functions.Embedding + +open import Cubical.Foundations.Path +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Pointed.Homogeneous + +open import Cubical.Homotopy.HomotopyGroup + +open import Cubical.Foundations.Univalence + +open import Cubical.Relation.Nullary + +open import Cubical.Data.Sum +open import Cubical.Data.Fin +open import Cubical.Data.Empty renaming (rec to ⊥-rec) +open import Cubical.Data.Sigma +open import Cubical.Data.Int hiding (_+'_) +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_ ; _·_ to _·ℕ_) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Unit +open import Cubical.Data.Bool +open import Cubical.Algebra.Group + renaming (ℤ to ℤGroup ; Unit to UnitGroup) hiding (Bool) +open import Cubical.Algebra.Group.ZModule + +open import Cubical.HITs.Pushout.Flattening +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.EilenbergSteenrod +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 renaming (_·_ to _*_) +open import Cubical.HITs.Truncation + renaming (rec to trRec ; elim to trElim ; elim2 to trElim2) +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap ; elim3 to sElim3) +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; elim to pElim) + +open import Cubical.Algebra.AbGroup + +open import Cubical.Homotopy.Loopspace + +open import Cubical.HITs.Join + +open import Cubical.Homotopy.Hopf + +open import Cubical.HITs.SetQuotients renaming (_/_ to _/'_) + + +open import Cubical.Experiments.Brunerie +open import Cubical.HITs.Hopf + +open import Cubical.HITs.Join + + +HopfMap : S₊∙ 3 →∙ S₊∙ 2 +fst HopfMap x = JoinS¹S¹→TotalHopf (Iso.inv (IsoSphereJoin 1 1) x) .fst +snd HopfMap = refl + +-- We use the Hopf fibration in order to connect it to the Gysin Sequence +module hopfS¹ = hopfBase S1-AssocHSpace (sphereElim2 _ (λ _ _ → squash) ∣ refl ∣) +S¹Hopf : S₊ 2 → Type +S¹Hopf = hopfS¹.Hopf + +TotalHopf' : Type _ +TotalHopf' = Σ (S₊ 2) S¹Hopf + +CP² : Type _ +CP² = hopfS¹.megaPush + +fibr : CP² → Type _ +fibr = hopfS¹.P + +-- Remove +hopf : join S¹ S¹ → S₊ 2 +hopf x = fst (JoinS¹S¹→TotalHopf x) + +E* : Type _ +E* = hopfS¹.totalSpaceMegaPush + +IsoE*join : Iso E* (join S¹ (join S¹ S¹)) +IsoE*join = hopfS¹.IsoJoin₁ + +IsoJoins : (join S¹ (join S¹ S¹)) ≡ join S¹ (S₊ 3) +IsoJoins = cong (join S¹) (isoToPath (IsoSphereJoin 1 1)) + +IsoTotalHopf' : Iso hopfS¹.TotalSpaceHopf' (join S¹ S¹) +IsoTotalHopf' = hopfS¹.joinIso₁ + +-- CP² is 1-connected +conCP² : (x y : CP²) → ∥ x ≡ y ∥₂ +conCP² x y = sRec2 squash₂ (λ p q → ∣ p ∙ sym q ∣₂) (conCP²' x) (conCP²' y) + where + conCP²' : (x : CP²) → ∥ x ≡ inl tt ∥₂ + conCP²' (inl x) = ∣ refl ∣₂ + conCP²' (inr x) = sphereElim 1 {A = λ x → ∥ inr x ≡ inl tt ∥₂} + (λ _ → squash₂) ∣ sym (push (inl base)) ∣₂ x + conCP²' (push a i₁) = main a i₁ + where + indLem : ∀ {ℓ} {A : hopfS¹.TotalSpaceHopf' → Type ℓ} → ((a : _) → isProp (A a)) + → A (inl base) + → ((a : hopfS¹.TotalSpaceHopf') → A a) + indLem {A = A} p b = + PushoutToProp p (sphereElim 0 (λ _ → p _) b) + (sphereElim 0 (λ _ → p _) (subst A (push (base , base)) b)) + + main : (a : hopfS¹.TotalSpaceHopf') + → PathP (λ i → ∥ Path CP² (push a i) (inl tt) ∥₂) + (conCP²' (inl tt)) (conCP²' (inr (hopfS¹.induced a))) + main = indLem (λ _ → isOfHLevelPathP' 1 squash₂ _ _) + λ j → ∣ (λ i → push (inl base) (~ i ∧ j)) ∣₂ + +module GysinS¹ = Gysin (CP² , inl tt) fibr conCP² 2 idIso refl + +E'S⁴Iso : Iso GysinS¹.E' (S₊ 5) +E'S⁴Iso = + compIso IsoE*join + (compIso (Iso→joinIso idIso (IsoSphereJoin 1 1)) + (IsoSphereJoin 1 3)) + +isContrH³E : isContr (coHom 3 (GysinS¹.E')) +isContrH³E = + subst isContr (sym (isoToPath + (fst (Hⁿ-Sᵐ≅0 2 4 + λ p → snotz (sym (cong (predℕ ∘ predℕ) p))))) + ∙ cong (coHom 3) (sym (isoToPath E'S⁴Iso))) + isContrUnit + +isContrH⁴E : isContr (coHom 4 (GysinS¹.E')) +isContrH⁴E = + subst isContr (sym (isoToPath + (fst (Hⁿ-Sᵐ≅0 3 4 + λ p → snotz (sym (cong (predℕ ∘ predℕ ∘ predℕ) p))))) + ∙ cong (coHom 4) (sym (isoToPath E'S⁴Iso))) + isContrUnit + +joinS¹S¹→Groupoid : ∀ {ℓ} (P : join S¹ S¹ → Type ℓ) + → ((x : _) → isGroupoid (P x)) + → P (inl base) + → (x : _) → P x +joinS¹S¹→Groupoid P grp b = + transport (λ i → (x : (isoToPath (invIso (IsoSphereJoin 1 1))) i) + → P (transp (λ j → isoToPath (invIso (IsoSphereJoin 1 1)) (i ∨ j)) i x )) + (sphereElim _ (λ _ → grp _) b) + +TotalHopf→Gpd : ∀ {ℓ} (P : hopfS¹.TotalSpaceHopf' → Type ℓ) + → ((x : _) → isGroupoid (P x)) + → P (inl base) + → (x : _) → P x +TotalHopf→Gpd P grp b = + transport (λ i → (x : sym (isoToPath IsoTotalHopf') i) + → P (transp (λ j → isoToPath IsoTotalHopf' (~ i ∧ ~ j)) i x)) + (joinS¹S¹→Groupoid _ (λ _ → grp _) b) + +TotalHopf→Gpd' : ∀ {ℓ} (P : Σ (S₊ 2) hopfS¹.Hopf → Type ℓ) + → ((x : _) → isGroupoid (P x)) + → P (north , base) + → (x : _) → P x +TotalHopf→Gpd' P grp b = + transport (λ i → (x : ua (_ , hopfS¹.isEquivTotalSpaceHopf'→TotalSpace) i) + → P (transp (λ j → ua (_ , hopfS¹.isEquivTotalSpaceHopf'→TotalSpace) (i ∨ j)) i x)) + (TotalHopf→Gpd _ (λ _ → grp _) b) + +CP²→Groupoid : ∀ {ℓ} {P : CP² → Type ℓ} + → ((x : _) → is2Groupoid (P x)) + → (b : P (inl tt)) + → (s2 : ((x : _) → P (inr x))) + → PathP (λ i → P (push (inl base) i)) b (s2 north) + → (x : _) → P x +CP²→Groupoid {P = P} grp b s2 pp (inl x) = b +CP²→Groupoid {P = P} grp b s2 pp (inr x) = s2 x +CP²→Groupoid {P = P} grp b s2 pp (push a i₁) = h3 a i₁ + where + h3 : (a : hopfS¹.TotalSpaceHopf') → PathP (λ i → P (push a i)) b (s2 _) + h3 = TotalHopf→Gpd _ (λ _ → isOfHLevelPathP' 3 (grp _) _ _) pp + +-- The function inducing the iso H²(S²) ≅ H²(CP²) +H²S²→H²CP²-raw : (S₊ 2 → coHomK 2) → (CP² → coHomK 2) +H²S²→H²CP²-raw f (inl x) = 0ₖ _ +H²S²→H²CP²-raw f (inr x) = f x -ₖ f north +H²S²→H²CP²-raw f (push a i) = + TotalHopf→Gpd (λ x → 0ₖ 2 ≡ f (hopfS¹.TotalSpaceHopf'→TotalSpace x .fst) -ₖ f north) + (λ _ → isOfHLevelTrunc 4 _ _) + (sym (rCancelₖ 2 (f north))) + a i + +H²S²→H²CP² : coHomGr 2 (S₊ 2) .fst → coHomGr 2 CP² .fst +H²S²→H²CP² = sMap H²S²→H²CP²-raw + +cancel-H²S²→H²CP² : (f : CP² → coHomK 2) → ∥ H²S²→H²CP²-raw (f ∘ inr) ≡ f ∥ +cancel-H²S²→H²CP² f = + pRec squash + (λ p → pRec squash + (λ f → ∣ funExt f ∣) + (cancelLem p)) + (connLem (f (inl tt))) + where + connLem : (x : coHomK 2) → ∥ x ≡ 0ₖ 2 ∥ + connLem = coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash) ∣ refl ∣ + + cancelLem : (p : f (inl tt) ≡ 0ₖ 2) → ∥ ((x : CP²) → H²S²→H²CP²-raw (f ∘ inr) x ≡ f x) ∥ + cancelLem p = trRec squash (λ pp → + ∣ CP²→Groupoid (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) + (sym p) + (λ x → (λ i → f (inr x) -ₖ f (push (inl base) (~ i))) + ∙∙ (λ i → f (inr x) -ₖ p i) + ∙∙ rUnitₖ 2 (f (inr x))) pp ∣) + help + where + help : hLevelTrunc 1 (PathP ((λ i₁ → H²S²→H²CP²-raw (f ∘ inr) (push (inl base) i₁) ≡ f (push (inl base) i₁))) + (sym p) + (((λ i₁ → f (inr north) -ₖ f (push (inl base) (~ i₁))) ∙∙ + (λ i₁ → f (inr north) -ₖ p i₁) ∙∙ rUnitₖ 2 (f (inr north))))) + help = isConnectedPathP 1 (isConnectedPath 2 (isConnectedKn 1) _ _) _ _ .fst + +H²CP²≅H²S² : GroupIso (coHomGr 2 CP²) (coHomGr 2 (S₊ 2)) +Iso.fun (fst H²CP²≅H²S²) = sMap (_∘ inr) +Iso.inv (fst H²CP²≅H²S²) = H²S²→H²CP² +Iso.rightInv (fst H²CP²≅H²S²) = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → trRec {B = Iso.fun (fst H²CP²≅H²S²) (Iso.inv (fst H²CP²≅H²S²) ∣ f ∣₂) ≡ ∣ f ∣₂} + (isOfHLevelPath 2 squash₂ _ _) + (λ p → cong ∣_∣₂ (funExt λ x → cong (λ y → (f x) -ₖ y) p ∙ rUnitₖ 2 (f x))) + (Iso.fun (PathIdTruncIso _) (isContr→isProp (isConnectedKn 1) ∣ f north ∣ ∣ 0ₖ 2 ∣)) +Iso.leftInv (fst H²CP²≅H²S²) = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → pRec (squash₂ _ _) (cong ∣_∣₂) (cancel-H²S²→H²CP² f) +snd H²CP²≅H²S² = + makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ f g → refl) + +H²CP²≅ℤ : GroupIso (coHomGr 2 CP²) ℤGroup +H²CP²≅ℤ = compGroupIso H²CP²≅H²S² (Hⁿ-Sⁿ≅ℤ 1) + +⌣Equiv : GroupEquiv (coHomGr 2 CP²) (coHomGr 4 CP²) +fst (fst ⌣Equiv) = GysinS¹.⌣-hom 2 .fst +snd (fst ⌣Equiv) = subst isEquiv (cong fst (GysinS¹.ϕ∘j≡ 2)) eq' + where + eq' : isEquiv (GysinS¹.ϕ∘j 2 .fst) + eq' = + SES→isEquiv + (GroupPath _ _ .fst (invGroupEquiv + (isContr→≃Unit isContrH³E + , makeIsGroupHom λ _ _ → refl))) + (GroupPath _ _ .fst (invGroupEquiv + (isContr→≃Unit isContrH⁴E + , makeIsGroupHom λ _ _ → refl))) + (GysinS¹.susp∘ϕ 1) + (GysinS¹.ϕ∘j 2) + (GysinS¹.p-hom 4) + (GysinS¹.Ker-ϕ∘j⊂Im-Susp∘ϕ _) + (GysinS¹.Ker-p⊂Im-ϕ∘j _) +snd ⌣Equiv = GysinS¹.⌣-hom 2 .snd + +genCP² : coHom 2 CP² +genCP² = ∣ CP²→Groupoid (λ _ → isOfHLevelTrunc 4) + (0ₖ _) + ∣_∣ + refl ∣₂ + +inrInjective : (f g : CP² → coHomK 2) → ∥ f ∘ inr ≡ g ∘ inr ∥ → Path (coHom 2 CP²) ∣ f ∣₂ ∣ g ∣₂ +inrInjective f g = pRec (squash₂ _ _) + (λ p → pRec (squash₂ _ _) (λ id → trRec (squash₂ _ _) + (λ pp → cong ∣_∣₂ + (funExt (CP²→Groupoid (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) + id + (funExt⁻ p) + (compPathR→PathP pp)))) + (conn2 (f (inl tt)) (g (inl tt)) id + (cong f (push (inl base)) ∙ (funExt⁻ p north) ∙ cong g (sym (push (inl base)))))) + (conn (f (inl tt)) (g (inl tt)))) + + where + conn : (x y : coHomK 2) → ∥ x ≡ y ∥ + conn = coHomK-elim _ (λ _ → isSetΠ λ _ → isOfHLevelSuc 1 squash) + (coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash) ∣ refl ∣) + + conn2 : (x y : coHomK 2) (p q : x ≡ y) → hLevelTrunc 1 (p ≡ q) + conn2 x y = λ p q → Iso.fun (PathIdTruncIso _) (isContr→isProp (isConnectedPath _ (isConnectedKn 1) x y) ∣ p ∣ ∣ q ∣) + +private + rUnit* : (x : S¹) → x * base ≡ x + rUnit* base = refl + rUnit* (loop i₁) = refl + + merid*-lem : (a x : S¹) → Path (Path (coHomK 2) _ _) (cong ∣_∣ₕ (merid (a * x) ∙ sym (merid base))) + ((cong ∣_∣ₕ (merid a ∙ sym (merid base))) ∙ (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) + merid*-lem = wedgeconFun _ _ (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) + (λ x → lUnit _ ∙ cong (_∙ cong ∣_∣ₕ (merid x ∙ sym (merid base))) (cong (cong ∣_∣ₕ) (sym (rCancel (merid base))))) + (λ x → (λ i → cong ∣_∣ₕ (merid (rUnit* x i) ∙ sym (merid base))) + ∙∙ rUnit _ + ∙∙ cong (cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙_) + (cong (cong ∣_∣ₕ) (sym (rCancel (merid base))))) + (sym (l (cong ∣_∣ₕ (merid base ∙ sym (merid base))) + (cong (cong ∣_∣ₕ) (sym (rCancel (merid base)))))) + where + l : ∀ {ℓ} {A : Type ℓ} {x : A} (p : x ≡ x) (P : refl ≡ p) + → lUnit p ∙ cong (_∙ p) P ≡ rUnit p ∙ cong (p ∙_) P + l p = J (λ p P → lUnit p ∙ cong (_∙ p) P ≡ rUnit p ∙ cong (p ∙_) P) refl + + lemmie : (x : S¹) → ptSn 1 ≡ x * (invLooper x) + lemmie base = refl + lemmie (loop i) j = + hcomp (λ r → λ {(i = i0) → base ; (i = i1) → base ; (j = i0) → base}) + base + + + meridInvLooperLem : (x : S¹) → Path (Path (coHomK 2) _ _) + (cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base))) + (cong ∣_∣ₕ (sym (merid x ∙ sym (merid base)))) + meridInvLooperLem x = (lUnit _ + ∙∙ cong (_∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base))) (sym (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base))))) ∙∙ sym (assoc _ _ _)) + ∙∙ cong (sym (cong ∣_∣ₕ (merid x ∙ sym (merid base))) ∙_) hh + ∙∙ (assoc _ _ _ + ∙∙ cong (_∙ (cong ∣_∣ₕ (sym (merid x ∙ sym (merid base))))) (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) + ∙∙ sym (lUnit _)) + where + hh : cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base)) + ≡ cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙ cong ∣_∣ₕ (sym (merid x ∙ sym (merid base))) + hh = sym (merid*-lem x (invLooper x)) ∙ ((λ i → cong ∣_∣ₕ (merid (lemmie x (~ i)) ∙ sym (merid base))) ∙ cong (cong ∣_∣ₕ) (rCancel (merid base))) ∙ sym (rCancel _) + + comm·S¹ : (a x : S¹) → a * x ≡ x * a + comm·S¹ = wedgeconFun _ _ (λ _ _ → isGroupoidS¹ _ _) + (sym ∘ rUnit*) + rUnit* + refl + + invLooperLem₁ : (a x : S¹) → (invEq (hopfS¹.μ-eq a) x) * a ≡ (invLooper a * x) * a + invLooperLem₁ a x = secEq (hopfS¹.μ-eq a) x ∙∙ cong (_* x) (lemmie a) ∙∙ assocHSpace.μ-assoc S1-AssocHSpace a (invLooper a) x ∙ comm·S¹ _ _ + + invLooperLem₂ : (a x : S¹) → invEq (hopfS¹.μ-eq a) x ≡ invLooper a * x + invLooperLem₂ a x = sym (retEq (hopfS¹.μ-eq a) (invEq (hopfS¹.μ-eq a) x)) + ∙∙ cong (invEq (hopfS¹.μ-eq a)) (invLooperLem₁ a x) + ∙∙ retEq (hopfS¹.μ-eq a) (invLooper a * x) + +-- We prove that the generator of CP² given by Gysin is the same one +-- as genCP², which is much easier to work with +Gysin-e≡genCP² : GysinS¹.e ≡ genCP² +Gysin-e≡genCP² = inrInjective _ _ ∣ funExt (λ x → funExt⁻ (cong fst (ll32 x)) south) ∣ + where + mainId : (x : Σ (S₊ 2) hopfS¹.Hopf) → Path (hLevelTrunc 4 _) ∣ fst x ∣ ∣ north ∣ + mainId = uncurry λ { north → λ y → cong ∣_∣ₕ (merid base ∙ sym (merid y)) + ; south → λ y → cong ∣_∣ₕ (sym (merid y)) + ; (merid a i) → main a i} + where + main : (a : S¹) → PathP (λ i → (y : hopfS¹.Hopf (merid a i)) → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a i ∣ ∣ north ∣) + (λ y → cong ∣_∣ₕ (merid base ∙ sym (merid y))) + λ y → cong ∣_∣ₕ (sym (merid y)) + main a = toPathP (funExt λ x → cong (transport ((λ i₁ → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a i₁ ∣ ∣ north ∣))) + ((λ i → (λ z → cong ∣_∣ₕ (merid base + ∙ sym (merid (transport (λ j → uaInvEquiv (hopfS¹.μ-eq a) (~ i) j) x))) z)) + ∙ λ i → cong ∣_∣ₕ (merid base + ∙ sym (merid (transportRefl (invEq (hopfS¹.μ-eq a) x) i)))) + ∙∙ (λ i → transp (λ i₁ → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a (i₁ ∨ i) ∣ ∣ north ∣) i + (compPath-filler' (cong ∣_∣ₕ (sym (merid a))) + (cong ∣_∣ₕ (merid base ∙ sym (merid (invLooperLem₂ a x i)))) i)) + ∙∙ cong ((cong ∣_∣ₕ) (sym (merid a)) ∙_) (cong (cong ∣_∣ₕ) (cong sym (symDistr (merid base) (sym (merid (invLooper a * x))))) + ∙ cong sym (merid*-lem (invLooper a) x) + ∙ symDistr ((cong ∣_∣ₕ) (merid (invLooper a) ∙ sym (merid base))) + ((cong ∣_∣ₕ) (merid x ∙ sym (merid base))) + ∙ isCommΩK 2 (sym (λ i₁ → ∣ (merid x ∙ (λ i₂ → merid base (~ i₂))) i₁ ∣)) + (sym (λ i₁ → ∣ (merid (invLooper a) ∙ (λ i₂ → merid base (~ i₂))) i₁ ∣)) + ∙ cong₂ _∙_ (cong sym (meridInvLooperLem a) ∙ cong-∙ ∣_∣ₕ (merid a) (sym (merid base))) + (cong (cong ∣_∣ₕ) (symDistr (merid x) (sym (merid base))) ∙ cong-∙ ∣_∣ₕ (merid base) (sym (merid x)))) + ∙∙ (λ j → (λ i₁ → ∣ merid a (~ i₁ ∨ j) ∣) ∙ + ((λ i₁ → ∣ merid a (i₁ ∨ j) ∣) ∙ + (λ i₁ → ∣ merid base (~ i₁ ∨ j) ∣)) + ∙ + (λ i₁ → ∣ merid base (i₁ ∨ j) ∣) ∙ + (λ i₁ → ∣ merid x (~ i₁) ∣ₕ)) + ∙∙ sym (lUnit _) + ∙∙ sym (assoc _ _ _) + ∙∙ (sym (lUnit _) ∙ sym (lUnit _) ∙ sym (lUnit _))) + + psst : (x : S₊ 2) → preThom.Q (CP² , inl tt) fibr (inr x) →∙ coHomK-ptd 2 + fst (psst x) north = ∣ north ∣ + fst (psst x) south = ∣ x ∣ + fst (psst x) (merid a i₁) = mainId (x , a) (~ i₁) + snd (psst x) = refl + + ll32-fst : GysinS¹.c (inr north) .fst ≡ psst north .fst + ll32-fst = cong fst (GysinS¹.cEq (inr north) ∣ push (inl base) ∣₂) + ∙ (funExt l) + where + l : (qb : _) → + ∣ (subst (fst ∘ preThom.Q (CP² , inl tt) fibr) (sym (push (inl base))) qb) ∣ + ≡ psst north .fst qb + l north = refl + l south = cong ∣_∣ₕ (sym (merid base)) + l (merid a i) j = + hcomp (λ k → λ { (i = i0) → ∣ merid a (~ k ∧ j) ∣ + ; (i = i1) → ∣ merid base (~ j) ∣ + ; (j = i0) → ∣ transportRefl (merid a i) (~ k) ∣ + ; (j = i1) → ∣ compPath-filler (merid base) (sym (merid a)) k (~ i) ∣ₕ}) + (hcomp (λ k → λ { (i = i0) → ∣ merid a (j ∨ ~ k) ∣ + ; (i = i1) → ∣ merid base (~ j ∨ ~ k) ∣ + ; (j = i0) → ∣ merid a (~ k ∨ i) ∣ + ; (j = i1) → ∣ merid base (~ i ∨ ~ k) ∣ₕ}) + ∣ south ∣) + + is-setHepl : (x : S₊ 2) → isSet (preThom.Q (CP² , inl tt) fibr (inr x) →∙ coHomK-ptd 2) + is-setHepl = sphereElim _ (λ _ → isProp→isOfHLevelSuc 1 (isPropIsOfHLevel 2)) + (isOfHLevel↑∙' 0 1) + + ll32 : (x : S₊ 2) → (GysinS¹.c (inr x) ≡ psst x) + ll32 = sphereElim _ (λ x → isOfHLevelPath 2 (is-setHepl x) _ _) + (→∙Homogeneous≡ (isHomogeneousKn _) ll32-fst) + +isGenerator≃ℤ : ∀ {ℓ} (G : Group ℓ) (g : fst G) + → Type _ +isGenerator≃ℤ G g = + Σ[ e ∈ GroupIso G ℤGroup ] abs (Iso.fun (fst e) g) ≡ 1 + +isGenerator≃ℤ-e : isGenerator≃ℤ (coHomGr 2 CP²) GysinS¹.e +isGenerator≃ℤ-e = + subst (isGenerator≃ℤ (coHomGr 2 CP²)) (sym Gysin-e≡genCP²) + (H²CP²≅ℤ , refl) + +HopfMap' : S₊ 3 → S₊ 2 +HopfMap' x = + hopfS¹.TotalSpaceHopf'→TotalSpace + (Iso.inv IsoTotalHopf' + (Iso.inv (IsoSphereJoin 1 1) x)) .fst + + +rotLoop² : (a : _) → Path (Path S¹ _ _) (λ i → rotLoop (rotLoop a i) (~ i)) refl +rotLoop² = + sphereElim 0 (λ _ → isGroupoidS¹ _ _ _ _) + λ i j → hcomp (λ {k → λ {(i = i1) → base ; (j = i0) → base ; (j = i1) → base}}) + base + +id1 : (x : _) → hopfS¹.TotalSpaceHopf'→TotalSpace x .fst ≡ JoinS¹S¹→TotalHopf (Iso.fun IsoTotalHopf' x) .fst +id1 (inl x) = refl +id1 (inr x) = refl +id1 (push (base , snd₁) i) = refl +id1 (push (loop i₁ , a) i) k = merid (rotLoop² a (~ k) i₁) i + +CP2' : Type _ +CP2' = Pushout {A = S₊ 3} (λ _ → tt) HopfMap' + +hopfMap≡HopfMap' : HopfMap ≡ (HopfMap' , refl) +hopfMap≡HopfMap' = + ΣPathP ((funExt (λ x → cong (λ x → JoinS¹S¹→TotalHopf x .fst) + (sym (Iso.rightInv IsoTotalHopf' (Iso.inv (IsoSphereJoin 1 1) x))) + ∙ sym (id1 (Iso.inv IsoTotalHopf' (Iso.inv (IsoSphereJoin 1 1) x))))) + , flipSquare (sym (rUnit refl) ◁ λ _ _ → north)) + +PushoutReplaceBase : + ∀ {ℓ ℓ' ℓ''} {A B : Type ℓ} {C : Type ℓ'} {D : Type ℓ''} {f : A → C} {g : A → D} + (e : B ≃ A) → Pushout (f ∘ fst e) (g ∘ fst e) ≡ Pushout f g +PushoutReplaceBase {f = f} {g = g} = + EquivJ (λ _ e → Pushout (f ∘ fst e) (g ∘ fst e) ≡ Pushout f g) + refl + +CP2≡CP2' : CP2' ≡ CP² +CP2≡CP2' = + PushoutReplaceBase + (isoToEquiv (compIso (invIso (IsoSphereJoin 1 1)) (invIso IsoTotalHopf'))) + +{- +⌣-pres1-CP² : ⌣-pres1 CP² 2 +fst ⌣-pres1-CP² = GysinS¹.e +fst (snd ⌣-pres1-CP2) = isGenerator≃ℤ-e +snd (snd ⌣-pres1-CP2) = ∣ p ∣ + where + p : Σ _ _ + fst p = compGroupIso (GroupEquiv→GroupIso (invGroupEquiv ⌣Equiv)) H²CP²≅ℤ + snd p = {!!} + ∙ sesIsoPresGen _ (invGroupEquiv (GroupIso→GroupEquiv H²CP²≅ℤ)) _ + (compGroupEquiv (invGroupEquiv (GroupIso→GroupEquiv H²CP²≅ℤ)) (invGroupEquiv (invGroupEquiv ⌣Equiv))) + GysinS¹.e {!!} ⌣Equiv +{- + GysinS¹.e + , (isGenerator≃ℤ-e + , ∣ compGroupIso (GroupEquiv→GroupIso (invGroupEquiv ⌣Equiv)) H²CP²≅ℤ + , {!!} ∣) + -} + +⌣-pres1-CP2' : ⌣-pres1 CP2' 2 +⌣-pres1-CP2' = subst (λ x → ⌣-pres1 x 2) (sym CP2≡CP2') ⌣-pres1-CP2 + -} +-- {- +-- Goal: snd (v' (pt A) (push a i₁)) ≡ +-- ua-gluePt (μ-eq (snd a)) i₁ (fst a) +-- ———— Boundary —————————————————————————————————————————————— +-- i₁ = i0 ⊢ HSpace.μₗ e (fst a) +-- i₁ = i1 ⊢ HSpace.μₗ e (HSpace.μ e (fst a) (snd a)) +-- -} + +-- -- help : (x : _) → (v' (pt A)) x ≡ TotalSpaceHopf'→TotalSpace x +-- -- help (inl x) = ΣPathP (refl , HSpace.μₗ e x) +-- -- help (inr x) = ΣPathP (refl , (HSpace.μₗ e x)) +-- -- help (push (x , y) i) j = +-- -- comp (λ _ → Σ (Susp (typ A)) Hopf) +-- -- (λ k → λ {(i = i0) → merid y i , HSpace.μₗ e x j +-- -- ; (i = i1) → merid y i , assocHSpace.μ-assoc-filler e-ass x y j k +-- -- ; (j = i0) → merid y i , hfill +-- -- (λ j → λ { (i = i0) → HSpace.μ e (pt A) x +-- -- ; (i = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j +-- -- }) +-- -- (inS (ua-gluePt (μ-eq y) i (HSpace.μ e (pt A) x))) +-- -- k +-- -- ; (j = i1) → merid y i , ua-gluePt (μ-eq y) i x}) +-- -- (merid y i , ua-gluePt (μ-eq y) i (HSpace.μₗ e x j)) +-- -- where +-- -- open import Cubical.Foundations.Path + +-- -- PPΣ : ∀ {ℓ} {A : Type ℓ} {f : A ≃ A} (p : f ≡ f) → {!!} +-- -- PPΣ = {!!} + +-- -- V : PathP (λ i₁ → hcomp +-- -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x +-- -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j +-- -- }) +-- -- (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x)) ≡ +-- -- ua-gluePt (μ-eq y) i₁ x) +-- -- (HSpace.μₗ e x) +-- -- (HSpace.μₗ e (HSpace.μ e x y)) -- (HSpace.μₗ e (HSpace.μ e (fst a) (snd a))) +-- -- V = transport (λ z → {!PathP (λ i₁ → hfill +-- -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x +-- -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j +-- -- }) +-- -- (inS (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x))) z ≡ ua-gluePt (μ-eq y) i₁ x) +-- -- ? ?!}) +-- -- {!hfill +-- -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x +-- -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j +-- -- }) +-- -- (inS (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x))) ?!} -- toPathP ({!!} ∙∙ {!!} ∙∙ {!!}) -- toPathP (flipSquare {!!}) -- hcomp {!!} {!!} + +-- -- P : Pushout {A = TotalSpaceHopf'} (λ _ → tt) induced → Type _ +-- -- P (inl x) = typ A +-- -- P (inr x) = Hopf x +-- -- P (push a i₁) = ua (v a) i₁ + +⌣equiv→pres1 : ∀ {ℓ} {G H : Type ℓ} → (G ≡ H) → (g₁ : coHom 2 G) (h₁ : coHom 2 H) + → (fstEq : (Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] abs (fst (fst ϕ) g₁) ≡ 1)) + → (sndEq : ((Σ[ ϕ ∈ GroupEquiv (coHomGr 2 H) ℤGroup ] abs (fst (fst ϕ) h₁) ≡ 1))) + → isEquiv {A = coHom 2 G} {B = coHom 4 G} (_⌣ g₁) + → (3rdEq : GroupEquiv (coHomGr 4 H) ℤGroup) + → abs (fst (fst 3rdEq) (h₁ ⌣ h₁)) ≡ 1 +⌣equiv→pres1 {G = G} = J (λ H _ → (g₁ : coHom 2 G) (h₁ : coHom 2 H) + → (fstEq : (Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] abs (fst (fst ϕ) g₁) ≡ 1)) + → (sndEq : ((Σ[ ϕ ∈ GroupEquiv (coHomGr 2 H) ℤGroup ] abs (fst (fst ϕ) h₁) ≡ 1))) + → isEquiv {A = coHom 2 G} {B = coHom 4 G} (_⌣ g₁) + → (3rdEq : GroupEquiv (coHomGr 4 H) ℤGroup) + → abs (fst (fst 3rdEq) (h₁ ⌣ h₁)) ≡ 1) + help + where + help : (g₁ h₁ : coHom 2 G) + → (fstEq : (Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] abs (fst (fst ϕ) g₁) ≡ 1)) + → (sndEq : ((Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] abs (fst (fst ϕ) h₁) ≡ 1))) + → isEquiv {A = coHom 2 G} {B = coHom 4 G} (_⌣ g₁) + → (3rdEq : GroupEquiv (coHomGr 4 G) ℤGroup) + → abs (fst (fst 3rdEq) (h₁ ⌣ h₁)) ≡ 1 + help g h (ϕ , idg) (ψ , idh) ⌣eq ξ = + ⊎→abs _ _ + (allisoPresGen _ + (compGroupEquiv main4 (compGroupEquiv (invGroupEquiv main4) ψ)) + h (abs→⊎ _ _ (cong abs (cong (fst (fst ψ)) (retEq (fst main4) h)) ∙ idh)) (compGroupEquiv main4 ξ)) + where + k1 : ((fst (fst ψ) h) ≡ 1) ⊎ (fst (fst ψ) h ≡ -1) + → abs (fst (fst ϕ) h) ≡ 1 + k1 p = ⊎→abs _ _ (allisoPresGen _ ψ h p ϕ) + + help2 : ((fst (fst ϕ) h) ≡ 1) ⊎ (fst (fst ϕ) h ≡ -1) + → ((fst (fst ϕ) g) ≡ 1) ⊎ (fst (fst ϕ) g ≡ -1) + → (h ≡ g) ⊎ (h ≡ (-ₕ g)) + help2 (inl x) (inl x₁) = + inl (sym (retEq (fst ϕ) h) + ∙∙ cong (invEq (fst ϕ)) (x ∙ sym x₁) + ∙∙ retEq (fst ϕ) g) + help2 (inl x) (inr x₁) = + inr (sym (retEq (fst ϕ) h) + ∙∙ cong (invEq (fst ϕ)) x + ∙ IsGroupHom.presinv (snd (invGroupEquiv ϕ)) (negsuc zero) + ∙∙ cong (-ₕ_) (cong (invEq (fst ϕ)) (sym x₁) ∙ (retEq (fst ϕ) g))) + help2 (inr x) (inl x₁) = + inr (sym (retEq (fst ϕ) h) + ∙∙ cong (invEq (fst ϕ)) x + ∙∙ (IsGroupHom.presinv (snd (invGroupEquiv ϕ)) 1 + ∙ cong (-ₕ_) (cong (invEq (fst ϕ)) (sym x₁) ∙ (retEq (fst ϕ) g)))) + help2 (inr x) (inr x₁) = + inl (sym (retEq (fst ϕ) h) + ∙∙ cong (invEq (fst ϕ)) (x ∙ sym x₁) + ∙∙ retEq (fst ϕ) g) + + -ₕeq : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → Iso (coHom n A) (coHom n A) + Iso.fun (-ₕeq n) = -ₕ_ + Iso.inv (-ₕeq n) = -ₕ_ + Iso.rightInv (-ₕeq n) = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt λ x → -ₖ^2 (f x)) + Iso.leftInv (-ₕeq n) = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt λ x → -ₖ^2 (f x)) + + theEq : coHom 2 G ≃ coHom 4 G + theEq = compEquiv (_ , ⌣eq) (isoToEquiv (-ₕeq 4)) + + main3 : (h ≡ g) ⊎ (h ≡ (-ₕ g)) → isEquiv {A = coHom 2 G} (_⌣ h) + main3 (inl x) = subst isEquiv (λ i → _⌣ (x (~ i))) ⌣eq + main3 (inr x) = + subst isEquiv (funExt (λ x → -ₕDistᵣ 2 2 x g) ∙ (λ i → _⌣ (x (~ i)))) + (theEq .snd) + + main4 : GroupEquiv (coHomGr 2 G) (coHomGr 4 G) + fst main4 = _ , (main3 (help2 (abs→⊎ _ _ (k1 (abs→⊎ _ _ idh))) (abs→⊎ _ _ idg))) + snd main4 = + makeIsGroupHom λ g1 g2 → rightDistr-⌣ _ _ g1 g2 h + +HopfInvariantPushElim : ∀ {ℓ} n → (f : _) → {P : HopfInvariantPush n f → Type ℓ} + → (isOfHLevel (suc (suc (suc (suc (n +ℕ n))))) (P (inl tt))) + → (e : P (inl tt)) + (g : (x : _) → P (inr x)) + (r : PathP (λ i → P (push north i)) e (g (f north))) + → (x : _) → P x +HopfInvariantPushElim n f {P = P} hlev e g r (inl x) = e +HopfInvariantPushElim n f {P = P} hlev e g r (inr x) = g x +HopfInvariantPushElim n f {P = P} hlev e g r (push a i₁) = help a i₁ + where + help : (a : Susp (Susp (S₊ (suc (n +ℕ n))))) → PathP (λ i → P (push a i)) e (g (f a)) + help = sphereElim _ (sphereElim _ (λ _ → isProp→isOfHLevelSuc ((suc (suc (n +ℕ n)))) (isPropIsOfHLevel _)) + (isOfHLevelPathP' (suc (suc (suc (n +ℕ n)))) + (subst (isOfHLevel (suc (suc (suc (suc (n +ℕ n)))))) + (cong P (push north)) + hlev) _ _)) r + +HopfInvariant-HopfMap' : abs (HopfInvariant zero (HopfMap' , λ _ → HopfMap' (snd (S₊∙ 3)))) ≡ suc zero +HopfInvariant-HopfMap' = + cong abs (cong (Iso.fun (fst (Hopfβ-Iso zero (HopfMap' , refl)))) + (transportRefl (⌣-α 0 (HopfMap' , refl)))) ∙ ⌣equiv→pres1 (sym CP2≡CP2') GysinS¹.e (Hopfα zero (HopfMap' , refl)) + (l isGenerator≃ℤ-e) + (GroupIso→GroupEquiv (Hopfα-Iso 0 (HopfMap' , refl)) , refl) + (snd (fst ⌣Equiv)) + (GroupIso→GroupEquiv (Hopfβ-Iso zero (HopfMap' , refl))) + where + l : Σ-syntax (GroupIso (coHomGr 2 CP²) ℤGroup) + (λ ϕ → abs (Iso.fun (fst ϕ) GysinS¹.e) ≡ 1) + → Σ-syntax (GroupEquiv (coHomGr 2 CP²) ℤGroup) + (λ ϕ → abs (fst (fst ϕ) GysinS¹.e) ≡ 1) + l p = (GroupIso→GroupEquiv (fst p)) , (snd p) + +HopfInvariant-HopfMap : abs (HopfInvariant zero HopfMap) ≡ suc zero +HopfInvariant-HopfMap = cong abs (cong (HopfInvariant zero) hopfMap≡HopfMap') + ∙ HopfInvariant-HopfMap' + + + +-- open import Cubical.HITs.Wedge +-- _∨→_ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} +-- → (f : A →∙ C) (g : B →∙ C) +-- → A ⋁ B → fst C +-- (f ∨→ g) (inl x) = fst f x +-- (f ∨→ g) (inr x) = fst g x +-- (f ∨→ g) (push a i₁) = (snd f ∙ sym (snd g)) i₁ + + +-- add2 : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → (f g : S₊∙ n →∙ A) → ((S₊∙ n) ⋁ (S₊∙ n) , inl (ptSn _)) →∙ A +-- fst (add2 {A = A} f g) = f ∨→ g +-- snd (add2 {A = A} f g) = snd f + +-- C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Type _ +-- C* n f g = Pushout (λ _ → tt) (fst (add2 f g)) + +-- θ : ∀ {ℓ} {A : Pointed ℓ} → Susp (fst A) → (Susp (fst A) , north) ⋁ (Susp (fst A) , north) +-- θ north = inl north +-- θ south = inr north +-- θ {A = A} (merid a i₁) = +-- ((λ i → inl ((merid a ∙ sym (merid (pt A))) i)) +-- ∙∙ push tt +-- ∙∙ λ i → inr ((merid a ∙ sym (merid (pt A))) i)) i₁ + + +-- help3 : ∀ {ℓ} {A : Type ℓ} {x y z w : A} (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) +-- → p ∙∙ q ∙∙ r ≡ p ∙ q ∙ r +-- help3 {x = x} {y = y} {z = z} {w = w} = +-- J (λ y p → (q : y ≡ z) (r : z ≡ w) → +-- (p ∙∙ q ∙∙ r) ≡ p ∙ q ∙ r) +-- λ q r → lUnit (q ∙ r) + +-- +≡ : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → (x : _) → ∙Π f g .fst x ≡ (f ∨→ g) (θ {A = S₊∙ _} x) +-- +≡ n f g north = sym (snd f) +-- +≡ n f g south = sym (snd g) +-- +≡ n f g (merid a i₁) j = main j i₁ +-- where + +-- help : cong (f ∨→ g) (cong (θ {A = S₊∙ _}) (merid a)) +-- ≡ (refl ∙∙ cong (fst f) (merid a ∙ sym (merid north)) ∙∙ snd f) +-- ∙ (sym (snd g) ∙∙ cong (fst g) (merid a ∙ sym (merid north)) ∙∙ refl) +-- help = cong-∙∙ (f ∨→ g) ((λ i → inl ((merid a ∙ sym (merid north)) i))) +-- (push tt) +-- (λ i → inr ((merid a ∙ sym (merid north)) i)) +-- ∙∙ help3 _ _ _ +-- ∙∙ cong (cong (f ∨→ g) +-- (λ i₂ → inl ((merid a ∙ (λ i₃ → merid north (~ i₃))) i₂)) ∙_) +-- (sym (assoc _ _ _)) +-- ∙∙ assoc _ _ _ +-- ∙∙ cong₂ _∙_ refl (compPath≡compPath' _ _) + +-- main : PathP (λ i → snd f (~ i) ≡ snd g (~ i)) (λ i₁ → ∙Π f g .fst (merid a i₁)) +-- λ i₁ → (f ∨→ g) (θ {A = S₊∙ _} (merid a i₁)) +-- main = (λ i → ((λ j → snd f (~ i ∧ ~ j)) ∙∙ cong (fst f) (merid a ∙ sym (merid north)) ∙∙ snd f) +-- ∙ (sym (snd g) ∙∙ cong (fst g) (merid a ∙ sym (merid north)) ∙∙ λ j → snd g (~ i ∧ j))) +-- ▷ sym help + +-- +≡' : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → ∙Π f g ≡ ((f ∨→ g) ∘ (θ {A = S₊∙ _}) , snd f) +-- +≡' n f g = ΣPathP ((funExt (+≡ n f g)) , (λ j i → snd f (i ∨ ~ j))) + +-- WedgeElim : (n : ℕ) → ∀ {ℓ} {P : S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) → Type ℓ} +-- → ((x : _) → isOfHLevel (3 +ℕ n) (P x)) +-- → P (inl north) +-- → (x : _) → P x +-- WedgeElim n {P = P} x s (inl x₁) = +-- sphereElim _ {A = P ∘ inl} +-- (λ _ → isOfHLevelPlus' {n = n} (3 +ℕ n) (x _)) s x₁ +-- WedgeElim n {P = P} x s (inr x₁) = +-- sphereElim _ {A = P ∘ inr} +-- (sphereElim _ (λ _ → isProp→isOfHLevelSuc ((suc (suc (n +ℕ n)))) +-- (isPropIsOfHLevel (suc (suc (suc (n +ℕ n)))))) +-- (subst (isOfHLevel ((3 +ℕ n) +ℕ n)) (cong P (push tt)) +-- (isOfHLevelPlus' {n = n} (3 +ℕ n) (x _)))) +-- (subst P (push tt) s) x₁ +-- WedgeElim n {P = P} x s (push a j) = transp (λ i → P (push tt (i ∧ j))) (~ j) s + + +-- H²-C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → GroupIso (coHomGr (2 +ℕ n) (C* n f g)) ℤGroup +-- H²-C* n f g = compGroupIso groupIso1 (Hⁿ-Sⁿ≅ℤ (suc n)) +-- where +-- help : (coHom (2 +ℕ n) (C* n f g)) → coHom (2 +ℕ n) (S₊ (2 +ℕ n)) +-- help = sMap λ f → f ∘ inr + + +-- invMapPrim : (S₊ (2 +ℕ n) → coHomK (2 +ℕ n)) → C* n f g → coHomK (2 +ℕ n) +-- invMapPrim h (inl x) = h (ptSn _) +-- invMapPrim h (inr x) = h x +-- invMapPrim h (push a i₁) = WedgeElim n {P = λ a → h north ≡ h (fst (add2 f g) a)} +-- (λ _ → isOfHLevelTrunc (4 +ℕ n) _ _) +-- (cong h (sym (snd f))) a i₁ + +-- invMap : coHom (2 +ℕ n) (S₊ (2 +ℕ n)) → (coHom (2 +ℕ n) (C* n f g)) +-- invMap = sMap invMapPrim + +-- ret1 : (x : coHom (2 +ℕ n) (S₊ (2 +ℕ n))) → help (invMap x) ≡ x +-- ret1 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) +-- λ a → refl + +-- ret2 : (x : (coHom (2 +ℕ n) (C* n f g))) → invMap (help x) ≡ x +-- ret2 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) +-- λ h → cong ∣_∣₂ (funExt λ { (inl x) → cong h ((λ i → inr ((snd f) (~ i))) +-- ∙ sym (push (inl north))) +-- ; (inr x) → refl +-- ; (push a i₁) → help2 h a i₁}) +-- where +-- help2 : (h : C* n f g → coHomK (2 +ℕ n)) +-- → (a : _) → PathP (λ i → invMapPrim (h ∘ inr) (push a i) ≡ h (push a i)) +-- (cong h ((λ i → inr ((snd f) (~ i))) ∙ sym (push (inl north)))) +-- refl +-- help2 h = WedgeElim n (λ _ → isOfHLevelPathP (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n) _ _) _ _) +-- help4 + +-- where +-- help4 : PathP (λ i → invMapPrim (h ∘ inr) (push (inl north) i) ≡ h (push (inl north) i)) +-- (cong h ((λ i → inr ((snd f) (~ i))) ∙ sym (push (inl north)))) +-- refl +-- help4 i j = h (hcomp (λ k → λ { (i = i1) → inr (fst f north) +-- ; (j = i0) → inr (snd f (~ i)) +-- ; (j = i1) → push (inl north) (i ∨ ~ k)}) +-- (inr (snd f (~ i ∧ ~ j)))) + +-- groupIso1 : GroupIso ((coHomGr (2 +ℕ n) (C* n f g))) (coHomGr (2 +ℕ n) (S₊ (2 +ℕ n))) +-- Iso.fun (fst groupIso1) = help +-- Iso.inv (fst groupIso1) = invMap +-- Iso.rightInv (fst groupIso1) = ret1 +-- Iso.leftInv (fst groupIso1) = ret2 +-- snd groupIso1 = makeIsGroupHom +-- (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) +-- λ f g → refl) + + +-- C+ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Type _ +-- C+ n f g = Pushout (λ _ → tt) (∙Π f g .fst) + +-- H⁴-C∨ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → GroupIso (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C+ n f g)) +-- ℤGroup +-- H⁴-C∨ n f g = compGroupIso +-- (transportCohomIso (cong (3 +ℕ_) (+-suc n n))) (Hopfβ-Iso n (∙Π f g)) + +-- lol : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) (x : ℤ) +-- → Iso.inv (fst (H⁴-C∨ n f g)) x +-- ≡ subst (λ m → coHom m (C+ n f g)) (sym (cong (3 +ℕ_) (+-suc n n))) (Iso.inv (fst (Hopfβ-Iso n (∙Π f g))) x) +-- lol n f g = λ _ → refl + +-- H²-C∨ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → GroupIso (coHomGr (2 +ℕ n) (C+ n f g)) +-- ℤGroup +-- H²-C∨ n f g = Hopfα-Iso n (∙Π f g) + + +-- H⁴-C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → GroupIso (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) +-- (DirProd ℤGroup ℤGroup) +-- H⁴-C* n f g = compGroupIso (GroupEquiv→GroupIso (invGroupEquiv fstIso)) +-- (compGroupIso (transportCohomIso (cong (2 +ℕ_) (+-suc n n))) +-- (compGroupIso (Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) _) +-- (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ _) (Hⁿ-Sⁿ≅ℤ _)))) +-- where +-- module Ms = MV _ _ _ (λ _ → tt) (fst (add2 f g)) +-- fstIso : GroupEquiv (coHomGr ((suc (suc (n +ℕ suc n)))) (S₊∙ (3 +ℕ (n +ℕ n)) ⋁ S₊∙ (3 +ℕ (n +ℕ n)))) +-- (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) +-- fst (fst fstIso) = Ms.d (suc (suc (n +ℕ suc n))) .fst +-- snd (fst fstIso) = +-- SES→isEquiv +-- (GroupPath _ _ .fst (compGroupEquiv (invEquiv (isContr→≃Unit ((tt , tt) , (λ _ → refl))) , makeIsGroupHom λ _ _ → refl) +-- (GroupEquivDirProd +-- (GroupIso→GroupEquiv (invGroupIso (Hⁿ-Unit≅0 _))) +-- (GroupIso→GroupEquiv +-- (invGroupIso +-- (Hⁿ-Sᵐ≅0 _ _ λ p → ¬lemHopf 0 ((λ _ → north) , refl) n n (cong suc (sym (+-suc n n)) ∙ p))))))) +-- (GroupPath _ _ .fst +-- (compGroupEquiv ((invEquiv (isContr→≃Unit ((tt , tt) , (λ _ → refl))) , makeIsGroupHom λ _ _ → refl)) +-- ((GroupEquivDirProd +-- (GroupIso→GroupEquiv (invGroupIso (Hⁿ-Unit≅0 _))) +-- (GroupIso→GroupEquiv +-- (invGroupIso (Hⁿ-Sᵐ≅0 _ _ λ p → ¬lemHopf 0 ((λ _ → north) , refl) n (suc n) (cong (2 +ℕ_) (sym (+-suc n n)) ∙ p)))))))) +-- (Ms.Δ (suc (suc (n +ℕ suc n)))) +-- (Ms.d (suc (suc (n +ℕ suc n)))) +-- (Ms.i (suc (suc (suc (n +ℕ suc n))))) +-- (Ms.Ker-d⊂Im-Δ _) +-- (Ms.Ker-i⊂Im-d _) +-- snd fstIso = Ms.d (suc (suc (n +ℕ suc n))) .snd + +-- Hopfie : {n : ℕ} → ∥ S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n) ∥₂ → ℤ +-- Hopfie = sRec isSetℤ (HopfInvariant _) + +-- subber : (n : ℕ) (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → _ +-- subber n f = subst (λ x → coHom x (HopfInvariantPush n (fst f))) +-- (λ i → suc (suc (suc (+-suc n n i)))) + +-- module _ (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) where + +-- C+' = C+ n f g + +-- βₗ : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) +-- βₗ = Iso.inv (fst (H⁴-C* n f g)) (1 , 0) + +-- βᵣ : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) +-- βᵣ = Iso.inv (fst (H⁴-C* n f g)) (0 , 1) + +-- βᵣ'-fun : C* n f g → coHomK ((4 +ℕ (n +ℕ n))) +-- βᵣ'-fun (inl x) = 0ₖ _ +-- βᵣ'-fun (inr x) = 0ₖ _ +-- βᵣ'-fun (push (inl x) i₁) = 0ₖ _ +-- βᵣ'-fun (push (inr x) i₁) = Kn→ΩKn+1 _ ∣ x ∣ i₁ +-- βᵣ'-fun (push (push a i₂) i₁) = Kn→ΩKn+10ₖ _ (~ i₂) i₁ + + +-- βₗ'-fun : C* n f g → coHomK (4 +ℕ (n +ℕ n)) +-- βₗ'-fun (inl x) = 0ₖ _ +-- βₗ'-fun (inr x) = 0ₖ _ +-- βₗ'-fun (push (inl x) i₁) = Kn→ΩKn+1 _ ∣ x ∣ i₁ +-- βₗ'-fun (push (inr x) i₁) = 0ₖ _ +-- βₗ'-fun (push (push a i₂) i₁) = Kn→ΩKn+10ₖ _ i₂ i₁ + +-- βₗ''-fun : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) +-- βₗ''-fun = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) ∣ βₗ'-fun ∣₂ + +-- βᵣ''-fun : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) +-- βᵣ''-fun = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) ∣ βᵣ'-fun ∣₂ + +-- α∨ : coHom (2 +ℕ n) (C* n f g) +-- α∨ = Iso.inv (fst (H²-C* n f g)) 1 + +-- private +-- pp : (a : _) → 0ₖ (suc (suc n)) ≡ ∣ (f ∨→ g) a ∣ +-- pp = WedgeElim _ (λ _ → isOfHLevelTrunc (4 +ℕ n) _ _) +-- (cong ∣_∣ₕ (sym (snd f))) + +-- pp-inr : pp (inr north) ≡ cong ∣_∣ₕ (sym (snd g)) +-- pp-inr = (λ j → transport (λ i → 0ₖ (2 +ℕ n) ≡ ∣ compPath-filler' (snd f) (sym (snd g)) (~ j) i ∣ₕ) +-- λ i → ∣ snd f (~ i ∨ j) ∣ₕ) +-- ∙ λ j → transp (λ i → 0ₖ (2 +ℕ n) ≡ ∣ snd g (~ i ∧ ~ j) ∣ₕ) j λ i → ∣ snd g (~ i ∨ ~ j) ∣ₕ + +-- α∨'-fun : C* n f g → coHomK (2 +ℕ n) +-- α∨'-fun (inl x) = 0ₖ _ +-- α∨'-fun (inr x) = ∣ x ∣ +-- α∨'-fun (push a i₁) = pp a i₁ + +-- α∨' : coHom (2 +ℕ n) (C* n f g) +-- α∨' = ∣ α∨'-fun ∣₂ + + +-- β+ : coHom ((2 +ℕ n) +' (2 +ℕ n)) C+' +-- β+ = Iso.inv (fst (H⁴-C∨ n f g)) 1 + +-- q : C+' → C* n f g +-- q (inl x) = inl x +-- q (inr x) = inr x +-- q (push a i₁) = (push (θ a) ∙ λ i → inr (+≡ n f g a (~ i))) i₁ + +-- jₗ : HopfInvariantPush n (fst f) → C* n f g +-- jₗ (inl x) = inl x +-- jₗ (inr x) = inr x +-- jₗ (push a i₁) = push (inl a) i₁ + +-- jᵣ : HopfInvariantPush n (fst g) → C* n f g +-- jᵣ (inl x) = inl x +-- jᵣ (inr x) = inr x +-- jᵣ (push a i₁) = push (inr a) i₁ + +-- α-∨→1 : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → Iso.fun (fst (H²-C* n f g)) (α∨' n f g) ≡ 1 +-- α-∨→1 n f g = sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (Hⁿ-Sⁿ≅ℤ-nice-generator n)) +-- ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) + +-- α-∨-lem : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → α∨ n f g ≡ α∨' n f g +-- α-∨-lem n f g = sym (Iso.leftInv (fst (H²-C* n f g)) _) +-- ∙∙ cong (Iso.inv (fst (H²-C* n f g))) help +-- ∙∙ Iso.leftInv (fst (H²-C* n f g)) _ +-- where +-- help : Iso.fun (fst (H²-C* n f g)) (α∨ n f g) ≡ Iso.fun (fst (H²-C* n f g)) (α∨' n f g) +-- help = (Iso.rightInv (fst (H²-C* n f g)) (pos 1)) ∙ sym (α-∨→1 n f g) + +-- q-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → coHomFun _ (q n f g) (α∨ n f g) ≡ Hopfα n (∙Π f g) +-- q-α n f g = (λ i → coHomFun _ (q n f g) (α-∨-lem n f g i)) +-- ∙ sym (Iso.leftInv is _) +-- ∙∙ cong (Iso.inv is) help +-- ∙∙ Iso.leftInv is _ +-- where +-- is = fst (Hopfα-Iso n (∙Π f g)) + +-- help : Iso.fun is (coHomFun _ (q n f g) (α∨' n f g)) +-- ≡ Iso.fun is (Hopfα n (∙Π f g)) +-- help = sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (Hⁿ-Sⁿ≅ℤ-nice-generator n)) +-- ∙∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) +-- ∙∙ sym (Hopfα-Iso-α n (∙Π f g)) + + +-- βₗ≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → βₗ n f g ≡ βₗ''-fun n f g +-- βₗ≡ n f g = cong (FF ∘ subber2) +-- (cong (Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) +-- (S₊∙ (suc (suc (suc (n +ℕ n))))) +-- (((suc (suc (n +ℕ n))))))))) help +-- ∙ help2) +-- ∙ funExt⁻ FF∘subber ∣ wedgeGen ∣₂ +-- ∙ cong subber3 (sym βₗ'-fun≡) +-- where +-- FF = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (n +ℕ suc n))) .fst +-- FF' = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (suc (n +ℕ n)))) .fst + +-- help : Iso.inv (fst (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (1 , 0) +-- ≡ (∣ ∣_∣ ∣₂ , 0ₕ _) +-- help = ΣPathP ((Hⁿ-Sⁿ≅ℤ-nice-generator _) , IsGroupHom.pres1 (snd (invGroupIso (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) + +-- subber3 = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + +-- subber2 = (subst (λ m → coHom m (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) +-- (sym (cong (2 +ℕ_) (+-suc n n)))) + +-- FF∘subber : FF ∘ subber2 ≡ subber3 ∘ FF' +-- FF∘subber = +-- funExt (λ a → (λ j → transp (λ i → coHom ((suc ∘ suc ∘ suc) (+-suc n n (~ i ∧ j))) (C* n f g)) (~ j) +-- (MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) ((suc (suc (+-suc n n j)))) .fst +-- (transp (λ i → coHom (2 +ℕ (+-suc n n (j ∨ ~ i))) +-- (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) j +-- a)))) + +-- wedgeGen : (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) → +-- coHomK (suc (suc (suc (n +ℕ n))))) +-- wedgeGen (inl x) = ∣ x ∣ +-- wedgeGen (inr x) = 0ₖ _ +-- wedgeGen (push a i₁) = 0ₖ _ + +-- βₗ'-fun≡ : ∣ βₗ'-fun n f g ∣₂ ≡ FF' ∣ wedgeGen ∣₂ +-- βₗ'-fun≡ = cong ∣_∣₂ (funExt λ { (inl x) → refl +-- ; (inr x) → refl +-- ; (push (inl x) i₁) → refl +-- ; (push (inr x) i₁) j → Kn→ΩKn+10ₖ _ (~ j) i₁ +-- ; (push (push a i₂) i₁) j → Kn→ΩKn+10ₖ _ (~ j ∧ i₂) i₁}) + +-- help2 : Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) (((suc (suc (n +ℕ n)))))))) +-- (∣ ∣_∣ ∣₂ , 0ₕ _) +-- ≡ ∣ wedgeGen ∣₂ +-- help2 = cong ∣_∣₂ (funExt λ { (inl x) → rUnitₖ (suc (suc (suc (n +ℕ n)))) ∣ x ∣ +-- ; (inr x) → refl +-- ; (push a i₁) → refl}) + +-- βᵣ≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → βᵣ n f g ≡ βᵣ''-fun n f g +-- βᵣ≡ n f g = +-- cong (FF ∘ subber2) +-- (cong (Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) +-- (S₊∙ (suc (suc (suc (n +ℕ n))))) +-- (((suc (suc (n +ℕ n))))))))) +-- help +-- ∙ help2) +-- ∙ funExt⁻ FF∘subber ∣ wedgeGen ∣₂ +-- ∙ cong (subber3) (sym βᵣ'-fun≡) +-- where +-- FF = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (n +ℕ suc n))) .fst +-- FF' = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (suc (n +ℕ n)))) .fst + +-- help : Iso.inv (fst (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (0 , 1) +-- ≡ (0ₕ _ , ∣ ∣_∣ ∣₂) +-- help = ΣPathP (IsGroupHom.pres1 (snd (invGroupIso (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) , (Hⁿ-Sⁿ≅ℤ-nice-generator _)) + +-- subber3 = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + +-- subber2 = (subst (λ m → coHom m (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) +-- (sym (cong (2 +ℕ_) (+-suc n n)))) + +-- FF∘subber : FF ∘ subber2 ≡ subber3 ∘ FF' +-- FF∘subber = +-- funExt (λ a → (λ j → transp (λ i → coHom ((suc ∘ suc ∘ suc) (+-suc n n (~ i ∧ j))) (C* n f g)) (~ j) +-- (MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) ((suc (suc (+-suc n n j)))) .fst +-- (transp (λ i → coHom (2 +ℕ (+-suc n n (j ∨ ~ i))) +-- (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) j +-- a)))) + +-- wedgeGen : (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) → +-- coHomK (suc (suc (suc (n +ℕ n))))) +-- wedgeGen (inl x) = 0ₖ _ +-- wedgeGen (inr x) = ∣ x ∣ +-- wedgeGen (push a i₁) = 0ₖ _ + + +-- help2 : Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) (((suc (suc (n +ℕ n)))))))) +-- (0ₕ _ , ∣ ∣_∣ ∣₂) +-- ≡ ∣ wedgeGen ∣₂ +-- help2 = cong ∣_∣₂ (funExt λ { (inl x) → refl +-- ; (inr x) → lUnitₖ (suc (suc (suc (n +ℕ n)))) ∣ x ∣ +-- ; (push a i₁) → refl}) + +-- βᵣ'-fun≡ : ∣ βᵣ'-fun n f g ∣₂ ≡ FF' ∣ wedgeGen ∣₂ +-- βᵣ'-fun≡ = cong ∣_∣₂ (funExt λ { (inl x) → refl +-- ; (inr x) → refl +-- ; (push (inl x) i₁) j → Kn→ΩKn+10ₖ _ (~ j) i₁ +-- ; (push (inr x) i₁) → refl +-- ; (push (push a i₂) i₁) j → Kn→ΩKn+10ₖ _ (~ j ∧ ~ i₂) i₁}) + +-- q-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → coHomFun _ (q n f g) (βₗ n f g) +-- ≡ β+ n f g +-- q-βₗ n f g = cong (coHomFun _ (q n f g)) (βₗ≡ n f g) +-- ∙ transportLem +-- ∙ cong (subst (λ m → coHom m (C+' n f g)) +-- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) +-- (sym (Iso.leftInv (fst (Hopfβ-Iso n (∙Π f g))) (Hopfβ n (∙Π f g))) +-- ∙ cong (Iso.inv ((fst (Hopfβ-Iso n (∙Π f g))))) (Hopfβ↦1 n (∙Π f g))) +-- where +-- transportLem : coHomFun (suc (suc (suc (n +ℕ suc n)))) (q n f g) +-- (βₗ''-fun n f g) +-- ≡ subst (λ m → coHom m (C+' n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- (Hopfβ n (∙Π f g)) +-- transportLem = +-- natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (q n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- ∙ cong (subst (λ m → coHom m (C+' n f g)) +-- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) +-- (cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → loll a i₁})) +-- where +-- open import Cubical.Foundations.Path +-- loll : (x : fst (S₊∙ (3 +ℕ (n +ℕ n)))) → +-- PathP +-- (λ i₁ → +-- βₗ'-fun n f g ((push (θ x) ∙ λ i → inr (+≡ n f g x (~ i))) i₁) ≡ +-- MV.d-pre Unit (fst (S₊∙ (2 +ℕ n))) (fst (S₊∙ (3 +ℕ n +ℕ n))) +-- (λ _ → tt) (fst (∙Π f g)) (3 +ℕ n +ℕ n) ∣_∣ (push x i₁)) +-- (λ _ → βₗ'-fun n f g (q n f g (inl tt))) +-- (λ _ → βₗ'-fun n f g (q n f g (inr (∙Π f g .fst x)))) +-- loll x = +-- flipSquare (cong-∙ (βₗ'-fun n f g) (push (θ x)) (λ i → inr (+≡ n f g x (~ i))) +-- ∙∙ sym (rUnit _) +-- ∙∙ lem₁ x) + +-- where +-- lem₁ : (x : _) → cong (βₗ'-fun n f g) (push (θ {A = S₊∙ _} x)) ≡ Kn→ΩKn+1 _ ∣ x ∣ +-- lem₁ north = refl +-- lem₁ south = sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)) +-- lem₁ (merid a j) k = lala k j +-- where +-- lala : PathP (λ k → Kn→ΩKn+1 _ ∣ north ∣ₕ ≡ (sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north))) k) +-- (λ j → cong (βₗ'-fun n f g) (push (θ {A = S₊∙ _} (merid a j)))) +-- (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) +-- lala = (cong-∙∙ (cong (βₗ'-fun n f g) ∘ push) +-- ((λ i₁ → inl ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁))) +-- (push tt) +-- ((λ i₁ → inr ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁))) +-- ∙ sym (compPath≡compPath' _ _) +-- ∙ (λ _ → (λ j → Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∣ (merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) j ∣ₕ) +-- ∙ Kn→ΩKn+10ₖ _) +-- ∙ cong (_∙ Kn→ΩKn+10ₖ _) (cong-∙ ((Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ)) +-- (merid a) (sym (merid north))) +-- ∙ sym (assoc _ _ _) +-- ∙ cong (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a) ∙_) +-- (sym (symDistr (sym ((Kn→ΩKn+10ₖ _))) (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)))))) +-- ◁ λ i j → compPath-filler (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) +-- (sym (sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) +-- (cong ∣_∣ₕ (merid north)))) +-- (~ i) j + +-- q-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → coHomFun _ (q n f g) (βᵣ n f g) +-- ≡ β+ n f g +-- q-βᵣ n f g = cong (coHomFun _ (q n f g)) (βᵣ≡ n f g) +-- ∙ transportLem +-- ∙ cong (subst (λ m → coHom m (C+' n f g)) +-- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) +-- (sym (Iso.leftInv (fst (Hopfβ-Iso n (∙Π f g))) (Hopfβ n (∙Π f g))) +-- ∙ cong (Iso.inv ((fst (Hopfβ-Iso n (∙Π f g))))) (Hopfβ↦1 n (∙Π f g))) +-- where +-- transportLem : coHomFun (suc (suc (suc (n +ℕ suc n)))) (q n f g) +-- (βᵣ''-fun n f g) +-- ≡ subst (λ m → coHom m (C+' n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- (Hopfβ n (∙Π f g)) +-- transportLem = +-- natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (q n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- ∙ cong (subst (λ m → coHom m (C+' n f g)) +-- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) +-- (cong ∣_∣₂ (funExt λ { (inl x) → refl +-- ; (inr x) → refl +-- ; (push a i₁) → loll a i₁})) +-- where +-- open import Cubical.Foundations.Path +-- loll : (x : fst (S₊∙ (3 +ℕ (n +ℕ n)))) → +-- PathP +-- (λ i₁ → +-- βᵣ'-fun n f g ((push (θ x) ∙ λ i → inr (+≡ n f g x (~ i))) i₁) ≡ +-- MV.d-pre Unit (fst (S₊∙ (2 +ℕ n))) (fst (S₊∙ (3 +ℕ n +ℕ n))) +-- (λ _ → tt) (fst (∙Π f g)) (3 +ℕ n +ℕ n) ∣_∣ (push x i₁)) +-- (λ _ → βᵣ'-fun n f g (q n f g (inl tt))) +-- (λ _ → βᵣ'-fun n f g (q n f g (inr (∙Π f g .fst x)))) +-- loll x = flipSquare (cong-∙ (βᵣ'-fun n f g) (push (θ x)) (λ i → inr (+≡ n f g x (~ i))) +-- ∙∙ sym (rUnit _) +-- ∙∙ lem₁ x) +-- where +-- lem₁ : (x : _) → cong (βᵣ'-fun n f g) (push (θ {A = S₊∙ _} x)) ≡ Kn→ΩKn+1 _ ∣ x ∣ +-- lem₁ north = sym (Kn→ΩKn+10ₖ _) +-- lem₁ south = cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)) +-- lem₁ (merid a j) k = lala k j -- lala k j +-- where +-- lala : PathP (λ k → (Kn→ΩKn+10ₖ _) (~ k) ≡ (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north))) k) +-- (λ j → cong (βᵣ'-fun n f g) (push (θ {A = S₊∙ _} (merid a j)))) +-- (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) +-- lala = (cong-∙∙ (cong (βᵣ'-fun n f g) ∘ push) +-- (λ i₁ → inl ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁)) +-- (push tt) +-- (λ i₁ → inr ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁)) +-- ∙∙ (cong (sym (Kn→ΩKn+10ₖ _) ∙_) (cong-∙ (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a) (sym (merid (ptSn _))))) +-- ∙∙ sym (help3 _ _ _)) +-- ◁ symP (doubleCompPath-filler +-- (sym (Kn→ΩKn+10ₖ _)) +-- (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) +-- (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (sym (merid north)))) +-- open import Cubical.Foundations.Path +-- jₗ-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → coHomFun _ (jₗ n f g) (α∨ n f g) +-- ≡ Hopfα n f +-- jₗ-α n f g = +-- cong (coHomFun _ (jₗ n f g)) (α-∨-lem n f g) +-- ∙ cong ∣_∣₂ (funExt (HopfInvariantPushElim n (fst f) +-- (isOfHLevelPath ((suc (suc (suc (suc (n +ℕ n)))))) +-- (isOfHLevelPlus' {n = n} (4 +ℕ n) (isOfHLevelTrunc (4 +ℕ n))) _ _) +-- refl +-- (λ _ → refl) +-- λ j → refl)) + +-- jₗ-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → coHomFun _ (jₗ n f g) (βₗ n f g) +-- ≡ subst (λ m → coHom m (HopfInvariantPush n (fst f))) +-- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- (Hopfβ n f) +-- jₗ-βₗ n f g = +-- cong (coHomFun _ (jₗ n f g)) (βₗ≡ n f g) +-- ∙∙ natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (jₗ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) +-- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) +-- (cong ∣_∣₂ +-- (funExt λ { (inl x) → refl +-- ; (inr x) → refl +-- ; (push a i₁) → refl})) + +-- jₗ-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → coHomFun _ (jₗ n f g) (βᵣ n f g) +-- ≡ 0ₕ _ +-- jₗ-βᵣ n f g = +-- cong (coHomFun _ (jₗ n f g)) (βᵣ≡ n f g) +-- ∙∙ natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (jₗ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) +-- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) +-- cool +-- where +-- cool : coHomFun (suc (suc (suc (suc (n +ℕ n))))) (jₗ n f g) +-- ∣ βᵣ'-fun n f g ∣₂ ≡ 0ₕ _ +-- cool = cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → refl}) + +-- jᵣ-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → coHomFun _ (jᵣ n f g) (α∨ n f g) +-- ≡ Hopfα n g +-- jᵣ-α n f g = cong (coHomFun _ (jᵣ n f g)) (α-∨-lem n f g) +-- ∙ cong ∣_∣₂ (funExt (HopfInvariantPushElim n (fst g) +-- (isOfHLevelPath ((suc (suc (suc (suc (n +ℕ n)))))) +-- (isOfHLevelPlus' {n = n} (4 +ℕ n) (isOfHLevelTrunc (4 +ℕ n))) _ _) +-- refl +-- (λ _ → refl) +-- (flipSquare (pp-inr n f g)))) + +-- jᵣ-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → coHomFun _ (jᵣ n f g) (βₗ n f g) ≡ 0ₕ _ +-- jᵣ-βₗ n f g = +-- cong (coHomFun _ (jᵣ n f g)) (βₗ≡ n f g) +-- ∙∙ natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (jᵣ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) +-- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) +-- cool +-- where +-- cool : coHomFun (suc (suc (suc (suc (n +ℕ n))))) (jᵣ n f g) +-- ∣ βₗ'-fun n f g ∣₂ ≡ 0ₕ _ +-- cool = cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → refl}) + + +-- jᵣ-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → coHomFun _ (jᵣ n f g) (βᵣ n f g) +-- ≡ subst (λ m → coHom m (HopfInvariantPush n (fst g))) +-- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- (Hopfβ n g) +-- jᵣ-βᵣ n f g = +-- cong (coHomFun _ (jᵣ n f g)) (βᵣ≡ n f g) +-- ∙∙ natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (jᵣ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst g))) +-- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) +-- (cong ∣_∣₂ +-- (funExt λ { (inl x) → refl +-- ; (inr x) → refl +-- ; (push a i₁) → refl})) + +-- gen₂ℤ×ℤ : gen₂-by (DirProd ℤGroup ℤGroup) (1 , 0) (0 , 1) +-- fst (gen₂ℤ×ℤ (x , y)) = x , y +-- snd (gen₂ℤ×ℤ (x , y)) = +-- ΣPathP ((cong₂ _+_ ((·Comm 1 x) ∙ cong fst (sym (distrLem 1 0 x))) ((·Comm 0 y) ∙ cong fst (sym (distrLem 0 1 y)))) +-- , +Comm y 0 ∙ cong₂ _+_ (·Comm 0 x ∙ cong snd (sym (distrLem 1 0 x))) (·Comm 1 y ∙ cong snd (sym (distrLem 0 1 y)))) +-- where +-- ℤ×ℤ = DirProd ℤGroup ℤGroup +-- _+''_ = GroupStr._·_ (snd ℤ×ℤ) + +-- ll3 : (x : ℤ) → - x ≡ 0 - x +-- ll3 (pos zero) = refl +-- ll3 (pos (suc zero)) = refl +-- ll3 (pos (suc (suc n))) = +-- cong predℤ (ll3 (pos (suc n))) +-- ll3 (negsuc zero) = refl +-- ll3 (negsuc (suc n)) = cong sucℤ (ll3 (negsuc n)) + +-- distrLem : (x y : ℤ) (z : ℤ) +-- → Path (ℤ × ℤ) (z ℤ[ ℤ×ℤ ]· (x , y)) (z · x , z · y) +-- distrLem x y (pos zero) = refl +-- distrLem x y (pos (suc n)) = +-- (cong ((x , y) +''_) (distrLem x y (pos n))) +-- distrLem x y (negsuc zero) = ΣPathP (sym (ll3 x) , sym (ll3 y)) +-- distrLem x y (negsuc (suc n)) = +-- cong₂ _+''_ (ΣPathP (sym (ll3 x) , sym (ll3 y))) +-- (distrLem x y (negsuc n)) + +-- genH²ⁿC* : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → gen₂-by (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) +-- (βₗ n f g) +-- (βᵣ n f g) +-- genH²ⁿC* n f g = +-- Iso-pres-gen₂ (DirProd ℤGroup ℤGroup) +-- (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) +-- (1 , 0) +-- (0 , 1) +-- gen₂ℤ×ℤ +-- (invGroupIso (H⁴-C* n f g)) + +-- private + +-- H⁴C* : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Group _ +-- H⁴C* n f g = coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) + +-- X : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → ℤ +-- X n f g = (genH²ⁿC* n f g) (α∨ n f g ⌣ α∨ n f g) .fst .fst +-- Y : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → ℤ +-- Y n f g = (genH²ⁿC* n f g) (α∨ n f g ⌣ α∨ n f g) .fst .snd + +-- α∨≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → α∨ n f g ⌣ α∨ n f g ≡ ((X n f g) ℤ[ H⁴C* n f g ]· βₗ n f g) +-- +ₕ ((Y n f g) ℤ[ H⁴C* n f g ]· βᵣ n f g) +-- α∨≡ n f g = (genH²ⁿC* n f g) (α∨ n f g ⌣ α∨ n f g) .snd + + +-- eq₁ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → (Hopfα n (∙Π f g)) ⌣ (Hopfα n (∙Π f g)) +-- ≡ ((X n f g + Y n f g) ℤ[ coHomGr _ _ ]· (β+ n f g)) +-- eq₁ n f g = +-- cong (λ x → x ⌣ x) (sym (q-α n f g) ∙ cong (coHomFun (suc (suc n)) (q n f g)) (α-∨-lem n f g)) +-- ∙ cong ((coHomFun _) (q _ f g)) (cong (λ x → x ⌣ x) (sym (α-∨-lem n f g)) +-- ∙ α∨≡ n f g) +-- ∙ IsGroupHom.pres· (coHomMorph _ (q n f g) .snd) _ _ +-- ∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (q n f g)) (βₗ n f g) (X n f g) +-- ∙ cong (λ z → (X n f g) ℤ[ coHomGr _ _ ]· z) +-- (q-βₗ n f g)) +-- (homPresℤ· (coHomMorph _ (q n f g)) (βᵣ n f g) (Y n f g) +-- ∙ cong (λ z → (Y n f g) ℤ[ coHomGr _ _ ]· z) +-- (q-βᵣ n f g)) +-- ∙ sym (distrℤ· (coHomGr _ _) (β+ n f g) (X n f g) (Y n f g)) + +-- coHomFun⌣ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) +-- → (n : ℕ) (x y : coHom n B) +-- → coHomFun _ f (x ⌣ y) ≡ coHomFun _ f x ⌣ coHomFun _ f y +-- coHomFun⌣ f n = sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl + +-- eq₂ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → Hopfα n f ⌣ Hopfα n f +-- ≡ (X n f g ℤ[ coHomGr _ _ ]· subst (λ m → coHom m (HopfInvariantPush n (fst f))) +-- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- (Hopfβ n f)) +-- eq₂ n f g = +-- cong (λ x → x ⌣ x) (sym (jₗ-α n f g)) +-- ∙∙ sym (coHomFun⌣ (jₗ n f g) _ _ _) +-- ∙∙ cong (coHomFun _ (jₗ n f g)) (α∨≡ n f g) +-- ∙∙ IsGroupHom.pres· (snd (coHomMorph _ (jₗ n f g))) _ _ +-- ∙∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (jₗ n f g)) (βₗ n f g) (X n f g)) +-- (homPresℤ· (coHomMorph _ (jₗ n f g)) (βᵣ n f g) (Y n f g) +-- ∙ cong (λ z → (Y n f g) ℤ[ coHomGr _ _ ]· z) +-- (jₗ-βᵣ n f g)) +-- ∙∙ cong₂ _+ₕ_ refl (rUnitℤ· (coHomGr _ _) (Y n f g)) +-- ∙∙ (rUnitₕ _ _ +-- ∙ cong (X n f g ℤ[ coHomGr _ _ ]·_) (jₗ-βₗ n f g)) + +-- eq₃ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +-- → Hopfα n g ⌣ Hopfα n g +-- ≡ (Y n f g ℤ[ coHomGr _ _ ]· subst (λ m → coHom m (HopfInvariantPush n (fst f))) +-- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) +-- (Hopfβ n g)) +-- eq₃ n f g = +-- cong (λ x → x ⌣ x) (sym (jᵣ-α n f g)) +-- ∙∙ sym (coHomFun⌣ (jᵣ n f g) _ _ _) +-- ∙∙ cong (coHomFun _ (jᵣ n f g)) (α∨≡ n f g) +-- ∙∙ IsGroupHom.pres· (snd (coHomMorph _ (jᵣ n f g))) _ _ +-- ∙∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (jᵣ n f g)) (βₗ n f g) (X n f g) +-- ∙ cong (λ z → (X n f g) ℤ[ coHomGr _ _ ]· z) +-- (jᵣ-βₗ n f g)) +-- (homPresℤ· (coHomMorph _ (jᵣ n f g)) (βᵣ n f g) (Y n f g)) +-- ∙∙ cong₂ _+ₕ_ (rUnitℤ· (coHomGr _ _) (X n f g)) refl +-- ∙∙ ((lUnitₕ _ _) ∙ cong (Y n f g ℤ[ coHomGr _ _ ]·_) (jᵣ-βᵣ n f g)) + +-- eq₂-eq₂ : (n : ℕ) (f g : S₊∙ (suc (suc (suc (n +ℕ n)))) →∙ S₊∙ (suc (suc n))) +-- → HopfInvariant n (∙Π f g) ≡ HopfInvariant n f + HopfInvariant n g +-- eq₂-eq₂ n f g = +-- mainL +-- ∙ sym (cong₂ _+_ (help n f g) (helpg n f g)) +-- where +-- transpLem : ∀ {ℓ} {G : ℕ → Group ℓ} +-- → (n m : ℕ) +-- → (x : ℤ) +-- → (p : m ≡ n) +-- → (g : fst (G n)) +-- → subst (fst ∘ G) p (x ℤ[ G m ]· subst (fst ∘ G) (sym p) g) ≡ (x ℤ[ G n ]· g) +-- transpLem {G = G} n m x = +-- J (λ n p → (g : fst (G n)) → subst (fst ∘ G) p (x ℤ[ G m ]· subst (fst ∘ G) (sym p) g) +-- ≡ (x ℤ[ G n ]· g)) +-- λ g → transportRefl _ ∙ cong (x ℤ[ G m ]·_) (transportRefl _) + +-- mainL : HopfInvariant n (∙Π f g) ≡ X n f g + Y n f g +-- mainL = cong (Iso.fun (fst (Hopfβ-Iso n (∙Π f g)))) +-- (cong (subst (λ x → coHom x (HopfInvariantPush n (fst (∙Π f g)))) +-- λ i₁ → suc (suc (suc (+-suc n n i₁)))) +-- (eq₁ n f g)) +-- ∙∙ cong (Iso.fun (fst (Hopfβ-Iso n (∙Π f g)))) +-- (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst (∙Π f g)))} _ _ +-- (X n f g + Y n f g) (λ i₁ → suc (suc (suc (+-suc n n i₁)))) +-- (Iso.inv (fst (Hopfβ-Iso n (∙Π f g))) 1)) +-- ∙∙ homPresℤ· (_ , snd (Hopfβ-Iso n (∙Π f g))) (Iso.inv (fst (Hopfβ-Iso n (∙Π f g))) 1) +-- (X n f g + Y n f g) +-- ∙∙ cong ((X n f g + Y n f g) ℤ[ ℤ , ℤGroup .snd ]·_) +-- (Iso.rightInv ((fst (Hopfβ-Iso n (∙Π f g)))) 1) +-- ∙∙ rUnitℤ·ℤ (X n f g + Y n f g) + + +-- help : (n : ℕ) (f g : _) → HopfInvariant n f ≡ X n f g +-- help n f g = cong (Iso.fun (fst (Hopfβ-Iso n f))) +-- (cong (subst (λ x → coHom x (HopfInvariantPush n (fst f))) +-- (λ i₁ → suc (suc (suc (+-suc n n i₁))))) (eq₂ n f g)) +-- ∙ cong (Iso.fun (fst (Hopfβ-Iso n f))) +-- (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst f))} _ _ +-- (X n f g) +-- ((cong (suc ∘ suc ∘ suc) (+-suc n n))) +-- (Hopfβ n f)) +-- ∙ homPresℤ· (_ , snd (Hopfβ-Iso n f)) (Hopfβ n f) (X n f g) +-- ∙ cong (X n f g ℤ[ ℤ , ℤGroup .snd ]·_) (Hopfβ↦1 n f) +-- ∙ rUnitℤ·ℤ (X n f g) + +-- helpg : (n : ℕ) (f g : _) → HopfInvariant n g ≡ Y n f g +-- helpg n f g = +-- cong (Iso.fun (fst (Hopfβ-Iso n g))) +-- (cong (subst (λ x → coHom x (HopfInvariantPush n (fst g))) +-- (λ i₁ → suc (suc (suc (+-suc n n i₁))))) +-- (eq₃ n f g)) +-- ∙∙ cong (Iso.fun (fst (Hopfβ-Iso n g))) +-- (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst g))} _ _ +-- (Y n f g) +-- ((cong (suc ∘ suc ∘ suc) (+-suc n n))) +-- (Hopfβ n g)) +-- ∙∙ homPresℤ· (_ , snd (Hopfβ-Iso n g)) (Hopfβ n g) (Y n f g) +-- ∙∙ cong (Y n f g ℤ[ ℤ , ℤGroup .snd ]·_) (Hopfβ↦1 n g) +-- ∙∙ rUnitℤ·ℤ (Y n f g) diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda index 7b206900e..f42c4248b 100644 --- a/Cubical/Homotopy/Loopspace.agda +++ b/Cubical/Homotopy/Loopspace.agda @@ -27,10 +27,6 @@ open Iso (Ω^ 0) p = p (Ω^ (suc n)) p = Ω ((Ω^ n) p) -{- homotopy Group -} -π : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Type ℓ -π n A = ∥ typ ((Ω^ n) A) ∥ 2 - {- loop space map -} Ω→ : ∀ {ℓA ℓB} {A : Pointed ℓA} {B : Pointed ℓB} (f : A →∙ B) → (Ω A →∙ Ω B) Ω→ (f , f∙) = (λ p → (sym f∙ ∙ cong f p) ∙ f∙) , cong (λ q → q ∙ f∙) (sym (rUnit (sym f∙))) ∙ lCancel f∙ diff --git a/Cubical/ZCohomology/Gysin.agda b/Cubical/ZCohomology/Gysin.agda index ac1cf61c4..fea5e95b2 100644 --- a/Cubical/ZCohomology/Gysin.agda +++ b/Cubical/ZCohomology/Gysin.agda @@ -12,15 +12,13 @@ open import Cubical.ZCohomology.Groups.Sn open import Cubical.ZCohomology.RingStructure.CupProduct open import Cubical.ZCohomology.RingStructure.RingLaws open import Cubical.ZCohomology.RingStructure.GradedCommutativity -open import Cubical.Foundations.Path - -open import Cubical.Data.Sum open import Cubical.Relation.Nullary +open import Cubical.Homotopy.HomotopyGroup open import Cubical.Functions.Embedding -open import Cubical.Data.Fin - +open import Cubical.Foundations.Path +open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.HLevels open import Cubical.Foundations.Transport open import Cubical.Foundations.Function @@ -36,6 +34,8 @@ open import Cubical.Foundations.Univalence open import Cubical.Relation.Nullary +open import Cubical.Data.Sum +open import Cubical.Data.Fin open import Cubical.Data.Empty renaming (rec to ⊥-rec) open import Cubical.Data.Sigma open import Cubical.Data.Int hiding (_+'_) @@ -69,267 +69,6 @@ open import Cubical.HITs.Join open import Cubical.Homotopy.Hopf -open import Cubical.HITs.SetQuotients renaming (_/_ to _/'_) - -isContr→≡UnitGroup : {G : Group ℓ-zero} → isContr (fst G) → UnitGroup ≡ G -isContr→≡UnitGroup c = - fst (GroupPath _ _) - (invGroupEquiv ((isContr→≃Unit c) - , (makeIsGroupHom (λ _ _ → refl)))) - -GroupIsoUnitGroup→isContr : {G : Group ℓ-zero} → GroupIso UnitGroup G → isContr (fst G) -GroupIsoUnitGroup→isContr is = - isOfHLevelRetractFromIso 0 (invIso (fst is)) isContrUnit - --- move to group stuff -SES→isEquiv : ∀ {ℓ ℓ'} {L R : Group ℓ-zero} - → {G : Group ℓ} {H : Group ℓ'} - → UnitGroup ≡ L - → UnitGroup ≡ R - → (lhom : GroupHom L G) (midhom : GroupHom G H) (rhom : GroupHom H R) - → ((x : _) → isInKer midhom x → isInIm lhom x) - → ((x : _) → isInKer rhom x → isInIm midhom x) - → isEquiv (fst midhom) -SES→isEquiv {R = R} {G = G} {H = H} = - J (λ L _ → UnitGroup ≡ R → - (lhom : GroupHom L G) (midhom : GroupHom G H) - (rhom : GroupHom H R) → - ((x : fst G) → isInKer midhom x → isInIm lhom x) → - ((x : fst H) → isInKer rhom x → isInIm midhom x) → - isEquiv (fst midhom)) - ((J (λ R _ → (lhom : GroupHom UnitGroup G) (midhom : GroupHom G H) - (rhom : GroupHom H R) → - ((x : fst G) → isInKer midhom x → isInIm lhom x) → - ((x : _) → isInKer rhom x → isInIm midhom x) → - isEquiv (fst midhom)) - main)) - where - main : (lhom : GroupHom UnitGroup G) (midhom : GroupHom G H) - (rhom : GroupHom H UnitGroup) → - ((x : fst G) → isInKer midhom x → isInIm lhom x) → - ((x : fst H) → isInKer rhom x → isInIm midhom x) → - isEquiv (fst midhom) - main lhom midhom rhom lexact rexact = - BijectionIsoToGroupEquiv {G = G} {H = H} - bijIso' .fst .snd - where - bijIso' : BijectionIso G H - BijectionIso.fun bijIso' = midhom - BijectionIso.inj bijIso' x inker = - pRec (GroupStr.is-set (snd G) _ _) - (λ s → sym (snd s) ∙ IsGroupHom.pres1 (snd lhom)) - (lexact _ inker) - BijectionIso.surj bijIso' x = rexact x refl - --- The pushout of the hopf invariant -HopfInvariantPush : (n : ℕ) - → (f : S₊ (3 +ℕ n +ℕ n) → S₊ (2 +ℕ n)) - → Type _ -HopfInvariantPush n f = Pushout (λ _ → tt) f - -Hopfα : (n : ℕ) - → (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) - → coHom (2 +ℕ n) (HopfInvariantPush n (fst f)) -Hopfα n f = ∣ (λ { (inl x) → 0ₖ _ - ; (inr x) → ∣ x ∣ - ; (push a i) → help a (~ i)}) ∣₂ - where - help : (a : S₊ (3 +ℕ n +ℕ n)) → ∣ fst f a ∣ ≡ 0ₖ (2 +ℕ n) - help = sphereElim _ (λ _ → isOfHLevelPlus' {n = n} (3 +ℕ n) - (isOfHLevelPath' (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n)) _ _)) (cong ∣_∣ₕ (snd f)) - -Hopfβ : (n : ℕ) - → (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) - → coHom (4 +ℕ (n +ℕ n)) (HopfInvariantPush n (fst f)) -Hopfβ n f = fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) ∣ ∣_∣ ∣₂ - --- To define the Hopf invariant, we need to establish the non-trivial isos of Hⁿ(HopfInvariant). -module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) where - module M = MV _ _ _ (λ _ → tt) (fst f) - ¬lem : (n m : ℕ) → ¬ suc (suc (m +ℕ n)) ≡ suc n - ¬lem zero zero p = snotz (cong predℕ p) - ¬lem (suc n) zero p = ¬lem n zero (cong predℕ p) - ¬lem zero (suc m) p = snotz (cong predℕ p) - ¬lem (suc n) (suc m) p = - ¬lem n (suc m) (cong (suc ∘ suc ) - (sym (+-suc m n)) ∙ (cong predℕ p)) - - SphereHopfCohomIso : - GroupEquiv - (coHomGr (3 +ℕ n +ℕ n) (S₊ (3 +ℕ n +ℕ n))) - ((coHomGr (suc (3 +ℕ n +ℕ n)) (HopfInvariantPush _ (fst f)))) - fst (fst SphereHopfCohomIso) = fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) - snd (fst SphereHopfCohomIso) = help - where - abstract - help : isEquiv (fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n))) - help = - SES→isEquiv - (isContr→≡UnitGroup - (isContrΣ (GroupIsoUnitGroup→isContr (invGroupIso (Hⁿ-Unit≅0 _))) - λ _ → GroupIsoUnitGroup→isContr - (invGroupIso (Hⁿ-Sᵐ≅0 _ _ (¬lem n n))))) - (isContr→≡UnitGroup - (isContrΣ (GroupIsoUnitGroup→isContr - (invGroupIso (Hⁿ-Unit≅0 _))) - λ _ → GroupIsoUnitGroup→isContr - (invGroupIso (Hⁿ-Sᵐ≅0 _ _ (¬lem n (suc n)))))) - (M.Δ (3 +ℕ n +ℕ n)) - (M.d (3 +ℕ n +ℕ n)) - (M.i (4 +ℕ n +ℕ n)) - (M.Ker-d⊂Im-Δ _) - (M.Ker-i⊂Im-d _) - snd SphereHopfCohomIso = s - where - -- Currently, type checking is very slow without the abstract flag. - -- TODO : Remove abstract - abstract - s : IsGroupHom (snd (coHomGr (3 +ℕ n +ℕ n) (S₊ (3 +ℕ n +ℕ n)))) - (fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n))) - (snd (coHomGr (suc (3 +ℕ n +ℕ n)) (Pushout (λ _ → tt) (fst f)))) - s = snd (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) - - Hopfβ-Iso : GroupIso (coHomGr (suc (3 +ℕ n +ℕ n)) (HopfInvariantPush _ (fst f))) - ℤGroup - Iso.fun (fst Hopfβ-Iso) x = - Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) - (invEq (fst SphereHopfCohomIso) x) - Iso.inv (fst Hopfβ-Iso) f = - (fst (fst SphereHopfCohomIso)) - (Iso.inv (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) f) - Iso.rightInv (fst Hopfβ-Iso) f = - cong (Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) - (retEq (fst SphereHopfCohomIso) (Iso.inv (fst (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) f)) - ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) f - Iso.leftInv (fst Hopfβ-Iso) x = - cong (fst (fst SphereHopfCohomIso)) - (Iso.leftInv (fst (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) - (invEq (fst SphereHopfCohomIso) x)) - ∙ secEq (fst SphereHopfCohomIso) x - snd Hopfβ-Iso = grHom - where - abstract - grHom : IsGroupHom (coHomGr (suc (3 +ℕ n +ℕ n)) (Pushout (λ _ → tt) (fst f)) .snd) - (λ x → Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) - (invEq (fst SphereHopfCohomIso) x)) - (ℤGroup .snd) - grHom = makeIsGroupHom λ x y → - cong (Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) - (IsGroupHom.pres· (isGroupHomInv SphereHopfCohomIso) x y) - ∙ IsGroupHom.pres· (snd (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) _ _ - -Hⁿ-Sⁿ≅ℤ-nice-generator : (n : ℕ) → Iso.inv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) 1 ≡ ∣ ∣_∣ ∣₂ -Hⁿ-Sⁿ≅ℤ-nice-generator zero = Iso.leftInv (fst (Hⁿ-Sⁿ≅ℤ (suc zero))) _ -Hⁿ-Sⁿ≅ℤ-nice-generator (suc n) = - (λ i → Iso.inv (fst (suspensionAx-Sn (suc n) (suc n))) (Hⁿ-Sⁿ≅ℤ-nice-generator n i)) - ∙ cong ∣_∣₂ (funExt λ { north → refl - ; south → cong ∣_∣ₕ (merid north) - ; (merid a i) j → ∣ compPath-filler (merid a) (sym (merid north)) (~ j) i ∣ₕ}) - -Hopfβ↦1 : (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) - → Iso.fun (fst (Hopfβ-Iso n f)) (Hopfβ n f) ≡ 1 -Hopfβ↦1 n f = sym (cong (Iso.fun (fst (Hopfβ-Iso n f))) lem) ∙ Iso.rightInv (fst (Hopfβ-Iso n f)) (pos 1) - where - lem : Iso.inv (fst (Hopfβ-Iso n f)) (pos 1) ≡ Hopfβ n f - lem = (λ i → fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) - (Hⁿ-Sⁿ≅ℤ-nice-generator _ i)) - ∙ cong ∣_∣₂ (funExt (λ { (inl x) → refl - ; (inr x) → refl - ; (push a i) → refl})) - -module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) where - private - 2+n = 2 +ℕ n - H = HopfInvariantPush n (fst f) - - H→Sphere : coHom 2+n H → coHom 2+n (S₊ (suc (suc n))) - H→Sphere = sMap (_∘ inr) - - grHom : IsGroupHom (snd (coHomGr 2+n H)) H→Sphere (snd (coHomGr 2+n (S₊ (suc (suc n))))) - grHom = - makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) - λ f g → refl) - - h' : (g : (S₊ (suc (suc n)) → coHomK 2+n)) → H → coHomK (2 +ℕ n) - h' g (inl x) = 0ₖ _ - h' g (inr x) = g x -ₖ g north - h' g (push a i) = h23 a i - where - h23 : (a : S₊ (suc (suc (suc (n +ℕ n))))) → 0ₖ (suc (suc n)) ≡ g (fst f a) -ₖ g north - h23 = sphereElim _ (λ x → isOfHLevelPlus' {n = n} (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n) _ _)) - (sym (rCancelₖ _ (g north)) ∙ cong (λ x → g x -ₖ g north) (sym (snd f))) - - Sphere→H : coHom 2+n (S₊ (suc (suc n))) → coHom 2+n H - Sphere→H = sMap h' - - conCohom2+n : (x : _) → ∥ x ≡ 0ₖ (suc (suc n)) ∥ - conCohom2+n = coHomK-elim _ (λ _ → isProp→isOfHLevelSuc (suc n) squash) ∣ refl ∣ - - HIPSphereCohomIso : Iso (coHom (2 +ℕ n) (HopfInvariantPush n (fst f))) - (coHom (2 +ℕ n) ((S₊ (suc (suc n))))) - Iso.fun HIPSphereCohomIso = H→Sphere - Iso.inv HIPSphereCohomIso = Sphere→H - Iso.rightInv HIPSphereCohomIso = - sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - λ g → pRec (squash₂ _ _) - (λ p → cong ∣_∣₂ (funExt λ x → cong (g x +ₖ_) (cong (-ₖ_) p) ∙ rUnitₖ _ (g x))) - (conCohom2+n (g north)) - Iso.leftInv HIPSphereCohomIso = - sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - λ g → pRec (squash₂ _ _) - (pRec (isPropΠ (λ _ → squash₂ _ _)) - (λ gn gtt → - trRec (isProp→isOfHLevelSuc n (squash₂ _ _)) - (λ r → cong ∣_∣₂ (funExt λ { - (inl x) → sym gtt - ; (inr x) → (λ i → g (inr x) -ₖ gn i) ∙ rUnitₖ _ (g (inr x)) - ; (push a i) - → sphereElim _ - {A = λ a → PathP (λ i → h' (λ x → g (inr x)) (push a i) ≡ g (push a i)) - (sym gtt) - ((λ i → g (inr (fst f a)) -ₖ gn i) ∙ rUnitₖ _ (g (inr (fst f a))))} - (λ _ → isOfHLevelPathP' (suc (suc (suc (n +ℕ n)))) - (isOfHLevelPath (suc (suc (suc (suc (n +ℕ n))))) - (isOfHLevelPlus' {n = n} (suc (suc (suc (suc n)))) - (isOfHLevelTrunc (suc (suc (suc (suc n)))))) _ _) _ _) - r a i})) - (l g gtt gn)) - (conCohom2+n (g (inr north)))) - (conCohom2+n (g (inl tt))) - where - l : (g : HopfInvariantPush n (fst f) → coHomK (suc (suc n))) - → (gtt : g (inl tt) ≡ 0ₖ (suc (suc n))) - → (gn : g (inr north) ≡ 0ₖ (suc (suc n))) - → hLevelTrunc (suc n) (PathP (λ i → h' (λ x → g (inr x)) (push north i) ≡ g (push north i)) - (sym gtt) ((λ i → g (inr (fst f north)) -ₖ gn i) ∙ rUnitₖ _ (g (inr (fst f north))))) - l g gtt gn = isConnectedPathP _ (isConnectedPath _ (isConnectedKn _) _ _) _ _ .fst - - Hopfα-Iso : GroupIso (coHomGr (2 +ℕ n) (HopfInvariantPush n (fst f))) ℤGroup - Hopfα-Iso = - compGroupIso - (HIPSphereCohomIso , grHom) - (Hⁿ-Sⁿ≅ℤ (suc n)) - -Hopfα-Iso-α : (n : ℕ) (f : _) - → Iso.fun (fst (Hopfα-Iso n f)) (Hopfα n f) - ≡ 1 -Hopfα-Iso-α n f = - sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (Hⁿ-Sⁿ≅ℤ-nice-generator n)) - ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) - where - hz : Iso.fun (HIPSphereCohomIso n f) (Hopfα n f) ≡ ∣ ∣_∣ ∣₂ - hz = refl - -⌣-α : (n : ℕ) → (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → _ -⌣-α n f = Hopfα n f ⌣ Hopfα n f - -HopfInvariant : (n : ℕ) - → (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → ℤ -HopfInvariant n f = Iso.fun (fst (Hopfβ-Iso n f)) - ((subst (λ x → coHom x (HopfInvariantPush n (fst f))) - (λ i → suc (suc (suc (+-suc n n i))))) (⌣-α n f)) open import Cubical.Algebra.AbGroup @@ -342,6 +81,52 @@ open import Cubical.Homotopy.Hopf open import Cubical.HITs.SetQuotients renaming (_/_ to _/'_) -- open import Cubical.ZCohomology.Gysin +-- There seems to be some problems with the termination checker. +-- Spelling out integer induction with 3 base cases like this +-- solves the issue. +private + Int-ind : ∀ {ℓ} (P : ℤ → Type ℓ) + → P (pos zero) → P (pos 1) + → P (negsuc zero) + → ((x y : ℤ) → P x → P y → P (x + y)) → (x : ℤ) → P x + Int-ind P z one min ind (pos zero) = z + Int-ind P z one min ind (pos (suc zero)) = one + Int-ind P z one min ind (pos (suc (suc n))) = + ind (pos (suc n)) 1 (Int-ind P z one min ind (pos (suc n))) one + Int-ind P z one min ind (negsuc zero) = min + Int-ind P z one min ind (negsuc (suc n)) = + ind (negsuc n) (negsuc zero) (Int-ind P z one min ind (negsuc n)) min + + +-- Move to embedding? +Whitehead : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → (f : A → B) + → isEmbedding f + → (∀ (b : B) → ∃[ a ∈ A ] f a ≡ b) + → isEquiv f +equiv-proof (Whitehead f emb p) b = + pRec isPropIsContr + (λ fib → fib , isEmbedding→hasPropFibers emb b fib) + (p b) + +-- move to cohom? +transportCohomIso : ∀ {ℓ} {A : Type ℓ} {n m : ℕ} + → (p : n ≡ m) + → GroupIso (coHomGr n A) (coHomGr m A) +Iso.fun (fst (transportCohomIso {A = A} p)) = subst (λ n → coHom n A) p +Iso.inv (fst (transportCohomIso {A = A} p)) = subst (λ n → coHom n A) (sym p) +Iso.rightInv (fst (transportCohomIso p)) = transportTransport⁻ _ +Iso.leftInv (fst (transportCohomIso p)) = transportTransport⁻ _ +snd (transportCohomIso {A = A} {n = n} {m = m} p) = + makeIsGroupHom (λ x y → help x y p) + where + help : (x y : coHom n A) (p : n ≡ m) → subst (λ n → coHom n A) p (x +ₕ y) + ≡ subst (λ n → coHom n A) p x +ₕ subst (λ n → coHom n A) p y + help x y = J (λ m p → subst (λ n → coHom n A) p (x +ₕ y) + ≡ subst (λ n → coHom n A) p x +ₕ subst (λ n → coHom n A) p y) + (transportRefl (x +ₕ y) ∙ cong₂ _+ₕ_ (sym (transportRefl x)) (sym (transportRefl y))) + + private _·₀ₕ_ : ∀ {ℓ} {n : ℕ} {A : Type ℓ} → ℤ → coHom n A → coHom n A _·₀ₕ_ {n = n} a b = a ℤ[ coHomGr n _ ]· b @@ -355,217 +140,35 @@ private cong (fst e) (ℤ·≡· a b) ∙ homPresℤ· {H = coHomGr n A} (_ , snd e) b a --- goal: remove -genH : ∀ {ℓ} {A : Type ℓ} (n : ℕ) - → (e : GroupIso (coHomGr n A) - ℤGroup) - → Σ[ x ∈ coHom n A ] - ((y : coHom n A) - → Σ[ a ∈ ℤ ] a ·₀ₕ x ≡ y) -genH {A = A} n e = (Iso.inv (fst e) 1) - , λ y → (Iso.fun (fst e) y) - , (sym (GroupHomℤ→ℤPres·₀ n (_ , snd (invGroupIso e)) (Iso.fun (fst e) y) (pos 1)) - ∙∙ cong (Iso.inv (fst e)) (·Comm (Iso.fun (fst e) y) (pos 1)) - ∙∙ Iso.leftInv (fst e) y) -- keep -genH' : ∀ {ℓ} {A : Type ℓ} (n : ℕ) +genH : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → (e : GroupIso (coHomGr n A) ℤGroup) → Σ[ x ∈ coHom n A ] gen₁-by (coHomGr n A) x -genH' {A = A} n e = (Iso.inv (fst e) 1) +genH {A = A} n e = (Iso.inv (fst e) 1) , λ h → (Iso.fun (fst e) h) , (sym (Iso.leftInv (fst e) h) ∙∙ sym (cong (Iso.inv (fst e)) (·Comm (Iso.fun (fst e) h) (pos 1))) ∙∙ (cong (Iso.inv (fst e)) (ℤ·≡· _ _) ∙ (homPresℤ· (_ , snd (invGroupIso e)) (pos 1) (Iso.fun (fst e) h)))) --- Move to intGroup - -⌣equiv→pres1 : ∀ {ℓ} {G H : Type ℓ} → (G ≡ H) → (g₁ : coHom 2 G) (h₁ : coHom 2 H) - → (fstEq : (Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] abs (fst (fst ϕ) g₁) ≡ 1)) - → (sndEq : ((Σ[ ϕ ∈ GroupEquiv (coHomGr 2 H) ℤGroup ] abs (fst (fst ϕ) h₁) ≡ 1))) - → isEquiv {A = coHom 2 G} {B = coHom 4 G} (_⌣ g₁) - → (3rdEq : GroupEquiv (coHomGr 4 H) ℤGroup) - → abs (fst (fst 3rdEq) (h₁ ⌣ h₁)) ≡ 1 -⌣equiv→pres1 {G = G} = J (λ H _ → (g₁ : coHom 2 G) (h₁ : coHom 2 H) - → (fstEq : (Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] abs (fst (fst ϕ) g₁) ≡ 1)) - → (sndEq : ((Σ[ ϕ ∈ GroupEquiv (coHomGr 2 H) ℤGroup ] abs (fst (fst ϕ) h₁) ≡ 1))) - → isEquiv {A = coHom 2 G} {B = coHom 4 G} (_⌣ g₁) - → (3rdEq : GroupEquiv (coHomGr 4 H) ℤGroup) - → abs (fst (fst 3rdEq) (h₁ ⌣ h₁)) ≡ 1) - help - where - help : (g₁ h₁ : coHom 2 G) - → (fstEq : (Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] abs (fst (fst ϕ) g₁) ≡ 1)) - → (sndEq : ((Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] abs (fst (fst ϕ) h₁) ≡ 1))) - → isEquiv {A = coHom 2 G} {B = coHom 4 G} (_⌣ g₁) - → (3rdEq : GroupEquiv (coHomGr 4 G) ℤGroup) - → abs (fst (fst 3rdEq) (h₁ ⌣ h₁)) ≡ 1 - help g h (ϕ , idg) (ψ , idh) ⌣eq ξ = - ⊎→abs _ _ - (allisoPresGen _ - (compGroupEquiv main4 (compGroupEquiv (invGroupEquiv main4) ψ)) - h (abs→⊎ _ _ (cong abs (cong (fst (fst ψ)) (retEq (fst main4) h)) ∙ idh)) (compGroupEquiv main4 ξ)) - where - k1 : ((fst (fst ψ) h) ≡ 1) ⊎ (fst (fst ψ) h ≡ -1) - → abs (fst (fst ϕ) h) ≡ 1 - k1 p = ⊎→abs _ _ (allisoPresGen _ ψ h p ϕ) - - help2 : ((fst (fst ϕ) h) ≡ 1) ⊎ (fst (fst ϕ) h ≡ -1) - → ((fst (fst ϕ) g) ≡ 1) ⊎ (fst (fst ϕ) g ≡ -1) - → (h ≡ g) ⊎ (h ≡ (-ₕ g)) - help2 (inl x) (inl x₁) = - inl (sym (retEq (fst ϕ) h) - ∙∙ cong (invEq (fst ϕ)) (x ∙ sym x₁) - ∙∙ retEq (fst ϕ) g) - help2 (inl x) (inr x₁) = - inr (sym (retEq (fst ϕ) h) - ∙∙ cong (invEq (fst ϕ)) x - ∙ IsGroupHom.presinv (snd (invGroupEquiv ϕ)) (negsuc zero) - ∙∙ cong (-ₕ_) (cong (invEq (fst ϕ)) (sym x₁) ∙ (retEq (fst ϕ) g))) - help2 (inr x) (inl x₁) = - inr (sym (retEq (fst ϕ) h) - ∙∙ cong (invEq (fst ϕ)) x - ∙∙ (IsGroupHom.presinv (snd (invGroupEquiv ϕ)) 1 - ∙ cong (-ₕ_) (cong (invEq (fst ϕ)) (sym x₁) ∙ (retEq (fst ϕ) g)))) - help2 (inr x) (inr x₁) = - inl (sym (retEq (fst ϕ) h) - ∙∙ cong (invEq (fst ϕ)) (x ∙ sym x₁) - ∙∙ retEq (fst ϕ) g) - - -ₕeq : ∀ {ℓ} {A : Type ℓ} (n : ℕ) → Iso (coHom n A) (coHom n A) - Iso.fun (-ₕeq n) = -ₕ_ - Iso.inv (-ₕeq n) = -ₕ_ - Iso.rightInv (-ₕeq n) = - sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - λ f → cong ∣_∣₂ (funExt λ x → -ₖ^2 (f x)) - Iso.leftInv (-ₕeq n) = - sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - λ f → cong ∣_∣₂ (funExt λ x → -ₖ^2 (f x)) - - theEq : coHom 2 G ≃ coHom 4 G - theEq = compEquiv (_ , ⌣eq) (isoToEquiv (-ₕeq 4)) - - main3 : (h ≡ g) ⊎ (h ≡ (-ₕ g)) → isEquiv {A = coHom 2 G} (_⌣ h) - main3 (inl x) = subst isEquiv (λ i → _⌣ (x (~ i))) ⌣eq - main3 (inr x) = - subst isEquiv (funExt (λ x → -ₕDistᵣ 2 2 x g) ∙ (λ i → _⌣ (x (~ i)))) - (theEq .snd) - - main4 : GroupEquiv (coHomGr 2 G) (coHomGr 4 G) - fst main4 = _ , (main3 (help2 (abs→⊎ _ _ (k1 (abs→⊎ _ _ idh))) (abs→⊎ _ _ idg))) - snd main4 = - makeIsGroupHom λ g1 g2 → rightDistr-⌣ _ _ g1 g2 h - --- move to S1 (not there already -- check Sn?) -characFunSpaceS¹ : ∀ {ℓ} {A : Type ℓ} → - Iso (S₊ 1 → A) (Σ[ x ∈ A ] x ≡ x) -Iso.fun characFunSpaceS¹ f = f base , cong f loop -Iso.inv characFunSpaceS¹ (x , p) base = x -Iso.inv characFunSpaceS¹ (x , p) (loop i) = p i -Iso.rightInv characFunSpaceS¹ _ = refl -Iso.leftInv characFunSpaceS¹ f = funExt λ { base → refl ; (loop i) → refl} - -characFunSpace : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} - → Iso (join A B → C) - (Σ[ f ∈ (A → C) ] Σ[ g ∈ (B → C) ] - ((a : A) (b : B) → f a ≡ g b)) -Iso.fun characFunSpace f = (f ∘ inl) , ((f ∘ inr) , (λ a b → cong f (push a b))) -Iso.inv characFunSpace (f , g , p) (inl x) = f x -Iso.inv characFunSpace (f , g , p) (inr x) = g x -Iso.inv characFunSpace (f , g , p) (push a b i) = p a b i -Iso.rightInv characFunSpace (f , g , p) = refl -Iso.leftInv characFunSpace f = - funExt λ { (inl x) → refl ; (inr x) → refl ; (push a b i) → refl} - -coHomS¹-ish : (n : ℕ) → Type _ -coHomS¹-ish n = hLevelTrunc 3 (S₊ 1 → coHomK (3 +ℕ n)) - -fib : (n : ℕ) → coHomS¹-ish n → Type _ -fib n x = - trRec {B = TypeOfHLevel ℓ-zero 2} (isOfHLevelTypeOfHLevel 2) - (λ f → ∥ (Σ[ g ∈ (S₊ 3 → coHomK (3 +ℕ n)) ] - ((a : S₊ 1) (b : S₊ 3) → f a ≡ g b)) ∥₂ , squash₂) x .fst - -Iso1 : (n : ℕ) → Iso (coHom (3 +ℕ n) (join S¹ (S₊ 3))) ∥ Σ (coHomS¹-ish n) (fib n) ∥₂ -Iso1 n = compIso (setTruncIso characFunSpace) Iso2 - where - Iso2 : Iso _ ∥ Σ (coHomS¹-ish n) (fib n) ∥₂ - Iso.fun Iso2 = sMap λ F → ∣ fst F ∣ , ∣ (fst (snd F)) , (snd (snd F)) ∣₂ - Iso.inv Iso2 = - sRec squash₂ - (uncurry (trElim (λ _ → isOfHLevelΠ 3 λ _ → isOfHLevelSuc 2 squash₂) - λ f → sRec squash₂ λ p → ∣ f , p ∣₂)) - Iso.rightInv Iso2 = - sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - (uncurry (trElim (λ _ → isOfHLevelΠ 3 λ _ → isProp→isOfHLevelSuc 2 (squash₂ _ _)) - λ f → sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl )) - Iso.leftInv Iso2 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl - -isContrcoHomS¹-ish : (n : ℕ) → isContr (coHomS¹-ish n) -isContrcoHomS¹-ish n = isOfHLevelRetractFromIso 0 (mapCompIso characFunSpaceS¹) isContrEnd - where - isContrEnd : isContr (hLevelTrunc 3 (Σ[ x ∈ coHomK (3 +ℕ n) ] x ≡ x)) - fst isContrEnd = ∣ 0ₖ _ , refl ∣ - snd isContrEnd = - trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) - (uncurry (coHomK-elim _ - (λ _ → isOfHLevelΠ (suc (suc (suc n))) - λ _ → isOfHLevelPlus' {n = n} 3 (isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _)) - (λ p → (trRec (isOfHLevelPlus' {n = n} 3 (isOfHLevelTrunc 3) _ _) - (λ p i → ∣ (0ₖ (3 +ℕ n) , p i) ∣) - (Iso.fun (PathIdTruncIso _) - (isContr→isProp - (isConnectedPath _ (isConnectedKn (2 +ℕ n)) (0ₖ _) (0ₖ _)) ∣ refl ∣ ∣ p ∣)))))) - -Iso2' : (n : ℕ) → Iso (∥ Σ (coHomS¹-ish n) (fib n) ∥₂) (fib n ∣ (λ _ → 0ₖ _) ∣) -Iso2' n = compIso (setTruncIso (Σ-contractFstIso centerChange)) - (setTruncIdempotentIso squash₂) - where - centerChange : isContr (coHomS¹-ish n) - fst centerChange = ∣ (λ _ → 0ₖ _) ∣ - snd centerChange y = isContr→isProp (isContrcoHomS¹-ish n) _ _ +-- The untruncated version (coHomRed n (S₊ n)), i.e. +-- (S₊∙ n →∙ coHomK-ptd n) is in fact equal to +-- (coHomRed n (S₊ n)). Its useful to formulate +-- (S₊∙ n →∙ coHomK-ptd n) as a group in its own right +-- and prove it equivalent to (coHomRed n (S₊ n)). -open import Cubical.Foundations.Equiv.HalfAdjoint -open import Cubical.Relation.Nullary -ll5 : (n : ℕ) → ¬ (n ≡ 2) → isContr (fib n ∣ (λ _ → 0ₖ _) ∣) -ll5 n p = - isOfHLevelRetractFromIso 0 - (compIso - (setTruncIso - (compIso (Σ-cong-iso-snd λ f → - (compIso characFunSpaceS¹ (invIso (Σ-cong-iso-fst (iso funExt⁻ funExt (λ _ → refl) λ _ → refl))))) - (compIso ΣΣ (Σ-contractFstIso (isContrSingl _))))) - (compIso (setTruncIso (iso funExt⁻ funExt (λ _ → refl) λ _ → refl)) - (compIso (setTruncIso (codomainIso (congIso (invIso (Iso-Kn-ΩKn+1 _))))) - ((compIso (invIso (fst (coHom≅coHomΩ _ (S₊ _)))) - (fst (Hⁿ-Sᵐ≅0 n 2 p))))))) - isContrUnit - -record ExactSeqℤ {ℓ : Level} (G : ℤ → Group ℓ) : Type ℓ where - field - maps : ∀ n → GroupHom (G n) (G (sucℤ n)) - im⊂ker : ∀ n → ∀ g → isInIm (maps n) g → isInKer (maps (sucℤ n)) g - ker⊂im : ∀ n → ∀ g → isInKer (maps (sucℤ n)) g → isInIm (maps n) g - -record ExactSeqℕ {ℓ : Level} (G : ℕ → Group ℓ) : Type ℓ where - field - maps : ∀ n → GroupHom (G n) (G (suc n)) - im⊂ker : ∀ n → ∀ g → isInIm (maps n) g → isInKer (maps (suc n)) g - ker⊂im : ∀ n → ∀ g → isInKer (maps (suc n)) g → isInIm (maps n) g - - -module _ {ℓ : Level} {A : Pointed ℓ} {n : ℕ} - where - funTyp : Type _ - funTyp = A →∙ coHomK-ptd n - - _++_ : funTyp → funTyp → funTyp +-- We start with the addition +private + _++_ : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → + (A →∙ coHomK-ptd n) → (A →∙ coHomK-ptd n) → (A →∙ coHomK-ptd n) fst (f ++ g) x = fst f x +ₖ fst g x snd (f ++ g) = cong₂ _+ₖ_ (snd f) (snd g) ∙ rUnitₖ _ (0ₖ _) -addAgree : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (x y : funTyp {A = A} {n = n}) +-- If we truncate the addition, we get back our usual +-- addition on (coHomRed n (S₊ n)) +addAgree : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (x y : A →∙ coHomK-ptd n) → Path (fst (coHomRedGrDir n A)) ∣ x ++ y ∣₂ (∣ x ∣₂ +ₕ∙ ∣ y ∣₂) @@ -576,48 +179,7 @@ addAgree {A = A} (suc zero) f g = addAgree {A = A} (suc (suc n)) f g = cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) refl) -ss : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Type ℓ''} - → isHomogeneous B - → (x y : (A →∙ B)) (f : C → x ≡ y) - → isEquiv (cong fst ∘ f) - → isEquiv f -ss homb x y f e = - isoToIsEquiv - (iso _ - (λ p → invEq (_ , e) (cong fst p)) - (λ p → →∙Homogeneous≡Path homb _ _ (secEq (_ , e) (cong fst p))) - (retEq (_ , e))) - -Whitehead : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} - → (f : A → B) - → isEmbedding f - → (∀ (b : B) → ∃[ a ∈ A ] f a ≡ b) - → isEquiv f -equiv-proof (Whitehead f emb p) b = - pRec isPropIsContr - (λ fib → fib , isEmbedding→hasPropFibers emb b fib) - (p b) - -Int-ind : ∀ {ℓ} (P : ℤ → Type ℓ) - → P (pos zero) → P (pos 1) - → P (negsuc zero) - → ((x y : ℤ) → P x → P y → P (x + y)) → (x : ℤ) → P x -Int-ind P z one min ind (pos zero) = z -Int-ind P z one min ind (pos (suc zero)) = one -Int-ind P z one min ind (pos (suc (suc n))) = - ind (pos (suc n)) 1 (Int-ind P z one min ind (pos (suc n))) one -Int-ind P z one min ind (negsuc zero) = min -Int-ind P z one min ind (negsuc (suc n)) = - ind (negsuc n) (negsuc zero) (Int-ind P z one min ind (negsuc n)) min - -genFunSpace : (n : ℕ) → S₊∙ n →∙ coHomK-ptd n -fst (genFunSpace zero) false = 1 -fst (genFunSpace zero) true = 0 -snd (genFunSpace zero) = refl -fst (genFunSpace (suc n)) = ∣_∣ -snd (genFunSpace (suc zero)) = refl -snd (genFunSpace (suc (suc n))) = refl - +-- We formulate (S₊∙ n →∙ coHomK-ptd n) as a group, πS. module _ where open import Cubical.Algebra.Monoid open import Cubical.Algebra.Semigroup @@ -650,6 +212,7 @@ module _ where isSetπS : (n : ℕ) → isSet (S₊∙ n →∙ coHomK-ptd n) isSetπS = λ n → is-set (snd (πS n)) +-- πS is equivalent to (coHomRed n (S₊∙ n)) K : (n : ℕ) → GroupIso (coHomRedGrDir n (S₊∙ n)) (πS n) fst (K n) = setTruncIdempotentIso (isSetπS n) snd (K zero) = @@ -665,57 +228,43 @@ snd (K (suc (suc n))) = (sElim2 (λ _ _ → isOfHLevelPath 2 (isSetπS _) _ _) λ f g → →∙Homogeneous≡ (isHomogeneousKn _) refl) -t : ∀ {ℓ} {A : Pointed ℓ} → Iso ((Bool , true) →∙ A) (typ A) -Iso.fun t f = fst f false -fst (Iso.inv t a) false = a -fst (Iso.inv (t {A = A}) a) true = pt A -snd (Iso.inv t a) = refl -Iso.rightInv t a = refl -Iso.leftInv t (f , p) = +-- πS has the following generator +genFunSpace : (n : ℕ) → S₊∙ n →∙ coHomK-ptd n +fst (genFunSpace zero) false = 1 +fst (genFunSpace zero) true = 0 +snd (genFunSpace zero) = refl +fst (genFunSpace (suc n)) = ∣_∣ +snd (genFunSpace (suc zero)) = refl +snd (genFunSpace (suc (suc n))) = refl + +π₀S→ℤ : ∀ {ℓ} {A : Pointed ℓ} → Iso ((Bool , true) →∙ A) (typ A) +Iso.fun π₀S→ℤ f = fst f false +fst (Iso.inv π₀S→ℤ a) false = a +fst (Iso.inv (π₀S→ℤ {A = A}) a) true = pt A +snd (Iso.inv π₀S→ℤ a) = refl +Iso.rightInv π₀S→ℤ a = refl +Iso.leftInv π₀S→ℤ (f , p) = ΣPathP ((funExt (λ { false → refl ; true → sym p})) , λ i j → p (~ i ∨ j)) -S1' : (n : ℕ) → GroupIso (πS n) ℤGroup -fst (S1' zero) = t -snd (S1' zero) = makeIsGroupHom λ _ _ → refl -S1' (suc n) = +πS≅ℤ : (n : ℕ) → GroupIso (πS n) ℤGroup +fst (πS≅ℤ zero) = π₀S→ℤ +snd (πS≅ℤ zero) = makeIsGroupHom λ _ _ → refl +πS≅ℤ (suc n) = compGroupIso (invGroupIso (K (suc n))) (compGroupIso (GroupEquiv→GroupIso (coHomGr≅coHomRedGr n (S₊∙ (suc n)))) (Hⁿ-Sⁿ≅ℤ n)) -S1 : (n : ℕ) → Iso (S₊∙ (suc n) →∙ coHomK-ptd (suc n)) ℤ -S1 n = compIso (invIso (setTruncIdempotentIso (isOfHLevel↑∙' 0 n))) + +Iso-πS-ℤ : (n : ℕ) → Iso (S₊∙ (suc n) →∙ coHomK-ptd (suc n)) ℤ +Iso-πS-ℤ n = compIso (invIso (setTruncIdempotentIso (isOfHLevel↑∙' 0 n))) (compIso (equivToIso (fst (coHomGr≅coHomRedGr n (S₊∙ (suc n))))) (fst (Hⁿ-Sⁿ≅ℤ n))) -connFunSpace : (n i : ℕ) → (x y : S₊∙ n →∙ coHomK-ptd (suc i +' n)) → ∥ x ≡ y ∥ -connFunSpace n i f g = Iso.fun PathIdTrunc₀Iso (isContr→isProp (lem n) ∣ f ∣₂ ∣ g ∣₂) - where - open import Cubical.Relation.Nullary - natLem : (n i : ℕ) → ¬ (suc (i +ℕ n) ≡ n) - natLem zero i = snotz - natLem (suc n) i p = natLem n i (sym (+-suc i n) ∙ (cong predℕ p)) - - lem : (n : ℕ) → isContr ∥ (S₊∙ n →∙ coHomK-ptd (suc i +' n)) ∥₂ - fst (lem zero) = 0ₕ∙ (suc i) - snd (lem zero) = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - λ f → pRec (squash₂ _ _) - (λ id → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn (suc i)) - (funExt λ { false → sym id ; true → sym (snd f)}))) - (help (f .fst false)) - where - help : (x : coHomK (suc i)) → ∥ x ≡ 0ₖ _ ∥ - help = coHomK-elim _ (λ _ → isProp→isOfHLevelSuc i squash) ∣ refl ∣ - lem (suc n) = - isOfHLevelRetractFromIso 0 - (compIso (equivToIso (fst (coHomGr≅coHomRedGr (suc (i +ℕ n)) (S₊∙ (suc n))))) - (fst (Hⁿ-Sᵐ≅0 (suc (i +ℕ n)) n (natLem n i)))) - isContrUnit - -S1Pres1 : (n : ℕ) → Iso.fun (fst (S1' (suc n))) (∣_∣ , refl) ≡ pos 1 -S1Pres1 zero = refl -S1Pres1 (suc n) = cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ n))) (lem n) ∙ S1Pres1 n +Iso-πS-ℤPres1 : (n : ℕ) → Iso.fun (fst (πS≅ℤ (suc n))) (∣_∣ , refl) ≡ pos 1 +Iso-πS-ℤPres1 zero = refl +Iso-πS-ℤPres1 (suc n) = cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ n))) (lem n) ∙ Iso-πS-ℤPres1 n where lem : (n : ℕ) → Iso.fun (fst (suspensionAx-Sn n n)) ∣ ∣_∣ ∣₂ ≡ ∣ ∣_∣ ∣₂ lem zero = cong ∣_∣₂ (funExt λ x → transportRefl (∣ x ∣ +ₖ (0ₖ 1)) ∙ rUnitₖ 1 ∣ x ∣) @@ -726,23 +275,26 @@ S1Pres1 (suc n) = cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ n))) (lem n) ∙ S1Pres1 n ∣)))) i) ∙ Iso.leftInv (Iso-Kn-ΩKn+1 (suc (suc n))) ∣ x ∣) -S1Pres1← : (n : ℕ) → Iso.inv (fst (S1' (suc n))) (pos 1) ≡ (∣_∣ , refl) -S1Pres1← n = sym (cong (Iso.inv (fst (S1' (suc n)))) (S1Pres1 n)) ∙ Iso.leftInv (fst (S1' (suc n))) (∣_∣ , refl) +Iso-πS-ℤPres1← : (n : ℕ) → Iso.inv (fst (πS≅ℤ (suc n))) (pos 1) ≡ (∣_∣ , refl) +Iso-πS-ℤPres1← n = sym (cong (Iso.inv (fst (πS≅ℤ (suc n)))) (Iso-πS-ℤPres1 n)) ∙ Iso.leftInv (fst (πS≅ℤ (suc n))) (∣_∣ , refl) IsoFunSpace : (n : ℕ) → Iso (S₊∙ n →∙ coHomK-ptd n) ℤ -IsoFunSpace n = fst (S1' n) +IsoFunSpace n = fst (πS≅ℤ n) +-- The first step of the Gysin sequence is to formulate +-- an equivalence g : Kᵢ ≃ (Sⁿ →∙ Kᵢ₊ₙ) module g-base where g : (n : ℕ) (i : ℕ) → coHomK i → (S₊∙ n →∙ coHomK-ptd (i +' n)) fst (g n i x) y = x ⌣ₖ (genFunSpace n) .fst y snd (g n i x) = cong (x ⌣ₖ_) ((genFunSpace n) .snd) ∙ ⌣ₖ-0ₖ i n x + -- We give another version of g which will be easier to work with now G : (n : ℕ) (i : ℕ) → coHomK i → (S₊∙ n →∙ coHomK-ptd (n +' i)) fst (G n i x) y = (genFunSpace n) .fst y ⌣ₖ x snd (G n i x) = cong (_⌣ₖ x) ((genFunSpace n) .snd) ∙ 0ₖ-⌣ₖ n i x - eq1 : (n : ℕ) (i : ℕ) → (S₊∙ n →∙ coHomK-ptd (i +' n)) ≃ (S₊∙ n →∙ coHomK-ptd (i +' n)) - eq1 n i = isoToEquiv (iso F F FF FF) + -ₖ^-Iso : (n : ℕ) (i : ℕ) → (S₊∙ n →∙ coHomK-ptd (i +' n)) ≃ (S₊∙ n →∙ coHomK-ptd (i +' n)) + -ₖ^-Iso n i = isoToEquiv (iso F F FF FF) where lem : (i n : ℕ) → (-ₖ^ i · n) (snd (coHomK-ptd (i +' n))) ≡ 0ₖ _ lem zero zero = refl @@ -764,69 +316,73 @@ module g-base where rCancel'' : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) → sym p ∙∙ refl ∙∙ p ≡ refl rCancel'' p = (λ j → (λ i → p (~ i ∨ j)) ∙∙ refl ∙∙ λ i → p (i ∨ j)) ∙ sym (rUnit refl) - transpPres0 : ∀ {k m : ℕ} (p : k ≡ m) → subst coHomK p (0ₖ k) ≡ 0ₖ m - transpPres0 {k = k} = J (λ m p → subst coHomK p (0ₖ k) ≡ 0ₖ m) (transportRefl _) + transpPres0ₖ : ∀ {k m : ℕ} (p : k ≡ m) → subst coHomK p (0ₖ k) ≡ 0ₖ m + transpPres0ₖ {k = k} = J (λ m p → subst coHomK p (0ₖ k) ≡ 0ₖ m) (transportRefl _) - eq2 : (n : ℕ) (i : ℕ) → (S₊∙ n →∙ coHomK-ptd (n +' i)) ≃ (S₊∙ n →∙ coHomK-ptd (i +' n)) - eq2 n i = + -- There will be some index swapping going on. We statet this explicitly, since we will + -- need to trace the maps later + indexSwap : (n : ℕ) (i : ℕ) → (S₊∙ n →∙ coHomK-ptd (n +' i)) ≃ (S₊∙ n →∙ coHomK-ptd (i +' n)) + indexSwap n i = isoToEquiv (iso (λ f → (λ x → subst coHomK (+'-comm n i) (fst f x)) , - cong (subst coHomK (+'-comm n i)) (snd f) ∙ transpPres0 (+'-comm n i)) + cong (subst coHomK (+'-comm n i)) (snd f) ∙ transpPres0ₖ (+'-comm n i)) (λ f → (λ x → subst coHomK (sym (+'-comm n i)) (fst f x)) - , (cong (subst coHomK (sym (+'-comm n i))) (snd f) ∙ transpPres0 (sym (+'-comm n i)))) + , (cong (subst coHomK (sym (+'-comm n i))) (snd f) ∙ transpPres0ₖ (sym (+'-comm n i)))) (λ f → →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ x → transportTransport⁻ _ (f .fst x))) λ f → →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ x → transportTransport⁻ _ (f .fst x))) - g≡ : (n : ℕ) (i : ℕ) → g n i ≡ λ x → fst (compEquiv (eq2 n i) (eq1 n i)) ((G n i) x) + -- g is a composition of G and our two previous equivs. + g≡ : (n : ℕ) (i : ℕ) → g n i ≡ λ x → fst (compEquiv (indexSwap n i) (-ₖ^-Iso n i)) ((G n i) x) g≡ n i = funExt (λ f → →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ y → gradedComm-⌣ₖ _ _ f (genFunSpace n .fst y))) - glIso3-fun : (n m : ℕ) → + -- We need a third Iso. + + suspKn-Iso-fun : (n m : ℕ) → (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) → (S₊ n → coHomK m) - glIso3-fun zero m (f , p) false = ΩKn+1→Kn _ (sym p ∙∙ cong f loop ∙∙ p) - glIso3-fun zero m (f , p) true = 0ₖ _ - glIso3-fun (suc n) m (f , p) x = + suspKn-Iso-fun zero m (f , p) false = ΩKn+1→Kn _ (sym p ∙∙ cong f loop ∙∙ p) + suspKn-Iso-fun zero m (f , p) true = 0ₖ _ + suspKn-Iso-fun (suc n) m (f , p) x = ΩKn+1→Kn _ (sym p ∙∙ cong f (merid x ∙ sym (merid (ptSn _))) ∙∙ p) - glIso3-fun∙ : (n m : ℕ) → (f : _) → glIso3-fun n m f (ptSn _) ≡ 0ₖ m - glIso3-fun∙ zero m = λ _ → refl - glIso3-fun∙ (suc n) m (f , p) = + suspKn-Iso-fun∙ : (n m : ℕ) → (f : _) → suspKn-Iso-fun n m f (ptSn _) ≡ 0ₖ m + suspKn-Iso-fun∙ zero m = λ _ → refl + suspKn-Iso-fun∙ (suc n) m (f , p) = cong (ΩKn+1→Kn m) (cong (sym p ∙∙_∙∙ p) (cong (cong f) (rCancel (merid (ptSn _))))) ∙∙ cong (ΩKn+1→Kn m) (rCancel'' p) ∙∙ ΩKn+1→Kn-refl m + suspKn-Iso-inv : (n m : ℕ) → (S₊∙ n →∙ coHomK-ptd m) → (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) + fst (suspKn-Iso-inv zero m (f , p)) base = 0ₖ _ + fst (suspKn-Iso-inv zero m (f , p)) (loop i) = Kn→ΩKn+1 m (f false) i + snd (suspKn-Iso-inv zero m (f , p)) = refl + fst (suspKn-Iso-inv (suc n) m (f , p)) north = 0ₖ _ + fst (suspKn-Iso-inv (suc n) m (f , p)) south = 0ₖ _ + fst (suspKn-Iso-inv (suc n) m (f , p)) (merid a i) = Kn→ΩKn+1 m (f a) i + snd (suspKn-Iso-inv (suc n) m (f , p)) = refl - glIso3-inv : (n m : ℕ) → (S₊∙ n →∙ coHomK-ptd m) → (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) - fst (glIso3-inv zero m (f , p)) base = 0ₖ _ - fst (glIso3-inv zero m (f , p)) (loop i) = Kn→ΩKn+1 m (f false) i - snd (glIso3-inv zero m (f , p)) = refl - fst (glIso3-inv (suc n) m (f , p)) north = 0ₖ _ - fst (glIso3-inv (suc n) m (f , p)) south = 0ₖ _ - fst (glIso3-inv (suc n) m (f , p)) (merid a i) = Kn→ΩKn+1 m (f a) i - snd (glIso3-inv (suc n) m (f , p)) = refl - - glIso3 : (n m : ℕ) → + suspKn-Iso : (n m : ℕ) → Iso (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) (S₊∙ n →∙ coHomK-ptd m) - Iso.fun (glIso3 n m) f = (glIso3-fun n m f) , (glIso3-fun∙ n m f) - Iso.inv (glIso3 n m) = glIso3-inv n m - Iso.rightInv (glIso3 zero m) (f , p) = + Iso.fun (suspKn-Iso n m) f = (suspKn-Iso-fun n m f) , (suspKn-Iso-fun∙ n m f) + Iso.inv (suspKn-Iso n m) = suspKn-Iso-inv n m + Iso.rightInv (suspKn-Iso zero m) (f , p) = →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ { false → cong (ΩKn+1→Kn m) (sym (rUnit _)) ∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f false) ; true → sym p}) - Iso.rightInv (glIso3 (suc n) m) (f , p) = + Iso.rightInv (suspKn-Iso (suc n) m) (f , p) = →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ x → (λ i → ΩKn+1→Kn m (sym (rUnit (cong-∙ (glIso3-inv (suc n) m (f , p) .fst) (merid x) (sym (merid (ptSn _))) i)) i)) + (funExt λ x → (λ i → ΩKn+1→Kn m (sym (rUnit (cong-∙ (suspKn-Iso-inv (suc n) m (f , p) .fst) (merid x) (sym (merid (ptSn _))) i)) i)) ∙∙ cong (ΩKn+1→Kn m) (cong (Kn→ΩKn+1 m (f x) ∙_) (cong sym (cong (Kn→ΩKn+1 m) p ∙ Kn→ΩKn+10ₖ m)) ∙ sym (rUnit _)) ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f x)) - Iso.leftInv (glIso3 zero m) (f , p) = →∙Homogeneous≡ (isHomogeneousKn _) + Iso.leftInv (suspKn-Iso zero m) (f , p) = →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ { base → sym p ; (loop i) j → h3 j i}) where h3 : PathP (λ i → p (~ i) ≡ p (~ i)) (Kn→ΩKn+1 m (ΩKn+1→Kn m (sym p ∙∙ cong f loop ∙∙ p))) (cong f loop) h3 = Iso.rightInv (Iso-Kn-ΩKn+1 _) _ ◁ λ i → doubleCompPath-filler (sym p) (cong f loop) p (~ i) - Iso.leftInv (glIso3 (suc n) m) (f , p) = + Iso.leftInv (suspKn-Iso (suc n) m) (f , p) = →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ { north → sym p ; south → sym p ∙ cong f (merid (ptSn _)) @@ -842,24 +398,24 @@ module g-base where (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) glIsoInvHom : (n m : ℕ) (x y : coHomK n) (z : S₊ (suc m)) - → Iso.inv (glIso3 _ _) (G m n (x +ₖ y)) .fst z - ≡ Iso.inv (glIso3 _ _) (G m n x) .fst z - +ₖ Iso.inv (glIso3 _ _) (G m n y) .fst z + → Iso.inv (suspKn-Iso _ _) (G m n (x +ₖ y)) .fst z + ≡ Iso.inv (suspKn-Iso _ _) (G m n x) .fst z + +ₖ Iso.inv (suspKn-Iso _ _) (G m n y) .fst z glIsoInvHom zero zero x y base = refl glIsoInvHom (suc n) zero x y base = refl glIsoInvHom zero zero x y (loop i) j = h3 j i where - h3 : (cong (Iso.inv (glIso3 _ _) (G zero zero (x + y)) .fst) loop) - ≡ cong₂ _+ₖ_ (cong (Iso.inv (glIso3 _ _) (G zero zero x) .fst) loop) - (cong (Iso.inv (glIso3 _ _) (G zero zero y) .fst) loop) + h3 : (cong (Iso.inv (suspKn-Iso _ _) (G zero zero (x + y)) .fst) loop) + ≡ cong₂ _+ₖ_ (cong (Iso.inv (suspKn-Iso _ _) (G zero zero x) .fst) loop) + (cong (Iso.inv (suspKn-Iso _ _) (G zero zero y) .fst) loop) h3 = Kn→ΩKn+1-hom 0 x y - ∙ ∙≡+₁ (cong (Iso.inv (glIso3 _ _) (G zero zero x) .fst) loop) - (cong (Iso.inv (glIso3 _ _) (G zero zero y) .fst) loop) + ∙ ∙≡+₁ (cong (Iso.inv (suspKn-Iso _ _) (G zero zero x) .fst) loop) + (cong (Iso.inv (suspKn-Iso _ _) (G zero zero y) .fst) loop) glIsoInvHom (suc n) zero x y (loop i) j = h3 j i where h3 : Kn→ΩKn+1 (suc n) ((pos (suc zero)) ·₀ (x +ₖ y)) - ≡ cong₂ _+ₖ_ (cong (Iso.inv (glIso3 zero (zero +' suc n)) (G zero (suc n) x) .fst) loop) - (cong (Iso.inv (glIso3 zero (zero +' suc n)) (G zero (suc n) y) .fst) loop) + ≡ cong₂ _+ₖ_ (cong (Iso.inv (suspKn-Iso zero (zero +' suc n)) (G zero (suc n) x) .fst) loop) + (cong (Iso.inv (suspKn-Iso zero (zero +' suc n)) (G zero (suc n) y) .fst) loop) h3 = cong (Kn→ΩKn+1 (suc n)) (lUnit⌣ₖ _ (x +ₖ y)) ∙∙ Kn→ΩKn+1-hom (suc n) x y ∙∙ (λ i → ∙≡+₂ n (Kn→ΩKn+1 (suc n) (lUnit⌣ₖ _ x (~ i))) @@ -886,22 +442,23 @@ module g-base where ∙∙ Kn→ΩKn+1-hom _ _ _ ∙∙ ∙≡+₂ _ _ _ - +'-suc : (n m : ℕ) → suc n +' m ≡ suc (n +' m) - +'-suc zero zero = refl - +'-suc (suc n) zero = refl - +'-suc zero (suc m) = refl - +'-suc (suc n) (suc m) = refl + private + +'-suc : (n m : ℕ) → suc n +' m ≡ suc (n +' m) + +'-suc zero zero = refl + +'-suc (suc n) zero = refl + +'-suc zero (suc m) = refl + +'-suc (suc n) (suc m) = refl - LEM : (i n : ℕ) (x : coHomK i) + decomposeG : (i n : ℕ) (x : coHomK i) → Path _ (G (suc n) i x) (subst (λ x → S₊∙ (suc n) →∙ coHomK-ptd x) (sym (+'-suc n i)) - ((Iso.inv (glIso3 n _)) (G n i x))) - LEM zero zero x = + ((Iso.inv (suspKn-Iso n _)) (G n i x))) + decomposeG zero zero x = →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ z → (λ i → x ·₀ ∣ z ∣) ∙ h3 x z ∙ sym (transportRefl _)) where - h3 : (x : ℤ) (z : S¹) → _·₀_ {n = 1} x ∣ z ∣ ≡ fst (Iso.inv (glIso3 0 zero) (G zero zero x)) z + h3 : (x : ℤ) (z : S¹) → _·₀_ {n = 1} x ∣ z ∣ ≡ fst (Iso.inv (suspKn-Iso 0 zero) (G zero zero x)) z h3 = Int-ind _ (λ { base → refl ; (loop i) j → Kn→ΩKn+10ₖ zero (~ j) i}) (λ { base → refl ; (loop i) j → rUnit (cong ∣_∣ₕ (lUnit loop j)) j i}) (λ { base → refl ; (loop i) j → rUnit (cong ∣_∣ₕ (sym loop)) j i}) @@ -909,24 +466,23 @@ module g-base where → rightDistr-⌣ₖ 0 1 x y ∣ z ∣ ∙∙ cong₂ _+ₖ_ (inx z) (iny z) ∙∙ sym (glIsoInvHom zero zero x y z) - LEM (suc i) zero x = + decomposeG (suc i) zero x = →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ z → h3 z ∙ sym (transportRefl - ((Iso.inv (glIso3 zero (suc i)) (G zero (suc i) x)) .fst z))) + ((Iso.inv (suspKn-Iso zero (suc i)) (G zero (suc i) x)) .fst z))) where - h3 : (z : S₊ 1) → _ ≡ Iso.inv (glIso3 zero (suc i)) (G zero (suc i) x) .fst z + h3 : (z : S₊ 1) → _ ≡ Iso.inv (suspKn-Iso zero (suc i)) (G zero (suc i) x) .fst z h3 base = refl h3 (loop k) j = Kn→ΩKn+1 (suc i) (lUnit⌣ₖ _ x (~ j)) k - LEM zero (suc n) x = + decomposeG zero (suc n) x = →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ z → h3 x z ∙ λ i → transportRefl (Iso.inv (glIso3 (suc n) (suc n)) (G (suc n) zero x)) (~ i) .fst z) + (funExt λ z → h3 x z ∙ λ i → transportRefl (Iso.inv (suspKn-Iso (suc n) (suc n)) (G (suc n) zero x)) (~ i) .fst z) where +merid : rUnitₖ (suc (suc n)) ∣ south ∣ ≡ cong ∣_∣ₕ (merid (ptSn _)) +merid = sym (lUnit _) ∙ cong (cong ∣_∣ₕ) λ j i → transp (λ _ → S₊ (suc (suc n))) (i ∨ j) (merid (ptSn (suc n)) i) - open import Cubical.Foundations.Path pp : (a : S₊ (suc n)) → PathP (λ i → ∣ merid a i ∣ₕ ≡ Kn→ΩKn+1 (suc n) (∣ a ∣ +ₖ (0ₖ _)) i) refl (sym (rUnitₖ (suc (suc n)) ∣ south ∣)) @@ -941,7 +497,7 @@ module g-base where h3 : (x : ℤ) (z : S₊ (suc (suc n))) → _⌣ₖ_ {n = suc (suc n)} {m = 0} ∣ z ∣ x - ≡ Iso.inv (glIso3 (suc n) (suc n)) (G (suc n) zero x) .fst z + ≡ Iso.inv (suspKn-Iso (suc n) (suc n)) (G (suc n) zero x) .fst z h3 = Int-ind _ (λ { north → refl ; south → refl ; (merid a i) j → Kn→ΩKn+10ₖ (suc n) (~ j) i}) (λ { north → refl ; south → refl @@ -956,16 +512,16 @@ module g-base where λ x y indx indy z → leftDistr-⌣ₖ _ _ ∣ z ∣ x y ∙ cong₂ _+ₖ_ (indx z) (indy z) ∙ sym (glIsoInvHom _ _ _ _ _) - LEM (suc i) (suc n) x = + decomposeG (suc i) (suc n) x = →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ z → h3 z - ∙ λ j → transportRefl ((Iso.inv (glIso3 (suc n) (suc (suc (n +ℕ i)))) + (funExt λ z → lem z + ∙ λ j → transportRefl ((Iso.inv (suspKn-Iso (suc n) (suc (suc (n +ℕ i)))) (G (suc n) (suc i) x))) (~ j) .fst z) where - h3 : (z : S₊ (suc (suc n))) → _ - h3 north = refl - h3 south = refl - h3 (merid a i) = refl + lem : (z : S₊ (suc (suc n))) → _ + lem north = refl + lem south = refl + lem (merid a i) = refl isEquivGzero : (i : ℕ) → isEquiv (G zero i) isEquivGzero i = @@ -983,17 +539,19 @@ module g-base where (funExt λ { false → rUnitₖ _ (f false) ; true → sym p})}) (lUnit⌣ₖ _)) isEquivG (suc n) i = - subst isEquiv (sym (funExt (LEM i n))) + subst isEquiv (sym (funExt (decomposeG i n))) (compEquiv (compEquiv (G n i , isEquivG n i) - (isoToEquiv (invIso (glIso3 n (n +' i))))) + (isoToEquiv (invIso (suspKn-Iso n (n +' i))))) (pathToEquiv (λ j → S₊∙ (suc n) →∙ coHomK-ptd (+'-suc n i (~ j)))) .snd) - isEquiv-g : (n i : ℕ) → isEquiv (g n i) isEquiv-g n i = subst isEquiv (sym (g≡ n i)) - (compEquiv (G n i , isEquivG n i) (compEquiv (eq2 n i) (eq1 n i)) .snd) + (compEquiv (G n i , isEquivG n i) (compEquiv (indexSwap n i) (-ₖ^-Iso n i)) .snd) + +-- We now generealise the equivalence g to also apply to arbitrary fibrations (Q : B → Type) +-- satisfying (Q * ≃∙ Sⁿ) module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) (conB : (x y : typ B) → ∥ x ≡ y ∥) (n : ℕ) (Q-is : Iso (typ (Q (pt B))) (S₊ n)) @@ -1005,22 +563,23 @@ module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) fst (g b i x) y = x ⌣ₖ c b .fst y snd (g b i x) = cong (x ⌣ₖ_) (c b .snd) ∙ ⌣ₖ-0ₖ i n x - g' : (b : typ B) (i : ℕ) → coHomK i → coHomK i → (Q b →∙ coHomK-ptd (i +' n)) - fst (g' b i x y) z = x ⌣ₖ c b .fst z +ₖ y ⌣ₖ c b .fst z - snd (g' b i x y) = cong₂ _+ₖ_ (cong (x ⌣ₖ_) (c b .snd) ∙ ⌣ₖ-0ₖ i n x) - (cong (y ⌣ₖ_) (c b .snd) ∙ ⌣ₖ-0ₖ i n y) ∙ rUnitₖ _ (0ₖ _) - - g-hom : (b : typ B) (i : ℕ) → (x y : coHomK i) → g b i (x +ₖ y) ≡ ((g b i x) ++ (g b i y)) + g-hom : (b : typ B) (i : ℕ) → (x y : coHomK i) + → g b i (x +ₖ y) ≡ ((g b i x) ++ (g b i y)) g-hom b i x y = →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ z → rightDistr-⌣ₖ i n x y (c b .fst z)) - gPathP' : (i : ℕ) → PathP (λ j → coHomK i → (isoToPath Q-is j , ua-gluePath (isoToEquiv Q-is) (Q-is-ptd) j) - →∙ coHomK-ptd (i +' n)) (g (pt B) i) (g-base.g n i) + gPathP' : (i : ℕ) + → PathP (λ j → coHomK i → (isoToPath Q-is j , ua-gluePath (isoToEquiv Q-is) (Q-is-ptd) j) + →∙ coHomK-ptd (i +' n)) + (g (pt B) i) (g-base.g n i) gPathP' i = toPathP (funExt λ x → →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ y → (λ i → transportRefl (transportRefl x i ⌣ₖ c (pt B) .fst (Iso.inv Q-is (transportRefl y i))) i) - ∙ cong (x ⌣ₖ_) (funExt⁻ c-pt (Iso.inv Q-is y) ∙ cong (genFunSpace n .fst) (Iso.rightInv Q-is y)))) + (funExt λ y + → (λ i → transportRefl (transportRefl x i ⌣ₖ c (pt B) .fst + (Iso.inv Q-is (transportRefl y i))) i) + ∙ cong (x ⌣ₖ_) (funExt⁻ c-pt (Iso.inv Q-is y) + ∙ cong (genFunSpace n .fst) (Iso.rightInv Q-is y)))) g-base : (i : ℕ) → isEquiv (g (pt B) i) g-base i = transport (λ j → isEquiv (gPathP' i (~ j))) (g-base.isEquiv-g n i) @@ -1043,7 +602,7 @@ module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) (J (λ b _ → isSet (Q b →∙ coHomK-ptd n)) (subst isSet (cong (_→∙ coHomK-ptd n) (ua∙ (isoToEquiv (invIso Q-is)) (cong (Iso.inv Q-is) (sym Q-is-ptd) ∙ Iso.leftInv Q-is _))) - (isOfHLevelRetractFromIso 2 (fst (S1' n)) isSetℤ))) + (isOfHLevelRetractFromIso 2 (fst (πS≅ℤ n)) isSetℤ))) (conB (pt B) b) @@ -1058,13 +617,16 @@ module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) isOfHLevelRetractFromIso 1 setTruncTrunc2Iso (isContr→isProp (isConnectedPath _ isConnB (pt B) (pt B))) + -- We construct a term in c* : (b : B) → (Q b →∙ Kₙ) + -- Which is equal to the generator of (Sⁿ →∙ Kₙ) over the base point. c* : Σ[ c ∈ ((b : typ B) → (Q b →∙ coHomK-ptd n)) ] (c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) fst c* b = sRec (is-setQ→K b) (J (λ b _ → Q b →∙ coHomK-ptd n) ((λ x → genFunSpace n .fst (Iso.fun Q-is x)) - , cong (genFunSpace n .fst) Q-is-ptd ∙ genFunSpace n .snd)) (conB (pt B) b) + , cong (genFunSpace n .fst) Q-is-ptd ∙ genFunSpace n .snd)) + (conB (pt B) b) snd c* = funExt λ x → (λ i → sRec (is-setQ→K (pt B)) (J (λ b _ → Q b →∙ coHomK-ptd n) @@ -1073,17 +635,19 @@ module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) (isPropPath (conB (pt B) (pt B)) ∣ refl ∣₂ i) .fst x) ∙ (λ i → transportRefl (genFunSpace n .fst (Iso.fun Q-is (transportRefl x i))) i) - p-help : {b : fst B} (p : pt B ≡ b) → (subst (fst ∘ Q) (sym p) (snd (Q b))) ≡ (snd (Q (pt B))) + p-help : {b : fst B} (p : pt B ≡ b) + → (subst (fst ∘ Q) (sym p) (snd (Q b))) ≡ (snd (Q (pt B))) p-help {b = b} = J (λ b p → subst (fst ∘ Q) (sym p) (snd (Q b)) ≡ snd (Q (pt B))) (transportRefl _) - cEquiv : (b : fst B) (p : ∥ pt B ≡ b ∥₂) + -- This form of c* will make things somewhat easier to work with later on. + c≡ : (b : fst B) (p : ∥ pt B ≡ b ∥₂) → (c* .fst b) ≡ sRec (is-setQ→K b) (λ pp → (λ qb → genFunSpace n .fst (Iso.fun Q-is (subst (fst ∘ Q) (sym pp) qb))) , cong (genFunSpace n .fst ∘ Iso.fun Q-is) (p-help pp) ∙ ((λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) p - cEquiv b = + c≡ b = sElim (λ _ → isOfHLevelPath 2 (is-setQ→K b) _ _) (J (λ b a → c* .fst b ≡ sRec (is-setQ→K b) (λ pp → @@ -1091,8 +655,7 @@ module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) genFunSpace n .fst (Iso.fun Q-is (subst (fst ∘ Q) (sym pp) qb))) , cong (genFunSpace n .fst ∘ Iso.fun Q-is) (p-help pp) ∙ - (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd) - ∣ a ∣₂) + (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd) ∣ a ∣₂) ((λ i → sRec (is-setQ→K (pt B)) (J (λ b₁ _ → Q b₁ →∙ coHomK-ptd n) ((λ x → genFunSpace n .fst (Iso.fun Q-is x)) , @@ -1102,7 +665,9 @@ module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) (transportRefl ((λ x → genFunSpace n .fst (Iso.fun Q-is x))) ∙ funExt λ x → cong (genFunSpace n .fst ∘ Iso.fun Q-is) (sym (transportRefl x))))) -module _ {ℓ ℓ'} (B : Pointed ℓ) (P : typ B → Type ℓ') where +-- We are now almost ready to define the Thom isomorphism. +-- The following module contains the types and functions occuring in it. +module preThom {ℓ ℓ'} (B : Pointed ℓ) (P : typ B → Type ℓ') where E : Type _ E = Σ (typ B) P @@ -1156,24 +721,24 @@ module _ {ℓ ℓ'} (B : Pointed ℓ) (P : typ B → Type ℓ') where compPath-filler' (push x) (λ i₁ → inr (x , merid a i₁)) (~ j) i Iso.leftInv IsoFE (push a i₁) k = push a (i₁ ∧ k) - - funDir1 : ∀ {ℓ} {A : Pointed ℓ} → (F̃ , inl tt) →∙ A → (b : typ B) → Q b →∙ A - fst (funDir1 {A = A , a} (f , p) b) north = f (inr (b , north)) - fst (funDir1 {A = A , a} (f , p) b) south = f (inr (b , south)) - fst (funDir1 {A = A , a} (f , p) b) (merid a₁ i₁) = f (inr (b , merid a₁ i₁)) - snd (funDir1 {A = A , a} (f , p) b) = sym (cong f (push b)) ∙ p - - funDir2 : ∀ {ℓ} {A : Pointed ℓ} → ((b : typ B) → Q b →∙ A) → (F̃ , inl tt) →∙ A - fst (funDir2 {A = A , a} f) (inl x) = a - fst (funDir2 {A = A , a} f) (inr (x , north)) = f x .fst north - fst (funDir2 {A = A , a} f) (inr (x , south)) = f x .fst south - fst (funDir2 {A = A , a} f) (inr (x , merid a₁ i₁)) = f x .fst (merid a₁ i₁) - fst (funDir2 {A = A , a} f) (push a₁ i₁) = snd (f a₁) (~ i₁) - snd (funDir2 {A = A , a} f) = refl - - funDir2-hom : (n : ℕ) → (f g : ((b : typ B) → Q b →∙ coHomK-ptd n)) - → funDir2 (λ b → f b ++ g b) ≡ (funDir2 f ++ funDir2 g) - funDir2-hom n f g = + + F̃→Q : ∀ {ℓ} {A : Pointed ℓ} → (F̃ , inl tt) →∙ A → (b : typ B) → Q b →∙ A + fst (F̃→Q {A = A , a} (f , p) b) north = f (inr (b , north)) + fst (F̃→Q {A = A , a} (f , p) b) south = f (inr (b , south)) + fst (F̃→Q {A = A , a} (f , p) b) (merid a₁ i₁) = f (inr (b , merid a₁ i₁)) + snd (F̃→Q {A = A , a} (f , p) b) = sym (cong f (push b)) ∙ p + + Q→F̃ : ∀ {ℓ} {A : Pointed ℓ} → ((b : typ B) → Q b →∙ A) → (F̃ , inl tt) →∙ A + fst (Q→F̃ {A = A , a} f) (inl x) = a + fst (Q→F̃ {A = A , a} f) (inr (x , north)) = f x .fst north + fst (Q→F̃ {A = A , a} f) (inr (x , south)) = f x .fst south + fst (Q→F̃ {A = A , a} f) (inr (x , merid a₁ i₁)) = f x .fst (merid a₁ i₁) + fst (Q→F̃ {A = A , a} f) (push a₁ i₁) = snd (f a₁) (~ i₁) + snd (Q→F̃ {A = A , a} f) = refl + + Q→F̃-hom : (n : ℕ) → (f g : ((b : typ B) → Q b →∙ coHomK-ptd n)) + → Q→F̃ (λ b → f b ++ g b) ≡ (Q→F̃ f ++ Q→F̃ g) + Q→F̃-hom n f g = →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ { (inl x) → sym (rUnitₖ _ (0ₖ _)) ; (inr (x , north)) → refl @@ -1182,46 +747,48 @@ module _ {ℓ ℓ'} (B : Pointed ℓ) (P : typ B → Type ℓ') where ; (push a j) i → compPath-filler (cong₂ _+ₖ_ (snd (f a)) (snd (g a))) (rUnitₖ n (0ₖ n)) (~ i) (~ j)}) - funDirSect : ∀ {ℓ} {A : Pointed ℓ} → (x : (b : typ B) → Q b →∙ A) (b : typ B) (q : Q b .fst) - → funDir1 (funDir2 x) b .fst q ≡ x b .fst q - funDirSect f b north = refl - funDirSect f b south = refl - funDirSect f b (merid a i₁) = refl - - funDirRetr : ∀ {ℓ} {A : Pointed ℓ} (f : F̃ → typ A) (p : _) - → (x : F̃) → fst (funDir2 {A = A} (funDir1 (f , p))) x ≡ f x - funDirRetr f p (inl x) = sym p - funDirRetr f p (inr (x , north)) = refl - funDirRetr f p (inr (x , south)) = refl - funDirRetr f p (inr (x , merid a i₁)) = refl - funDirRetr f p (push a i₁) j = compPath-filler (sym (cong f (push a))) p (~ j) (~ i₁) - - Iso2 : ∀ {ℓ} {A : Pointed ℓ} + Q→F̃→Q : ∀ {ℓ} {A : Pointed ℓ} → (x : (b : typ B) → Q b →∙ A) (b : typ B) (q : Q b .fst) + → F̃→Q (Q→F̃ x) b .fst q ≡ x b .fst q + Q→F̃→Q f b north = refl + Q→F̃→Q f b south = refl + Q→F̃→Q f b (merid a i₁) = refl + + F̃→Q→F̃ : ∀ {ℓ} {A : Pointed ℓ} (f : F̃ → typ A) (p : _) + → (x : F̃) → fst (Q→F̃ {A = A} (F̃→Q (f , p))) x ≡ f x + F̃→Q→F̃ f p (inl x) = sym p + F̃→Q→F̃ f p (inr (x , north)) = refl + F̃→Q→F̃ f p (inr (x , south)) = refl + F̃→Q→F̃ f p (inr (x , merid a i₁)) = refl + F̃→Q→F̃ f p (push a i₁) j = compPath-filler (sym (cong f (push a))) p (~ j) (~ i₁) + + IsoF̃Q : ∀ {ℓ} {A : Pointed ℓ} → Iso ((F̃ , inl tt) →∙ A) ((b : typ B) → Q b →∙ A) - Iso.fun (Iso2 {A = A , a}) = funDir1 - Iso.inv (Iso2 {A = A , a}) = funDir2 - Iso.rightInv (Iso2 {A = A , a}) f = - funExt λ b → ΣPathP (funExt (funDirSect f b) + Iso.fun (IsoF̃Q {A = A , a}) = F̃→Q + Iso.inv (IsoF̃Q {A = A , a}) = Q→F̃ + Iso.rightInv (IsoF̃Q {A = A , a}) f = + funExt λ b → ΣPathP (funExt (Q→F̃→Q f b) , sym (rUnit (snd (f b)))) - Iso.leftInv (Iso2 {A = A , a}) (f , p) = - ΣPathP ((funExt (funDirRetr f p)) + Iso.leftInv (IsoF̃Q {A = A , a}) (f , p) = + ΣPathP ((funExt (F̃→Q→F̃ f p)) , λ i j → p (~ i ∨ j)) + -- The main result ι : (k : ℕ) → Iso ((b : typ B) → Q b →∙ coHomK-ptd k) ((Ẽ , inl tt) →∙ coHomK-ptd k) - ι k = compIso (invIso Iso2) h3 + ι k = compIso (invIso IsoF̃Q) IsoFE-extend where - h3 : Iso ((F̃ , inl tt) →∙ coHomK-ptd k) + + IsoFE-extend : Iso ((F̃ , inl tt) →∙ coHomK-ptd k) ((Ẽ , inl tt) →∙ coHomK-ptd k) - Iso.fun h3 G = (λ x → G .fst (Iso.inv IsoFE x)) + Iso.fun IsoFE-extend G = (λ x → G .fst (Iso.inv IsoFE x)) , (snd G) - Iso.inv h3 G = (λ x → G .fst (Iso.fun IsoFE x)) + Iso.inv IsoFE-extend G = (λ x → G .fst (Iso.fun IsoFE x)) , (snd G) - Iso.rightInv h3 G = + Iso.rightInv IsoFE-extend G = →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ x → cong (G .fst) (Iso.rightInv IsoFE x)) - Iso.leftInv h3 G = + Iso.leftInv IsoFE-extend G = →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ x → cong (G .fst) (Iso.leftInv IsoFE x)) @@ -1230,72 +797,71 @@ module _ {ℓ ℓ'} (B : Pointed ℓ) (P : typ B → Type ℓ') where ≡ (Iso.fun (ι k) f ++ Iso.fun (ι k) g) ι-hom k f g = →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ x → funExt⁻ (cong fst (funDir2-hom _ f g)) (invFE x)) - -codomainIsoDep : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : A → Type ℓ'} {C : A → Type ℓ''} - → ((a : A) → Iso (B a) (C a)) - → Iso ((a : A) → B a) ((a : A) → C a) -Iso.fun (codomainIsoDep is) f a = Iso.fun (is a) (f a) -Iso.inv (codomainIsoDep is) f a = Iso.inv (is a) (f a) -Iso.rightInv (codomainIsoDep is) f = funExt λ a → Iso.rightInv (is a) (f a) -Iso.leftInv (codomainIsoDep is) f = funExt λ a → Iso.leftInv (is a) (f a) + (funExt λ x → funExt⁻ (cong fst (Q→F̃-hom _ f g)) (invFE x)) +-- Packing everything up gives us the Thom Isomorphism between +-- the nᵗʰ cohomology of B and the (n+i)ᵗʰ reduced cohomology of Ẽ, as defined above module Thom {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) (conB : (x y : typ B) → ∥ x ≡ y ∥) - (n : ℕ) (Q-is : Iso (typ (Q B P (pt B))) (S₊ n)) - (Q-is-ptd : Iso.fun Q-is (pt (Q B P (pt B))) ≡ snd (S₊∙ n)) - (c : (b : typ B) → (Q B P b →∙ coHomK-ptd n)) + (n : ℕ) (Q-is : Iso (typ (preThom.Q B P (pt B))) (S₊ n)) + (Q-is-ptd : Iso.fun Q-is (pt (preThom.Q B P (pt B))) ≡ snd (S₊∙ n)) + (c : (b : typ B) → (preThom.Q B P b →∙ coHomK-ptd n)) (c-pt : c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) where - ϕ : (i : ℕ) → GroupEquiv (coHomGr i (typ B)) - (coHomRedGrDir (i +' n) (Ẽ B P , inl tt)) + ϕ : (i : ℕ) → GroupEquiv (coHomGr i (typ B)) (coHomRedGrDir (i +' n) (preThom.Ẽ B P , inl tt)) fst (ϕ i) = isoToEquiv (setTruncIso (compIso (codomainIsoDep - λ b → equivToIso ((g B (Q B P) conB n Q-is Q-is-ptd c c-pt b i) - , g-equiv B (Q B P) conB n Q-is Q-is-ptd c c-pt b i)) - (ι B P (i +' n)))) + λ b → equivToIso ((g B (preThom.Q B P) conB n Q-is Q-is-ptd c c-pt b i) + , g-equiv B (preThom.Q B P) conB n Q-is Q-is-ptd c c-pt b i)) + (preThom.ι B P (i +' n)))) snd (ϕ i) = makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) - λ F G → cong ∣_∣₂ (cong (Iso.fun (ι B P (i +' n))) - (funExt (λ a → g-hom B (Q B P) conB n Q-is Q-is-ptd c c-pt a i (F a) (G a))) - ∙ ι-hom B P (i +' n) _ _) + λ F G → cong ∣_∣₂ (cong (Iso.fun (preThom.ι B P (i +' n))) + (funExt (λ a → g-hom B (preThom.Q B P) + conB n Q-is Q-is-ptd c c-pt a i (F a) (G a))) + ∙ preThom.ι-hom B P (i +' n) _ _) ∙ addAgree (i +' n) _ _) +-- We finally get the Gysin sequence module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) (conB : (x y : typ B) → ∥ x ≡ y ∥₂) - (n : ℕ) (Q-is : Iso (typ (Q B P (pt B))) (S₊ n)) - (Q-is-ptd : Iso.fun Q-is (pt (Q B P (pt B))) ≡ snd (S₊∙ n)) where + (n : ℕ) (Q-is : Iso (typ (preThom.Q B P (pt B))) (S₊ n)) + (Q-is-ptd : Iso.fun Q-is (pt (preThom.Q B P (pt B))) ≡ snd (S₊∙ n)) where 0-connB : (x y : typ B) → ∥ x ≡ y ∥ 0-connB x y = sRec (isProp→isSet squash) (∥_∥.∣_∣) (conB x y) - c = (c* B (Q B P) conB n Q-is Q-is-ptd .fst) - c-ptd = (c* B (Q B P) conB n Q-is Q-is-ptd .snd) - cEq = cEquiv B (Q B P) conB n Q-is Q-is-ptd + c = (c* B (preThom.Q B P) conB n Q-is Q-is-ptd .fst) + c-ptd = (c* B (preThom.Q B P) conB n Q-is Q-is-ptd .snd) + cEq = c≡ B (preThom.Q B P) conB n Q-is Q-is-ptd module w = Thom B P 0-connB n Q-is Q-is-ptd c c-ptd ϕ = w.ϕ - E' = E B P - - E'̃ = Ẽ B P + E' = preThom.E B P - p : E' → typ B - p = fst + E'̃ = preThom.Ẽ B P + -- The generator of coHom n (typ B) e : coHom n (typ B) e = ∣ (λ b → c b .fst south) ∣₂ + -- The maps of interest are ⌣, p, E-susp and j*. In reality, we are interested + -- in the composition of ϕ and j* (which is just the cup product), but it's easier + -- to first give an exact sequence involving p, E-susp and j* ⌣-hom : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr (i +' n) (typ B)) fst (⌣-hom i) x = x ⌣ e snd (⌣-hom i) = makeIsGroupHom λ f g → rightDistr-⌣ _ _ f g e + p : E' → typ B + p = fst + p-hom : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr i E') p-hom i = coHomMorph _ p @@ -1360,14 +926,15 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) (funExt λ { (inl x) → refl ; (inr x) → sym (Kn→ΩKn+1 _ (g x)) ; (push a j) k → Kn→ΩKn+1 i (g (fst a)) (~ k ∧ j)}))))) - open import Cubical.Foundations.Path + Ker-Susp⊂Im-p : (i : ℕ) (x : _) → isInKer (E-susp i) x → isInIm (p-hom i) x Ker-Susp⊂Im-p i = sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) λ f ker → pRec squash (λ id → ∣ ∣ (λ x → ΩKn+1→Kn i (sym (funExt⁻ (cong fst id) (inr x)))) ∣₂ , cong ∣_∣₂ (funExt (λ { (a , b) → - cong (ΩKn+1→Kn i) (lUnit _ ∙ cong (_∙ sym (funExt⁻ (λ i₁ → fst (id i₁)) (inr a))) (sym (flipSquare (cong snd id)) + cong (ΩKn+1→Kn i) (lUnit _ ∙ cong (_∙ sym (funExt⁻ (λ i₁ → fst (id i₁)) (inr a))) + (sym (flipSquare (cong snd id)) ∙∙ (PathP→compPathR (cong (funExt⁻ (cong fst id)) (push (a , b)))) ∙∙ assoc _ _ _ ∙ sym (rUnit _)) @@ -1401,49 +968,51 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) where h3 : (f : (E'̃ , inl tt) →∙ coHomK-ptd (suc i)) → (p : (fst f ∘ inr) ≡ (λ _ → 0ₖ (suc i))) - → (a : E B P) + → (a : preThom.E B P) → PathP (λ i → snd f (~ i) ≡ p (~ i) (fst a)) (Kn→ΩKn+1 i (ΩKn+1→Kn i (sym (snd f) ∙∙ cong (fst f) (push a) ∙∙ funExt⁻ p (fst a)))) (cong (fst f) (push a)) h3 f p a = Iso.rightInv (Iso-Kn-ΩKn+1 i) _ ◁ λ i j → doubleCompPath-filler (sym (snd f)) (cong (fst f) (push a)) (funExt⁻ p (fst a)) (~ i) j + -- We compose ϕ and j*. The above exact sequence will induce another one for this + -- composition. ϕ∘j : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr (i +' n) (typ B)) ϕ∘j i = compGroupHom (fst (fst (ϕ i)) , snd (ϕ i)) (cofibSeq.j* (i +' n)) - +'-suc : (i n : ℕ) → (suc i +' n) ≡ suc (i +' n) - +'-suc zero zero = refl - +'-suc (suc i₁) zero = refl - +'-suc zero (suc n) = refl - +'-suc (suc i₁) (suc n) = refl - private - h3 : ∀ {ℓ ℓ'} {n m : ℕ} (G : ℕ → Group ℓ) (H : Group ℓ') (p : n ≡ m) + +'-suc : (i n : ℕ) → (suc i +' n) ≡ suc (i +' n) + +'-suc zero zero = refl + +'-suc (suc i₁) zero = refl + +'-suc zero (suc n) = refl + +'-suc (suc i₁) (suc n) = refl + + Path→GroupPath : ∀ {ℓ ℓ'} {n m : ℕ} (G : ℕ → Group ℓ) (H : Group ℓ') (p : n ≡ m) → GroupEquiv (G n) H → GroupEquiv (G m) H - h3 {n = n} G H = + Path→GroupPath {n = n} G H = J (λ m _ → GroupEquiv (G n) H → GroupEquiv (G m) H) λ p → p h-ret : ∀ {ℓ ℓ'} {n m : ℕ} (G : ℕ → Group ℓ) (H : Group ℓ') → (e : GroupEquiv (G n) H) → (p : n ≡ m) - → (x : G m .fst) → invEq (fst e) (fst (fst (h3 G H p e)) x) ≡ subst (λ x → G x .fst) (sym p) x + → (x : G m .fst) → invEq (fst e) (fst (fst (Path→GroupPath G H p e)) x) ≡ subst (λ x → G x .fst) (sym p) x h-ret G H e = - J (λ m p → ((x : G m .fst) → invEq (fst e) (fst (fst (h3 G H p e)) x) ≡ subst (λ x → G x .fst) (sym p) x)) + J (λ m p → ((x : G m .fst) → invEq (fst e) (fst (fst (Path→GroupPath G H p e)) x) ≡ subst (λ x → G x .fst) (sym p) x)) λ x → cong (invEq (fst e) ) (λ i → transportRefl (transportRefl (fst (fst e) (transportRefl (transportRefl x i) i)) i) i) ∙∙ retEq (fst e) x ∙∙ sym (transportRefl _) - isEquivϕ' : (i : ℕ) → GroupEquiv (coHomRedGrDir (suc (i +' n)) (E'̃ , inl tt)) + thomIso' : (i : ℕ) → GroupEquiv (coHomRedGrDir (suc (i +' n)) (E'̃ , inl tt)) (coHomGr (suc i) (typ B)) - isEquivϕ' i = (h3 (λ n → coHomRedGrDir n (E'̃ , inl tt)) _ (+'-suc i n) + thomIso' i = (Path→GroupPath (λ n → coHomRedGrDir n (E'̃ , inl tt)) _ (+'-suc i n) (invGroupEquiv (ϕ (suc i)))) ϕ' : (i : ℕ) → GroupHom (coHomRedGrDir (suc (i +' n)) (E'̃ , inl tt)) (coHomGr (suc i) (typ B)) - ϕ' i = fst (fst (isEquivϕ' i)) , snd (isEquivϕ' i) + ϕ' i = fst (fst (thomIso' i)) , snd (thomIso' i) susp∘ϕ : (i : ℕ) → GroupHom (coHomGr (i +' n) E') (coHomGr (suc i) (typ B)) susp∘ϕ i = compGroupHom (E-susp (i +' n)) (ϕ' i) @@ -1465,9 +1034,9 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) KerSusp∘ϕ⊂Im-p : (i : ℕ) (x : _) → isInKer (susp∘ϕ i) x → isInIm (p-hom _) x KerSusp∘ϕ⊂Im-p i x p = - Ker-Susp⊂Im-p _ x (sym (retEq (fst (isEquivϕ' _)) _) - ∙ (cong (invEq (fst (isEquivϕ' _))) p - ∙ IsGroupHom.pres1 (snd (invGroupEquiv (isEquivϕ' _))))) + Ker-Susp⊂Im-p _ x (sym (retEq (fst (thomIso' _)) _) + ∙ (cong (invEq (fst (thomIso' _))) p + ∙ IsGroupHom.pres1 (snd (invGroupEquiv (thomIso' _))))) Im-Susp∘ϕ⊂Ker-ϕ∘j : (i : ℕ) → (x : _) → isInIm (susp∘ϕ i) x → isInKer (ϕ∘j (suc i)) x Im-Susp∘ϕ⊂Ker-ϕ∘j i x = @@ -1494,14 +1063,15 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) → isInKer (ϕ∘j (suc i)) x → isInIm (susp∘ϕ i) x Ker-ϕ∘j⊂Im-Susp∘ϕ i x p = pRec squash - (uncurry (λ f p → ∣ f , cong (fst (fst (isEquivϕ' i))) p ∙ secEq (fst (isEquivϕ' _)) x ∣)) - (Ker-j⊂Im-Susp _ (invEq (fst (isEquivϕ' _)) x) + (uncurry (λ f p → ∣ f , cong (fst (fst (thomIso' i))) p ∙ secEq (fst (thomIso' _)) x ∣)) + (Ker-j⊂Im-Susp _ (invEq (fst (thomIso' _)) x) ((cong (cofibSeq.j* (suc (i +' n)) .fst ) lem2 ∙ natTranspLem _ (λ n → cofibSeq.j* n .fst) (+'-suc i n)) ∙∙ cong (transport (λ j → (coHomGr (+'-suc i n j) (typ B) .fst))) p ∙∙ h2 i n)) where - lem2 : invEq (fst (isEquivϕ' i)) x ≡ transport (λ j → coHomRed (+'-suc i n j) (E'̃ , inl tt)) (fst (fst (ϕ _)) x) + lem2 : invEq (fst (thomIso' i)) x + ≡ transport (λ j → coHomRed (+'-suc i n j) (E'̃ , inl tt)) (fst (fst (ϕ _)) x) lem2 = cong (transport (λ j → coHomRed (+'-suc i n j) (E'̃ , inl tt))) (transportRefl _ ∙ cong (fst (fst (ϕ _))) λ i → transportRefl (transportRefl x i) i) @@ -1513,1796 +1083,10 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) h2 (suc i₁) zero = refl h2 (suc i₁) (suc n) = refl - + -- Finally, we have that ϕ∘j is just the cup product, and we have arrived at an exact + -- sequence involving it. ϕ∘j≡ : (i : ℕ) → ϕ∘j i ≡ ⌣-hom i ϕ∘j≡ i = Σ≡Prop (λ _ → isPropIsGroupHom _ _) (funExt (sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl)) - -open import Cubical.Experiments.Brunerie -open import Cubical.HITs.Hopf - -open import Cubical.HITs.Join - -module fibS1 = hopfBase S1-AssocHSpace (sphereElim2 _ (λ _ _ → squash) ∣ refl ∣) - -S¹Hopf : S₊ 2 → Type -S¹Hopf = fibS1.Hopf - -TotalHopf' : Type _ -TotalHopf' = Σ (S₊ 2) S¹Hopf - -CP2 : Type _ -CP2 = fibS1.megaPush - -fibr : CP2 → Type _ -fibr = fibS1.P - -hopf : join S¹ S¹ → S₊ 2 -hopf x = fst (JoinS¹S¹→TotalHopf x) - -E* : Type _ -E* = fibS1.totalSpaceMegaPush - -IsoE' : Iso E* (join S¹ (join S¹ S¹)) -IsoE' = fibS1.IsoJoin₁ - -IsoE2 : (join S¹ (join S¹ S¹)) ≡ join S¹ (S₊ 3) -IsoE2 = cong (join S¹) (sym S³≡joinS¹S¹ ∙ isoToPath IsoS³S3) - -IsoTotalHopf' : Iso fibS1.TotalSpaceHopf' (join S¹ S¹) -IsoTotalHopf' = fibS1.joinIso₁ - -CP' : Type _ -CP' = Pushout (λ _ → tt) hopf - -conCP2 : (x y : CP2) → ∥ x ≡ y ∥₂ -conCP2 x y = sRec2 squash₂ (λ p q → ∣ p ∙ sym q ∣₂) (conCP2' x) (conCP2' y) - where - conCP2' : (x : CP2) → ∥ x ≡ inl tt ∥₂ - conCP2' (inl x) = ∣ refl ∣₂ - conCP2' (inr x) = sphereElim 1 {A = λ x → ∥ inr x ≡ inl tt ∥₂} (λ _ → squash₂) ∣ sym (push (inl base)) ∣₂ x - conCP2' (push a i₁) = ll a i₁ - where - h2 : ∀ {ℓ} {A : fibS1.TotalSpaceHopf' → Type ℓ} → ((a : _) → isProp (A a)) - → A (inl base) - → ((a : fibS1.TotalSpaceHopf') → A a) - h2 {A = A} p b = - PushoutToProp p (sphereElim 0 (λ _ → p _) b) - (sphereElim 0 (λ _ → p _) (subst A (push (base , base)) b)) - - ll : (a : fibS1.TotalSpaceHopf') → PathP (λ i → ∥ Path CP2 (push a i) (inl tt) ∥₂) (conCP2' (inl tt)) (conCP2' (inr (fibS1.induced a))) - ll = h2 (λ _ → isOfHLevelPathP' 1 squash₂ _ _) λ j → ∣ (λ i → push (inl base) (~ i ∧ j)) ∣₂ - -module GysinS1 = Gysin (CP2 , inl tt) fibr conCP2 2 idIso refl - -PushoutReplaceBase : - ∀ {ℓ ℓ' ℓ''} {A B : Type ℓ} {C : Type ℓ'} {D : Type ℓ''} {f : A → C} {g : A → D} - (e : B ≃ A) → Pushout (f ∘ fst e) (g ∘ fst e) ≡ Pushout f g -PushoutReplaceBase {f = f} {g = g} = - EquivJ (λ _ e → Pushout (f ∘ fst e) (g ∘ fst e) ≡ Pushout f g) - refl - -isContrH³E : isContr (coHom 3 (GysinS1.E')) -isContrH³E = - subst isContr - (sym (isoToPath (compIso (Iso1 0) (Iso2' 0))) - ∙ cong (coHom 3) (sym (isoToPath IsoE' ∙ IsoE2))) - (ll5 0 (snotz ∘ sym)) - -isContrH⁴E : isContr (coHom 4 (GysinS1.E')) -isContrH⁴E = - subst isContr - (sym (isoToPath (compIso (Iso1 1) (Iso2' 1))) - ∙ cong (coHom 4) (sym (isoToPath IsoE' ∙ IsoE2))) - (ll5 1 λ p → snotz (sym (cong predℕ p))) - -genH2 = GysinS1.e - -S³→Groupoid : ∀ {ℓ} (P : join S¹ S¹ → Type ℓ) - → ((x : _) → isGroupoid (P x)) - → P (inl base) - → (x : _) → P x -S³→Groupoid P grp b = - transport (λ i → (x : (sym (isoToPath S³IsojoinS¹S¹) ∙ isoToPath IsoS³S3) (~ i)) - → P (transp (λ j → (sym (isoToPath S³IsojoinS¹S¹) ∙ isoToPath IsoS³S3) (~ i ∧ ~ j)) i x)) - (sphereElim _ (λ _ → grp _) b) - -TotalHopf→Gpd : ∀ {ℓ} (P : fibS1.TotalSpaceHopf' → Type ℓ) - → ((x : _) → isGroupoid (P x)) - → P (inl base) - → (x : _) → P x -TotalHopf→Gpd P grp b = - transport (λ i → (x : sym (isoToPath IsoTotalHopf') i) - → P (transp (λ j → isoToPath IsoTotalHopf' (~ i ∧ ~ j)) i x)) - (S³→Groupoid _ (λ _ → grp _) b) - -TotalHopf→Gpd' : ∀ {ℓ} (P : Σ (S₊ 2) fibS1.Hopf → Type ℓ) - → ((x : _) → isGroupoid (P x)) - → P (north , base) - → (x : _) → P x -TotalHopf→Gpd' P grp b = - transport (λ i → (x : ua (_ , fibS1.isEquivTotalSpaceHopf'→TotalSpace) i) - → P (transp (λ j → ua (_ , fibS1.isEquivTotalSpaceHopf'→TotalSpace) (i ∨ j)) i x)) - (TotalHopf→Gpd _ (λ _ → grp _) b) - -CP2→Groupoid : ∀ {ℓ} {P : CP2 → Type ℓ} - → ((x : _) → is2Groupoid (P x)) - → (b : P (inl tt)) - → (s2 : ((x : _) → P (inr x))) - → PathP (λ i → P (push (inl base) i)) b (s2 north) - → (x : _) → P x -CP2→Groupoid {P = P} grp b s2 pp (inl x) = b -CP2→Groupoid {P = P} grp b s2 pp (inr x) = s2 x -CP2→Groupoid {P = P} grp b s2 pp (push a i₁) = h3 a i₁ - where - h3 : (a : fibS1.TotalSpaceHopf') → PathP (λ i → P (push a i)) b (s2 _) - h3 = TotalHopf→Gpd _ (λ _ → isOfHLevelPathP' 3 (grp _) _ _) pp - -mm : (S₊ 2 → coHomK 2) → (CP2 → coHomK 2) -mm f = λ { (inl x) → 0ₖ _ - ; (inr x) → f x -ₖ f north - ; (push a i) → TotalHopf→Gpd (λ x → 0ₖ 2 ≡ f (fibS1.TotalSpaceHopf'→TotalSpace x .fst) -ₖ f north) - (λ _ → isOfHLevelTrunc 4 _ _) - (sym (rCancelₖ 2 (f north))) - a i} - -e-inv : coHomGr 2 (S₊ 2) .fst → coHomGr 2 CP2 .fst -e-inv = sMap mm - -cancel : (f : CP2 → coHomK 2) → ∥ mm (f ∘ inr) ≡ f ∥ -cancel f = - pRec squash - (λ p → pRec squash - (λ f → ∣ funExt f ∣) - (zz2 p)) - (h1 (f (inl tt))) - where - h1 : (x : coHomK 2) → ∥ x ≡ 0ₖ 2 ∥ - h1 = coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash) ∣ refl ∣ - - zz2 : (p : f (inl tt) ≡ 0ₖ 2) → ∥ ((x : CP2) → mm (f ∘ inr) x ≡ f x) ∥ - zz2 p = trRec squash (λ pp → - ∣ CP2→Groupoid (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) - (sym p) - (λ x → (λ i → f (inr x) -ₖ f (push (inl base) (~ i))) - ∙∙ (λ i → f (inr x) -ₖ p i) - ∙∙ rUnitₖ 2 (f (inr x))) - pp ∣) - l - where - l : hLevelTrunc 1 (PathP ((λ i₁ → mm (f ∘ inr) (push (inl base) i₁) ≡ f (push (inl base) i₁))) - (sym p) - (((λ i₁ → f (inr north) -ₖ f (push (inl base) (~ i₁))) ∙∙ - (λ i₁ → f (inr north) -ₖ p i₁) ∙∙ rUnitₖ 2 (f (inr north))))) - l = isConnectedPathP 1 (isConnectedPath 2 (isConnectedKn 1) _ _) _ _ .fst - -e' : GroupIso (coHomGr 2 CP2) (coHomGr 2 (S₊ 2)) -Iso.fun (fst e') = sMap (_∘ inr) -Iso.inv (fst e') = e-inv -Iso.rightInv (fst e') = - sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - λ f → trRec {B = Iso.fun (fst e') (Iso.inv (fst e') ∣ f ∣₂) ≡ ∣ f ∣₂} - (isOfHLevelPath 2 squash₂ _ _) - (λ p → cong ∣_∣₂ (funExt λ x → cong (λ y → (f x) -ₖ y) p ∙ rUnitₖ 2 (f x))) - (Iso.fun (PathIdTruncIso _) (isContr→isProp (isConnectedKn 1) ∣ f north ∣ ∣ 0ₖ 2 ∣)) -Iso.leftInv (fst e') = - sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - λ f → pRec (squash₂ _ _) (cong ∣_∣₂) (cancel f) -snd e' = - makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ f g → refl) - -GroupEquivℤ-pres1g : GroupIso (coHomGr 2 CP2) ℤGroup -GroupEquivℤ-pres1g = compGroupIso e' (Hⁿ-Sⁿ≅ℤ 1) - -⌣hom : GroupEquiv (coHomGr 2 CP2) (coHomGr 4 CP2) -fst (fst ⌣hom) = GysinS1.⌣-hom 2 .fst -snd (fst ⌣hom) = subst isEquiv (cong fst (GysinS1.ϕ∘j≡ 2)) h3 - where - h3 : isEquiv (GysinS1.ϕ∘j 2 .fst) - h3 = - SES→isEquiv - (GroupPath _ _ .fst (invGroupEquiv - (isContr→≃Unit isContrH³E - , makeIsGroupHom λ _ _ → refl))) - (GroupPath _ _ .fst (invGroupEquiv - (isContr→≃Unit isContrH⁴E - , makeIsGroupHom λ _ _ → refl))) - (GysinS1.susp∘ϕ 1) - (GysinS1.ϕ∘j 2) - (GysinS1.p-hom 4) - (GysinS1.Ker-ϕ∘j⊂Im-Susp∘ϕ _) - (GysinS1.Ker-p⊂Im-ϕ∘j _) -snd ⌣hom = GysinS1.⌣-hom 2 .snd - -isGenerator≃ℤ : ∀ {ℓ} (G : Group ℓ) (g : fst G) - → Type _ -isGenerator≃ℤ G g = - Σ[ e ∈ GroupIso G ℤGroup ] abs (Iso.fun (fst e) g) ≡ 1 - - -⌣-pres1 : ∀ {ℓ} (G : Type ℓ) (n : ℕ) → Type _ -⌣-pres1 G n = - Σ[ x ∈ coHomGr n G .fst ] - isGenerator≃ℤ (coHomGr n G) x × isGenerator≃ℤ (coHomGr (n +' n) G) (x ⌣ x) - -genr : coHom 2 CP2 -genr = ∣ CP2→Groupoid (λ _ → isOfHLevelTrunc 4) - (0ₖ _) - ∣_∣ - refl ∣₂ - -≡CP2 : (f g : CP2 → coHomK 2) → ∥ f ∘ inr ≡ g ∘ inr ∥ → Path (coHom 2 CP2) ∣ f ∣₂ ∣ g ∣₂ -≡CP2 f g = pRec (squash₂ _ _) - (λ p → pRec (squash₂ _ _) (λ id → trRec (squash₂ _ _) - (λ pp → cong ∣_∣₂ - (funExt (CP2→Groupoid (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) - id - (funExt⁻ p) - (compPathR→PathP pp)))) - (pp2 (f (inl tt)) (g (inl tt)) id - (cong f (push (inl base)) ∙ (funExt⁻ p north) ∙ cong g (sym (push (inl base)))))) - (conn (f (inl tt)) (g (inl tt)))) - - where - conn : (x y : coHomK 2) → ∥ x ≡ y ∥ - conn = coHomK-elim _ (λ _ → isSetΠ λ _ → isOfHLevelSuc 1 squash) - (coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash) ∣ refl ∣) - - pp2 : (x y : coHomK 2) (p q : x ≡ y) → hLevelTrunc 1 (p ≡ q) - pp2 x y = λ p q → Iso.fun (PathIdTruncIso _) (isContr→isProp (isConnectedPath _ (isConnectedKn 1) x y) ∣ p ∣ ∣ q ∣) - -rUnit* : (x : S¹) → x * base ≡ x -rUnit* base = refl -rUnit* (loop i₁) = refl - -meridP : (a x : S¹) → Path (Path (coHomK 2) _ _) (cong ∣_∣ₕ (merid (a * x) ∙ sym (merid base))) - ((cong ∣_∣ₕ (merid a ∙ sym (merid base))) ∙ (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) -meridP = wedgeconFun _ _ (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) - (λ x → lUnit _ ∙ cong (_∙ cong ∣_∣ₕ (merid x ∙ sym (merid base))) (cong (cong ∣_∣ₕ) (sym (rCancel (merid base))))) - (λ x → (λ i → cong ∣_∣ₕ (merid (rUnit* x i) ∙ sym (merid base))) - ∙∙ rUnit _ - ∙∙ cong (cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙_) - (cong (cong ∣_∣ₕ) (sym (rCancel (merid base))))) - (sym (l (cong ∣_∣ₕ (merid base ∙ sym (merid base))) - (cong (cong ∣_∣ₕ) (sym (rCancel (merid base)))))) - where - l : ∀ {ℓ} {A : Type ℓ} {x : A} (p : x ≡ x) (P : refl ≡ p) - → lUnit p ∙ cong (_∙ p) P ≡ rUnit p ∙ cong (p ∙_) P - l p = J (λ p P → lUnit p ∙ cong (_∙ p) P ≡ rUnit p ∙ cong (p ∙_) P) refl - -lemmie : (x : S¹) → ptSn 1 ≡ x * (invLooper x) -lemmie base = refl -lemmie (loop i) j = - hcomp (λ r → λ {(i = i0) → base ; (i = i1) → base ; (j = i0) → base}) - base - -mmInv : (x : S¹) → Path (Path (coHomK 2) _ _) - (cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base))) - (cong ∣_∣ₕ (sym (merid x ∙ sym (merid base)))) -mmInv x = (lUnit _ - ∙∙ cong (_∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base))) (sym (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base))))) ∙∙ sym (assoc _ _ _)) - ∙∙ cong (sym (cong ∣_∣ₕ (merid x ∙ sym (merid base))) ∙_) hh - ∙∙ (assoc _ _ _ - ∙∙ cong (_∙ (cong ∣_∣ₕ (sym (merid x ∙ sym (merid base))))) (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) - ∙∙ sym (lUnit _)) - where - hh : cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base)) - ≡ cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙ cong ∣_∣ₕ (sym (merid x ∙ sym (merid base))) - hh = sym (meridP x (invLooper x)) ∙ ((λ i → cong ∣_∣ₕ (merid (lemmie x (~ i)) ∙ sym (merid base))) ∙ cong (cong ∣_∣ₕ) (rCancel (merid base))) ∙ sym (rCancel _) - -commS1 : (a x : S¹) → a * x ≡ x * a -commS1 = wedgeconFun _ _ (λ _ _ → isGroupoidS¹ _ _) - (sym ∘ rUnit*) - rUnit* - refl - -s233 : (a x : S¹) → (invEq (fibS1.μ-eq a) x) * a ≡ (invLooper a * x) * a -s233 a x = secEq (fibS1.μ-eq a) x ∙∙ cong (_* x) (lemmie a) ∙∙ assocHSpace.μ-assoc S1-AssocHSpace a (invLooper a) x ∙ commS1 _ _ - -ss23 : (a x : S¹) → invEq (fibS1.μ-eq a) x ≡ invLooper a * x -ss23 a x = sym (retEq (fibS1.μ-eq a) (invEq (fibS1.μ-eq a) x)) - ∙∙ cong (invEq (fibS1.μ-eq a)) (s233 a x) - ∙∙ retEq (fibS1.μ-eq a) (invLooper a * x) - -ll : GysinS1.e ≡ genr -ll = ≡CP2 _ _ ∣ funExt (λ x → funExt⁻ (cong fst (ll32 x)) south) ∣ - where - kr : (x : Σ (S₊ 2) fibS1.Hopf) → Path (hLevelTrunc 4 _) ∣ fst x ∣ ∣ north ∣ - kr = uncurry λ { north → λ y → cong ∣_∣ₕ (merid base ∙ sym (merid y)) - ; south → λ y → cong ∣_∣ₕ (sym (merid y)) - ; (merid a i) → lem a i} - where - lem : (a : S¹) → PathP (λ i → (y : fibS1.Hopf (merid a i)) → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a i ∣ ∣ north ∣) - (λ y → cong ∣_∣ₕ (merid base ∙ sym (merid y))) - λ y → cong ∣_∣ₕ (sym (merid y)) - lem a = toPathP (funExt λ x → cong (transport ((λ i₁ → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a i₁ ∣ ∣ north ∣))) - ((λ i → (λ z → cong ∣_∣ₕ (merid base - ∙ sym (merid (transport (λ j → uaInvEquiv (fibS1.μ-eq a) (~ i) j) x))) z)) - ∙ λ i → cong ∣_∣ₕ (merid base - ∙ sym (merid (transportRefl (invEq (fibS1.μ-eq a) x) i)))) - ∙∙ (λ i → transp (λ i₁ → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a (i₁ ∨ i) ∣ ∣ north ∣) i - (compPath-filler' (cong ∣_∣ₕ (sym (merid a))) - (cong ∣_∣ₕ (merid base ∙ sym (merid (ss23 a x i)))) i)) - ∙∙ cong ((cong ∣_∣ₕ) (sym (merid a)) ∙_) (cong (cong ∣_∣ₕ) (cong sym (symDistr (merid base) (sym (merid (invLooper a * x))))) - ∙ cong sym (meridP (invLooper a) x) - ∙ symDistr ((cong ∣_∣ₕ) (merid (invLooper a) ∙ sym (merid base))) - ((cong ∣_∣ₕ) (merid x ∙ sym (merid base))) - ∙ isCommΩK 2 (sym (λ i₁ → ∣ (merid x ∙ (λ i₂ → merid base (~ i₂))) i₁ ∣)) - (sym (λ i₁ → ∣ (merid (invLooper a) ∙ (λ i₂ → merid base (~ i₂))) i₁ ∣)) - ∙ cong₂ _∙_ (cong sym (mmInv a) ∙ cong-∙ ∣_∣ₕ (merid a) (sym (merid base))) - (cong (cong ∣_∣ₕ) (symDistr (merid x) (sym (merid base))) ∙ cong-∙ ∣_∣ₕ (merid base) (sym (merid x)))) - ∙∙ (λ j → (λ i₁ → ∣ merid a (~ i₁ ∨ j) ∣) ∙ - ((λ i₁ → ∣ merid a (i₁ ∨ j) ∣) ∙ - (λ i₁ → ∣ merid base (~ i₁ ∨ j) ∣)) - ∙ - (λ i₁ → ∣ merid base (i₁ ∨ j) ∣) ∙ - (λ i₁ → ∣ merid x (~ i₁) ∣ₕ)) - ∙∙ sym (lUnit _) - ∙∙ sym (assoc _ _ _) - ∙∙ (sym (lUnit _) ∙ sym (lUnit _) ∙ sym (lUnit _))) - - psst : (x : S₊ 2) → Q (CP2 , inl tt) fibr (inr x) →∙ coHomK-ptd 2 - fst (psst x) north = ∣ north ∣ - fst (psst x) south = ∣ x ∣ - fst (psst x) (merid a i₁) = kr (x , a) (~ i₁) - snd (psst x) = refl - - ll32-fst : GysinS1.c (inr north) .fst ≡ psst north .fst - ll32-fst = cong fst (GysinS1.cEq (inr north) ∣ push (inl base) ∣₂) - ∙ (funExt l) - where - l : (qb : _) → - ∣ (subst (fst ∘ Q (CP2 , inl tt) fibr) (sym (push (inl base))) qb) ∣ - ≡ psst north .fst qb - l north = refl - l south = cong ∣_∣ₕ (sym (merid base)) - l (merid a i) j = - hcomp (λ k → λ { (i = i0) → ∣ merid a (~ k ∧ j) ∣ - ; (i = i1) → ∣ merid base (~ j) ∣ - ; (j = i0) → ∣ transportRefl (merid a i) (~ k) ∣ - ; (j = i1) → ∣ compPath-filler (merid base) (sym (merid a)) k (~ i) ∣ₕ}) - (hcomp (λ k → λ { (i = i0) → ∣ merid a (j ∨ ~ k) ∣ - ; (i = i1) → ∣ merid base (~ j ∨ ~ k) ∣ - ; (j = i0) → ∣ merid a (~ k ∨ i) ∣ - ; (j = i1) → ∣ merid base (~ i ∨ ~ k) ∣ₕ}) - ∣ south ∣) - - is-setHepl : (x : S₊ 2) → isSet (Q (CP2 , inl tt) fibr (inr x) →∙ coHomK-ptd 2) - is-setHepl = sphereElim _ (λ _ → isProp→isOfHLevelSuc 1 (isPropIsOfHLevel 2)) - (isOfHLevel↑∙' 0 1) - - ll32 : (x : S₊ 2) → (GysinS1.c (inr x) ≡ psst x) - ll32 = sphereElim _ (λ x → isOfHLevelPath 2 (is-setHepl x) _ _) - (→∙Homogeneous≡ (isHomogeneousKn _) ll32-fst) - -isGenerator≃ℤ-e : isGenerator≃ℤ (coHomGr 2 CP2) GysinS1.e -isGenerator≃ℤ-e = - subst (isGenerator≃ℤ (coHomGr 2 CP2)) (sym ll) - (GroupEquivℤ-pres1g , refl) - -hopfFun : S₊ 3 → S₊ 2 -hopfFun x = fibS1.TotalSpaceHopf'→TotalSpace - (Iso.inv IsoTotalHopf' - (Iso.fun S³IsojoinS¹S¹ (Iso.inv IsoS³S3 x))) .fst - -CP2' : Type _ -CP2' = Pushout {A = S₊ 3} (λ _ → tt) hopfFun - -CP2≡CP2' : CP2' ≡ CP2 -CP2≡CP2' = - PushoutReplaceBase - (isoToEquiv (compIso (invIso IsoS³S3) - (compIso S³IsojoinS¹S¹ (invIso IsoTotalHopf')))) - - -{- -⌣-pres1-CP2 : ⌣-pres1 CP2 2 -fst ⌣-pres1-CP2 = GysinS1.e -fst (snd ⌣-pres1-CP2) = isGenerator≃ℤ-e -snd (snd ⌣-pres1-CP2) = ∣ p ∣ - where - p : Σ _ _ - fst p = compGroupIso (GroupEquiv→GroupIso (invGroupEquiv ⌣hom)) GroupEquivℤ-pres1g - snd p = {!!} - ∙ sesIsoPresGen _ (invGroupEquiv (GroupIso→GroupEquiv GroupEquivℤ-pres1g)) _ - (compGroupEquiv (invGroupEquiv (GroupIso→GroupEquiv GroupEquivℤ-pres1g)) (invGroupEquiv (invGroupEquiv ⌣hom))) - GysinS1.e {!!} ⌣hom -{- - GysinS1.e - , (isGenerator≃ℤ-e - , ∣ compGroupIso (GroupEquiv→GroupIso (invGroupEquiv ⌣hom)) GroupEquivℤ-pres1g - , {!!} ∣) - -} - -⌣-pres1-CP2' : ⌣-pres1 CP2' 2 -⌣-pres1-CP2' = subst (λ x → ⌣-pres1 x 2) (sym CP2≡CP2') ⌣-pres1-CP2 - -} --- {- --- Goal: snd (v' (pt A) (push a i₁)) ≡ --- ua-gluePt (μ-eq (snd a)) i₁ (fst a) --- ———— Boundary —————————————————————————————————————————————— --- i₁ = i0 ⊢ HSpace.μₗ e (fst a) --- i₁ = i1 ⊢ HSpace.μₗ e (HSpace.μ e (fst a) (snd a)) --- -} - --- -- help : (x : _) → (v' (pt A)) x ≡ TotalSpaceHopf'→TotalSpace x --- -- help (inl x) = ΣPathP (refl , HSpace.μₗ e x) --- -- help (inr x) = ΣPathP (refl , (HSpace.μₗ e x)) --- -- help (push (x , y) i) j = --- -- comp (λ _ → Σ (Susp (typ A)) Hopf) --- -- (λ k → λ {(i = i0) → merid y i , HSpace.μₗ e x j --- -- ; (i = i1) → merid y i , assocHSpace.μ-assoc-filler e-ass x y j k --- -- ; (j = i0) → merid y i , hfill --- -- (λ j → λ { (i = i0) → HSpace.μ e (pt A) x --- -- ; (i = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j --- -- }) --- -- (inS (ua-gluePt (μ-eq y) i (HSpace.μ e (pt A) x))) --- -- k --- -- ; (j = i1) → merid y i , ua-gluePt (μ-eq y) i x}) --- -- (merid y i , ua-gluePt (μ-eq y) i (HSpace.μₗ e x j)) --- -- where --- -- open import Cubical.Foundations.Path - --- -- PPΣ : ∀ {ℓ} {A : Type ℓ} {f : A ≃ A} (p : f ≡ f) → {!!} --- -- PPΣ = {!!} - --- -- V : PathP (λ i₁ → hcomp --- -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x --- -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j --- -- }) --- -- (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x)) ≡ --- -- ua-gluePt (μ-eq y) i₁ x) --- -- (HSpace.μₗ e x) --- -- (HSpace.μₗ e (HSpace.μ e x y)) -- (HSpace.μₗ e (HSpace.μ e (fst a) (snd a))) --- -- V = transport (λ z → {!PathP (λ i₁ → hfill --- -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x --- -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j --- -- }) --- -- (inS (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x))) z ≡ ua-gluePt (μ-eq y) i₁ x) --- -- ? ?!}) --- -- {!hfill --- -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x --- -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j --- -- }) --- -- (inS (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x))) ?!} -- toPathP ({!!} ∙∙ {!!} ∙∙ {!!}) -- toPathP (flipSquare {!!}) -- hcomp {!!} {!!} - --- -- P : Pushout {A = TotalSpaceHopf'} (λ _ → tt) induced → Type _ --- -- P (inl x) = typ A --- -- P (inr x) = Hopf x --- -- P (push a i₁) = ua (v a) i₁ - -HopfInvariantPushElim : ∀ {ℓ} n → (f : _) → {P : HopfInvariantPush n f → Type ℓ} - → (isOfHLevel (suc (suc (suc (suc (n +ℕ n))))) (P (inl tt))) - → (e : P (inl tt)) - (g : (x : _) → P (inr x)) - (r : PathP (λ i → P (push north i)) e (g (f north))) - → (x : _) → P x -HopfInvariantPushElim n f {P = P} hlev e g r (inl x) = e -HopfInvariantPushElim n f {P = P} hlev e g r (inr x) = g x -HopfInvariantPushElim n f {P = P} hlev e g r (push a i₁) = help a i₁ - where - help : (a : Susp (Susp (S₊ (suc (n +ℕ n))))) → PathP (λ i → P (push a i)) e (g (f a)) - help = sphereElim _ (sphereElim _ (λ _ → isProp→isOfHLevelSuc ((suc (suc (n +ℕ n)))) (isPropIsOfHLevel _)) - (isOfHLevelPathP' (suc (suc (suc (n +ℕ n)))) - (subst (isOfHLevel (suc (suc (suc (suc (n +ℕ n)))))) - (cong P (push north)) - hlev) _ _)) r - -transportCohomIso : ∀ {ℓ} {A : Type ℓ} {n m : ℕ} - → (p : n ≡ m) - → GroupIso (coHomGr n A) (coHomGr m A) -Iso.fun (fst (transportCohomIso {A = A} p)) = subst (λ n → coHom n A) p -Iso.inv (fst (transportCohomIso {A = A} p)) = subst (λ n → coHom n A) (sym p) -Iso.rightInv (fst (transportCohomIso p)) = transportTransport⁻ _ -Iso.leftInv (fst (transportCohomIso p)) = transportTransport⁻ _ -snd (transportCohomIso {A = A} {n = n} {m = m} p) = - makeIsGroupHom (λ x y → help x y p) - where - help : (x y : coHom n A) (p : n ≡ m) → subst (λ n → coHom n A) p (x +ₕ y) - ≡ subst (λ n → coHom n A) p x +ₕ subst (λ n → coHom n A) p y - help x y = J (λ m p → subst (λ n → coHom n A) p (x +ₕ y) - ≡ subst (λ n → coHom n A) p x +ₕ subst (λ n → coHom n A) p y) - (transportRefl (x +ₕ y) ∙ cong₂ _+ₕ_ (sym (transportRefl x)) (sym (transportRefl y))) - -α-hopfMap : abs (HopfInvariant zero (hopfFun , λ _ → hopfFun (snd (S₊∙ 3)))) ≡ suc zero -α-hopfMap = cong abs (cong (Iso.fun (fst (Hopfβ-Iso zero (hopfFun , refl)))) - (transportRefl (⌣-α 0 (hopfFun , refl)))) ∙ ⌣equiv→pres1 (sym CP2≡CP2') GysinS1.e (Hopfα zero (hopfFun , refl)) - (l isGenerator≃ℤ-e) - (GroupIso→GroupEquiv (Hopfα-Iso 0 (hopfFun , refl)) , refl) - (snd (fst ⌣hom)) - (GroupIso→GroupEquiv (Hopfβ-Iso zero (hopfFun , refl))) - where - l : Σ-syntax (GroupIso (coHomGr 2 CP2) ℤGroup) - (λ ϕ → abs (Iso.fun (fst ϕ) GysinS1.e) ≡ 1) - → Σ-syntax (GroupEquiv (coHomGr 2 CP2) ℤGroup) - (λ ϕ → abs (fst (fst ϕ) GysinS1.e) ≡ 1) - l p = (GroupIso→GroupEquiv (fst p)) , (snd p) - -·Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (S₊∙ n →∙ A) - → (S₊∙ n →∙ A) - → (S₊∙ n →∙ A) -·Π {A = A} {n = zero} p q = (λ _ → pt A) , refl -fst (·Π {A = A} {n = suc zero} (f , p) (g , q)) base = pt A -fst (·Π {A = A} {n = suc zero} (f , p) (g , q)) (loop j) = - ((sym p ∙∙ cong f loop ∙∙ p) ∙ (sym q ∙∙ cong g loop ∙∙ q)) j -snd (·Π {A = A} {n = suc zero} (f , p) (g , q)) = refl -fst (·Π {A = A} {n = suc (suc n)} (f , p) (g , q)) north = pt A -fst (·Π {A = A} {n = suc (suc n)} (f , p) (g , q)) south = pt A -fst (·Π {A = A} {n = suc (suc n)} (f , p) (g , q)) (merid a j) = - ((sym p ∙∙ cong f (merid a ∙ sym (merid (ptSn (suc n)))) ∙∙ p) - ∙ (sym q ∙∙ cong g (merid a ∙ sym (merid (ptSn (suc n)))) ∙∙ q)) j -snd (·Π {A = A} {n = suc (suc n)} (f , p) (g , q)) = refl - --Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (S₊∙ n →∙ A) - → (S₊∙ n →∙ A) --Π {n = zero} f = f -fst (-Π {A = A} {n = suc zero} f) base = fst f base -fst (-Π {A = A} {n = suc zero} f) (loop j) = fst f (loop (~ j)) -snd (-Π {A = A} {n = suc zero} f) = snd f -fst (-Π {A = A} {n = suc (suc n)} f) north = fst f north -fst (-Π {A = A} {n = suc (suc n)} f) south = fst f north -fst (-Π {A = A} {n = suc (suc n)} f) (merid a j) = fst f ((merid a ∙ sym (merid (ptSn _))) (~ j)) -snd (-Π {A = A} {n = suc (suc n)} f) = snd f - -open import Cubical.Foundations.Equiv.HalfAdjoint -module _ {ℓ : Level} {A : Pointed ℓ} where -flipLoopSpace' : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → ((Ω^ (suc n)) A) ≡ (Ω^ n) (Ω A) -flipLoopSpace' {A = A} zero = refl -flipLoopSpace' {A = A} (suc n) = cong Ω (flipLoopSpace' {A = A} n) - -flipLoopSpace : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → Iso (fst ((Ω^ (suc n)) A)) (fst ((Ω^ n) (Ω A))) -flipLoopSpace {A = A} n = pathToIso (cong fst (flipLoopSpace' n)) - -Ω'∙ : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → fst ((Ω^ n) (Ω A)) → fst ((Ω^ n) (Ω A)) → fst ((Ω^ n) (Ω A)) -Ω'∙ zero p q = p ∙ q -Ω'∙ (suc n) p q = p ∙ q - -flipLoopSpacepres· : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → (f g : fst ((Ω^ n) (Ω A))) - → Iso.inv (flipLoopSpace n) (Ω'∙ n f g) - ≡ (Iso.inv (flipLoopSpace n) f) ∙ (Iso.inv (flipLoopSpace n) g) -flipLoopSpacepres· zero f g = transportRefl (f ∙ g) ∙ cong₂ _∙_ (sym (transportRefl f)) (sym (transportRefl g)) -flipLoopSpacepres· {A = A} (suc n) f g i = - transp (λ j → flipLoopSpace' {A = A} n (~ i ∧ ~ j) .snd ≡ flipLoopSpace' n (~ i ∧ ~ j) .snd) i - (transp (λ j → flipLoopSpace' {A = A} n (~ i ∨ ~ j) .snd ≡ flipLoopSpace' n (~ i ∨ ~ j) .snd) (~ i) f - ∙ transp (λ j → flipLoopSpace' {A = A} n (~ i ∨ ~ j) .snd ≡ flipLoopSpace' n (~ i ∨ ~ j) .snd) (~ i) g) - -module _ {ℓ : Level} {A : Pointed ℓ} where - convertLoopFun→ : (n : ℕ) → (S₊∙ (suc n) →∙ Ω A) → (S₊∙ (suc (suc n)) →∙ A) - fst (convertLoopFun→ n f) north = pt A - fst (convertLoopFun→ n f) south = pt A - fst (convertLoopFun→ n f) (merid a i₁) = fst f a i₁ - snd (convertLoopFun→ n f) = refl - - convertLoopFun-pres∙ : - (n : ℕ) (f g : (S₊∙ (suc n) →∙ Ω A)) - → convertLoopFun→ n (·Π f g) ≡ ·Π (convertLoopFun→ n f) (convertLoopFun→ n g) - convertLoopFun-pres∙ n f g = ΣPathP ((funExt l) , refl) - where - J2 : ∀ {ℓ ℓ'} {A : Type ℓ} {x : A} - → (P : (p q : x ≡ x) → (refl ≡ p) → (refl ≡ q) → Type ℓ') - → P refl refl refl refl - → (p q : x ≡ x) (P' : _) (Q : _) → P p q P' Q - J2 P b p q = - J (λ p P' → (Q₁ : refl ≡ q) → P p q P' Q₁) - (J (λ q Q → P refl q refl Q) b) - - J4 : ∀ {ℓ ℓ'} {A : Type ℓ} {x : A} - → (P : (p q r s : x ≡ x) → (refl ≡ p) → (p ≡ q) → (refl ≡ r) → (r ≡ s) → Type ℓ') - → P refl refl refl refl refl refl refl refl - → (p q r s : x ≡ x) (P' : _) (Q : _) (R : _) (S : _) → P p q r s P' Q R S - J4 P b p q r s = - J (λ p P' → (Q₁ : p ≡ q) (R : refl ≡ r) (S₁ : r ≡ s) → P p q r s P' Q₁ R S₁) - (J (λ q Q₁ → (R : refl ≡ r) (S₁ : r ≡ s) → P refl q r s refl Q₁ R S₁) - (J (λ r R → (S₁ : r ≡ s) → P refl refl r s refl refl R S₁) - (J (λ s S₁ → P refl refl refl s refl refl refl S₁) - b))) - - looper-help : (n : ℕ) (f g : (S₊∙ (suc n) →∙ Ω A)) (a : _) - → fst (·Π f g) a ≡ fst f a ∙ fst g a - looper-help zero f g base = rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g)) - looper-help zero f g (loop i) k = - lazy-fill _ _ (sym (snd f)) (sym (snd g)) (cong (fst f) loop) (cong (fst g) loop) k i - where - lazy-fill : ∀ {ℓ} {A : Type ℓ} {x : A} (p q : x ≡ x) - (prefl : refl ≡ p) (qrefl : refl ≡ q) (fl : p ≡ p) (fr : q ≡ q) - → PathP (λ i → (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i - ≡ (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i) - ((prefl ∙∙ fl ∙∙ sym prefl) ∙ (qrefl ∙∙ fr ∙∙ sym qrefl)) - (cong₂ _∙_ fl fr) - lazy-fill {x = x} = - J2 (λ p q prefl qrefl → (fl : p ≡ p) (fr : q ≡ q) - → PathP (λ i → (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i - ≡ (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i) - ((prefl ∙∙ fl ∙∙ sym prefl) ∙ (qrefl ∙∙ fr ∙∙ sym qrefl)) - (cong₂ _∙_ fl fr)) - λ fl fr → cong₂ _∙_ (sym (rUnit fl)) (sym (rUnit fr)) - ◁ flipSquare (sym (rUnit (rUnit refl)) - ◁ (flipSquare ((λ j → cong (λ x → rUnit x j) fl - ∙ cong (λ x → lUnit x j) fr) - ▷ sym (cong₂Funct _∙_ fl fr)) - ▷ (rUnit (rUnit refl)))) - looper-help (suc n) f g north = rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g)) - looper-help (suc n) f g south = (rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) - ∙ cong₂ _∙_ (cong (fst f) (merid (ptSn _))) (cong (fst g) (merid (ptSn _))) - looper-help (suc n) f g (merid a j) i = help i j - where - help : PathP (λ i → (rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) i - ≡ ((rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) - ∙ cong₂ _∙_ (cong (fst f) (merid (ptSn _))) (cong (fst g) (merid (ptSn _)))) i) - (cong (fst (·Π f g)) (merid a)) - (cong₂ _∙_ (cong (fst f) (merid a)) (cong (fst g) (merid a))) - help = (cong₂ _∙_ (cong (sym (snd f) ∙∙_∙∙ snd f) (cong-∙ (fst f) (merid a) (sym (merid (ptSn _))))) - ((cong (sym (snd g) ∙∙_∙∙ snd g) (cong-∙ (fst g) (merid a) (sym (merid (ptSn _)))))) - ◁ lazy-help _ _ _ _ (sym (snd f)) (cong (fst f) (merid (ptSn _))) - (sym (snd g)) (cong (fst g) (merid (ptSn _))) - (cong (fst f) (merid a)) (cong (fst g) (merid a))) - where - lazy-help : ∀ {ℓ} {A : Type ℓ} {x : A} (p-north p-south q-north q-south : x ≡ x) - (prefl : refl ≡ p-north) (p-merid : p-north ≡ p-south) - (qrefl : refl ≡ q-north) (q-merid : q-north ≡ q-south) - (fl : p-north ≡ p-south) (fr : q-north ≡ q-south) - → PathP (λ i → (rUnit refl ∙ cong₂ _∙_ prefl qrefl) i - ≡ ((rUnit refl ∙ cong₂ _∙_ prefl qrefl) - ∙ cong₂ _∙_ p-merid q-merid) i) - ((prefl ∙∙ fl ∙ sym p-merid ∙∙ sym prefl) - ∙ (qrefl ∙∙ fr ∙ sym q-merid ∙∙ sym qrefl)) - (cong₂ _∙_ fl fr) - lazy-help {x = x} = - J4 (λ p-north p-south q-north q-south prefl p-merid qrefl q-merid - → (fl : p-north ≡ p-south) (fr : q-north ≡ q-south) → - PathP (λ i₁ → (rUnit refl ∙ cong₂ _∙_ prefl qrefl) i₁ ≡ - ((rUnit refl ∙ cong₂ _∙_ prefl qrefl) ∙ cong₂ _∙_ p-merid q-merid) i₁) - ((prefl ∙∙ fl ∙ sym p-merid ∙∙ sym prefl) ∙ - (qrefl ∙∙ fr ∙ sym q-merid ∙∙ sym qrefl)) - (cong₂ _∙_ fl fr)) - λ fl fr → (cong₂ _∙_ (sym (rUnit (fl ∙ refl)) ∙ sym (rUnit fl)) - (sym (rUnit (fr ∙ refl)) ∙ sym (rUnit fr)) - ◁ flipSquare ((sym (rUnit _) - ◁ flipSquare - ((λ j → cong (λ x → rUnit x j) fl - ∙ cong (λ x → lUnit x j) fr) - ▷ sym (cong₂Funct _∙_ fl fr))) - ▷ (rUnit _ ∙ rUnit _))) - l : (x : fst (S₊∙ (suc (suc n)))) → - fst (convertLoopFun→ n (·Π f g)) x ≡ - fst (·Π (convertLoopFun→ n f) (convertLoopFun→ n g)) x - l north = refl - l south = refl - l (merid a i₁) j = help j i₁ - where - help : fst (·Π f g) a ≡ cong (fst (·Π (convertLoopFun→ n f) (convertLoopFun→ n g))) (merid a) - help = looper-help n f g a - ∙ cong₂ _∙_ (sym (cong-∙ (fst (convertLoopFun→ n f)) (merid a) (sym (merid (ptSn _))) - ∙∙ cong (fst f a ∙_) (cong sym (snd f)) - ∙∙ sym (rUnit _))) - (sym (cong-∙ (fst (convertLoopFun→ n g)) (merid a) (sym (merid (ptSn _))) - ∙∙ cong (fst g a ∙_) (cong sym (snd g)) - ∙∙ sym (rUnit _))) - ∙ λ i → rUnit (cong (fst (convertLoopFun→ n f)) (merid a ∙ sym (merid (ptSn _)))) i - ∙ rUnit (cong (fst (convertLoopFun→ n g)) (merid a ∙ sym (merid (ptSn _)))) i - - myFiller : (n : ℕ) → (S₊∙ (suc (suc n)) →∙ A) → I → I → I → fst A - myFiller n f i j k = - hfill (λ k → λ { (i = i0) → doubleCompPath-filler - (sym (snd f)) - (cong (fst f) (merid (ptSn _) ∙ sym (merid (ptSn _)))) - (snd f) k j - ; (i = i1) → snd f k - ; (j = i0) → snd f k - ; (j = i1) → snd f k}) - (inS ((fst f) (rCancel' (merid (ptSn _)) i j))) - k - - convertLoopFun← : (n : ℕ) → (S₊∙ (suc (suc n)) →∙ A) → (S₊∙ (suc n) →∙ Ω A) - fst (convertLoopFun← n f) a = - sym (snd f) ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f - snd (convertLoopFun← n f) i j = myFiller n f i j i1 - - convertLoopFun : (n : ℕ) → Iso (S₊∙ (suc n) →∙ Ω A) (S₊∙ (suc (suc n)) →∙ A) - Iso.fun (convertLoopFun n) f = convertLoopFun→ n f - Iso.inv (convertLoopFun n) f = convertLoopFun← n f - Iso.rightInv (convertLoopFun n) (f , p) = - ΣPathP (funExt help , λ i j → p (~ i ∨ j)) - where - help : (x : _) → fst (convertLoopFun→ n (convertLoopFun← n (f , p))) x ≡ f x - help north = sym p - help south = sym p ∙ cong f (merid (ptSn _)) - help (merid a j) i = - hcomp (λ k → λ { (i = i1) → f (merid a j) - ; (j = i0) → p (~ i ∧ k) - ; (j = i1) → compPath-filler' (sym p) (cong f (merid (ptSn _))) k i}) - (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) - Iso.leftInv (convertLoopFun n) f = - ΣPathP (funExt help34 , help2) - where - - sillyFill₂ : (x : S₊ (suc n)) → I → I → I → fst A - sillyFill₂ x i j k = - hfill (λ k → λ { (i = i0) → convertLoopFun→ n f .fst - (compPath-filler' (merid x) (sym (merid (ptSn _))) k j) - ; (i = i1) → snd A - ; (j = i0) → fst f x (i ∨ ~ k) - ; (j = i1) → snd A}) - (inS (snd f i (~ j))) - k - - sillyFill : (x : S₊ (suc n)) → I → I → I → fst A - sillyFill x i j k = - hfill (λ k → λ { (i = i0) → doubleCompPath-filler - refl - (cong (convertLoopFun→ n f .fst) (merid x ∙ sym (merid (ptSn _)))) - refl k j - ; (i = i1) → fst f x (j ∨ ~ k) - ; (j = i0) → fst f x (i ∧ ~ k) - ; (j = i1) → snd A}) - (inS (sillyFill₂ x i j i1)) - k - - compPath-filler'' : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) - → I → I → I → A - compPath-filler'' {z = z} p q k j i = - hfill (λ k → λ { (i = i0) → p (~ j) - ; (i = i1) → q k - ; (j = i0) → q (i ∧ k) }) - (inS (p (i ∨ ~ j))) - k - - help34 : (x : _) → fst (convertLoopFun← n (convertLoopFun→ n f)) x ≡ fst f x - help34 x i j = sillyFill x i j i1 - - sideCube : Cube (λ j k → snd f j (~ k)) (λ j k → fst (convertLoopFun→ n f) (rCancel' (merid (ptSn _)) j k)) - (λ r k → convertLoopFun→ n f .fst (compPath-filler' (merid (ptSn _)) (sym (merid (ptSn _))) r k)) - (λ _ _ → pt A) - (λ r j → snd f j (~ r)) λ _ _ → pt A -- I → I → I → fst A -- r j k - sideCube r j k = - hcomp (λ i → λ {(r = i0) → snd f j (~ k ∨ ~ i) - ; (r = i1) → fst (convertLoopFun→ n f) (rCancel-filler' (merid (ptSn _)) (~ i) j k) - ; (j = i0) → convertLoopFun→ n f .fst - (compPath-filler'' (merid (ptSn _)) (sym (merid (ptSn _))) i r k) - ; (j = i1) → snd f (~ r) (~ i ∧ k) - ; (k = i0) → snd f j (~ r) - ; (k = i1) → snd f (~ r ∧ j) (~ i)}) - (hcomp (λ i → λ {(r = i0) → pt A - ; (r = i1) → snd f (~ i) k - ; (j = i0) → snd f (~ i) (k ∨ ~ r) - ; (j = i1) → snd f (~ r ∨ ~ i) k - ; (k = i0) → snd f (j ∨ ~ i) (~ r) - ; (k = i1) → pt A}) - (pt A)) - - help2 : PathP _ _ _ - help2 i j k = - hcomp (λ r → λ {(i = i0) → myFiller n (convertLoopFun→ n f) j k r - ; (i = i1) → snd f j (k ∨ ~ r) - ; (j = i0) → sillyFill (ptSn _) i k r - ; (j = i1) → snd A - ; (k = i0) → snd f j (i ∧ ~ r) - ; (k = i1) → snd A}) - (hcomp ((λ r → λ {(i = i0) → sideCube r j k - ; (i = i1) → pt A - ; (j = i0) → sillyFill₂ (ptSn _) i k r - ; (j = i1) → pt A - ; (k = i0) → snd f j (i ∨ ~ r) - ; (k = i1) → pt A})) - (snd f (i ∨ j) (~ k))) - - convertLoopFun←-pres∙ : (n : ℕ) → (f g : S₊∙ (suc (suc n)) →∙ A) - → convertLoopFun← n (·Π f g) - ≡ ·Π (convertLoopFun← n f) (convertLoopFun← n g) - convertLoopFun←-pres∙ n f g = - sym (Iso.leftInv (convertLoopFun n) (convertLoopFun← n (·Π f g))) - ∙∙ cong (convertLoopFun← n) - (Iso.rightInv (convertLoopFun n) (·Π f g) - ∙∙ cong₂ ·Π (sym (Iso.rightInv (convertLoopFun n) f)) (sym (Iso.rightInv (convertLoopFun n) g)) - ∙∙ sym (convertLoopFun-pres∙ _ (convertLoopFun← n f) (convertLoopFun← n g))) - ∙∙ (Iso.leftInv (convertLoopFun n) (·Π (convertLoopFun← n f) (convertLoopFun← n g))) - -→' : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) → (S₊∙ (suc n) →∙ A) → (fst ((Ω^ (suc n)) A)) -→' zero (f , p) = sym p ∙∙ cong f loop ∙∙ p -→' {A = A} (suc n) (f , p) = Iso.inv (flipLoopSpace _) (→' {A = Ω A} n (Iso.inv (convertLoopFun _) (f , p))) - -→'pres∙ : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → (p q : S₊∙ (suc n) →∙ A) → →' n (·Π p q) ≡ →' n p ∙ →' n q -→'pres∙ zero f g = sym (rUnit _) -→'pres∙ (suc n) f g = - cong (Iso.inv (flipLoopSpace (suc n))) - (cong (→' n) (convertLoopFun←-pres∙ n f g) ∙ →'pres∙ n (convertLoopFun← n f) ((convertLoopFun← n g))) - ∙ flipLoopSpacepres· (suc n) _ _ - -←' : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) → (fst ((Ω^ (suc n)) A)) → (S₊∙ (suc n) →∙ A) -fst (←' {A = A} zero f) base = pt A -fst (←' {A = A} zero f) (loop i₁) = f i₁ -snd (←' {A = A} zero f) = refl -←' {A = A} (suc n) f = Iso.fun (convertLoopFun _) (←' n (Iso.fun (flipLoopSpace _) f)) - -cancel-r : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → (x : fst ((Ω^ (suc n)) A)) - → →' n (←' n x) ≡ x -cancel-r zero x = sym (rUnit _) -cancel-r (suc n) x = - cong (Iso.inv (flipLoopSpace (suc n)) ∘ →' n) - (Iso.leftInv (convertLoopFun _) - (←' n (Iso.fun (flipLoopSpace (suc n)) x))) - ∙∙ cong (Iso.inv (flipLoopSpace (suc n))) - (cancel-r n (Iso.fun (flipLoopSpace (suc n)) x)) - ∙∙ Iso.leftInv (flipLoopSpace (suc n)) x - -cancel-l : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → (x : (S₊∙ (suc n) →∙ A)) - → ←' n (→' n x) ≡ x -cancel-l zero (f , p) = ΣPathP (funExt fstp , λ i j → p (~ i ∨ j)) - where - fstp : (x : _) → _ ≡ f x - fstp base = sym p - fstp (loop i) j = doubleCompPath-filler (sym p) (cong f loop) p (~ j) i -cancel-l (suc n) f = - cong (Iso.fun (convertLoopFun n) ∘ ←' n) (Iso.rightInv (flipLoopSpace (suc n)) - (→' n (Iso.inv (convertLoopFun n) f))) - ∙∙ cong (Iso.fun (convertLoopFun n)) (cancel-l n (Iso.inv (convertLoopFun n) f)) - ∙∙ Iso.rightInv (convertLoopFun _) _ - -→∙Ω : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → Iso (S₊∙ n →∙ A) (fst ((Ω^ n) A)) -Iso.fun (→∙Ω {A = A} zero) (f , p) = f false -fst (Iso.inv (→∙Ω {A = A} zero) a) false = a -fst (Iso.inv (→∙Ω {A = A} zero) a) true = pt A -snd (Iso.inv (→∙Ω {A = A} zero) a) = refl -Iso.rightInv (→∙Ω {A = A} zero) a = refl -Iso.leftInv (→∙Ω {A = A} zero) (f , p) = - ΣPathP ((funExt (λ { false → refl ; true → sym p})) , λ i j → p (~ i ∨ j)) -Iso.fun (→∙Ω {A = A} (suc n)) = →' n -Iso.inv (→∙Ω {A = A} (suc n)) = ←' n -Iso.rightInv (→∙Ω {A = A} (suc n)) = cancel-r n -Iso.leftInv (→∙Ω {A = A} (suc n)) = cancel-l n - -→∙Ω-pres·Π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f g : S₊∙ (suc n) →∙ A) - → →' n (·Π f g) ≡ →' n f ∙ →' n g -→∙Ω-pres·Π zero f g = sym (rUnit _) -→∙Ω-pres·Π (suc n) f g = - cong (Iso.inv (flipLoopSpace (suc n))) (cong (→' n) (convertLoopFun←-pres∙ n f g) - ∙ →'pres∙ n (convertLoopFun← n f) (convertLoopFun← n g)) - ∙ (flipLoopSpacepres· (suc n) (→' n (convertLoopFun← n f)) (→' n (convertLoopFun← n g))) - -0Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → (S₊∙ n →∙ A) -fst (0Π {A = A}) _ = pt A -snd (0Π {A = A}) = refl - -rUnit·Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f : S₊∙ (suc n) →∙ A) - → ·Π f 0Π ≡ f -fst (rUnit·Π {A = A} {n = zero} f i) base = snd f (~ i) -fst (rUnit·Π {A = A} {n = zero} f i) (loop j) = help i j - where - help : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) - (((sym (snd f)) ∙∙ (cong (fst f) loop) ∙∙ snd f) - ∙ (refl ∙ refl)) - (cong (fst f) loop) - help = (cong ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) ∙_) (sym (rUnit refl)) ∙ sym (rUnit _)) - ◁ λ i j → doubleCompPath-filler (sym (snd f)) (cong (fst f) loop) (snd f) (~ i) j -snd (rUnit·Π {A = A} {n = zero} f i) j = snd f (~ i ∨ j) -fst (rUnit·Π {A = A} {n = suc n} f i) north = snd f (~ i) -fst (rUnit·Π {A = A} {n = suc n} f i) south = (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i -fst (rUnit·Π {A = A} {n = suc n} f i) (merid a j) = help i j - where - help : PathP (λ i → snd f (~ i) ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) - (((sym (snd f)) ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) ∙∙ snd f) - ∙ (refl ∙ refl)) - (cong (fst f) (merid a)) - help = (cong (((sym (snd f)) ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) ∙∙ snd f) ∙_) - (sym (rUnit refl)) - ∙ sym (rUnit _)) - ◁ λ i j → hcomp (λ k → λ { (j = i0) → snd f (~ i ∧ k) - ; (j = i1) → compPath-filler' (sym (snd f)) (cong (fst f) (merid (ptSn (suc n)))) k i - ; (i = i0) → doubleCompPath-filler (sym (snd f)) - (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) - (snd f) k j - ; (i = i1) → fst f (merid a j)}) - (fst f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) -snd (rUnit·Π {A = A} {n = suc n} f i) j = snd f (~ i ∨ j) - -lUnit·Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f : S₊∙ (suc n) →∙ A) - → ·Π 0Π f ≡ f -fst (lUnit·Π {n = zero} f i) base = snd f (~ i) -fst (lUnit·Π {n = zero} f i) (loop j) = s i j - where - s : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) - ((refl ∙ refl) ∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) - (cong (fst f) loop) - s = (cong (_∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) (sym (rUnit refl)) ∙ sym (lUnit _)) - ◁ λ i j → doubleCompPath-filler (sym (snd f)) (cong (fst f) loop) (snd f) (~ i) j -snd (lUnit·Π {n = zero} f i) j = snd f (~ i ∨ j) -fst (lUnit·Π {n = suc n} f i) north = snd f (~ i) -fst (lUnit·Π {n = suc n} f i) south = (sym (snd f) ∙ cong (fst f) (merid (ptSn _))) i -fst (lUnit·Π {n = suc n} f i) (merid a j) = help i j - where - help : PathP (λ i → snd f (~ i) ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) - ((refl ∙ refl) ∙ ((sym (snd f)) ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) ∙∙ snd f)) - (cong (fst f) (merid a)) - help = (cong (_∙ ((sym (snd f)) ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) ∙∙ snd f)) - (sym (rUnit refl)) - ∙ sym (lUnit _)) - ◁ λ i j → hcomp (λ k → λ { (j = i0) → snd f (~ i ∧ k) - ; (j = i1) → compPath-filler' (sym (snd f)) (cong (fst f) (merid (ptSn (suc n)))) k i - ; (i = i0) → doubleCompPath-filler (sym (snd f)) - (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) - (snd f) k j - ; (i = i1) → fst f (merid a j)}) - (fst f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) -snd (lUnit·Π {n = suc n} f i) j = snd f (~ i ∨ j) - -rCancel·Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f : S₊∙ (suc n) →∙ A) - → ·Π f (-Π f) ≡ 0Π -fst (rCancel·Π {A = A} {n = zero} f i) base = pt A -fst (rCancel·Π {A = A} {n = zero} f i) (loop j) = rCancel (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) i j -snd (rCancel·Π {A = A} {n = zero} f i) = refl -fst (rCancel·Π {A = A} {n = suc n} f i) north = pt A -fst (rCancel·Π {A = A} {n = suc n} f i) south = pt A -fst (rCancel·Π {A = A} {n = suc n} f i) (merid a i₁) = sl i i₁ - where - pl = (sym (snd f) ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f) - sl : pl - ∙ ((sym (snd f) ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f)) ≡ refl - sl = cong (pl ∙_) (cong (sym (snd f) ∙∙_∙∙ (snd f)) - (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) - ∙∙ cong₂ _∙_ refl - (cong (cong (fst f)) (rCancel (merid (ptSn _)))) - ∙∙ sym (rUnit _))) - ∙ rCancel pl -snd (rCancel·Π {A = A} {n = suc n} f i) = refl - -lCancel·Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f : S₊∙ (suc n) →∙ A) - → ·Π (-Π f) f ≡ 0Π -fst (lCancel·Π {A = A} {n = zero} f i) base = pt A -fst (lCancel·Π {A = A} {n = zero} f i) (loop j) = - rCancel (sym (snd f) ∙∙ cong (fst f) (sym loop) ∙∙ snd f) i j -fst (lCancel·Π {A = A} {n = suc n} f i) north = pt A -fst (lCancel·Π {A = A} {n = suc n} f i) south = pt A -fst (lCancel·Π {A = A} {n = suc n} f i) (merid a j) = sl i j - where - pl = (sym (snd f) ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f) - - sl : ((sym (snd f) ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f)) ∙ pl ≡ refl - sl = cong (_∙ pl) (cong (sym (snd f) ∙∙_∙∙ (snd f)) - (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) - ∙∙ cong₂ _∙_ refl (cong (cong (fst f)) (rCancel (merid (ptSn _)))) - ∙∙ sym (rUnit _))) - ∙ lCancel pl -snd (lCancel·Π {A = A} {n = zero} f i) = refl -snd (lCancel·Π {A = A} {n = suc n} f i) = refl - -assoc·Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f g h : S₊∙ (suc n) →∙ A) - → ·Π f (·Π g h) ≡ ·Π (·Π f g) h -assoc·Π {n = n} f g h = - sym (Iso.leftInv (→∙Ω (suc n)) (·Π f (·Π g h))) - ∙∙ cong (←' n) (→'pres∙ n f (·Π g h) - ∙∙ cong (→' n f ∙_) (→'pres∙ n g h) - ∙∙ assoc _ _ _ - ∙∙ cong (_∙ →' n h) (sym (→'pres∙ n f g)) - ∙∙ sym (→'pres∙ n (·Π f g) h)) - ∙∙ Iso.leftInv (→∙Ω (suc n)) (·Π (·Π f g) h) - -comm·Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f g : S₊∙ (suc (suc n)) →∙ A) - → ·Π f g ≡ ·Π g f -comm·Π {A = A} {n = n} f g = - sym (Iso.leftInv (→∙Ω (suc (suc n))) (·Π f g)) - ∙∙ cong (←' (suc n)) (→'pres∙ (suc n) f g - ∙∙ EH _ _ _ - ∙∙ sym (→'pres∙ (suc n) g f)) - ∙∙ Iso.leftInv (→∙Ω (suc (suc n))) (·Π g f) - -module _ {ℓ} {A : Pointed ℓ} {n : ℕ} where - _·π_ : ∥ S₊∙ n →∙ A ∥₂ - → ∥ S₊∙ n →∙ A ∥₂ - → ∥ S₊∙ n →∙ A ∥₂ - _·π_ = sRec2 squash₂ λ f g → ∣ ·Π f g ∣₂ - - -π : ∥ S₊∙ n →∙ A ∥₂ → ∥ S₊∙ n →∙ A ∥₂ - -π = sMap -Π - - 0π : ∥ S₊∙ n →∙ A ∥₂ - 0π = ∣ 0Π ∣₂ - - -module _ {ℓ} {A : Pointed ℓ} {n : ℕ} where - 2path : ∀ {ℓ} {A : Type ℓ} {x y : ∥ A ∥₂} → isSet (x ≡ y) - 2path = isProp→isSet (squash₂ _ _) - - rUnit·π : (f : ∥ S₊∙ (suc n) →∙ A ∥₂) → f ·π 0π ≡ f - rUnit·π = sElim (λ _ → 2path) λ f i → ∣ rUnit·Π f i ∣₂ - - lUnit·π : (f : ∥ S₊∙ (suc n) →∙ A ∥₂) → 0π ·π f ≡ f - lUnit·π = sElim (λ _ → 2path) λ f i → ∣ lUnit·Π f i ∣₂ - - rCancel·π : (f : ∥ S₊∙ (suc n) →∙ A ∥₂) → f ·π -π f ≡ 0π - rCancel·π = sElim (λ _ → 2path) λ f i → ∣ rCancel·Π f i ∣₂ - - lCancel·π : (f : ∥ S₊∙ (suc n) →∙ A ∥₂) → -π f ·π f ≡ 0π - lCancel·π = sElim (λ _ → 2path) λ f i → ∣ lCancel·Π f i ∣₂ - - assoc·π : (f g h : ∥ S₊∙ (suc n) →∙ A ∥₂) → f ·π (g ·π h) ≡ (f ·π g) ·π h - assoc·π = sElim3 (λ _ _ _ → 2path) λ f g h i → ∣ assoc·Π f g h i ∣₂ - - comm·π : (f g : ∥ S₊∙ (suc (suc n)) →∙ A ∥₂) → f ·π g ≡ g ·π f - comm·π = sElim2 (λ _ _ → 2path) λ f g i → ∣ comm·Π f g i ∣₂ - - -open import Cubical.HITs.Wedge -_∨→_ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} - → (f : A →∙ C) (g : B →∙ C) - → A ⋁ B → fst C -(f ∨→ g) (inl x) = fst f x -(f ∨→ g) (inr x) = fst g x -(f ∨→ g) (push a i₁) = (snd f ∙ sym (snd g)) i₁ - - -add2 : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → (f g : S₊∙ n →∙ A) → ((S₊∙ n) ⋁ (S₊∙ n) , inl (ptSn _)) →∙ A -fst (add2 {A = A} f g) = f ∨→ g -snd (add2 {A = A} f g) = snd f - -C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Type _ -C* n f g = Pushout (λ _ → tt) (fst (add2 f g)) - -θ : ∀ {ℓ} {A : Pointed ℓ} → Susp (fst A) → (Susp (fst A) , north) ⋁ (Susp (fst A) , north) -θ north = inl north -θ south = inr north -θ {A = A} (merid a i₁) = - ((λ i → inl ((merid a ∙ sym (merid (pt A))) i)) - ∙∙ push tt - ∙∙ λ i → inr ((merid a ∙ sym (merid (pt A))) i)) i₁ - - -help3 : ∀ {ℓ} {A : Type ℓ} {x y z w : A} (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) - → p ∙∙ q ∙∙ r ≡ p ∙ q ∙ r -help3 {x = x} {y = y} {z = z} {w = w} = - J (λ y p → (q : y ≡ z) (r : z ≡ w) → - (p ∙∙ q ∙∙ r) ≡ p ∙ q ∙ r) - λ q r → lUnit (q ∙ r) - -+≡ : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → (x : _) → ·Π f g .fst x ≡ (f ∨→ g) (θ {A = S₊∙ _} x) -+≡ n f g north = sym (snd f) -+≡ n f g south = sym (snd g) -+≡ n f g (merid a i₁) j = main j i₁ - where - - help : cong (f ∨→ g) (cong (θ {A = S₊∙ _}) (merid a)) - ≡ (refl ∙∙ cong (fst f) (merid a ∙ sym (merid north)) ∙∙ snd f) - ∙ (sym (snd g) ∙∙ cong (fst g) (merid a ∙ sym (merid north)) ∙∙ refl) - help = cong-∙∙ (f ∨→ g) ((λ i → inl ((merid a ∙ sym (merid north)) i))) - (push tt) - (λ i → inr ((merid a ∙ sym (merid north)) i)) - ∙∙ help3 _ _ _ - ∙∙ cong (cong (f ∨→ g) - (λ i₂ → inl ((merid a ∙ (λ i₃ → merid north (~ i₃))) i₂)) ∙_) - (sym (assoc _ _ _)) - ∙∙ assoc _ _ _ - ∙∙ cong₂ _∙_ refl (compPath≡compPath' _ _) - - main : PathP (λ i → snd f (~ i) ≡ snd g (~ i)) (λ i₁ → ·Π f g .fst (merid a i₁)) - λ i₁ → (f ∨→ g) (θ {A = S₊∙ _} (merid a i₁)) - main = (λ i → ((λ j → snd f (~ i ∧ ~ j)) ∙∙ cong (fst f) (merid a ∙ sym (merid north)) ∙∙ snd f) - ∙ (sym (snd g) ∙∙ cong (fst g) (merid a ∙ sym (merid north)) ∙∙ λ j → snd g (~ i ∧ j))) - ▷ sym help - -+≡' : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → ·Π f g ≡ ((f ∨→ g) ∘ (θ {A = S₊∙ _}) , snd f) -+≡' n f g = ΣPathP ((funExt (+≡ n f g)) , (λ j i → snd f (i ∨ ~ j))) - -WedgeElim : (n : ℕ) → ∀ {ℓ} {P : S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) → Type ℓ} - → ((x : _) → isOfHLevel (3 +ℕ n) (P x)) - → P (inl north) - → (x : _) → P x -WedgeElim n {P = P} x s (inl x₁) = - sphereElim _ {A = P ∘ inl} - (λ _ → isOfHLevelPlus' {n = n} (3 +ℕ n) (x _)) s x₁ -WedgeElim n {P = P} x s (inr x₁) = - sphereElim _ {A = P ∘ inr} - (sphereElim _ (λ _ → isProp→isOfHLevelSuc ((suc (suc (n +ℕ n)))) - (isPropIsOfHLevel (suc (suc (suc (n +ℕ n)))))) - (subst (isOfHLevel ((3 +ℕ n) +ℕ n)) (cong P (push tt)) - (isOfHLevelPlus' {n = n} (3 +ℕ n) (x _)))) - (subst P (push tt) s) x₁ -WedgeElim n {P = P} x s (push a j) = transp (λ i → P (push tt (i ∧ j))) (~ j) s - - -H²-C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → GroupIso (coHomGr (2 +ℕ n) (C* n f g)) ℤGroup -H²-C* n f g = compGroupIso groupIso1 (Hⁿ-Sⁿ≅ℤ (suc n)) - where - help : (coHom (2 +ℕ n) (C* n f g)) → coHom (2 +ℕ n) (S₊ (2 +ℕ n)) - help = sMap λ f → f ∘ inr - - - invMapPrim : (S₊ (2 +ℕ n) → coHomK (2 +ℕ n)) → C* n f g → coHomK (2 +ℕ n) - invMapPrim h (inl x) = h (ptSn _) - invMapPrim h (inr x) = h x - invMapPrim h (push a i₁) = WedgeElim n {P = λ a → h north ≡ h (fst (add2 f g) a)} - (λ _ → isOfHLevelTrunc (4 +ℕ n) _ _) - (cong h (sym (snd f))) a i₁ - - invMap : coHom (2 +ℕ n) (S₊ (2 +ℕ n)) → (coHom (2 +ℕ n) (C* n f g)) - invMap = sMap invMapPrim - - ret1 : (x : coHom (2 +ℕ n) (S₊ (2 +ℕ n))) → help (invMap x) ≡ x - ret1 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - λ a → refl - - ret2 : (x : (coHom (2 +ℕ n) (C* n f g))) → invMap (help x) ≡ x - ret2 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - λ h → cong ∣_∣₂ (funExt λ { (inl x) → cong h ((λ i → inr ((snd f) (~ i))) - ∙ sym (push (inl north))) - ; (inr x) → refl - ; (push a i₁) → help2 h a i₁}) - where - help2 : (h : C* n f g → coHomK (2 +ℕ n)) - → (a : _) → PathP (λ i → invMapPrim (h ∘ inr) (push a i) ≡ h (push a i)) - (cong h ((λ i → inr ((snd f) (~ i))) ∙ sym (push (inl north)))) - refl - help2 h = WedgeElim n (λ _ → isOfHLevelPathP (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n) _ _) _ _) - help4 - - where - help4 : PathP (λ i → invMapPrim (h ∘ inr) (push (inl north) i) ≡ h (push (inl north) i)) - (cong h ((λ i → inr ((snd f) (~ i))) ∙ sym (push (inl north)))) - refl - help4 i j = h (hcomp (λ k → λ { (i = i1) → inr (fst f north) - ; (j = i0) → inr (snd f (~ i)) - ; (j = i1) → push (inl north) (i ∨ ~ k)}) - (inr (snd f (~ i ∧ ~ j)))) - - groupIso1 : GroupIso ((coHomGr (2 +ℕ n) (C* n f g))) (coHomGr (2 +ℕ n) (S₊ (2 +ℕ n))) - Iso.fun (fst groupIso1) = help - Iso.inv (fst groupIso1) = invMap - Iso.rightInv (fst groupIso1) = ret1 - Iso.leftInv (fst groupIso1) = ret2 - snd groupIso1 = makeIsGroupHom - (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) - λ f g → refl) - - -C+ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Type _ -C+ n f g = Pushout (λ _ → tt) (·Π f g .fst) - -H⁴-C∨ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → GroupIso (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C+ n f g)) - ℤGroup -H⁴-C∨ n f g = compGroupIso - (transportCohomIso (cong (3 +ℕ_) (+-suc n n))) (Hopfβ-Iso n (·Π f g)) - -lol : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) (x : ℤ) - → Iso.inv (fst (H⁴-C∨ n f g)) x - ≡ subst (λ m → coHom m (C+ n f g)) (sym (cong (3 +ℕ_) (+-suc n n))) (Iso.inv (fst (Hopfβ-Iso n (·Π f g))) x) -lol n f g = λ _ → refl - -H²-C∨ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → GroupIso (coHomGr (2 +ℕ n) (C+ n f g)) - ℤGroup -H²-C∨ n f g = Hopfα-Iso n (·Π f g) - - -H⁴-C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → GroupIso (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) - (DirProd ℤGroup ℤGroup) -H⁴-C* n f g = compGroupIso (GroupEquiv→GroupIso (invGroupEquiv fstIso)) - (compGroupIso (transportCohomIso (cong (2 +ℕ_) (+-suc n n))) - (compGroupIso (Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) _) - (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ _) (Hⁿ-Sⁿ≅ℤ _)))) - where - module Ms = MV _ _ _ (λ _ → tt) (fst (add2 f g)) - fstIso : GroupEquiv (coHomGr ((suc (suc (n +ℕ suc n)))) (S₊∙ (3 +ℕ (n +ℕ n)) ⋁ S₊∙ (3 +ℕ (n +ℕ n)))) - (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) - fst (fst fstIso) = Ms.d (suc (suc (n +ℕ suc n))) .fst - snd (fst fstIso) = - SES→isEquiv - (GroupPath _ _ .fst (compGroupEquiv (invEquiv (isContr→≃Unit ((tt , tt) , (λ _ → refl))) , makeIsGroupHom λ _ _ → refl) - (GroupEquivDirProd - (GroupIso→GroupEquiv (invGroupIso (Hⁿ-Unit≅0 _))) - (GroupIso→GroupEquiv - (invGroupIso - (Hⁿ-Sᵐ≅0 _ _ λ p → ¬lem 0 ((λ _ → north) , refl) n n (cong suc (sym (+-suc n n)) ∙ p))))))) - (GroupPath _ _ .fst - (compGroupEquiv ((invEquiv (isContr→≃Unit ((tt , tt) , (λ _ → refl))) , makeIsGroupHom λ _ _ → refl)) - ((GroupEquivDirProd - (GroupIso→GroupEquiv (invGroupIso (Hⁿ-Unit≅0 _))) - (GroupIso→GroupEquiv - (invGroupIso (Hⁿ-Sᵐ≅0 _ _ λ p → ¬lem 0 ((λ _ → north) , refl) n (suc n) (cong (2 +ℕ_) (sym (+-suc n n)) ∙ p)))))))) - (Ms.Δ (suc (suc (n +ℕ suc n)))) - (Ms.d (suc (suc (n +ℕ suc n)))) - (Ms.i (suc (suc (suc (n +ℕ suc n))))) - (Ms.Ker-d⊂Im-Δ _) - (Ms.Ker-i⊂Im-d _) - snd fstIso = Ms.d (suc (suc (n +ℕ suc n))) .snd - -Hopfie : {n : ℕ} → ∥ S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n) ∥₂ → ℤ -Hopfie = sRec isSetℤ (HopfInvariant _) - -subber : (n : ℕ) (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → _ -subber n f = subst (λ x → coHom x (HopfInvariantPush n (fst f))) - (λ i → suc (suc (suc (+-suc n n i)))) - -module _ (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) where - - C+' = C+ n f g - - βₗ : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) - βₗ = Iso.inv (fst (H⁴-C* n f g)) (1 , 0) - - βᵣ : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) - βᵣ = Iso.inv (fst (H⁴-C* n f g)) (0 , 1) - - βᵣ'-fun : C* n f g → coHomK ((4 +ℕ (n +ℕ n))) - βᵣ'-fun (inl x) = 0ₖ _ - βᵣ'-fun (inr x) = 0ₖ _ - βᵣ'-fun (push (inl x) i₁) = 0ₖ _ - βᵣ'-fun (push (inr x) i₁) = Kn→ΩKn+1 _ ∣ x ∣ i₁ - βᵣ'-fun (push (push a i₂) i₁) = Kn→ΩKn+10ₖ _ (~ i₂) i₁ - - - βₗ'-fun : C* n f g → coHomK (4 +ℕ (n +ℕ n)) - βₗ'-fun (inl x) = 0ₖ _ - βₗ'-fun (inr x) = 0ₖ _ - βₗ'-fun (push (inl x) i₁) = Kn→ΩKn+1 _ ∣ x ∣ i₁ - βₗ'-fun (push (inr x) i₁) = 0ₖ _ - βₗ'-fun (push (push a i₂) i₁) = Kn→ΩKn+10ₖ _ i₂ i₁ - - βₗ''-fun : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) - βₗ''-fun = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) ∣ βₗ'-fun ∣₂ - - βᵣ''-fun : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) - βᵣ''-fun = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) ∣ βᵣ'-fun ∣₂ - - α∨ : coHom (2 +ℕ n) (C* n f g) - α∨ = Iso.inv (fst (H²-C* n f g)) 1 - - private - pp : (a : _) → 0ₖ (suc (suc n)) ≡ ∣ (f ∨→ g) a ∣ - pp = WedgeElim _ (λ _ → isOfHLevelTrunc (4 +ℕ n) _ _) - (cong ∣_∣ₕ (sym (snd f))) - - pp-inr : pp (inr north) ≡ cong ∣_∣ₕ (sym (snd g)) - pp-inr = (λ j → transport (λ i → 0ₖ (2 +ℕ n) ≡ ∣ compPath-filler' (snd f) (sym (snd g)) (~ j) i ∣ₕ) - λ i → ∣ snd f (~ i ∨ j) ∣ₕ) - ∙ λ j → transp (λ i → 0ₖ (2 +ℕ n) ≡ ∣ snd g (~ i ∧ ~ j) ∣ₕ) j λ i → ∣ snd g (~ i ∨ ~ j) ∣ₕ - - α∨'-fun : C* n f g → coHomK (2 +ℕ n) - α∨'-fun (inl x) = 0ₖ _ - α∨'-fun (inr x) = ∣ x ∣ - α∨'-fun (push a i₁) = pp a i₁ - - α∨' : coHom (2 +ℕ n) (C* n f g) - α∨' = ∣ α∨'-fun ∣₂ - - - β+ : coHom ((2 +ℕ n) +' (2 +ℕ n)) C+' - β+ = Iso.inv (fst (H⁴-C∨ n f g)) 1 - - q : C+' → C* n f g - q (inl x) = inl x - q (inr x) = inr x - q (push a i₁) = (push (θ a) ∙ λ i → inr (+≡ n f g a (~ i))) i₁ - - jₗ : HopfInvariantPush n (fst f) → C* n f g - jₗ (inl x) = inl x - jₗ (inr x) = inr x - jₗ (push a i₁) = push (inl a) i₁ - - jᵣ : HopfInvariantPush n (fst g) → C* n f g - jᵣ (inl x) = inl x - jᵣ (inr x) = inr x - jᵣ (push a i₁) = push (inr a) i₁ - -α-∨→1 : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → Iso.fun (fst (H²-C* n f g)) (α∨' n f g) ≡ 1 -α-∨→1 n f g = sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (Hⁿ-Sⁿ≅ℤ-nice-generator n)) - ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) - -α-∨-lem : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → α∨ n f g ≡ α∨' n f g -α-∨-lem n f g = sym (Iso.leftInv (fst (H²-C* n f g)) _) - ∙∙ cong (Iso.inv (fst (H²-C* n f g))) help - ∙∙ Iso.leftInv (fst (H²-C* n f g)) _ - where - help : Iso.fun (fst (H²-C* n f g)) (α∨ n f g) ≡ Iso.fun (fst (H²-C* n f g)) (α∨' n f g) - help = (Iso.rightInv (fst (H²-C* n f g)) (pos 1)) ∙ sym (α-∨→1 n f g) - -q-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → coHomFun _ (q n f g) (α∨ n f g) ≡ Hopfα n (·Π f g) -q-α n f g = (λ i → coHomFun _ (q n f g) (α-∨-lem n f g i)) - ∙ sym (Iso.leftInv is _) - ∙∙ cong (Iso.inv is) help - ∙∙ Iso.leftInv is _ - where - is = fst (Hopfα-Iso n (·Π f g)) - - help : Iso.fun is (coHomFun _ (q n f g) (α∨' n f g)) - ≡ Iso.fun is (Hopfα n (·Π f g)) - help = sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (Hⁿ-Sⁿ≅ℤ-nice-generator n)) - ∙∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) - ∙∙ sym (Hopfα-Iso-α n (·Π f g)) - - -βₗ≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → βₗ n f g ≡ βₗ''-fun n f g -βₗ≡ n f g = cong (FF ∘ subber2) - (cong (Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) - (S₊∙ (suc (suc (suc (n +ℕ n))))) - (((suc (suc (n +ℕ n))))))))) help - ∙ help2) - ∙ funExt⁻ FF∘subber ∣ wedgeGen ∣₂ - ∙ cong subber3 (sym βₗ'-fun≡) - where - FF = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (n +ℕ suc n))) .fst - FF' = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (suc (n +ℕ n)))) .fst - - help : Iso.inv (fst (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (1 , 0) - ≡ (∣ ∣_∣ ∣₂ , 0ₕ _) - help = ΣPathP ((Hⁿ-Sⁿ≅ℤ-nice-generator _) , IsGroupHom.pres1 (snd (invGroupIso (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) - - subber3 = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - - subber2 = (subst (λ m → coHom m (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) - (sym (cong (2 +ℕ_) (+-suc n n)))) - - FF∘subber : FF ∘ subber2 ≡ subber3 ∘ FF' - FF∘subber = - funExt (λ a → (λ j → transp (λ i → coHom ((suc ∘ suc ∘ suc) (+-suc n n (~ i ∧ j))) (C* n f g)) (~ j) - (MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) ((suc (suc (+-suc n n j)))) .fst - (transp (λ i → coHom (2 +ℕ (+-suc n n (j ∨ ~ i))) - (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) j - a)))) - - wedgeGen : (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) → - coHomK (suc (suc (suc (n +ℕ n))))) - wedgeGen (inl x) = ∣ x ∣ - wedgeGen (inr x) = 0ₖ _ - wedgeGen (push a i₁) = 0ₖ _ - - βₗ'-fun≡ : ∣ βₗ'-fun n f g ∣₂ ≡ FF' ∣ wedgeGen ∣₂ - βₗ'-fun≡ = cong ∣_∣₂ (funExt λ { (inl x) → refl - ; (inr x) → refl - ; (push (inl x) i₁) → refl - ; (push (inr x) i₁) j → Kn→ΩKn+10ₖ _ (~ j) i₁ - ; (push (push a i₂) i₁) j → Kn→ΩKn+10ₖ _ (~ j ∧ i₂) i₁}) - - help2 : Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) (((suc (suc (n +ℕ n)))))))) - (∣ ∣_∣ ∣₂ , 0ₕ _) - ≡ ∣ wedgeGen ∣₂ - help2 = cong ∣_∣₂ (funExt λ { (inl x) → rUnitₖ (suc (suc (suc (n +ℕ n)))) ∣ x ∣ - ; (inr x) → refl - ; (push a i₁) → refl}) - -βᵣ≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → βᵣ n f g ≡ βᵣ''-fun n f g -βᵣ≡ n f g = - cong (FF ∘ subber2) - (cong (Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) - (S₊∙ (suc (suc (suc (n +ℕ n))))) - (((suc (suc (n +ℕ n))))))))) - help - ∙ help2) - ∙ funExt⁻ FF∘subber ∣ wedgeGen ∣₂ - ∙ cong (subber3) (sym βᵣ'-fun≡) - where - FF = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (n +ℕ suc n))) .fst - FF' = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (suc (n +ℕ n)))) .fst - - help : Iso.inv (fst (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (0 , 1) - ≡ (0ₕ _ , ∣ ∣_∣ ∣₂) - help = ΣPathP (IsGroupHom.pres1 (snd (invGroupIso (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) , (Hⁿ-Sⁿ≅ℤ-nice-generator _)) - - subber3 = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - - subber2 = (subst (λ m → coHom m (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) - (sym (cong (2 +ℕ_) (+-suc n n)))) - - FF∘subber : FF ∘ subber2 ≡ subber3 ∘ FF' - FF∘subber = - funExt (λ a → (λ j → transp (λ i → coHom ((suc ∘ suc ∘ suc) (+-suc n n (~ i ∧ j))) (C* n f g)) (~ j) - (MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) ((suc (suc (+-suc n n j)))) .fst - (transp (λ i → coHom (2 +ℕ (+-suc n n (j ∨ ~ i))) - (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) j - a)))) - - wedgeGen : (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) → - coHomK (suc (suc (suc (n +ℕ n))))) - wedgeGen (inl x) = 0ₖ _ - wedgeGen (inr x) = ∣ x ∣ - wedgeGen (push a i₁) = 0ₖ _ - - - help2 : Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) (((suc (suc (n +ℕ n)))))))) - (0ₕ _ , ∣ ∣_∣ ∣₂) - ≡ ∣ wedgeGen ∣₂ - help2 = cong ∣_∣₂ (funExt λ { (inl x) → refl - ; (inr x) → lUnitₖ (suc (suc (suc (n +ℕ n)))) ∣ x ∣ - ; (push a i₁) → refl}) - - βᵣ'-fun≡ : ∣ βᵣ'-fun n f g ∣₂ ≡ FF' ∣ wedgeGen ∣₂ - βᵣ'-fun≡ = cong ∣_∣₂ (funExt λ { (inl x) → refl - ; (inr x) → refl - ; (push (inl x) i₁) j → Kn→ΩKn+10ₖ _ (~ j) i₁ - ; (push (inr x) i₁) → refl - ; (push (push a i₂) i₁) j → Kn→ΩKn+10ₖ _ (~ j ∧ ~ i₂) i₁}) - -q-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → coHomFun _ (q n f g) (βₗ n f g) - ≡ β+ n f g -q-βₗ n f g = cong (coHomFun _ (q n f g)) (βₗ≡ n f g) - ∙ transportLem - ∙ cong (subst (λ m → coHom m (C+' n f g)) - (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) - (sym (Iso.leftInv (fst (Hopfβ-Iso n (·Π f g))) (Hopfβ n (·Π f g))) - ∙ cong (Iso.inv ((fst (Hopfβ-Iso n (·Π f g))))) (Hopfβ↦1 n (·Π f g))) - where - transportLem : coHomFun (suc (suc (suc (n +ℕ suc n)))) (q n f g) - (βₗ''-fun n f g) - ≡ subst (λ m → coHom m (C+' n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - (Hopfβ n (·Π f g)) - transportLem = - natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (q n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - ∙ cong (subst (λ m → coHom m (C+' n f g)) - (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) - (cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → loll a i₁})) - where - open import Cubical.Foundations.Path - loll : (x : fst (S₊∙ (3 +ℕ (n +ℕ n)))) → - PathP - (λ i₁ → - βₗ'-fun n f g ((push (θ x) ∙ λ i → inr (+≡ n f g x (~ i))) i₁) ≡ - MV.d-pre Unit (fst (S₊∙ (2 +ℕ n))) (fst (S₊∙ (3 +ℕ n +ℕ n))) - (λ _ → tt) (fst (·Π f g)) (3 +ℕ n +ℕ n) ∣_∣ (push x i₁)) - (λ _ → βₗ'-fun n f g (q n f g (inl tt))) - (λ _ → βₗ'-fun n f g (q n f g (inr (·Π f g .fst x)))) - loll x = - flipSquare (cong-∙ (βₗ'-fun n f g) (push (θ x)) (λ i → inr (+≡ n f g x (~ i))) - ∙∙ sym (rUnit _) - ∙∙ lem₁ x) - - where - lem₁ : (x : _) → cong (βₗ'-fun n f g) (push (θ {A = S₊∙ _} x)) ≡ Kn→ΩKn+1 _ ∣ x ∣ - lem₁ north = refl - lem₁ south = sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)) - lem₁ (merid a j) k = lala k j - where - lala : PathP (λ k → Kn→ΩKn+1 _ ∣ north ∣ₕ ≡ (sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north))) k) - (λ j → cong (βₗ'-fun n f g) (push (θ {A = S₊∙ _} (merid a j)))) - (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) - lala = (cong-∙∙ (cong (βₗ'-fun n f g) ∘ push) - ((λ i₁ → inl ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁))) - (push tt) - ((λ i₁ → inr ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁))) - ∙ sym (compPath≡compPath' _ _) - ∙ (λ _ → (λ j → Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∣ (merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) j ∣ₕ) - ∙ Kn→ΩKn+10ₖ _) - ∙ cong (_∙ Kn→ΩKn+10ₖ _) (cong-∙ ((Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ)) - (merid a) (sym (merid north))) - ∙ sym (assoc _ _ _) - ∙ cong (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a) ∙_) - (sym (symDistr (sym ((Kn→ΩKn+10ₖ _))) (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)))))) - ◁ λ i j → compPath-filler (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) - (sym (sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) - (cong ∣_∣ₕ (merid north)))) - (~ i) j - -q-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → coHomFun _ (q n f g) (βᵣ n f g) - ≡ β+ n f g -q-βᵣ n f g = cong (coHomFun _ (q n f g)) (βᵣ≡ n f g) - ∙ transportLem - ∙ cong (subst (λ m → coHom m (C+' n f g)) - (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) - (sym (Iso.leftInv (fst (Hopfβ-Iso n (·Π f g))) (Hopfβ n (·Π f g))) - ∙ cong (Iso.inv ((fst (Hopfβ-Iso n (·Π f g))))) (Hopfβ↦1 n (·Π f g))) - where - transportLem : coHomFun (suc (suc (suc (n +ℕ suc n)))) (q n f g) - (βᵣ''-fun n f g) - ≡ subst (λ m → coHom m (C+' n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - (Hopfβ n (·Π f g)) - transportLem = - natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (q n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - ∙ cong (subst (λ m → coHom m (C+' n f g)) - (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) - (cong ∣_∣₂ (funExt λ { (inl x) → refl - ; (inr x) → refl - ; (push a i₁) → loll a i₁})) - where - open import Cubical.Foundations.Path - loll : (x : fst (S₊∙ (3 +ℕ (n +ℕ n)))) → - PathP - (λ i₁ → - βᵣ'-fun n f g ((push (θ x) ∙ λ i → inr (+≡ n f g x (~ i))) i₁) ≡ - MV.d-pre Unit (fst (S₊∙ (2 +ℕ n))) (fst (S₊∙ (3 +ℕ n +ℕ n))) - (λ _ → tt) (fst (·Π f g)) (3 +ℕ n +ℕ n) ∣_∣ (push x i₁)) - (λ _ → βᵣ'-fun n f g (q n f g (inl tt))) - (λ _ → βᵣ'-fun n f g (q n f g (inr (·Π f g .fst x)))) - loll x = flipSquare (cong-∙ (βᵣ'-fun n f g) (push (θ x)) (λ i → inr (+≡ n f g x (~ i))) - ∙∙ sym (rUnit _) - ∙∙ lem₁ x) - where - lem₁ : (x : _) → cong (βᵣ'-fun n f g) (push (θ {A = S₊∙ _} x)) ≡ Kn→ΩKn+1 _ ∣ x ∣ - lem₁ north = sym (Kn→ΩKn+10ₖ _) - lem₁ south = cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)) - lem₁ (merid a j) k = lala k j -- lala k j - where - lala : PathP (λ k → (Kn→ΩKn+10ₖ _) (~ k) ≡ (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north))) k) - (λ j → cong (βᵣ'-fun n f g) (push (θ {A = S₊∙ _} (merid a j)))) - (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) - lala = (cong-∙∙ (cong (βᵣ'-fun n f g) ∘ push) - (λ i₁ → inl ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁)) - (push tt) - (λ i₁ → inr ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁)) - ∙∙ (cong (sym (Kn→ΩKn+10ₖ _) ∙_) (cong-∙ (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a) (sym (merid (ptSn _))))) - ∙∙ sym (help3 _ _ _)) - ◁ symP (doubleCompPath-filler - (sym (Kn→ΩKn+10ₖ _)) - (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) - (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (sym (merid north)))) -open import Cubical.Foundations.Path -jₗ-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → coHomFun _ (jₗ n f g) (α∨ n f g) - ≡ Hopfα n f -jₗ-α n f g = - cong (coHomFun _ (jₗ n f g)) (α-∨-lem n f g) - ∙ cong ∣_∣₂ (funExt (HopfInvariantPushElim n (fst f) - (isOfHLevelPath ((suc (suc (suc (suc (n +ℕ n)))))) - (isOfHLevelPlus' {n = n} (4 +ℕ n) (isOfHLevelTrunc (4 +ℕ n))) _ _) - refl - (λ _ → refl) - λ j → refl)) - -jₗ-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → coHomFun _ (jₗ n f g) (βₗ n f g) - ≡ subst (λ m → coHom m (HopfInvariantPush n (fst f))) - (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - (Hopfβ n f) -jₗ-βₗ n f g = - cong (coHomFun _ (jₗ n f g)) (βₗ≡ n f g) - ∙∙ natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (jₗ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) - (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) - (cong ∣_∣₂ - (funExt λ { (inl x) → refl - ; (inr x) → refl - ; (push a i₁) → refl})) - -jₗ-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → coHomFun _ (jₗ n f g) (βᵣ n f g) - ≡ 0ₕ _ -jₗ-βᵣ n f g = - cong (coHomFun _ (jₗ n f g)) (βᵣ≡ n f g) - ∙∙ natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (jₗ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) - (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) - cool - where - cool : coHomFun (suc (suc (suc (suc (n +ℕ n))))) (jₗ n f g) - ∣ βᵣ'-fun n f g ∣₂ ≡ 0ₕ _ - cool = cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → refl}) - -jᵣ-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → coHomFun _ (jᵣ n f g) (α∨ n f g) - ≡ Hopfα n g -jᵣ-α n f g = cong (coHomFun _ (jᵣ n f g)) (α-∨-lem n f g) - ∙ cong ∣_∣₂ (funExt (HopfInvariantPushElim n (fst g) - (isOfHLevelPath ((suc (suc (suc (suc (n +ℕ n)))))) - (isOfHLevelPlus' {n = n} (4 +ℕ n) (isOfHLevelTrunc (4 +ℕ n))) _ _) - refl - (λ _ → refl) - (flipSquare (pp-inr n f g)))) - -jᵣ-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → coHomFun _ (jᵣ n f g) (βₗ n f g) ≡ 0ₕ _ -jᵣ-βₗ n f g = - cong (coHomFun _ (jᵣ n f g)) (βₗ≡ n f g) - ∙∙ natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (jᵣ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) - (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) - cool - where - cool : coHomFun (suc (suc (suc (suc (n +ℕ n))))) (jᵣ n f g) - ∣ βₗ'-fun n f g ∣₂ ≡ 0ₕ _ - cool = cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → refl}) - - -jᵣ-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → coHomFun _ (jᵣ n f g) (βᵣ n f g) - ≡ subst (λ m → coHom m (HopfInvariantPush n (fst g))) - (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - (Hopfβ n g) -jᵣ-βᵣ n f g = - cong (coHomFun _ (jᵣ n f g)) (βᵣ≡ n f g) - ∙∙ natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (jᵣ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst g))) - (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) - (cong ∣_∣₂ - (funExt λ { (inl x) → refl - ; (inr x) → refl - ; (push a i₁) → refl})) - -gen₂ℤ×ℤ : gen₂-by (DirProd ℤGroup ℤGroup) (1 , 0) (0 , 1) -fst (gen₂ℤ×ℤ (x , y)) = x , y -snd (gen₂ℤ×ℤ (x , y)) = - ΣPathP ((cong₂ _+_ ((·Comm 1 x) ∙ cong fst (sym (distrLem 1 0 x))) ((·Comm 0 y) ∙ cong fst (sym (distrLem 0 1 y)))) - , +Comm y 0 ∙ cong₂ _+_ (·Comm 0 x ∙ cong snd (sym (distrLem 1 0 x))) (·Comm 1 y ∙ cong snd (sym (distrLem 0 1 y)))) - where - ℤ×ℤ = DirProd ℤGroup ℤGroup - _+''_ = GroupStr._·_ (snd ℤ×ℤ) - - ll3 : (x : ℤ) → - x ≡ 0 - x - ll3 (pos zero) = refl - ll3 (pos (suc zero)) = refl - ll3 (pos (suc (suc n))) = - cong predℤ (ll3 (pos (suc n))) - ll3 (negsuc zero) = refl - ll3 (negsuc (suc n)) = cong sucℤ (ll3 (negsuc n)) - - distrLem : (x y : ℤ) (z : ℤ) - → Path (ℤ × ℤ) (z ℤ[ ℤ×ℤ ]· (x , y)) (z · x , z · y) - distrLem x y (pos zero) = refl - distrLem x y (pos (suc n)) = - (cong ((x , y) +''_) (distrLem x y (pos n))) - distrLem x y (negsuc zero) = ΣPathP (sym (ll3 x) , sym (ll3 y)) - distrLem x y (negsuc (suc n)) = - cong₂ _+''_ (ΣPathP (sym (ll3 x) , sym (ll3 y))) - (distrLem x y (negsuc n)) - -genH²ⁿC* : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → gen₂-by (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) - (βₗ n f g) - (βᵣ n f g) -genH²ⁿC* n f g = - Iso-pres-gen₂ (DirProd ℤGroup ℤGroup) - (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) - (1 , 0) - (0 , 1) - gen₂ℤ×ℤ - (invGroupIso (H⁴-C* n f g)) - -private - - H⁴C* : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Group _ - H⁴C* n f g = coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) - - X : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → ℤ - X n f g = (genH²ⁿC* n f g) (α∨ n f g ⌣ α∨ n f g) .fst .fst - Y : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → ℤ - Y n f g = (genH²ⁿC* n f g) (α∨ n f g ⌣ α∨ n f g) .fst .snd - - α∨≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → α∨ n f g ⌣ α∨ n f g ≡ ((X n f g) ℤ[ H⁴C* n f g ]· βₗ n f g) - +ₕ ((Y n f g) ℤ[ H⁴C* n f g ]· βᵣ n f g) - α∨≡ n f g = (genH²ⁿC* n f g) (α∨ n f g ⌣ α∨ n f g) .snd - - -eq₁ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → (Hopfα n (·Π f g)) ⌣ (Hopfα n (·Π f g)) - ≡ ((X n f g + Y n f g) ℤ[ coHomGr _ _ ]· (β+ n f g)) -eq₁ n f g = - cong (λ x → x ⌣ x) (sym (q-α n f g) ∙ cong (coHomFun (suc (suc n)) (q n f g)) (α-∨-lem n f g)) - ∙ cong ((coHomFun _) (q _ f g)) (cong (λ x → x ⌣ x) (sym (α-∨-lem n f g)) - ∙ α∨≡ n f g) - ∙ IsGroupHom.pres· (coHomMorph _ (q n f g) .snd) _ _ - ∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (q n f g)) (βₗ n f g) (X n f g) - ∙ cong (λ z → (X n f g) ℤ[ coHomGr _ _ ]· z) - (q-βₗ n f g)) - (homPresℤ· (coHomMorph _ (q n f g)) (βᵣ n f g) (Y n f g) - ∙ cong (λ z → (Y n f g) ℤ[ coHomGr _ _ ]· z) - (q-βᵣ n f g)) - ∙ sym (distrℤ· (coHomGr _ _) (β+ n f g) (X n f g) (Y n f g)) - -coHomFun⌣ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) - → (n : ℕ) (x y : coHom n B) - → coHomFun _ f (x ⌣ y) ≡ coHomFun _ f x ⌣ coHomFun _ f y -coHomFun⌣ f n = sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl - -eq₂ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → Hopfα n f ⌣ Hopfα n f - ≡ (X n f g ℤ[ coHomGr _ _ ]· subst (λ m → coHom m (HopfInvariantPush n (fst f))) - (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - (Hopfβ n f)) -eq₂ n f g = - cong (λ x → x ⌣ x) (sym (jₗ-α n f g)) - ∙∙ sym (coHomFun⌣ (jₗ n f g) _ _ _) - ∙∙ cong (coHomFun _ (jₗ n f g)) (α∨≡ n f g) - ∙∙ IsGroupHom.pres· (snd (coHomMorph _ (jₗ n f g))) _ _ - ∙∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (jₗ n f g)) (βₗ n f g) (X n f g)) - (homPresℤ· (coHomMorph _ (jₗ n f g)) (βᵣ n f g) (Y n f g) - ∙ cong (λ z → (Y n f g) ℤ[ coHomGr _ _ ]· z) - (jₗ-βᵣ n f g)) - ∙∙ cong₂ _+ₕ_ refl (rUnitℤ· (coHomGr _ _) (Y n f g)) - ∙∙ (rUnitₕ _ _ - ∙ cong (X n f g ℤ[ coHomGr _ _ ]·_) (jₗ-βₗ n f g)) - -eq₃ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → Hopfα n g ⌣ Hopfα n g - ≡ (Y n f g ℤ[ coHomGr _ _ ]· subst (λ m → coHom m (HopfInvariantPush n (fst f))) - (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - (Hopfβ n g)) -eq₃ n f g = - cong (λ x → x ⌣ x) (sym (jᵣ-α n f g)) - ∙∙ sym (coHomFun⌣ (jᵣ n f g) _ _ _) - ∙∙ cong (coHomFun _ (jᵣ n f g)) (α∨≡ n f g) - ∙∙ IsGroupHom.pres· (snd (coHomMorph _ (jᵣ n f g))) _ _ - ∙∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (jᵣ n f g)) (βₗ n f g) (X n f g) - ∙ cong (λ z → (X n f g) ℤ[ coHomGr _ _ ]· z) - (jᵣ-βₗ n f g)) - (homPresℤ· (coHomMorph _ (jᵣ n f g)) (βᵣ n f g) (Y n f g)) - ∙∙ cong₂ _+ₕ_ (rUnitℤ· (coHomGr _ _) (X n f g)) refl - ∙∙ ((lUnitₕ _ _) ∙ cong (Y n f g ℤ[ coHomGr _ _ ]·_) (jᵣ-βᵣ n f g)) - -eq₂-eq₂ : (n : ℕ) (f g : S₊∙ (suc (suc (suc (n +ℕ n)))) →∙ S₊∙ (suc (suc n))) - → HopfInvariant n (·Π f g) ≡ HopfInvariant n f + HopfInvariant n g -eq₂-eq₂ n f g = - mainL - ∙ sym (cong₂ _+_ (help n f g) (helpg n f g)) - where - transpLem : ∀ {ℓ} {G : ℕ → Group ℓ} - → (n m : ℕ) - → (x : ℤ) - → (p : m ≡ n) - → (g : fst (G n)) - → subst (fst ∘ G) p (x ℤ[ G m ]· subst (fst ∘ G) (sym p) g) ≡ (x ℤ[ G n ]· g) - transpLem {G = G} n m x = - J (λ n p → (g : fst (G n)) → subst (fst ∘ G) p (x ℤ[ G m ]· subst (fst ∘ G) (sym p) g) - ≡ (x ℤ[ G n ]· g)) - λ g → transportRefl _ ∙ cong (x ℤ[ G m ]·_) (transportRefl _) - - mainL : HopfInvariant n (·Π f g) ≡ X n f g + Y n f g - mainL = cong (Iso.fun (fst (Hopfβ-Iso n (·Π f g)))) - (cong (subst (λ x → coHom x (HopfInvariantPush n (fst (·Π f g)))) - λ i₁ → suc (suc (suc (+-suc n n i₁)))) - (eq₁ n f g)) - ∙∙ cong (Iso.fun (fst (Hopfβ-Iso n (·Π f g)))) - (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst (·Π f g)))} _ _ - (X n f g + Y n f g) (λ i₁ → suc (suc (suc (+-suc n n i₁)))) - (Iso.inv (fst (Hopfβ-Iso n (·Π f g))) 1)) - ∙∙ homPresℤ· (_ , snd (Hopfβ-Iso n (·Π f g))) (Iso.inv (fst (Hopfβ-Iso n (·Π f g))) 1) - (X n f g + Y n f g) - ∙∙ cong ((X n f g + Y n f g) ℤ[ ℤ , ℤGroup .snd ]·_) - (Iso.rightInv ((fst (Hopfβ-Iso n (·Π f g)))) 1) - ∙∙ rUnitℤ·ℤ (X n f g + Y n f g) - - - help : (n : ℕ) (f g : _) → HopfInvariant n f ≡ X n f g - help n f g = cong (Iso.fun (fst (Hopfβ-Iso n f))) - (cong (subst (λ x → coHom x (HopfInvariantPush n (fst f))) - (λ i₁ → suc (suc (suc (+-suc n n i₁))))) (eq₂ n f g)) - ∙ cong (Iso.fun (fst (Hopfβ-Iso n f))) - (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst f))} _ _ - (X n f g) - ((cong (suc ∘ suc ∘ suc) (+-suc n n))) - (Hopfβ n f)) - ∙ homPresℤ· (_ , snd (Hopfβ-Iso n f)) (Hopfβ n f) (X n f g) - ∙ cong (X n f g ℤ[ ℤ , ℤGroup .snd ]·_) (Hopfβ↦1 n f) - ∙ rUnitℤ·ℤ (X n f g) - - helpg : (n : ℕ) (f g : _) → HopfInvariant n g ≡ Y n f g - helpg n f g = - cong (Iso.fun (fst (Hopfβ-Iso n g))) - (cong (subst (λ x → coHom x (HopfInvariantPush n (fst g))) - (λ i₁ → suc (suc (suc (+-suc n n i₁))))) - (eq₃ n f g)) - ∙∙ cong (Iso.fun (fst (Hopfβ-Iso n g))) - (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst g))} _ _ - (Y n f g) - ((cong (suc ∘ suc ∘ suc) (+-suc n n))) - (Hopfβ n g)) - ∙∙ homPresℤ· (_ , snd (Hopfβ-Iso n g)) (Hopfβ n g) (Y n f g) - ∙∙ cong (Y n f g ℤ[ ℤ , ℤGroup .snd ]·_) (Hopfβ↦1 n g) - ∙∙ rUnitℤ·ℤ (Y n f g) From 806bc8f5ad90fb6b34e35542098b42366d294ded Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Sun, 10 Oct 2021 04:20:53 +0200 Subject: [PATCH 71/89] wow --- Cubical/HITs/Wedge/Base.agda | 11 +- Cubical/Homotopy/HopfInvariant/Base.agda | 6 + .../Homotopy/HopfInvariant/Homomorphism.agda | 842 ++++++++++++++++++ Cubical/Homotopy/HopfInvariant/HopfMap.agda | 374 ++++---- 4 files changed, 1010 insertions(+), 223 deletions(-) create mode 100644 Cubical/Homotopy/HopfInvariant/Homomorphism.agda diff --git a/Cubical/HITs/Wedge/Base.agda b/Cubical/HITs/Wedge/Base.agda index 7fb25a501..4bda0d609 100644 --- a/Cubical/HITs/Wedge/Base.agda +++ b/Cubical/HITs/Wedge/Base.agda @@ -10,14 +10,21 @@ _⋁_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Type (ℓ-max ℓ _⋁_ (A , ptA) (B , ptB) = Pushout {A = Unit} {B = A} {C = B} (λ _ → ptA) (λ _ → ptB) --- pointed versions - +-- Pointed versions _⋁∙ₗ_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Pointed (ℓ-max ℓ ℓ') A ⋁∙ₗ B = (A ⋁ B) , (inl (snd A)) _⋁∙ᵣ_ : ∀ {ℓ ℓ'} → Pointed ℓ → Pointed ℓ' → Pointed (ℓ-max ℓ ℓ') A ⋁∙ᵣ B = (A ⋁ B) , (inr (snd B)) +-- Wedge sums of functions +_∨→_ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + → (f : A →∙ C) (g : B →∙ C) + → A ⋁ B → fst C +(f ∨→ g) (inl x) = fst f x +(f ∨→ g) (inr x) = fst g x +(f ∨→ g) (push a i₁) = (snd f ∙ sym (snd g)) i₁ + -- Wedge sum of Units is contractible isContr-Unit⋁Unit : isContr ((Unit , tt) ⋁ (Unit , tt)) fst isContr-Unit⋁Unit = inl tt diff --git a/Cubical/Homotopy/HopfInvariant/Base.agda b/Cubical/Homotopy/HopfInvariant/Base.agda index cdfb55efe..df81c6afd 100644 --- a/Cubical/Homotopy/HopfInvariant/Base.agda +++ b/Cubical/Homotopy/HopfInvariant/Base.agda @@ -332,3 +332,9 @@ HopfInvariant : (n : ℕ) HopfInvariant n f = Iso.fun (fst (Hopfβ-Iso n f)) ((subst (λ x → coHom x (HopfInvariantPush n (fst f))) (λ i → suc (suc (suc (+-suc n n i))))) (⌣-α n f)) + +HopfInvariant-π' : (n : ℕ) → π' (3 +ℕ (n +ℕ n)) (S₊∙ (2 +ℕ n)) → ℤ +HopfInvariant-π' n = sRec isSetℤ (HopfInvariant n) + +HopfInvariant-π : (n : ℕ) → π (3 +ℕ (n +ℕ n)) (S₊∙ (2 +ℕ n)) → ℤ +HopfInvariant-π n = sRec isSetℤ λ x → HopfInvariant-π' n ∣ Ω→SphereMap _ x ∣₂ diff --git a/Cubical/Homotopy/HopfInvariant/Homomorphism.agda b/Cubical/Homotopy/HopfInvariant/Homomorphism.agda new file mode 100644 index 000000000..6a5255fd1 --- /dev/null +++ b/Cubical/Homotopy/HopfInvariant/Homomorphism.agda @@ -0,0 +1,842 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Homotopy.HopfInvariant.Homomorphism where + +open import Cubical.Homotopy.HopfInvariant.HopfMap +open import Cubical.Homotopy.HopfInvariant.Base +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.Groups.Connected +open import Cubical.ZCohomology.GroupStructure +open import Cubical.ZCohomology.Properties +open import Cubical.ZCohomology.MayerVietorisUnreduced +open import Cubical.ZCohomology.Groups.Unit +open import Cubical.ZCohomology.Groups.Wedge +open import Cubical.ZCohomology.Groups.Sn +open import Cubical.ZCohomology.RingStructure.CupProduct +open import Cubical.ZCohomology.RingStructure.RingLaws +open import Cubical.ZCohomology.RingStructure.GradedCommutativity +open import Cubical.Relation.Nullary +open import Cubical.ZCohomology.Gysin + +open import Cubical.Functions.Embedding + +open import Cubical.Foundations.Path +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Pointed.Homogeneous + +open import Cubical.Homotopy.HomotopyGroup + +open import Cubical.Foundations.Univalence + +open import Cubical.Relation.Nullary + +open import Cubical.Data.Sum +open import Cubical.Data.Fin +open import Cubical.Data.Empty renaming (rec to ⊥-rec) +open import Cubical.Data.Sigma +open import Cubical.Data.Int hiding (_+'_) +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_ ; _·_ to _·ℕ_) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Unit +open import Cubical.Data.Bool +open import Cubical.Algebra.Group + renaming (ℤ to ℤGroup ; Unit to UnitGroup) hiding (Bool) +open import Cubical.Algebra.Group.ZModule + +open import Cubical.HITs.Pushout.Flattening +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.EilenbergSteenrod +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 renaming (_·_ to _*_) +open import Cubical.HITs.Truncation + renaming (rec to trRec ; elim to trElim ; elim2 to trElim2) +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap ; elim3 to sElim3) +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; elim to pElim) + +open import Cubical.Algebra.AbGroup + +open import Cubical.Homotopy.Loopspace + +open import Cubical.HITs.Join + +open import Cubical.Homotopy.Hopf + +open import Cubical.HITs.SetQuotients renaming (_/_ to _/'_) + + +open import Cubical.Experiments.Brunerie +open import Cubical.HITs.Hopf + +open import Cubical.HITs.Join + +open import Cubical.HITs.Wedge + +-- probably to be moved to Wedge + +∨→∙ : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f g : S₊∙ n →∙ A) → ((S₊∙ n) ⋁ (S₊∙ n) , inl (ptSn _)) →∙ A +fst (∨→∙ {A = A} f g) = f ∨→ g +snd (∨→∙ {A = A} f g) = snd f + +C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Type _ +C* n f g = Pushout (λ _ → tt) (fst (∨→∙ f g)) + +C·Π : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Type _ +C·Π n f g = Pushout (λ _ → tt) (∙Π f g .fst) + +θ : ∀ {ℓ} {A : Pointed ℓ} → Susp (fst A) + → (Susp (fst A) , north) ⋁ (Susp (fst A) , north) +θ north = inl north +θ south = inr north +θ {A = A} (merid a i₁) = + ((λ i → inl ((merid a ∙ sym (merid (pt A))) i)) + ∙∙ push tt + ∙∙ λ i → inr ((merid a ∙ sym (merid (pt A))) i)) i₁ + +-- move to path + +doubleCompPath≡compPath : ∀ {ℓ} {A : Type ℓ} {x y z w : A} (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) + → p ∙∙ q ∙∙ r ≡ p ∙ q ∙ r +doubleCompPath≡compPath {x = x} {y = y} {z = z} {w = w} = + J (λ y p → (q : y ≡ z) (r : z ≡ w) → + (p ∙∙ q ∙∙ r) ≡ p ∙ q ∙ r) + λ q r → lUnit (q ∙ r) + + +∙Π≡∨→∘θ : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → (x : _) → ∙Π f g .fst x ≡ (f ∨→ g) (θ {A = S₊∙ _} x) +∙Π≡∨→∘θ n f g north = sym (snd f) +∙Π≡∨→∘θ n f g south = sym (snd g) +∙Π≡∨→∘θ n f g (merid a i₁) j = main j i₁ + where + help : cong (f ∨→ g) (cong (θ {A = S₊∙ _}) (merid a)) + ≡ (refl ∙∙ cong (fst f) (merid a ∙ sym (merid north)) ∙∙ snd f) + ∙ (sym (snd g) ∙∙ cong (fst g) (merid a ∙ sym (merid north)) ∙∙ refl) + help = cong-∙∙ (f ∨→ g) ((λ i → inl ((merid a ∙ sym (merid north)) i))) + (push tt) + (λ i → inr ((merid a ∙ sym (merid north)) i)) + ∙∙ doubleCompPath≡compPath _ _ _ + ∙∙ cong (cong (f ∨→ g) + (λ i₂ → inl ((merid a ∙ (λ i₃ → merid north (~ i₃))) i₂)) ∙_) + (sym (assoc _ _ _)) + ∙∙ assoc _ _ _ + ∙∙ cong₂ _∙_ refl (compPath≡compPath' _ _) + + main : PathP (λ i → snd f (~ i) ≡ snd g (~ i)) (λ i₁ → ∙Π f g .fst (merid a i₁)) + λ i₁ → (f ∨→ g) (θ {A = S₊∙ _} (merid a i₁)) + main = (λ i → ((λ j → snd f (~ i ∧ ~ j)) ∙∙ cong (fst f) (merid a ∙ sym (merid north)) ∙∙ snd f) + ∙ (sym (snd g) ∙∙ cong (fst g) (merid a ∙ sym (merid north)) ∙∙ λ j → snd g (~ i ∧ j))) + ▷ sym help + +private + WedgeElim : ∀ {ℓ} (n : ℕ) + {P : S₊∙ (3 +ℕ (n +ℕ n)) ⋁ S₊∙ (3 +ℕ (n +ℕ n)) → Type ℓ} + → ((x : _) → isOfHLevel (3 +ℕ n) (P x)) + → P (inl north) + → (x : _) → P x + WedgeElim n {P = P} x s (inl x₁) = + sphereElim _ {A = P ∘ inl} + (λ _ → isOfHLevelPlus' {n = n} (3 +ℕ n) (x _)) s x₁ + WedgeElim n {P = P} x s (inr x₁) = + sphereElim _ {A = P ∘ inr} + (sphereElim _ (λ _ → isProp→isOfHLevelSuc ((suc (suc (n +ℕ n)))) + (isPropIsOfHLevel (suc (suc (suc (n +ℕ n)))))) + (subst (isOfHLevel ((3 +ℕ n) +ℕ n)) (cong P (push tt)) + (isOfHLevelPlus' {n = n} (3 +ℕ n) (x _)))) + (subst P (push tt) s) x₁ + WedgeElim n {P = P} x s (push a j) = transp (λ i → P (push tt (i ∧ j))) (~ j) s + +H²C*≅ℤ : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → GroupIso (coHomGr (2 +ℕ n) (C* n f g)) ℤGroup +H²C*≅ℤ n f g = compGroupIso groupIso1 (Hⁿ-Sⁿ≅ℤ (suc n)) + where + help : (coHom (2 +ℕ n) (C* n f g)) → coHom (2 +ℕ n) (S₊ (2 +ℕ n)) + help = sMap λ f → f ∘ inr + + + invMapPrim : (S₊ (2 +ℕ n) → coHomK (2 +ℕ n)) → C* n f g → coHomK (2 +ℕ n) + invMapPrim h (inl x) = h (ptSn _) + invMapPrim h (inr x) = h x + invMapPrim h (push a i₁) = WedgeElim n {P = λ a → h north ≡ h (fst (∨→∙ f g) a)} + (λ _ → isOfHLevelTrunc (4 +ℕ n) _ _) + (cong h (sym (snd f))) a i₁ + + invMap : coHom (2 +ℕ n) (S₊ (2 +ℕ n)) → (coHom (2 +ℕ n) (C* n f g)) + invMap = sMap invMapPrim + + ret1 : (x : coHom (2 +ℕ n) (S₊ (2 +ℕ n))) → help (invMap x) ≡ x + ret1 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ a → refl + + ret2 : (x : (coHom (2 +ℕ n) (C* n f g))) → invMap (help x) ≡ x + ret2 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ h → cong ∣_∣₂ (funExt λ { (inl x) → cong h ((λ i → inr ((snd f) (~ i))) + ∙ sym (push (inl north))) + ; (inr x) → refl + ; (push a i₁) → help2 h a i₁}) + where + help2 : (h : C* n f g → coHomK (2 +ℕ n)) + → (a : _) → PathP (λ i → invMapPrim (h ∘ inr) (push a i) ≡ h (push a i)) + (cong h ((λ i → inr ((snd f) (~ i))) ∙ sym (push (inl north)))) + refl + help2 h = WedgeElim n (λ _ → isOfHLevelPathP (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n) _ _) _ _) + help4 + + where + help4 : PathP (λ i → invMapPrim (h ∘ inr) (push (inl north) i) ≡ h (push (inl north) i)) + (cong h ((λ i → inr ((snd f) (~ i))) ∙ sym (push (inl north)))) + refl + help4 i j = h (hcomp (λ k → λ { (i = i1) → inr (fst f north) + ; (j = i0) → inr (snd f (~ i)) + ; (j = i1) → push (inl north) (i ∨ ~ k)}) + (inr (snd f (~ i ∧ ~ j)))) + + groupIso1 : GroupIso ((coHomGr (2 +ℕ n) (C* n f g))) (coHomGr (2 +ℕ n) (S₊ (2 +ℕ n))) + Iso.fun (fst groupIso1) = help + Iso.inv (fst groupIso1) = invMap + Iso.rightInv (fst groupIso1) = ret1 + Iso.leftInv (fst groupIso1) = ret2 + snd groupIso1 = makeIsGroupHom + (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f g → refl) + +H⁴-C·Π : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → GroupIso (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C·Π n f g)) + ℤGroup +H⁴-C·Π n f g = compGroupIso + (transportCohomIso (cong (3 +ℕ_) (+-suc n n))) (Hopfβ-Iso n (∙Π f g)) + +H²-C·Π : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → GroupIso (coHomGr (2 +ℕ n) (C·Π n f g)) + ℤGroup +H²-C·Π n f g = Hopfα-Iso n (∙Π f g) + + +H⁴-C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → GroupIso (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) + (DirProd ℤGroup ℤGroup) +H⁴-C* n f g = + compGroupIso (GroupEquiv→GroupIso (invGroupEquiv fstIso)) + (compGroupIso (transportCohomIso (cong (2 +ℕ_) (+-suc n n))) + (compGroupIso (Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) _) + (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ _) (Hⁿ-Sⁿ≅ℤ _)))) + where + module Ms = MV _ _ _ (λ _ → tt) (fst (∨→∙ f g)) + fstIso : GroupEquiv (coHomGr ((suc (suc (n +ℕ suc n)))) (S₊∙ (3 +ℕ (n +ℕ n)) ⋁ S₊∙ (3 +ℕ (n +ℕ n)))) + (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) + fst (fst fstIso) = Ms.d (suc (suc (n +ℕ suc n))) .fst + snd (fst fstIso) = + SES→isEquiv + (GroupPath _ _ .fst (compGroupEquiv (invEquiv (isContr→≃Unit ((tt , tt) , (λ _ → refl))) , makeIsGroupHom λ _ _ → refl) + (GroupEquivDirProd + (GroupIso→GroupEquiv (invGroupIso (Hⁿ-Unit≅0 _))) + (GroupIso→GroupEquiv + (invGroupIso + (Hⁿ-Sᵐ≅0 _ _ λ p → ¬lemHopf 0 ((λ _ → north) , refl) n n (cong suc (sym (+-suc n n)) ∙ p))))))) + (GroupPath _ _ .fst + (compGroupEquiv ((invEquiv (isContr→≃Unit ((tt , tt) , (λ _ → refl))) , makeIsGroupHom λ _ _ → refl)) + ((GroupEquivDirProd + (GroupIso→GroupEquiv (invGroupIso (Hⁿ-Unit≅0 _))) + (GroupIso→GroupEquiv + (invGroupIso (Hⁿ-Sᵐ≅0 _ _ λ p → ¬lemHopf 0 ((λ _ → north) , refl) n (suc n) (cong (2 +ℕ_) (sym (+-suc n n)) ∙ p)))))))) + (Ms.Δ (suc (suc (n +ℕ suc n)))) + (Ms.d (suc (suc (n +ℕ suc n)))) + (Ms.i (suc (suc (suc (n +ℕ suc n))))) + (Ms.Ker-d⊂Im-Δ _) + (Ms.Ker-i⊂Im-d _) + snd fstIso = Ms.d (suc (suc (n +ℕ suc n))) .snd + +module _ (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) where + + C·Π' = C·Π n f g + + -- The generators of the cohomology groups of C* and C·Π + + βₗ : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) + βₗ = Iso.inv (fst (H⁴-C* n f g)) (1 , 0) + + βᵣ : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) + βᵣ = Iso.inv (fst (H⁴-C* n f g)) (0 , 1) + + β·Π : coHom ((2 +ℕ n) +' (2 +ℕ n)) C·Π' + β·Π = Iso.inv (fst (H⁴-C·Π n f g)) 1 + + α* : coHom (2 +ℕ n) (C* n f g) + α* = Iso.inv (fst (H²C*≅ℤ n f g)) 1 + + -- They can be difficult to work with in this form. + -- We give them simpler forms and prove that these are equivalent + βᵣ'-fun : C* n f g → coHomK ((4 +ℕ (n +ℕ n))) + βᵣ'-fun (inl x) = 0ₖ _ + βᵣ'-fun (inr x) = 0ₖ _ + βᵣ'-fun (push (inl x) i₁) = 0ₖ _ + βᵣ'-fun (push (inr x) i₁) = Kn→ΩKn+1 _ ∣ x ∣ i₁ + βᵣ'-fun (push (push a i₂) i₁) = Kn→ΩKn+10ₖ _ (~ i₂) i₁ + + βₗ'-fun : C* n f g → coHomK (4 +ℕ (n +ℕ n)) + βₗ'-fun (inl x) = 0ₖ _ + βₗ'-fun (inr x) = 0ₖ _ + βₗ'-fun (push (inl x) i₁) = Kn→ΩKn+1 _ ∣ x ∣ i₁ + βₗ'-fun (push (inr x) i₁) = 0ₖ _ + βₗ'-fun (push (push a i₂) i₁) = Kn→ΩKn+10ₖ _ i₂ i₁ + + βₗ''-fun : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) + βₗ''-fun = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∣ βₗ'-fun ∣₂ + + βᵣ''-fun : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) + βᵣ''-fun = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∣ βᵣ'-fun ∣₂ + + private + pp : (a : _) → 0ₖ (suc (suc n)) ≡ ∣ (f ∨→ g) a ∣ + pp = WedgeElim _ (λ _ → isOfHLevelTrunc (4 +ℕ n) _ _) + (cong ∣_∣ₕ (sym (snd f))) + + pp-inr : pp (inr north) ≡ cong ∣_∣ₕ (sym (snd g)) + pp-inr = (λ j → transport (λ i → 0ₖ (2 +ℕ n) ≡ ∣ compPath-filler' (snd f) (sym (snd g)) (~ j) i ∣ₕ) + λ i → ∣ snd f (~ i ∨ j) ∣ₕ) + ∙ λ j → transp (λ i → 0ₖ (2 +ℕ n) ≡ ∣ snd g (~ i ∧ ~ j) ∣ₕ) j λ i → ∣ snd g (~ i ∨ ~ j) ∣ₕ + + α*'-fun : C* n f g → coHomK (2 +ℕ n) + α*'-fun (inl x) = 0ₖ _ + α*'-fun (inr x) = ∣ x ∣ + α*'-fun (push a i₁) = pp a i₁ + + α*' : coHom (2 +ℕ n) (C* n f g) + α*' = ∣ α*'-fun ∣₂ + + -- We also need the following three maps + jₗ : HopfInvariantPush n (fst f) → C* n f g + jₗ (inl x) = inl x + jₗ (inr x) = inr x + jₗ (push a i₁) = push (inl a) i₁ + + jᵣ : HopfInvariantPush n (fst g) → C* n f g + jᵣ (inl x) = inl x + jᵣ (inr x) = inr x + jᵣ (push a i₁) = push (inr a) i₁ + + q : C·Π' → C* n f g + q (inl x) = inl x + q (inr x) = inr x + q (push a i₁) = (push (θ a) ∙ λ i → inr (∙Π≡∨→∘θ n f g a (~ i))) i₁ + +α*↦1 : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → Iso.fun (fst (H²C*≅ℤ n f g)) (α*' n f g) ≡ 1 +α*↦1 n f g = sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (Hⁿ-Sⁿ≅ℤ-nice-generator n)) + ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) + +α*-lem : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → α* n f g ≡ α*' n f g +α*-lem n f g = sym (Iso.leftInv (fst (H²C*≅ℤ n f g)) _) + ∙∙ cong (Iso.inv (fst (H²C*≅ℤ n f g))) help + ∙∙ Iso.leftInv (fst (H²C*≅ℤ n f g)) _ + where + help : Iso.fun (fst (H²C*≅ℤ n f g)) (α* n f g) ≡ Iso.fun (fst (H²C*≅ℤ n f g)) (α*' n f g) + help = (Iso.rightInv (fst (H²C*≅ℤ n f g)) (pos 1)) ∙ sym (α*↦1 n f g) + +q-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (q n f g) (α* n f g) ≡ Hopfα n (∙Π f g) +q-α n f g = (λ i → coHomFun _ (q n f g) (α*-lem n f g i)) + ∙ sym (Iso.leftInv is _) + ∙∙ cong (Iso.inv is) help + ∙∙ Iso.leftInv is _ + where + is = fst (Hopfα-Iso n (∙Π f g)) + + help : Iso.fun is (coHomFun _ (q n f g) (α*' n f g)) + ≡ Iso.fun is (Hopfα n (∙Π f g)) + help = sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (Hⁿ-Sⁿ≅ℤ-nice-generator n)) + ∙∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) + ∙∙ sym (Hopfα-Iso-α n (∙Π f g)) + + +βₗ≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → βₗ n f g ≡ βₗ''-fun n f g +βₗ≡ n f g = cong (FF ∘ subber2) + (cong (Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) + (S₊∙ (suc (suc (suc (n +ℕ n))))) + (((suc (suc (n +ℕ n))))))))) help + ∙ help2) + ∙ funExt⁻ FF∘subber ∣ wedgeGen ∣₂ + ∙ cong subber3 (sym βₗ'-fun≡) + where + FF = MV.d _ _ _ (λ _ → tt) (fst (∨→∙ f g)) (suc (suc (n +ℕ suc n))) .fst + FF' = MV.d _ _ _ (λ _ → tt) (fst (∨→∙ f g)) (suc (suc (suc (n +ℕ n)))) .fst + + help : Iso.inv (fst (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (1 , 0) + ≡ (∣ ∣_∣ ∣₂ , 0ₕ _) + help = ΣPathP ((Hⁿ-Sⁿ≅ℤ-nice-generator _) , IsGroupHom.pres1 (snd (invGroupIso (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) + + subber3 = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + + subber2 = (subst (λ m → coHom m (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) + (sym (cong (2 +ℕ_) (+-suc n n)))) + + FF∘subber : FF ∘ subber2 ≡ subber3 ∘ FF' + FF∘subber = + funExt (λ a → (λ j → transp (λ i → coHom ((suc ∘ suc ∘ suc) (+-suc n n (~ i ∧ j))) (C* n f g)) (~ j) + (MV.d _ _ _ (λ _ → tt) (fst (∨→∙ f g)) ((suc (suc (+-suc n n j)))) .fst + (transp (λ i → coHom (2 +ℕ (+-suc n n (j ∨ ~ i))) + (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) j + a)))) + + wedgeGen : (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) → + coHomK (suc (suc (suc (n +ℕ n))))) + wedgeGen (inl x) = ∣ x ∣ + wedgeGen (inr x) = 0ₖ _ + wedgeGen (push a i₁) = 0ₖ _ + + βₗ'-fun≡ : ∣ βₗ'-fun n f g ∣₂ ≡ FF' ∣ wedgeGen ∣₂ + βₗ'-fun≡ = cong ∣_∣₂ (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push (inl x) i₁) → refl + ; (push (inr x) i₁) j → Kn→ΩKn+10ₖ _ (~ j) i₁ + ; (push (push a i₂) i₁) j → Kn→ΩKn+10ₖ _ (~ j ∧ i₂) i₁}) + + help2 : Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) (((suc (suc (n +ℕ n)))))))) + (∣ ∣_∣ ∣₂ , 0ₕ _) + ≡ ∣ wedgeGen ∣₂ + help2 = cong ∣_∣₂ (funExt λ { (inl x) → rUnitₖ (suc (suc (suc (n +ℕ n)))) ∣ x ∣ + ; (inr x) → refl + ; (push a i₁) → refl}) + +βᵣ≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → βᵣ n f g ≡ βᵣ''-fun n f g +βᵣ≡ n f g = + cong (FF ∘ subber2) + (cong (Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) + (S₊∙ (suc (suc (suc (n +ℕ n))))) + (((suc (suc (n +ℕ n))))))))) + help + ∙ help2) + ∙ funExt⁻ FF∘subber ∣ wedgeGen ∣₂ + ∙ cong (subber3) (sym βᵣ'-fun≡) + where + FF = MV.d _ _ _ (λ _ → tt) (fst (∨→∙ f g)) (suc (suc (n +ℕ suc n))) .fst + FF' = MV.d _ _ _ (λ _ → tt) (fst (∨→∙ f g)) (suc (suc (suc (n +ℕ n)))) .fst + + help : Iso.inv (fst (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (0 , 1) + ≡ (0ₕ _ , ∣ ∣_∣ ∣₂) + help = ΣPathP (IsGroupHom.pres1 (snd (invGroupIso (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) , (Hⁿ-Sⁿ≅ℤ-nice-generator _)) + + subber3 = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + + subber2 = (subst (λ m → coHom m (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) + (sym (cong (2 +ℕ_) (+-suc n n)))) + + FF∘subber : FF ∘ subber2 ≡ subber3 ∘ FF' + FF∘subber = + funExt (λ a → (λ j → transp (λ i → coHom ((suc ∘ suc ∘ suc) (+-suc n n (~ i ∧ j))) (C* n f g)) (~ j) + (MV.d _ _ _ (λ _ → tt) (fst (∨→∙ f g)) ((suc (suc (+-suc n n j)))) .fst + (transp (λ i → coHom (2 +ℕ (+-suc n n (j ∨ ~ i))) + (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) j + a)))) + + wedgeGen : (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) → + coHomK (suc (suc (suc (n +ℕ n))))) + wedgeGen (inl x) = 0ₖ _ + wedgeGen (inr x) = ∣ x ∣ + wedgeGen (push a i₁) = 0ₖ _ + + + help2 : Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) + (S₊∙ (suc (suc (suc (n +ℕ n))))) (((suc (suc (n +ℕ n)))))))) + (0ₕ _ , ∣ ∣_∣ ∣₂) + ≡ ∣ wedgeGen ∣₂ + help2 = cong ∣_∣₂ (funExt λ { (inl x) → refl + ; (inr x) → lUnitₖ (suc (suc (suc (n +ℕ n)))) ∣ x ∣ + ; (push a i₁) → refl}) + + βᵣ'-fun≡ : ∣ βᵣ'-fun n f g ∣₂ ≡ FF' ∣ wedgeGen ∣₂ + βᵣ'-fun≡ = cong ∣_∣₂ (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push (inl x) i₁) j → Kn→ΩKn+10ₖ _ (~ j) i₁ + ; (push (inr x) i₁) → refl + ; (push (push a i₂) i₁) j → Kn→ΩKn+10ₖ _ (~ j ∧ ~ i₂) i₁}) + +q-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (q n f g) (βₗ n f g) + ≡ β·Π n f g +q-βₗ n f g = cong (coHomFun _ (q n f g)) (βₗ≡ n f g) + ∙ transportLem + ∙ cong (subst (λ m → coHom m (C·Π' n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + (sym (Iso.leftInv (fst (Hopfβ-Iso n (∙Π f g))) (Hopfβ n (∙Π f g))) + ∙ cong (Iso.inv ((fst (Hopfβ-Iso n (∙Π f g))))) (Hopfβ↦1 n (∙Π f g))) + where + transportLem : coHomFun (suc (suc (suc (n +ℕ suc n)))) (q n f g) + (βₗ''-fun n f g) + ≡ subst (λ m → coHom m (C·Π' n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + (Hopfβ n (∙Π f g)) + transportLem = + natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (q n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∙ cong (subst (λ m → coHom m (C·Π' n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + (cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → loll a i₁})) + where + open import Cubical.Foundations.Path + loll : (x : fst (S₊∙ (3 +ℕ (n +ℕ n)))) → + PathP + (λ i₁ → + βₗ'-fun n f g ((push (θ x) ∙ λ i → inr (∙Π≡∨→∘θ n f g x (~ i))) i₁) ≡ + MV.d-pre Unit (fst (S₊∙ (2 +ℕ n))) (fst (S₊∙ (3 +ℕ n +ℕ n))) + (λ _ → tt) (fst (∙Π f g)) (3 +ℕ n +ℕ n) ∣_∣ (push x i₁)) + (λ _ → βₗ'-fun n f g (q n f g (inl tt))) + (λ _ → βₗ'-fun n f g (q n f g (inr (∙Π f g .fst x)))) + loll x = + flipSquare (cong-∙ (βₗ'-fun n f g) (push (θ x)) (λ i → inr (∙Π≡∨→∘θ n f g x (~ i))) + ∙∙ sym (rUnit _) + ∙∙ lem₁ x) + + where + lem₁ : (x : _) → cong (βₗ'-fun n f g) (push (θ {A = S₊∙ _} x)) ≡ Kn→ΩKn+1 _ ∣ x ∣ + lem₁ north = refl + lem₁ south = sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)) + lem₁ (merid a j) k = lala k j + where + lala : PathP (λ k → Kn→ΩKn+1 _ ∣ north ∣ₕ ≡ (sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north))) k) + (λ j → cong (βₗ'-fun n f g) (push (θ {A = S₊∙ _} (merid a j)))) + (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) + lala = (cong-∙∙ (cong (βₗ'-fun n f g) ∘ push) + ((λ i₁ → inl ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁))) + (push tt) + ((λ i₁ → inr ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁))) + ∙ sym (compPath≡compPath' _ _) + ∙ (λ _ → (λ j → Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∣ (merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) j ∣ₕ) + ∙ Kn→ΩKn+10ₖ _) + ∙ cong (_∙ Kn→ΩKn+10ₖ _) (cong-∙ ((Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ)) + (merid a) (sym (merid north))) + ∙ sym (assoc _ _ _) + ∙ cong (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a) ∙_) + (sym (symDistr (sym ((Kn→ΩKn+10ₖ _))) (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)))))) + ◁ λ i j → compPath-filler (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) + (sym (sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) + (cong ∣_∣ₕ (merid north)))) + (~ i) j + +q-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (q n f g) (βᵣ n f g) + ≡ β·Π n f g +q-βᵣ n f g = cong (coHomFun _ (q n f g)) (βᵣ≡ n f g) + ∙ transportLem + ∙ cong (subst (λ m → coHom m (C·Π' n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + (sym (Iso.leftInv (fst (Hopfβ-Iso n (∙Π f g))) (Hopfβ n (∙Π f g))) + ∙ cong (Iso.inv ((fst (Hopfβ-Iso n (∙Π f g))))) (Hopfβ↦1 n (∙Π f g))) + where + transportLem : coHomFun (suc (suc (suc (n +ℕ suc n)))) (q n f g) + (βᵣ''-fun n f g) + ≡ subst (λ m → coHom m (C·Π' n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + (Hopfβ n (∙Π f g)) + transportLem = + natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (q n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∙ cong (subst (λ m → coHom m (C·Π' n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + (cong ∣_∣₂ (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a i₁) → loll a i₁})) + where + open import Cubical.Foundations.Path + loll : (x : fst (S₊∙ (3 +ℕ (n +ℕ n)))) → + PathP + (λ i₁ → + βᵣ'-fun n f g ((push (θ x) ∙ λ i → inr (∙Π≡∨→∘θ n f g x (~ i))) i₁) ≡ + MV.d-pre Unit (fst (S₊∙ (2 +ℕ n))) (fst (S₊∙ (3 +ℕ n +ℕ n))) + (λ _ → tt) (fst (∙Π f g)) (3 +ℕ n +ℕ n) ∣_∣ (push x i₁)) + (λ _ → βᵣ'-fun n f g (q n f g (inl tt))) + (λ _ → βᵣ'-fun n f g (q n f g (inr (∙Π f g .fst x)))) + loll x = flipSquare (cong-∙ (βᵣ'-fun n f g) (push (θ x)) (λ i → inr (∙Π≡∨→∘θ n f g x (~ i))) + ∙∙ sym (rUnit _) + ∙∙ lem₁ x) + where + lem₁ : (x : _) → cong (βᵣ'-fun n f g) (push (θ {A = S₊∙ _} x)) ≡ Kn→ΩKn+1 _ ∣ x ∣ + lem₁ north = sym (Kn→ΩKn+10ₖ _) + lem₁ south = cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)) + lem₁ (merid a j) k = lala k j -- lala k j + where + lala : PathP (λ k → (Kn→ΩKn+10ₖ _) (~ k) ≡ (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north))) k) + (λ j → cong (βᵣ'-fun n f g) (push (θ {A = S₊∙ _} (merid a j)))) + (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) + lala = (cong-∙∙ (cong (βᵣ'-fun n f g) ∘ push) + (λ i₁ → inl ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁)) + (push tt) + (λ i₁ → inr ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁)) + ∙∙ (cong (sym (Kn→ΩKn+10ₖ _) ∙_) (cong-∙ (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a) (sym (merid (ptSn _))))) + ∙∙ sym (doubleCompPath≡compPath _ _ _)) + ◁ symP (doubleCompPath-filler + (sym (Kn→ΩKn+10ₖ _)) + (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) + (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (sym (merid north)))) +open import Cubical.Foundations.Path +jₗ-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (jₗ n f g) (α* n f g) + ≡ Hopfα n f +jₗ-α n f g = + cong (coHomFun _ (jₗ n f g)) (α*-lem n f g) + ∙ cong ∣_∣₂ (funExt (HopfInvariantPushElim n (fst f) + (isOfHLevelPath ((suc (suc (suc (suc (n +ℕ n)))))) + (isOfHLevelPlus' {n = n} (4 +ℕ n) (isOfHLevelTrunc (4 +ℕ n))) _ _) + refl + (λ _ → refl) + λ j → refl)) + +jₗ-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (jₗ n f g) (βₗ n f g) + ≡ subst (λ m → coHom m (HopfInvariantPush n (fst f))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + (Hopfβ n f) +jₗ-βₗ n f g = + cong (coHomFun _ (jₗ n f g)) (βₗ≡ n f g) + ∙∙ natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (jₗ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + (cong ∣_∣₂ + (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a i₁) → refl})) + +jₗ-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (jₗ n f g) (βᵣ n f g) + ≡ 0ₕ _ +jₗ-βᵣ n f g = + cong (coHomFun _ (jₗ n f g)) (βᵣ≡ n f g) + ∙∙ natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (jₗ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + cool + where + cool : coHomFun (suc (suc (suc (suc (n +ℕ n))))) (jₗ n f g) + ∣ βᵣ'-fun n f g ∣₂ ≡ 0ₕ _ + cool = cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → refl}) + +jᵣ-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (jᵣ n f g) (α* n f g) + ≡ Hopfα n g +jᵣ-α n f g = cong (coHomFun _ (jᵣ n f g)) (α*-lem n f g) + ∙ cong ∣_∣₂ (funExt (HopfInvariantPushElim n (fst g) + (isOfHLevelPath ((suc (suc (suc (suc (n +ℕ n)))))) + (isOfHLevelPlus' {n = n} (4 +ℕ n) (isOfHLevelTrunc (4 +ℕ n))) _ _) + refl + (λ _ → refl) + (flipSquare (pp-inr n f g)))) + +jᵣ-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (jᵣ n f g) (βₗ n f g) ≡ 0ₕ _ +jᵣ-βₗ n f g = + cong (coHomFun _ (jᵣ n f g)) (βₗ≡ n f g) + ∙∙ natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (jᵣ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + cool + where + cool : coHomFun (suc (suc (suc (suc (n +ℕ n))))) (jᵣ n f g) + ∣ βₗ'-fun n f g ∣₂ ≡ 0ₕ _ + cool = cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → refl}) + + +jᵣ-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → coHomFun _ (jᵣ n f g) (βᵣ n f g) + ≡ subst (λ m → coHom m (HopfInvariantPush n (fst g))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + (Hopfβ n g) +jᵣ-βᵣ n f g = + cong (coHomFun _ (jᵣ n f g)) (βᵣ≡ n f g) + ∙∙ natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (jᵣ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst g))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) + (cong ∣_∣₂ + (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a i₁) → refl})) + +gen₂ℤ×ℤ : gen₂-by (DirProd ℤGroup ℤGroup) (1 , 0) (0 , 1) +fst (gen₂ℤ×ℤ (x , y)) = x , y +snd (gen₂ℤ×ℤ (x , y)) = + ΣPathP ((cong₂ _+_ ((·Comm 1 x) ∙ cong fst (sym (distrLem 1 0 x))) ((·Comm 0 y) ∙ cong fst (sym (distrLem 0 1 y)))) + , +Comm y 0 ∙ cong₂ _+_ (·Comm 0 x ∙ cong snd (sym (distrLem 1 0 x))) (·Comm 1 y ∙ cong snd (sym (distrLem 0 1 y)))) + where + ℤ×ℤ = DirProd ℤGroup ℤGroup + _+''_ = GroupStr._·_ (snd ℤ×ℤ) + + ll3 : (x : ℤ) → - x ≡ 0 - x + ll3 (pos zero) = refl + ll3 (pos (suc zero)) = refl + ll3 (pos (suc (suc n))) = + cong predℤ (ll3 (pos (suc n))) + ll3 (negsuc zero) = refl + ll3 (negsuc (suc n)) = cong sucℤ (ll3 (negsuc n)) + + distrLem : (x y : ℤ) (z : ℤ) + → Path (ℤ × ℤ) (z ℤ[ ℤ×ℤ ]· (x , y)) (z · x , z · y) + distrLem x y (pos zero) = refl + distrLem x y (pos (suc n)) = + (cong ((x , y) +''_) (distrLem x y (pos n))) + distrLem x y (negsuc zero) = ΣPathP (sym (ll3 x) , sym (ll3 y)) + distrLem x y (negsuc (suc n)) = + cong₂ _+''_ (ΣPathP (sym (ll3 x) , sym (ll3 y))) + (distrLem x y (negsuc n)) + +genH²ⁿC* : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → gen₂-by (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) + (βₗ n f g) + (βᵣ n f g) +genH²ⁿC* n f g = + Iso-pres-gen₂ (DirProd ℤGroup ℤGroup) + (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) + (1 , 0) + (0 , 1) + gen₂ℤ×ℤ + (invGroupIso (H⁴-C* n f g)) + +private + + H⁴C* : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Group _ + H⁴C* n f g = coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) + + X : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → ℤ + X n f g = (genH²ⁿC* n f g) (α* n f g ⌣ α* n f g) .fst .fst + Y : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → ℤ + Y n f g = (genH²ⁿC* n f g) (α* n f g ⌣ α* n f g) .fst .snd + + α*≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → α* n f g ⌣ α* n f g ≡ ((X n f g) ℤ[ H⁴C* n f g ]· βₗ n f g) + +ₕ ((Y n f g) ℤ[ H⁴C* n f g ]· βᵣ n f g) + α*≡ n f g = (genH²ⁿC* n f g) (α* n f g ⌣ α* n f g) .snd + + +eq₁ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → (Hopfα n (∙Π f g)) ⌣ (Hopfα n (∙Π f g)) + ≡ ((X n f g + Y n f g) ℤ[ coHomGr _ _ ]· (β·Π n f g)) +eq₁ n f g = + cong (λ x → x ⌣ x) (sym (q-α n f g) ∙ cong (coHomFun (suc (suc n)) (q n f g)) (α*-lem n f g)) + ∙ cong ((coHomFun _) (q _ f g)) (cong (λ x → x ⌣ x) (sym (α*-lem n f g)) + ∙ α*≡ n f g) + ∙ IsGroupHom.pres· (coHomMorph _ (q n f g) .snd) _ _ + ∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (q n f g)) (βₗ n f g) (X n f g) + ∙ cong (λ z → (X n f g) ℤ[ coHomGr _ _ ]· z) + (q-βₗ n f g)) + (homPresℤ· (coHomMorph _ (q n f g)) (βᵣ n f g) (Y n f g) + ∙ cong (λ z → (Y n f g) ℤ[ coHomGr _ _ ]· z) + (q-βᵣ n f g)) + ∙ sym (distrℤ· (coHomGr _ _) (β·Π n f g) (X n f g) (Y n f g)) + +coHomFun⌣ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) + → (n : ℕ) (x y : coHom n B) + → coHomFun _ f (x ⌣ y) ≡ coHomFun _ f x ⌣ coHomFun _ f y +coHomFun⌣ f n = sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl + +eq₂ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → Hopfα n f ⌣ Hopfα n f + ≡ (X n f g ℤ[ coHomGr _ _ ]· subst (λ m → coHom m (HopfInvariantPush n (fst f))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + (Hopfβ n f)) +eq₂ n f g = + cong (λ x → x ⌣ x) (sym (jₗ-α n f g)) + ∙∙ sym (coHomFun⌣ (jₗ n f g) _ _ _) + ∙∙ cong (coHomFun _ (jₗ n f g)) (α*≡ n f g) + ∙∙ IsGroupHom.pres· (snd (coHomMorph _ (jₗ n f g))) _ _ + ∙∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (jₗ n f g)) (βₗ n f g) (X n f g)) + (homPresℤ· (coHomMorph _ (jₗ n f g)) (βᵣ n f g) (Y n f g) + ∙ cong (λ z → (Y n f g) ℤ[ coHomGr _ _ ]· z) + (jₗ-βᵣ n f g)) + ∙∙ cong₂ _+ₕ_ refl (rUnitℤ· (coHomGr _ _) (Y n f g)) + ∙∙ (rUnitₕ _ _ + ∙ cong (X n f g ℤ[ coHomGr _ _ ]·_) (jₗ-βₗ n f g)) + +eq₃ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + → Hopfα n g ⌣ Hopfα n g + ≡ (Y n f g ℤ[ coHomGr _ _ ]· subst (λ m → coHom m (HopfInvariantPush n (fst f))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + (Hopfβ n g)) +eq₃ n f g = + cong (λ x → x ⌣ x) (sym (jᵣ-α n f g)) + ∙∙ sym (coHomFun⌣ (jᵣ n f g) _ _ _) + ∙∙ cong (coHomFun _ (jᵣ n f g)) (α*≡ n f g) + ∙∙ IsGroupHom.pres· (snd (coHomMorph _ (jᵣ n f g))) _ _ + ∙∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (jᵣ n f g)) (βₗ n f g) (X n f g) + ∙ cong (λ z → (X n f g) ℤ[ coHomGr _ _ ]· z) + (jᵣ-βₗ n f g)) + (homPresℤ· (coHomMorph _ (jᵣ n f g)) (βᵣ n f g) (Y n f g)) + ∙∙ cong₂ _+ₕ_ (rUnitℤ· (coHomGr _ _) (X n f g)) refl + ∙∙ ((lUnitₕ _ _) ∙ cong (Y n f g ℤ[ coHomGr _ _ ]·_) (jᵣ-βᵣ n f g)) + +eq₂-eq₂ : (n : ℕ) (f g : S₊∙ (suc (suc (suc (n +ℕ n)))) →∙ S₊∙ (suc (suc n))) + → HopfInvariant n (∙Π f g) ≡ HopfInvariant n f + HopfInvariant n g +eq₂-eq₂ n f g = + mainL + ∙ sym (cong₂ _+_ (help n f g) (helpg n f g)) + where + transpLem : ∀ {ℓ} {G : ℕ → Group ℓ} + → (n m : ℕ) + → (x : ℤ) + → (p : m ≡ n) + → (g : fst (G n)) + → subst (fst ∘ G) p (x ℤ[ G m ]· subst (fst ∘ G) (sym p) g) ≡ (x ℤ[ G n ]· g) + transpLem {G = G} n m x = + J (λ n p → (g : fst (G n)) → subst (fst ∘ G) p (x ℤ[ G m ]· subst (fst ∘ G) (sym p) g) + ≡ (x ℤ[ G n ]· g)) + λ g → transportRefl _ ∙ cong (x ℤ[ G m ]·_) (transportRefl _) + + mainL : HopfInvariant n (∙Π f g) ≡ X n f g + Y n f g + mainL = cong (Iso.fun (fst (Hopfβ-Iso n (∙Π f g)))) + (cong (subst (λ x → coHom x (HopfInvariantPush n (fst (∙Π f g)))) + λ i₁ → suc (suc (suc (+-suc n n i₁)))) + (eq₁ n f g)) + ∙∙ cong (Iso.fun (fst (Hopfβ-Iso n (∙Π f g)))) + (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst (∙Π f g)))} _ _ + (X n f g + Y n f g) (λ i₁ → suc (suc (suc (+-suc n n i₁)))) + (Iso.inv (fst (Hopfβ-Iso n (∙Π f g))) 1)) + ∙∙ homPresℤ· (_ , snd (Hopfβ-Iso n (∙Π f g))) (Iso.inv (fst (Hopfβ-Iso n (∙Π f g))) 1) + (X n f g + Y n f g) + ∙∙ cong ((X n f g + Y n f g) ℤ[ ℤ , ℤGroup .snd ]·_) + (Iso.rightInv ((fst (Hopfβ-Iso n (∙Π f g)))) 1) + ∙∙ rUnitℤ·ℤ (X n f g + Y n f g) + + + help : (n : ℕ) (f g : _) → HopfInvariant n f ≡ X n f g + help n f g = cong (Iso.fun (fst (Hopfβ-Iso n f))) + (cong (subst (λ x → coHom x (HopfInvariantPush n (fst f))) + (λ i₁ → suc (suc (suc (+-suc n n i₁))))) (eq₂ n f g)) + ∙ cong (Iso.fun (fst (Hopfβ-Iso n f))) + (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst f))} _ _ + (X n f g) + ((cong (suc ∘ suc ∘ suc) (+-suc n n))) + (Hopfβ n f)) + ∙ homPresℤ· (_ , snd (Hopfβ-Iso n f)) (Hopfβ n f) (X n f g) + ∙ cong (X n f g ℤ[ ℤ , ℤGroup .snd ]·_) (Hopfβ↦1 n f) + ∙ rUnitℤ·ℤ (X n f g) + + helpg : (n : ℕ) (f g : _) → HopfInvariant n g ≡ Y n f g + helpg n f g = + cong (Iso.fun (fst (Hopfβ-Iso n g))) + (cong (subst (λ x → coHom x (HopfInvariantPush n (fst g))) + (λ i₁ → suc (suc (suc (+-suc n n i₁))))) + (eq₃ n f g)) + ∙∙ cong (Iso.fun (fst (Hopfβ-Iso n g))) + (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst g))} _ _ + (Y n f g) + ((cong (suc ∘ suc ∘ suc) (+-suc n n))) + (Hopfβ n g)) + ∙∙ homPresℤ· (_ , snd (Hopfβ-Iso n g)) (Hopfβ n g) (Y n f g) + ∙∙ cong (Y n f g ℤ[ ℤ , ℤGroup .snd ]·_) (Hopfβ↦1 n g) + ∙∙ rUnitℤ·ℤ (Y n f g) + + +-- GroupHom-HopfInvariant-π' : (n : ℕ) → GroupHom (π'Gr (suc (suc (n +ℕ n))) (S₊∙ (suc (suc n)))) ℤGroup +-- GroupHom-HopfInvariant-π' n = HopfInvariant-π' n , makeIsGroupHom {!eq₂-eq₂ n!} diff --git a/Cubical/Homotopy/HopfInvariant/HopfMap.agda b/Cubical/Homotopy/HopfInvariant/HopfMap.agda index 06eb05722..e554a9c85 100644 --- a/Cubical/Homotopy/HopfInvariant/HopfMap.agda +++ b/Cubical/Homotopy/HopfInvariant/HopfMap.agda @@ -91,34 +91,19 @@ snd HopfMap = refl -- We use the Hopf fibration in order to connect it to the Gysin Sequence module hopfS¹ = hopfBase S1-AssocHSpace (sphereElim2 _ (λ _ _ → squash) ∣ refl ∣) -S¹Hopf : S₊ 2 → Type S¹Hopf = hopfS¹.Hopf - -TotalHopf' : Type _ -TotalHopf' = Σ (S₊ 2) S¹Hopf - -CP² : Type _ +E* = hopfS¹.totalSpaceMegaPush +IsoE*join = hopfS¹.IsoJoin₁ +IsoTotalHopf' = hopfS¹.joinIso₁ CP² = hopfS¹.megaPush - -fibr : CP² → Type _ fibr = hopfS¹.P --- Remove -hopf : join S¹ S¹ → S₊ 2 -hopf x = fst (JoinS¹S¹→TotalHopf x) - -E* : Type _ -E* = hopfS¹.totalSpaceMegaPush - -IsoE*join : Iso E* (join S¹ (join S¹ S¹)) -IsoE*join = hopfS¹.IsoJoin₁ +TotalHopf' : Type _ +TotalHopf' = Σ (S₊ 2) S¹Hopf IsoJoins : (join S¹ (join S¹ S¹)) ≡ join S¹ (S₊ 3) IsoJoins = cong (join S¹) (isoToPath (IsoSphereJoin 1 1)) -IsoTotalHopf' : Iso hopfS¹.TotalSpaceHopf' (join S¹ S¹) -IsoTotalHopf' = hopfS¹.joinIso₁ - -- CP² is 1-connected conCP² : (x y : CP²) → ∥ x ≡ y ∥₂ conCP² x y = sRec2 squash₂ (λ p q → ∣ p ∙ sym q ∣₂) (conCP²' x) (conCP²' y) @@ -142,15 +127,15 @@ conCP² x y = sRec2 squash₂ (λ p q → ∣ p ∙ sym q ∣₂) (conCP²' x) ( main = indLem (λ _ → isOfHLevelPathP' 1 squash₂ _ _) λ j → ∣ (λ i → push (inl base) (~ i ∧ j)) ∣₂ -module GysinS¹ = Gysin (CP² , inl tt) fibr conCP² 2 idIso refl +module GysinS² = Gysin (CP² , inl tt) fibr conCP² 2 idIso refl -E'S⁴Iso : Iso GysinS¹.E' (S₊ 5) +E'S⁴Iso : Iso GysinS².E' (S₊ 5) E'S⁴Iso = compIso IsoE*join (compIso (Iso→joinIso idIso (IsoSphereJoin 1 1)) (IsoSphereJoin 1 3)) -isContrH³E : isContr (coHom 3 (GysinS¹.E')) +isContrH³E : isContr (coHom 3 (GysinS².E')) isContrH³E = subst isContr (sym (isoToPath (fst (Hⁿ-Sᵐ≅0 2 4 @@ -158,7 +143,7 @@ isContrH³E = ∙ cong (coHom 3) (sym (isoToPath E'S⁴Iso))) isContrUnit -isContrH⁴E : isContr (coHom 4 (GysinS¹.E')) +isContrH⁴E : isContr (coHom 4 (GysinS².E')) isContrH⁴E = subst isContr (sym (isoToPath (fst (Hⁿ-Sᵐ≅0 3 4 @@ -166,6 +151,7 @@ isContrH⁴E = ∙ cong (coHom 4) (sym (isoToPath E'S⁴Iso))) isContrUnit +-- We will need a bunch of elimination principles joinS¹S¹→Groupoid : ∀ {ℓ} (P : join S¹ S¹ → Type ℓ) → ((x : _) → isGroupoid (P x)) → P (inl base) @@ -201,10 +187,10 @@ CP²→Groupoid : ∀ {ℓ} {P : CP² → Type ℓ} → (x : _) → P x CP²→Groupoid {P = P} grp b s2 pp (inl x) = b CP²→Groupoid {P = P} grp b s2 pp (inr x) = s2 x -CP²→Groupoid {P = P} grp b s2 pp (push a i₁) = h3 a i₁ +CP²→Groupoid {P = P} grp b s2 pp (push a i₁) = lem a i₁ where - h3 : (a : hopfS¹.TotalSpaceHopf') → PathP (λ i → P (push a i)) b (s2 _) - h3 = TotalHopf→Gpd _ (λ _ → isOfHLevelPathP' 3 (grp _) _ _) pp + lem : (a : hopfS¹.TotalSpaceHopf') → PathP (λ i → P (push a i)) b (s2 _) + lem = TotalHopf→Gpd _ (λ _ → isOfHLevelPathP' 3 (grp _) _ _) pp -- The function inducing the iso H²(S²) ≅ H²(CP²) H²S²→H²CP²-raw : (S₊ 2 → coHomK 2) → (CP² → coHomK 2) @@ -264,11 +250,12 @@ snd H²CP²≅H²S² = H²CP²≅ℤ : GroupIso (coHomGr 2 CP²) ℤGroup H²CP²≅ℤ = compGroupIso H²CP²≅H²S² (Hⁿ-Sⁿ≅ℤ 1) +-- ⌣ gives a groupEquiv H²(CP²) ≃ H⁴(CP²) ⌣Equiv : GroupEquiv (coHomGr 2 CP²) (coHomGr 4 CP²) -fst (fst ⌣Equiv) = GysinS¹.⌣-hom 2 .fst -snd (fst ⌣Equiv) = subst isEquiv (cong fst (GysinS¹.ϕ∘j≡ 2)) eq' +fst (fst ⌣Equiv) = GysinS².⌣-hom 2 .fst +snd (fst ⌣Equiv) = subst isEquiv (cong fst (GysinS².ϕ∘j≡ 2)) eq' where - eq' : isEquiv (GysinS¹.ϕ∘j 2 .fst) + eq' : isEquiv (GysinS².ϕ∘j 2 .fst) eq' = SES→isEquiv (GroupPath _ _ .fst (invGroupEquiv @@ -277,13 +264,14 @@ snd (fst ⌣Equiv) = subst isEquiv (cong fst (GysinS¹.ϕ∘j≡ 2)) eq' (GroupPath _ _ .fst (invGroupEquiv (isContr→≃Unit isContrH⁴E , makeIsGroupHom λ _ _ → refl))) - (GysinS¹.susp∘ϕ 1) - (GysinS¹.ϕ∘j 2) - (GysinS¹.p-hom 4) - (GysinS¹.Ker-ϕ∘j⊂Im-Susp∘ϕ _) - (GysinS¹.Ker-p⊂Im-ϕ∘j _) -snd ⌣Equiv = GysinS¹.⌣-hom 2 .snd - + (GysinS².susp∘ϕ 1) + (GysinS².ϕ∘j 2) + (GysinS².p-hom 4) + (GysinS².Ker-ϕ∘j⊂Im-Susp∘ϕ _) + (GysinS².Ker-p⊂Im-ϕ∘j _) +snd ⌣Equiv = GysinS².⌣-hom 2 .snd + +-- The generator of H²(CP²) genCP² : coHom 2 CP² genCP² = ∣ CP²→Groupoid (λ _ → isOfHLevelTrunc 4) (0ₖ _) @@ -292,15 +280,15 @@ genCP² = ∣ CP²→Groupoid (λ _ → isOfHLevelTrunc 4) inrInjective : (f g : CP² → coHomK 2) → ∥ f ∘ inr ≡ g ∘ inr ∥ → Path (coHom 2 CP²) ∣ f ∣₂ ∣ g ∣₂ inrInjective f g = pRec (squash₂ _ _) - (λ p → pRec (squash₂ _ _) (λ id → trRec (squash₂ _ _) - (λ pp → cong ∣_∣₂ - (funExt (CP²→Groupoid (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) - id - (funExt⁻ p) - (compPathR→PathP pp)))) - (conn2 (f (inl tt)) (g (inl tt)) id - (cong f (push (inl base)) ∙ (funExt⁻ p north) ∙ cong g (sym (push (inl base)))))) - (conn (f (inl tt)) (g (inl tt)))) + (λ p → pRec (squash₂ _ _) (λ id → trRec (squash₂ _ _) + (λ pp → cong ∣_∣₂ + (funExt (CP²→Groupoid (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) + id + (funExt⁻ p) + (compPathR→PathP pp)))) + (conn2 (f (inl tt)) (g (inl tt)) id + (cong f (push (inl base)) ∙ (funExt⁻ p north) ∙ cong g (sym (push (inl base)))))) + (conn (f (inl tt)) (g (inl tt)))) where conn : (x y : coHomK 2) → ∥ x ≡ y ∥ @@ -310,6 +298,7 @@ inrInjective f g = pRec (squash₂ _ _) conn2 : (x y : coHomK 2) (p q : x ≡ y) → hLevelTrunc 1 (p ≡ q) conn2 x y = λ p q → Iso.fun (PathIdTruncIso _) (isContr→isProp (isConnectedPath _ (isConnectedKn 1) x y) ∣ p ∣ ∣ q ∣) +-- A couple of basic lemma concerning the hSpace structure on S¹ private rUnit* : (x : S¹) → x * base ≡ x rUnit* base = refl @@ -341,15 +330,19 @@ private (cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base))) (cong ∣_∣ₕ (sym (merid x ∙ sym (merid base)))) meridInvLooperLem x = (lUnit _ - ∙∙ cong (_∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base))) (sym (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base))))) ∙∙ sym (assoc _ _ _)) - ∙∙ cong (sym (cong ∣_∣ₕ (merid x ∙ sym (merid base))) ∙_) hh + ∙∙ cong (_∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base))) + (sym (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base))))) ∙∙ sym (assoc _ _ _)) + ∙∙ cong (sym (cong ∣_∣ₕ (merid x ∙ sym (merid base))) ∙_) lem ∙∙ (assoc _ _ _ - ∙∙ cong (_∙ (cong ∣_∣ₕ (sym (merid x ∙ sym (merid base))))) (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) + ∙∙ cong (_∙ (cong ∣_∣ₕ (sym (merid x ∙ sym (merid base))))) + (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) ∙∙ sym (lUnit _)) where - hh : cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base)) + lem : cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base)) ≡ cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙ cong ∣_∣ₕ (sym (merid x ∙ sym (merid base))) - hh = sym (merid*-lem x (invLooper x)) ∙ ((λ i → cong ∣_∣ₕ (merid (lemmie x (~ i)) ∙ sym (merid base))) ∙ cong (cong ∣_∣ₕ) (rCancel (merid base))) ∙ sym (rCancel _) + lem = sym (merid*-lem x (invLooper x)) + ∙ ((λ i → cong ∣_∣ₕ (merid (lemmie x (~ i)) ∙ sym (merid base))) + ∙ cong (cong ∣_∣ₕ) (rCancel (merid base))) ∙ sym (rCancel _) comm·S¹ : (a x : S¹) → a * x ≡ x * a comm·S¹ = wedgeconFun _ _ (λ _ _ → isGroupoidS¹ _ _) @@ -358,65 +351,82 @@ private refl invLooperLem₁ : (a x : S¹) → (invEq (hopfS¹.μ-eq a) x) * a ≡ (invLooper a * x) * a - invLooperLem₁ a x = secEq (hopfS¹.μ-eq a) x ∙∙ cong (_* x) (lemmie a) ∙∙ assocHSpace.μ-assoc S1-AssocHSpace a (invLooper a) x ∙ comm·S¹ _ _ + invLooperLem₁ a x = + secEq (hopfS¹.μ-eq a) x + ∙∙ cong (_* x) (lemmie a) + ∙∙ assocHSpace.μ-assoc S1-AssocHSpace a (invLooper a) x + ∙ comm·S¹ _ _ invLooperLem₂ : (a x : S¹) → invEq (hopfS¹.μ-eq a) x ≡ invLooper a * x invLooperLem₂ a x = sym (retEq (hopfS¹.μ-eq a) (invEq (hopfS¹.μ-eq a) x)) ∙∙ cong (invEq (hopfS¹.μ-eq a)) (invLooperLem₁ a x) ∙∙ retEq (hopfS¹.μ-eq a) (invLooper a * x) + rotLoop² : (a : S¹) → Path (a ≡ a) (λ i → rotLoop (rotLoop a i) (~ i)) refl + rotLoop² = + sphereElim 0 (λ _ → isGroupoidS¹ _ _ _ _) + λ i j → hcomp (λ {k → λ { (i = i1) → base + ; (j = i0) → base + ; (j = i1) → base}}) + base + + -- We prove that the generator of CP² given by Gysin is the same one -- as genCP², which is much easier to work with -Gysin-e≡genCP² : GysinS¹.e ≡ genCP² -Gysin-e≡genCP² = inrInjective _ _ ∣ funExt (λ x → funExt⁻ (cong fst (ll32 x)) south) ∣ +Gysin-e≡genCP² : GysinS².e ≡ genCP² +Gysin-e≡genCP² = inrInjective _ _ ∣ funExt (λ x → funExt⁻ (cong fst (main x)) south) ∣ where mainId : (x : Σ (S₊ 2) hopfS¹.Hopf) → Path (hLevelTrunc 4 _) ∣ fst x ∣ ∣ north ∣ mainId = uncurry λ { north → λ y → cong ∣_∣ₕ (merid base ∙ sym (merid y)) - ; south → λ y → cong ∣_∣ₕ (sym (merid y)) - ; (merid a i) → main a i} + ; south → λ y → cong ∣_∣ₕ (sym (merid y)) + ; (merid a i) → main a i} where - main : (a : S¹) → PathP (λ i → (y : hopfS¹.Hopf (merid a i)) → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a i ∣ ∣ north ∣) + main : (a : S¹) → PathP (λ i → (y : hopfS¹.Hopf (merid a i)) + → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a i ∣ ∣ north ∣) (λ y → cong ∣_∣ₕ (merid base ∙ sym (merid y))) λ y → cong ∣_∣ₕ (sym (merid y)) - main a = toPathP (funExt λ x → cong (transport ((λ i₁ → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a i₁ ∣ ∣ north ∣))) - ((λ i → (λ z → cong ∣_∣ₕ (merid base - ∙ sym (merid (transport (λ j → uaInvEquiv (hopfS¹.μ-eq a) (~ i) j) x))) z)) - ∙ λ i → cong ∣_∣ₕ (merid base - ∙ sym (merid (transportRefl (invEq (hopfS¹.μ-eq a) x) i)))) - ∙∙ (λ i → transp (λ i₁ → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a (i₁ ∨ i) ∣ ∣ north ∣) i - (compPath-filler' (cong ∣_∣ₕ (sym (merid a))) - (cong ∣_∣ₕ (merid base ∙ sym (merid (invLooperLem₂ a x i)))) i)) - ∙∙ cong ((cong ∣_∣ₕ) (sym (merid a)) ∙_) (cong (cong ∣_∣ₕ) (cong sym (symDistr (merid base) (sym (merid (invLooper a * x))))) - ∙ cong sym (merid*-lem (invLooper a) x) - ∙ symDistr ((cong ∣_∣ₕ) (merid (invLooper a) ∙ sym (merid base))) - ((cong ∣_∣ₕ) (merid x ∙ sym (merid base))) - ∙ isCommΩK 2 (sym (λ i₁ → ∣ (merid x ∙ (λ i₂ → merid base (~ i₂))) i₁ ∣)) - (sym (λ i₁ → ∣ (merid (invLooper a) ∙ (λ i₂ → merid base (~ i₂))) i₁ ∣)) - ∙ cong₂ _∙_ (cong sym (meridInvLooperLem a) ∙ cong-∙ ∣_∣ₕ (merid a) (sym (merid base))) - (cong (cong ∣_∣ₕ) (symDistr (merid x) (sym (merid base))) ∙ cong-∙ ∣_∣ₕ (merid base) (sym (merid x)))) - ∙∙ (λ j → (λ i₁ → ∣ merid a (~ i₁ ∨ j) ∣) ∙ - ((λ i₁ → ∣ merid a (i₁ ∨ j) ∣) ∙ - (λ i₁ → ∣ merid base (~ i₁ ∨ j) ∣)) - ∙ - (λ i₁ → ∣ merid base (i₁ ∨ j) ∣) ∙ - (λ i₁ → ∣ merid x (~ i₁) ∣ₕ)) - ∙∙ sym (lUnit _) - ∙∙ sym (assoc _ _ _) - ∙∙ (sym (lUnit _) ∙ sym (lUnit _) ∙ sym (lUnit _))) - - psst : (x : S₊ 2) → preThom.Q (CP² , inl tt) fibr (inr x) →∙ coHomK-ptd 2 - fst (psst x) north = ∣ north ∣ - fst (psst x) south = ∣ x ∣ - fst (psst x) (merid a i₁) = mainId (x , a) (~ i₁) - snd (psst x) = refl - - ll32-fst : GysinS¹.c (inr north) .fst ≡ psst north .fst - ll32-fst = cong fst (GysinS¹.cEq (inr north) ∣ push (inl base) ∣₂) + main a = + toPathP + (funExt λ x → + cong (transport (λ i₁ → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a i₁ ∣ ∣ north ∣)) + ((λ i → (λ z → cong ∣_∣ₕ + (merid base ∙ sym (merid (transport (λ j → uaInvEquiv (hopfS¹.μ-eq a) (~ i) j) x))) z)) + ∙ λ i → cong ∣_∣ₕ (merid base + ∙ sym (merid (transportRefl (invEq (hopfS¹.μ-eq a) x) i)))) + ∙∙ (λ i → transp (λ i₁ → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a (i₁ ∨ i) ∣ ∣ north ∣) i + (compPath-filler' (cong ∣_∣ₕ (sym (merid a))) + (cong ∣_∣ₕ (merid base ∙ sym (merid (invLooperLem₂ a x i)))) i)) + ∙∙ cong ((cong ∣_∣ₕ) (sym (merid a)) ∙_) + (cong (cong ∣_∣ₕ) (cong sym (symDistr (merid base) (sym (merid (invLooper a * x))))) + ∙ cong sym (merid*-lem (invLooper a) x) + ∙ symDistr ((cong ∣_∣ₕ) (merid (invLooper a) ∙ sym (merid base))) + ((cong ∣_∣ₕ) (merid x ∙ sym (merid base))) + ∙ isCommΩK 2 (sym (λ i₁ → ∣ (merid x ∙ (λ i₂ → merid base (~ i₂))) i₁ ∣)) + (sym (λ i₁ → ∣ (merid (invLooper a) ∙ (λ i₂ → merid base (~ i₂))) i₁ ∣)) + ∙ cong₂ _∙_ (cong sym (meridInvLooperLem a) + ∙ cong-∙ ∣_∣ₕ (merid a) (sym (merid base))) + (cong (cong ∣_∣ₕ) (symDistr (merid x) (sym (merid base))) + ∙ cong-∙ ∣_∣ₕ (merid base) (sym (merid x)))) + ∙∙ (λ j → (λ i₁ → ∣ merid a (~ i₁ ∨ j) ∣) ∙ ((λ i₁ → ∣ merid a (i₁ ∨ j) ∣) + ∙ (λ i₁ → ∣ merid base (~ i₁ ∨ j) ∣)) ∙ (λ i₁ → ∣ merid base (i₁ ∨ j) ∣) + ∙ (λ i₁ → ∣ merid x (~ i₁) ∣ₕ)) + ∙∙ sym (lUnit _) + ∙∙ sym (assoc _ _ _) + ∙∙ (sym (lUnit _) ∙ sym (lUnit _) ∙ sym (lUnit _))) + + gen' : (x : S₊ 2) → preThom.Q (CP² , inl tt) fibr (inr x) →∙ coHomK-ptd 2 + fst (gen' x) north = ∣ north ∣ + fst (gen' x) south = ∣ x ∣ + fst (gen' x) (merid a i₁) = mainId (x , a) (~ i₁) + snd (gen' x) = refl + + gen'Id : GysinS².c (inr north) .fst ≡ gen' north .fst + gen'Id = cong fst (GysinS².cEq (inr north) ∣ push (inl base) ∣₂) ∙ (funExt l) where l : (qb : _) → ∣ (subst (fst ∘ preThom.Q (CP² , inl tt) fibr) (sym (push (inl base))) qb) ∣ - ≡ psst north .fst qb + ≡ gen' north .fst qb l north = refl l south = cong ∣_∣ₕ (sym (merid base)) l (merid a i) j = @@ -425,146 +435,64 @@ Gysin-e≡genCP² = inrInjective _ _ ∣ funExt (λ x → funExt⁻ (cong fst (l ; (j = i0) → ∣ transportRefl (merid a i) (~ k) ∣ ; (j = i1) → ∣ compPath-filler (merid base) (sym (merid a)) k (~ i) ∣ₕ}) (hcomp (λ k → λ { (i = i0) → ∣ merid a (j ∨ ~ k) ∣ - ; (i = i1) → ∣ merid base (~ j ∨ ~ k) ∣ - ; (j = i0) → ∣ merid a (~ k ∨ i) ∣ - ; (j = i1) → ∣ merid base (~ i ∨ ~ k) ∣ₕ}) - ∣ south ∣) + ; (i = i1) → ∣ merid base (~ j ∨ ~ k) ∣ + ; (j = i0) → ∣ merid a (~ k ∨ i) ∣ + ; (j = i1) → ∣ merid base (~ i ∨ ~ k) ∣ₕ}) + ∣ south ∣) - is-setHepl : (x : S₊ 2) → isSet (preThom.Q (CP² , inl tt) fibr (inr x) →∙ coHomK-ptd 2) - is-setHepl = sphereElim _ (λ _ → isProp→isOfHLevelSuc 1 (isPropIsOfHLevel 2)) + setHelp : (x : S₊ 2) → isSet (preThom.Q (CP² , inl tt) fibr (inr x) →∙ coHomK-ptd 2) + setHelp = sphereElim _ (λ _ → isProp→isOfHLevelSuc 1 (isPropIsOfHLevel 2)) (isOfHLevel↑∙' 0 1) - ll32 : (x : S₊ 2) → (GysinS¹.c (inr x) ≡ psst x) - ll32 = sphereElim _ (λ x → isOfHLevelPath 2 (is-setHepl x) _ _) - (→∙Homogeneous≡ (isHomogeneousKn _) ll32-fst) + main : (x : S₊ 2) → (GysinS².c (inr x) ≡ gen' x) + main = sphereElim _ (λ x → isOfHLevelPath 2 (setHelp x) _ _) + (→∙Homogeneous≡ (isHomogeneousKn _) gen'Id) -isGenerator≃ℤ : ∀ {ℓ} (G : Group ℓ) (g : fst G) - → Type _ +isGenerator≃ℤ : ∀ {ℓ} (G : Group ℓ) (g : fst G) → Type ℓ isGenerator≃ℤ G g = Σ[ e ∈ GroupIso G ℤGroup ] abs (Iso.fun (fst e) g) ≡ 1 -isGenerator≃ℤ-e : isGenerator≃ℤ (coHomGr 2 CP²) GysinS¹.e +isGenerator≃ℤ-e : isGenerator≃ℤ (coHomGr 2 CP²) GysinS².e isGenerator≃ℤ-e = subst (isGenerator≃ℤ (coHomGr 2 CP²)) (sym Gysin-e≡genCP²) (H²CP²≅ℤ , refl) +-- Alternative definition of the hopfMap HopfMap' : S₊ 3 → S₊ 2 HopfMap' x = hopfS¹.TotalSpaceHopf'→TotalSpace (Iso.inv IsoTotalHopf' (Iso.inv (IsoSphereJoin 1 1) x)) .fst - -rotLoop² : (a : _) → Path (Path S¹ _ _) (λ i → rotLoop (rotLoop a i) (~ i)) refl -rotLoop² = - sphereElim 0 (λ _ → isGroupoidS¹ _ _ _ _) - λ i j → hcomp (λ {k → λ {(i = i1) → base ; (j = i0) → base ; (j = i1) → base}}) - base - -id1 : (x : _) → hopfS¹.TotalSpaceHopf'→TotalSpace x .fst ≡ JoinS¹S¹→TotalHopf (Iso.fun IsoTotalHopf' x) .fst -id1 (inl x) = refl -id1 (inr x) = refl -id1 (push (base , snd₁) i) = refl -id1 (push (loop i₁ , a) i) k = merid (rotLoop² a (~ k) i₁) i - -CP2' : Type _ -CP2' = Pushout {A = S₊ 3} (λ _ → tt) HopfMap' - hopfMap≡HopfMap' : HopfMap ≡ (HopfMap' , refl) hopfMap≡HopfMap' = ΣPathP ((funExt (λ x → cong (λ x → JoinS¹S¹→TotalHopf x .fst) (sym (Iso.rightInv IsoTotalHopf' (Iso.inv (IsoSphereJoin 1 1) x))) - ∙ sym (id1 (Iso.inv IsoTotalHopf' (Iso.inv (IsoSphereJoin 1 1) x))))) + ∙ sym (lem (Iso.inv IsoTotalHopf' (Iso.inv (IsoSphereJoin 1 1) x))))) , flipSquare (sym (rUnit refl) ◁ λ _ _ → north)) + where + lem : (x : _) → hopfS¹.TotalSpaceHopf'→TotalSpace x .fst + ≡ JoinS¹S¹→TotalHopf (Iso.fun IsoTotalHopf' x) .fst + lem (inl x) = refl + lem (inr x) = refl + lem (push (base , snd₁) i) = refl + lem (push (loop i₁ , a) i) k = merid (rotLoop² a (~ k) i₁) i + +CP²' : Type _ +CP²' = Pushout {A = S₊ 3} (λ _ → tt) HopfMap' PushoutReplaceBase : ∀ {ℓ ℓ' ℓ''} {A B : Type ℓ} {C : Type ℓ'} {D : Type ℓ''} {f : A → C} {g : A → D} (e : B ≃ A) → Pushout (f ∘ fst e) (g ∘ fst e) ≡ Pushout f g PushoutReplaceBase {f = f} {g = g} = - EquivJ (λ _ e → Pushout (f ∘ fst e) (g ∘ fst e) ≡ Pushout f g) - refl + EquivJ (λ _ e → Pushout (f ∘ fst e) (g ∘ fst e) ≡ Pushout f g) refl -CP2≡CP2' : CP2' ≡ CP² -CP2≡CP2' = +CP2≡CP²' : CP²' ≡ CP² +CP2≡CP²' = PushoutReplaceBase (isoToEquiv (compIso (invIso (IsoSphereJoin 1 1)) (invIso IsoTotalHopf'))) -{- -⌣-pres1-CP² : ⌣-pres1 CP² 2 -fst ⌣-pres1-CP² = GysinS¹.e -fst (snd ⌣-pres1-CP2) = isGenerator≃ℤ-e -snd (snd ⌣-pres1-CP2) = ∣ p ∣ - where - p : Σ _ _ - fst p = compGroupIso (GroupEquiv→GroupIso (invGroupEquiv ⌣Equiv)) H²CP²≅ℤ - snd p = {!!} - ∙ sesIsoPresGen _ (invGroupEquiv (GroupIso→GroupEquiv H²CP²≅ℤ)) _ - (compGroupEquiv (invGroupEquiv (GroupIso→GroupEquiv H²CP²≅ℤ)) (invGroupEquiv (invGroupEquiv ⌣Equiv))) - GysinS¹.e {!!} ⌣Equiv -{- - GysinS¹.e - , (isGenerator≃ℤ-e - , ∣ compGroupIso (GroupEquiv→GroupIso (invGroupEquiv ⌣Equiv)) H²CP²≅ℤ - , {!!} ∣) - -} - -⌣-pres1-CP2' : ⌣-pres1 CP2' 2 -⌣-pres1-CP2' = subst (λ x → ⌣-pres1 x 2) (sym CP2≡CP2') ⌣-pres1-CP2 - -} --- {- --- Goal: snd (v' (pt A) (push a i₁)) ≡ --- ua-gluePt (μ-eq (snd a)) i₁ (fst a) --- ———— Boundary —————————————————————————————————————————————— --- i₁ = i0 ⊢ HSpace.μₗ e (fst a) --- i₁ = i1 ⊢ HSpace.μₗ e (HSpace.μ e (fst a) (snd a)) --- -} - --- -- help : (x : _) → (v' (pt A)) x ≡ TotalSpaceHopf'→TotalSpace x --- -- help (inl x) = ΣPathP (refl , HSpace.μₗ e x) --- -- help (inr x) = ΣPathP (refl , (HSpace.μₗ e x)) --- -- help (push (x , y) i) j = --- -- comp (λ _ → Σ (Susp (typ A)) Hopf) --- -- (λ k → λ {(i = i0) → merid y i , HSpace.μₗ e x j --- -- ; (i = i1) → merid y i , assocHSpace.μ-assoc-filler e-ass x y j k --- -- ; (j = i0) → merid y i , hfill --- -- (λ j → λ { (i = i0) → HSpace.μ e (pt A) x --- -- ; (i = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j --- -- }) --- -- (inS (ua-gluePt (μ-eq y) i (HSpace.μ e (pt A) x))) --- -- k --- -- ; (j = i1) → merid y i , ua-gluePt (μ-eq y) i x}) --- -- (merid y i , ua-gluePt (μ-eq y) i (HSpace.μₗ e x j)) --- -- where --- -- open import Cubical.Foundations.Path - --- -- PPΣ : ∀ {ℓ} {A : Type ℓ} {f : A ≃ A} (p : f ≡ f) → {!!} --- -- PPΣ = {!!} - --- -- V : PathP (λ i₁ → hcomp --- -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x --- -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j --- -- }) --- -- (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x)) ≡ --- -- ua-gluePt (μ-eq y) i₁ x) --- -- (HSpace.μₗ e x) --- -- (HSpace.μₗ e (HSpace.μ e x y)) -- (HSpace.μₗ e (HSpace.μ e (fst a) (snd a))) --- -- V = transport (λ z → {!PathP (λ i₁ → hfill --- -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x --- -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j --- -- }) --- -- (inS (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x))) z ≡ ua-gluePt (μ-eq y) i₁ x) --- -- ? ?!}) --- -- {!hfill --- -- (λ { j (i₁ = i0) → HSpace.μ e (pt A) x --- -- ; j (i₁ = i1) → assocHSpace.μ-assoc e-ass (pt A) x y j --- -- }) --- -- (inS (ua-gluePt (μ-eq y) i₁ (HSpace.μ e (pt A) x))) ?!} -- toPathP ({!!} ∙∙ {!!} ∙∙ {!!}) -- toPathP (flipSquare {!!}) -- hcomp {!!} {!!} - --- -- P : Pushout {A = TotalSpaceHopf'} (λ _ → tt) induced → Type _ --- -- P (inl x) = typ A --- -- P (inr x) = Hopf x --- -- P (push a i₁) = ua (v a) i₁ - +-- packaging everything up: ⌣equiv→pres1 : ∀ {ℓ} {G H : Type ℓ} → (G ≡ H) → (g₁ : coHom 2 G) (h₁ : coHom 2 H) → (fstEq : (Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] abs (fst (fst ϕ) g₁) ≡ 1)) → (sndEq : ((Σ[ ϕ ∈ GroupEquiv (coHomGr 2 H) ℤGroup ] abs (fst (fst ϕ) h₁) ≡ 1))) @@ -641,6 +569,30 @@ snd (snd ⌣-pres1-CP2) = ∣ p ∣ snd main4 = makeIsGroupHom λ g1 g2 → rightDistr-⌣ _ _ g1 g2 h +-- The hopf invariant is ±1 for both definitions of the hopf map +HopfInvariant-HopfMap' : abs (HopfInvariant zero (HopfMap' , λ _ → HopfMap' (snd (S₊∙ 3)))) + ≡ suc zero +HopfInvariant-HopfMap' = + cong abs (cong (Iso.fun (fst (Hopfβ-Iso zero (HopfMap' , refl)))) + (transportRefl (⌣-α 0 (HopfMap' , refl)))) + ∙ ⌣equiv→pres1 (sym CP2≡CP²') GysinS².e (Hopfα zero (HopfMap' , refl)) + (l isGenerator≃ℤ-e) + (GroupIso→GroupEquiv (Hopfα-Iso 0 (HopfMap' , refl)) , refl) + (snd (fst ⌣Equiv)) + (GroupIso→GroupEquiv (Hopfβ-Iso zero (HopfMap' , refl))) + where + l : Σ[ ϕ ∈ GroupIso (coHomGr 2 CP²) ℤGroup ] + (abs (Iso.fun (fst ϕ) GysinS².e) ≡ 1) + → Σ[ ϕ ∈ GroupEquiv (coHomGr 2 CP²) ℤGroup ] + (abs (fst (fst ϕ) GysinS².e) ≡ 1) + l p = (GroupIso→GroupEquiv (fst p)) , (snd p) + +HopfInvariant-HopfMap : abs (HopfInvariant zero HopfMap) ≡ suc zero +HopfInvariant-HopfMap = cong abs (cong (HopfInvariant zero) hopfMap≡HopfMap') + ∙ HopfInvariant-HopfMap' + + + HopfInvariantPushElim : ∀ {ℓ} n → (f : _) → {P : HopfInvariantPush n f → Type ℓ} → (isOfHLevel (suc (suc (suc (suc (n +ℕ n))))) (P (inl tt))) → (e : P (inl tt)) @@ -658,26 +610,6 @@ HopfInvariantPushElim n f {P = P} hlev e g r (push a i₁) = help a i₁ (cong P (push north)) hlev) _ _)) r -HopfInvariant-HopfMap' : abs (HopfInvariant zero (HopfMap' , λ _ → HopfMap' (snd (S₊∙ 3)))) ≡ suc zero -HopfInvariant-HopfMap' = - cong abs (cong (Iso.fun (fst (Hopfβ-Iso zero (HopfMap' , refl)))) - (transportRefl (⌣-α 0 (HopfMap' , refl)))) ∙ ⌣equiv→pres1 (sym CP2≡CP2') GysinS¹.e (Hopfα zero (HopfMap' , refl)) - (l isGenerator≃ℤ-e) - (GroupIso→GroupEquiv (Hopfα-Iso 0 (HopfMap' , refl)) , refl) - (snd (fst ⌣Equiv)) - (GroupIso→GroupEquiv (Hopfβ-Iso zero (HopfMap' , refl))) - where - l : Σ-syntax (GroupIso (coHomGr 2 CP²) ℤGroup) - (λ ϕ → abs (Iso.fun (fst ϕ) GysinS¹.e) ≡ 1) - → Σ-syntax (GroupEquiv (coHomGr 2 CP²) ℤGroup) - (λ ϕ → abs (fst (fst ϕ) GysinS¹.e) ≡ 1) - l p = (GroupIso→GroupEquiv (fst p)) , (snd p) - -HopfInvariant-HopfMap : abs (HopfInvariant zero HopfMap) ≡ suc zero -HopfInvariant-HopfMap = cong abs (cong (HopfInvariant zero) hopfMap≡HopfMap') - ∙ HopfInvariant-HopfMap' - - -- open import Cubical.HITs.Wedge -- _∨→_ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} From 8a26ba81b11c594b859d7003f8a5afabaddf3a78 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 26 Oct 2021 16:09:27 +0200 Subject: [PATCH 72/89] stuff --- Cubical/Foundations/Prelude.agda | 13 +- Cubical/HITs/Wedge/Base.agda | 6 + Cubical/Homotopy/Freudenthal.agda | 1 - Cubical/Homotopy/Group/Base.agda | 724 +++++++++++ Cubical/Homotopy/Group/S3.agda | 686 ++++++++++ Cubical/Homotopy/HomotopyGroup.agda | 261 ++-- Cubical/Homotopy/HopfInvariant/Base.agda | 205 +-- .../Homotopy/HopfInvariant/Homomorphism.agda | 861 +++++++------ Cubical/Homotopy/HopfInvariant/HopfMap.agda | 1109 +++-------------- Cubical/ZCohomology/Gysin.agda | 576 +++++---- Cubical/ZCohomology/Properties.agda | 24 +- 11 files changed, 2707 insertions(+), 1759 deletions(-) create mode 100644 Cubical/Homotopy/Group/Base.agda create mode 100644 Cubical/Homotopy/Group/S3.agda diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index 8d5731cf4..c4368711a 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -485,9 +485,6 @@ PathP→compPathR {p = p} {q = q} {r = r} {s = s} P j i = ; (j = i1) → doubleCompPath-filler∙ p s (sym q) (~ k) i}) (P j i) - --- otherdir - compPathL→PathP : {a b c d : A} {p : a ≡ c} {q : b ≡ d} {r : a ≡ b} {s : c ≡ d} → sym p ∙ r ∙ q ≡ s → PathP (λ i → p i ≡ q i) r s @@ -517,3 +514,13 @@ compPathR→PathP∙∙ {p = p} {q = q} {r = r} {s = s} P j i = ; (j = i0) → r i ; (j = i1) → doubleCompPath-filler p s (sym q) (~ k) i}) (P j i) + +doubleCompPath≡compPath : + ∀ {ℓ} {A : Type ℓ} {x y z w : A} + (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) + → p ∙∙ q ∙∙ r ≡ p ∙ q ∙ r +doubleCompPath≡compPath p q r i j = + hcomp (λ k → λ { (i = i1) → compPath-filler' p (q ∙ r) k j + ; (j = i0) → p (~ k) + ; (j = i1) → r (i ∨ k)}) + (compPath-filler q r i j) diff --git a/Cubical/HITs/Wedge/Base.agda b/Cubical/HITs/Wedge/Base.agda index 4bda0d609..8a5c1cd36 100644 --- a/Cubical/HITs/Wedge/Base.agda +++ b/Cubical/HITs/Wedge/Base.agda @@ -25,6 +25,12 @@ _∨→_ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointe (f ∨→ g) (inr x) = fst g x (f ∨→ g) (push a i₁) = (snd f ∙ sym (snd g)) i₁ +-- Pointed version +∨→∙ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + → (f : A →∙ C) (g : B →∙ C) → ((A ⋁∙ₗ B) →∙ C) +fst (∨→∙ {A = A} f g) = f ∨→ g +snd (∨→∙ {A = A} f g) = snd f + -- Wedge sum of Units is contractible isContr-Unit⋁Unit : isContr ((Unit , tt) ⋁ (Unit , tt)) fst isContr-Unit⋁Unit = inl tt diff --git a/Cubical/Homotopy/Freudenthal.agda b/Cubical/Homotopy/Freudenthal.agda index 9a097f81d..e319b5172 100644 --- a/Cubical/Homotopy/Freudenthal.agda +++ b/Cubical/Homotopy/Freudenthal.agda @@ -7,7 +7,6 @@ Freudenthal suspension theorem module Cubical.Homotopy.Freudenthal where open import Cubical.Foundations.Everything --- open import Cubical.Data.HomotopyGroup open import Cubical.Data.Nat open import Cubical.Data.Sigma open import Cubical.HITs.Nullification diff --git a/Cubical/Homotopy/Group/Base.agda b/Cubical/Homotopy/Group/Base.agda new file mode 100644 index 000000000..5d619d011 --- /dev/null +++ b/Cubical/Homotopy/Group/Base.agda @@ -0,0 +1,724 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +{- +This file contains +1. The definition of πₙ as a truncated loop space +2. The definition of πₙ as a truncated function space (Sⁿ →∙ A) +3. A structure preserving equivalence Ωⁿ A ≃ (Sⁿ →∙ A) +4. A proof that the two constructions of homotopy groups are isomorphic +-} +module Cubical.Homotopy.Group.Base where + +open import Cubical.Homotopy.Loopspace + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Path +open import Cubical.Foundations.Transport + +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim + ; elim2 to sElim2 ; elim3 to sElim3 ; map to sMap) +open import Cubical.HITs.Sn +open import Cubical.Data.Bool +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 +open import Cubical.Data.Sigma +open import Cubical.Data.Nat + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid + +open Iso +open IsGroup +open IsSemigroup +open IsMonoid +open GroupStr + + +{- homotopy Group -} +π : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Type ℓ +π n A = ∥ typ ((Ω^ n) A) ∥₂ + +π' : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Type ℓ +π' n A = ∥ S₊∙ n →∙ A ∥₂ + + +{- π as a group -} +1π : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π n A +1π zero {A = A} = ∣ pt A ∣₂ +1π (suc n) = ∣ refl ∣₂ + +·π : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π (suc n) A → π (suc n) A → π (suc n) A +·π n = sRec2 squash₂ λ p q → ∣ p ∙ q ∣₂ + +-π : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π (suc n) A → π (suc n) A +-π n = sMap sym + +π-rUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π (suc n) A) + → (·π n x (1π (suc n))) ≡ x +π-rUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ rUnit p (~ i) ∣₂ + +π-lUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π (suc n) A) + → (·π n (1π (suc n)) x) ≡ x +π-lUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ lUnit p (~ i) ∣₂ + +π-rCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π (suc n) A) + → (·π n x (-π n x)) ≡ 1π (suc n) +π-rCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ rCancel p i ∣₂ + +π-lCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π (suc n) A) + → (·π n (-π n x) x) ≡ 1π (suc n) +π-lCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ lCancel p i ∣₂ + +π-assoc : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y z : π (suc n) A) + → ·π n x (·π n y z) ≡ ·π n (·π n x y) z +π-assoc n = sElim3 (λ _ _ _ → isSetPathImplicit) λ p q r i → ∣ ∙assoc p q r i ∣₂ + +π-comm : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y : π (suc (suc n)) A) + → ·π (suc n) x y ≡ ·π (suc n) y x +π-comm n = sElim2 (λ _ _ → isSetPathImplicit) λ p q i → ∣ EH n p q i ∣₂ + +πGr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Group ℓ +fst (πGr n A) = π (suc n) A +1g (snd (πGr n A)) = 1π (suc n) +GroupStr._·_ (snd (πGr n A)) = ·π n +inv (snd (πGr n A)) = -π n +is-set (isSemigroup (isMonoid (isGroup (snd (πGr n A))))) = squash₂ +assoc (isSemigroup (isMonoid (isGroup (snd (πGr n A))))) = π-assoc n +identity (isMonoid (isGroup (snd (πGr n A)))) x = (π-rUnit n x) , (π-lUnit n x) +inverse (isGroup (snd (πGr n A))) x = (π-rCancel n x) , (π-lCancel n x) + +-- We now define the group structure on π'. +-- We first define the corresponding structure on the untruncated +-- (S₊∙ n →∙ A). We first give multiplication and inversion and use this +-- to give a structure preserving equivalence (S₊∙ n →∙ A) ≃ Ωⁿ A. + +∙Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (S₊∙ n →∙ A) + → (S₊∙ n →∙ A) + → (S₊∙ n →∙ A) +∙Π {A = A} {n = zero} p q = (λ _ → pt A) , refl +fst (∙Π {A = A} {n = suc zero} (f , p) (g , q)) base = pt A +fst (∙Π {A = A} {n = suc zero} (f , p) (g , q)) (loop j) = + ((sym p ∙∙ cong f loop ∙∙ p) ∙ (sym q ∙∙ cong g loop ∙∙ q)) j +snd (∙Π {A = A} {n = suc zero} (f , p) (g , q)) = refl +fst (∙Π {A = A} {n = suc (suc n)} (f , p) (g , q)) north = pt A +fst (∙Π {A = A} {n = suc (suc n)} (f , p) (g , q)) south = pt A +fst (∙Π {A = A} {n = suc (suc n)} (f , p) (g , q)) (merid a j) = + ((sym p ∙∙ cong f (merid a ∙ sym (merid (ptSn (suc n)))) ∙∙ p) + ∙ (sym q ∙∙ cong g (merid a ∙ sym (merid (ptSn (suc n)))) ∙∙ q)) j +snd (∙Π {A = A} {n = suc (suc n)} (f , p) (g , q)) = refl + +-Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (S₊∙ n →∙ A) + → (S₊∙ n →∙ A) +-Π {n = zero} f = f +fst (-Π {A = A} {n = suc zero} f) base = fst f base +fst (-Π {A = A} {n = suc zero} f) (loop j) = fst f (loop (~ j)) +snd (-Π {A = A} {n = suc zero} f) = snd f +fst (-Π {A = A} {n = suc (suc n)} f) north = fst f north +fst (-Π {A = A} {n = suc (suc n)} f) south = fst f north +fst (-Π {A = A} {n = suc (suc n)} f) (merid a j) = + fst f ((merid a ∙ sym (merid (ptSn _))) (~ j)) +snd (-Π {A = A} {n = suc (suc n)} f) = snd f + +module _ {ℓ : Level} {A : Pointed ℓ} where +flipΩPath : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → ((Ω^ (suc n)) A) ≡ (Ω^ n) (Ω A) +flipΩPath {A = A} zero = refl +flipΩPath {A = A} (suc n) = cong Ω (flipΩPath {A = A} n) + +flipΩIso : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → Iso (fst ((Ω^ (suc n)) A)) (fst ((Ω^ n) (Ω A))) +flipΩIso {A = A} n = pathToIso (cong fst (flipΩPath n)) + +flipΩIsopres· : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (f g : fst ((Ω^ (suc n)) (Ω A))) + → Iso.inv (flipΩIso (suc n)) (f ∙ g) + ≡ (Iso.inv (flipΩIso (suc n)) f) + ∙ (Iso.inv (flipΩIso (suc n)) g) +flipΩIsopres· {A = A} n f g i = + transp (λ j → flipΩPath {A = A} n (~ i ∧ ~ j) .snd + ≡ flipΩPath n (~ i ∧ ~ j) .snd) i + (transp (λ j → flipΩPath {A = A} n (~ i ∨ ~ j) .snd + ≡ flipΩPath n (~ i ∨ ~ j) .snd) (~ i) f + ∙ transp (λ j → flipΩPath {A = A} n (~ i ∨ ~ j) .snd + ≡ flipΩPath n (~ i ∨ ~ j) .snd) (~ i) g) + +module _ {ℓ : Level} {A : Pointed ℓ} where + suspension→ : (n : ℕ) → (S₊∙ (suc n) →∙ Ω A) → (S₊∙ (suc (suc n)) →∙ A) + fst (suspension→ n f) north = pt A + fst (suspension→ n f) south = pt A + fst (suspension→ n f) (merid a i₁) = fst f a i₁ + snd (suspension→ n f) = refl + + suspension-pres∙ : + (n : ℕ) (f g : (S₊∙ (suc n) →∙ Ω A)) + → suspension→ n (∙Π f g) ≡ ∙Π (suspension→ n f) (suspension→ n g) + suspension-pres∙ n f g = ΣPathP ((funExt main) , refl) + where + J2 : ∀ {ℓ ℓ'} {A : Type ℓ} {x : A} + → (P : (p q : x ≡ x) → (refl ≡ p) → (refl ≡ q) → Type ℓ') + → P refl refl refl refl + → (p q : x ≡ x) (P' : _) (Q : _) → P p q P' Q + J2 P b p q = + J (λ p P' → (Q₁ : refl ≡ q) → P p q P' Q₁) + (J (λ q Q → P refl q refl Q) b) + + J4 : ∀ {ℓ ℓ'} {A : Type ℓ} {x : A} + → (P : (p q r s : x ≡ x) + → (refl ≡ p) → (p ≡ q) → (refl ≡ r) → (r ≡ s) → Type ℓ') + → P refl refl refl refl refl refl refl refl + → (p q r s : x ≡ x) (P' : _) (Q : _) (R : _) (S : _) → P p q r s P' Q R S + J4 P b p q r s = + J (λ p P' → (Q₁ : p ≡ q) (R : refl ≡ r) (S₁ : r ≡ s) + → P p q r s P' Q₁ R S₁) + (J (λ q Q₁ → (R : refl ≡ r) (S₁ : r ≡ s) + → P refl q r s refl Q₁ R S₁) + (J (λ r R → (S₁ : r ≡ s) + → P refl refl r s refl refl R S₁) + (J (λ s S₁ → P refl refl refl s refl refl refl S₁) + b))) + + looper-help : (n : ℕ) (f g : (S₊∙ (suc n) →∙ Ω A)) (a : _) + → fst (∙Π f g) a ≡ fst f a ∙ fst g a + looper-help zero f g base = + rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g)) + looper-help zero f g (loop i) k = + lazy-fill _ _ + (sym (snd f)) (sym (snd g)) + (cong (fst f) loop) (cong (fst g) loop) k i + where + lazy-fill : + ∀ {ℓ} {A : Type ℓ} {x : A} (p q : x ≡ x) + (prefl : refl ≡ p) (qrefl : refl ≡ q) (fl : p ≡ p) (fr : q ≡ q) + → PathP (λ i → (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i + ≡ (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i) + ((prefl ∙∙ fl ∙∙ sym prefl) ∙ (qrefl ∙∙ fr ∙∙ sym qrefl)) + (cong₂ _∙_ fl fr) + lazy-fill {x = x} = + J2 (λ p q prefl qrefl → (fl : p ≡ p) (fr : q ≡ q) + → PathP (λ i → (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i + ≡ (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i) + ((prefl ∙∙ fl ∙∙ sym prefl) ∙ (qrefl ∙∙ fr ∙∙ sym qrefl)) + (cong₂ _∙_ fl fr)) + λ fl fr → cong₂ _∙_ (sym (rUnit fl)) (sym (rUnit fr)) + ◁ flipSquare (sym (rUnit (rUnit refl)) + ◁ (flipSquare ((λ j → cong (λ x → rUnit x j) fl + ∙ cong (λ x → lUnit x j) fr) + ▷ sym (cong₂Funct _∙_ fl fr)) + ▷ (rUnit (rUnit refl)))) + looper-help (suc n) f g north = + rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g)) + looper-help (suc n) f g south = + (rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) + ∙ cong₂ _∙_ (cong (fst f) (merid (ptSn _))) + (cong (fst g) (merid (ptSn _))) + looper-help (suc n) f g (merid a j) i = help i j + where + help : + PathP (λ i → (rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) i + ≡ ((rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) + ∙ cong₂ _∙_ (cong (fst f) (merid (ptSn _))) + (cong (fst g) (merid (ptSn _)))) i) + (cong (fst (∙Π f g)) (merid a)) + (cong₂ _∙_ (cong (fst f) (merid a)) + (cong (fst g) (merid a))) + help = (cong₂ _∙_ (cong (sym (snd f) ∙∙_∙∙ snd f) + (cong-∙ (fst f) + (merid a) (sym (merid (ptSn _))))) + ((cong (sym (snd g) ∙∙_∙∙ snd g) + (cong-∙ (fst g) + (merid a) (sym (merid (ptSn _)))))) + ◁ lazy-help _ _ _ _ (sym (snd f)) (cong (fst f) (merid (ptSn _))) + (sym (snd g)) (cong (fst g) (merid (ptSn _))) + (cong (fst f) (merid a)) (cong (fst g) (merid a))) + where + lazy-help : + ∀ {ℓ} {A : Type ℓ} {x : A} + (p-north p-south q-north q-south : x ≡ x) + (prefl : refl ≡ p-north) (p-merid : p-north ≡ p-south) + (qrefl : refl ≡ q-north) (q-merid : q-north ≡ q-south) + (fl : p-north ≡ p-south) (fr : q-north ≡ q-south) + → PathP (λ i → (rUnit refl ∙ cong₂ _∙_ prefl qrefl) i + ≡ ((rUnit refl ∙ cong₂ _∙_ prefl qrefl) + ∙ cong₂ _∙_ p-merid q-merid) i) + ((prefl ∙∙ fl ∙ sym p-merid ∙∙ sym prefl) + ∙ (qrefl ∙∙ fr ∙ sym q-merid ∙∙ sym qrefl)) + (cong₂ _∙_ fl fr) + lazy-help {x = x} = + J4 (λ p-north p-south q-north q-south prefl p-merid qrefl q-merid + → (fl : p-north ≡ p-south) (fr : q-north ≡ q-south) → + PathP (λ i₁ → (rUnit refl ∙ cong₂ _∙_ prefl qrefl) i₁ ≡ + ((rUnit refl ∙ cong₂ _∙_ prefl qrefl) + ∙ cong₂ _∙_ p-merid q-merid) i₁) + ((prefl ∙∙ fl ∙ sym p-merid ∙∙ sym prefl) ∙ + (qrefl ∙∙ fr ∙ sym q-merid ∙∙ sym qrefl)) + (cong₂ _∙_ fl fr)) + λ fl fr → (cong₂ _∙_ (sym (rUnit (fl ∙ refl)) ∙ sym (rUnit fl)) + (sym (rUnit (fr ∙ refl)) ∙ sym (rUnit fr)) + ◁ flipSquare ((sym (rUnit _) + ◁ flipSquare + ((λ j → cong (λ x → rUnit x j) fl + ∙ cong (λ x → lUnit x j) fr) + ▷ sym (cong₂Funct _∙_ fl fr))) + ▷ (rUnit _ ∙ rUnit _))) + main : (x : fst (S₊∙ (suc (suc n)))) → + fst (suspension→ n (∙Π f g)) x ≡ + fst (∙Π (suspension→ n f) (suspension→ n g)) x + main north = refl + main south = refl + main (merid a i₁) j = help j i₁ + where + help : fst (∙Π f g) a + ≡ cong (fst (∙Π (suspension→ n f) (suspension→ n g))) (merid a) + help = looper-help n f g a + ∙ cong₂ _∙_ (sym (cong-∙ (fst (suspension→ n f)) + (merid a) + (sym (merid (ptSn _))) + ∙∙ cong (fst f a ∙_) (cong sym (snd f)) + ∙∙ sym (rUnit _))) + (sym (cong-∙ (fst (suspension→ n g)) + (merid a) + (sym (merid (ptSn _))) + ∙∙ cong (fst g a ∙_) (cong sym (snd g)) + ∙∙ sym (rUnit _))) + ∙ λ i → rUnit (cong (fst (suspension→ n f)) + (merid a ∙ sym (merid (ptSn _)))) i + ∙ rUnit (cong (fst (suspension→ n g)) + (merid a ∙ sym (merid (ptSn _)))) i + + suspension←filler : (n : ℕ) → (S₊∙ (suc (suc n)) →∙ A) + → I → I → I → fst A + suspension←filler n f i j k = + hfill (λ k → λ { (i = i0) → doubleCompPath-filler + (sym (snd f)) + (cong (fst f) (merid (ptSn _) + ∙ sym (merid (ptSn _)))) + (snd f) k j + ; (i = i1) → snd f k + ; (j = i0) → snd f k + ; (j = i1) → snd f k}) + (inS ((fst f) (rCancel' (merid (ptSn _)) i j))) + k + + suspension← : (n : ℕ) → (S₊∙ (suc (suc n)) →∙ A) → (S₊∙ (suc n) →∙ Ω A) + fst (suspension← n f) a = + sym (snd f) ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f + snd (suspension← n f) i j = suspension←filler n f i j i1 + + suspensionIso : (n : ℕ) → Iso (S₊∙ (suc n) →∙ Ω A) (S₊∙ (suc (suc n)) →∙ A) + Iso.fun (suspensionIso n) f = suspension→ n f + Iso.inv (suspensionIso n) f = suspension← n f + Iso.rightInv (suspensionIso n) (f , p) = + ΣPathP (funExt help , λ i j → p (~ i ∨ j)) + where + help : (x : _) → fst (suspension→ n (suspension← n (f , p))) x ≡ f x + help north = sym p + help south = sym p ∙ cong f (merid (ptSn _)) + help (merid a j) i = + hcomp (λ k → λ { (i = i1) → f (merid a j) + ; (j = i0) → p (~ i ∧ k) + ; (j = i1) → compPath-filler' (sym p) + (cong f (merid (ptSn _))) k i}) + (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) + Iso.leftInv (suspensionIso n) f = + ΣPathP (funExt lem , PathP-filler) + where + + side₁ : (x : S₊ (suc n)) → I → I → I → fst A + side₁ x i j k = + hfill (λ k → λ { (i = i0) → suspension→ n f .fst + (compPath-filler' (merid x) + (sym (merid (ptSn _))) k j) + ; (i = i1) → snd A + ; (j = i0) → fst f x (i ∨ ~ k) + ; (j = i1) → snd A}) + (inS (snd f i (~ j))) + k + + side₂ : (x : S₊ (suc n)) → I → I → I → fst A + side₂ x i j k = + hfill (λ k → λ { (i = i0) → doubleCompPath-filler + refl + (cong (suspension→ n f .fst) + (merid x ∙ sym (merid (ptSn _)))) + refl k j + ; (i = i1) → fst f x (j ∨ ~ k) + ; (j = i0) → fst f x (i ∧ ~ k) + ; (j = i1) → snd A}) + (inS (side₁ x i j i1)) + k + + compPath-filler'' : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) + → I → I → I → A + compPath-filler'' {z = z} p q k j i = + hfill (λ k → λ { (i = i0) → p (~ j) + ; (i = i1) → q k + ; (j = i0) → q (i ∧ k) }) + (inS (p (i ∨ ~ j))) + k + + lem : (x : _) → fst (suspension← n (suspension→ n f)) x ≡ fst f x + lem x i j = side₂ x i j i1 + + sideCube : Cube (λ j k → snd f j (~ k)) + (λ j k → fst (suspension→ n f) + (rCancel' (merid (ptSn _)) j k)) + (λ r k → suspension→ n f .fst + (compPath-filler' (merid (ptSn _)) + (sym (merid (ptSn _))) r k)) + (λ _ _ → pt A) + (λ r j → snd f j (~ r)) λ _ _ → pt A + sideCube r j k = + hcomp (λ i → λ {(r = i0) → snd f j (~ k ∨ ~ i) + ; (r = i1) → fst (suspension→ n f) + (rCancel-filler' (merid (ptSn _)) (~ i) j k) + ; (j = i0) → suspension→ n f .fst + (compPath-filler'' (merid (ptSn _)) + (sym (merid (ptSn _))) i r k) + ; (j = i1) → snd f (~ r) (~ i ∧ k) + ; (k = i0) → snd f j (~ r) + ; (k = i1) → snd f (~ r ∧ j) (~ i)}) + (hcomp (λ i → λ {(r = i0) → pt A + ; (r = i1) → snd f (~ i) k + ; (j = i0) → snd f (~ i) (k ∨ ~ r) + ; (j = i1) → snd f (~ r ∨ ~ i) k + ; (k = i0) → snd f (j ∨ ~ i) (~ r) + ; (k = i1) → pt A}) + (pt A)) + + PathP-filler : PathP _ _ _ + PathP-filler i j k = + hcomp (λ r → λ {(i = i0) → suspension←filler n (suspension→ n f) j k r + ; (i = i1) → snd f j (k ∨ ~ r) + ; (j = i0) → side₂ (ptSn _) i k r + ; (j = i1) → snd A + ; (k = i0) → snd f j (i ∧ ~ r) + ; (k = i1) → snd A}) + (hcomp ((λ r → λ {(i = i0) → sideCube r j k + ; (i = i1) → pt A + ; (j = i0) → side₁ (ptSn _) i k r + ; (j = i1) → pt A + ; (k = i0) → snd f j (i ∨ ~ r) + ; (k = i1) → pt A})) + (snd f (i ∨ j) (~ k))) + + suspension←-pres∙ : (n : ℕ) → (f g : S₊∙ (suc (suc n)) →∙ A) + → suspension← n (∙Π f g) + ≡ ∙Π (suspension← n f) (suspension← n g) + suspension←-pres∙ n f g = + sym (Iso.leftInv (suspensionIso n) (suspension← n (∙Π f g))) + ∙∙ cong (suspension← n) + (Iso.rightInv (suspensionIso n) (∙Π f g) + ∙∙ cong₂ ∙Π (sym (Iso.rightInv (suspensionIso n) f)) + (sym (Iso.rightInv (suspensionIso n) g)) + ∙∙ sym (suspension-pres∙ _ (suspension← n f) (suspension← n g))) + ∙∙ (Iso.leftInv (suspensionIso n) (∙Π (suspension← n f) (suspension← n g))) + +SphereMap→Ω : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (S₊∙ (suc n) →∙ A) → (fst ((Ω^ (suc n)) A)) +SphereMap→Ω zero (f , p) = sym p ∙∙ cong f loop ∙∙ p +SphereMap→Ω {A = A} (suc n) (f , p) = + Iso.inv (flipΩIso _) + (SphereMap→Ω {A = Ω A} n (Iso.inv (suspensionIso _) (f , p))) + +SphereMap→Ωpres∙ : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (p q : S₊∙ (suc n) →∙ A) + → SphereMap→Ω n (∙Π p q) ≡ SphereMap→Ω n p ∙ SphereMap→Ω n q +SphereMap→Ωpres∙ zero f g = sym (rUnit _) +SphereMap→Ωpres∙ (suc n) f g = + cong (Iso.inv (flipΩIso (suc n))) + (cong (SphereMap→Ω n) (suspension←-pres∙ n f g) + ∙ SphereMap→Ωpres∙ n (suspension← n f) (suspension← n g)) + ∙ flipΩIsopres· n _ _ + +Ω→SphereMap : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (fst ((Ω^ (suc n)) A)) → (S₊∙ (suc n) →∙ A) +fst (Ω→SphereMap {A = A} zero f) base = pt A +fst (Ω→SphereMap {A = A} zero f) (loop i₁) = f i₁ +snd (Ω→SphereMap {A = A} zero f) = refl +Ω→SphereMap {A = A} (suc n) f = + Iso.fun (suspensionIso _) (Ω→SphereMap n (Iso.fun (flipΩIso _) f)) + +Ω→SphereMap→Ω : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (x : fst ((Ω^ (suc n)) A)) + → SphereMap→Ω n (Ω→SphereMap n x) ≡ x +Ω→SphereMap→Ω zero x = sym (rUnit _) +Ω→SphereMap→Ω (suc n) x = + cong (Iso.inv (flipΩIso (suc n)) ∘ SphereMap→Ω n) + (Iso.leftInv (suspensionIso _) + (Ω→SphereMap n (Iso.fun (flipΩIso (suc n)) x))) + ∙∙ cong (Iso.inv (flipΩIso (suc n))) + (Ω→SphereMap→Ω n (Iso.fun (flipΩIso (suc n)) x)) + ∙∙ Iso.leftInv (flipΩIso (suc n)) x + +SphereMap→Ω→SphereMap : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (x : (S₊∙ (suc n) →∙ A)) + → Ω→SphereMap n (SphereMap→Ω n x) ≡ x +SphereMap→Ω→SphereMap zero (f , p) = + ΣPathP (funExt fstp , λ i j → p (~ i ∨ j)) + where + fstp : (x : _) → _ ≡ f x + fstp base = sym p + fstp (loop i) j = doubleCompPath-filler (sym p) (cong f loop) p (~ j) i +SphereMap→Ω→SphereMap (suc n) f = + cong (Iso.fun (suspensionIso n) ∘ Ω→SphereMap n) + (Iso.rightInv (flipΩIso (suc n)) + (SphereMap→Ω n (Iso.inv (suspensionIso n) f))) + ∙∙ cong (Iso.fun (suspensionIso n)) + (SphereMap→Ω→SphereMap n (Iso.inv (suspensionIso n) f)) + ∙∙ Iso.rightInv (suspensionIso _) _ + + +-- We finally arrive at the main result +IsoSphereMapΩ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (S₊∙ n →∙ A) (fst ((Ω^ n) A)) +Iso.fun (IsoSphereMapΩ {A = A} zero) (f , p) = f false +fst (Iso.inv (IsoSphereMapΩ {A = A} zero) a) false = a +fst (Iso.inv (IsoSphereMapΩ {A = A} zero) a) true = pt A +snd (Iso.inv (IsoSphereMapΩ {A = A} zero) a) = refl +Iso.rightInv (IsoSphereMapΩ {A = A} zero) a = refl +Iso.leftInv (IsoSphereMapΩ {A = A} zero) (f , p) = + ΣPathP ((funExt (λ { false → refl ; true → sym p})) + , λ i j → p (~ i ∨ j)) +Iso.fun (IsoSphereMapΩ {A = A} (suc n)) = SphereMap→Ω n +Iso.inv (IsoSphereMapΩ {A = A} (suc n)) = Ω→SphereMap n +Iso.rightInv (IsoSphereMapΩ {A = A} (suc n)) = Ω→SphereMap→Ω n +Iso.leftInv (IsoSphereMapΩ {A = A} (suc n)) = SphereMap→Ω→SphereMap n + +-- The iso is structure preserving +IsoSphereMapΩ-pres∙Π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f g : S₊∙ (suc n) →∙ A) + → SphereMap→Ω n (∙Π f g) ≡ SphereMap→Ω n f ∙ SphereMap→Ω n g +IsoSphereMapΩ-pres∙Π zero f g = sym (rUnit _) +IsoSphereMapΩ-pres∙Π (suc n) f g = + cong (Iso.inv (flipΩIso (suc n))) + (cong (SphereMap→Ω n) (suspension←-pres∙ n f g) + ∙ SphereMap→Ωpres∙ n (suspension← n f) (suspension← n g)) + ∙ (flipΩIsopres· n (SphereMap→Ω n (suspension← n f)) + (SphereMap→Ω n (suspension← n g))) + +-- It is useful to define the ``Group Structure'' on (S₊∙ n →∙ A) +-- before doing it on π'. These will be the equivalents of the +-- usual groupoid laws on Ω A. +1Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → (S₊∙ n →∙ A) +fst (1Π {A = A}) _ = pt A +snd (1Π {A = A}) = refl + +∙Π-rUnit : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ∙Π f 1Π ≡ f +fst (∙Π-rUnit {A = A} {n = zero} f i) base = snd f (~ i) +fst (∙Π-rUnit {A = A} {n = zero} f i) (loop j) = help i j + where + help : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) + (((sym (snd f)) ∙∙ (cong (fst f) loop) ∙∙ snd f) + ∙ (refl ∙ refl)) + (cong (fst f) loop) + help = (cong ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) ∙_) + (sym (rUnit refl)) ∙ sym (rUnit _)) + ◁ λ i j → doubleCompPath-filler (sym (snd f)) + (cong (fst f) loop) (snd f) (~ i) j +snd (∙Π-rUnit {A = A} {n = zero} f i) j = snd f (~ i ∨ j) +fst (∙Π-rUnit {A = A} {n = suc n} f i) north = snd f (~ i) +fst (∙Π-rUnit {A = A} {n = suc n} f i) south = + (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i +fst (∙Π-rUnit {A = A} {n = suc n} f i) (merid a j) = help i j + where + help : PathP (λ i → snd f (~ i) + ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) + (((sym (snd f)) + ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + ∙∙ snd f) + ∙ (refl ∙ refl)) + (cong (fst f) (merid a)) + help = (cong (((sym (snd f)) + ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + ∙∙ snd f) ∙_) + (sym (rUnit refl)) + ∙ sym (rUnit _)) + ◁ λ i j → hcomp (λ k → + λ { (j = i0) → snd f (~ i ∧ k) + ; (j = i1) → compPath-filler' (sym (snd f)) + (cong (fst f) (merid (ptSn (suc n)))) k i + ; (i = i0) → doubleCompPath-filler (sym (snd f)) + (cong (fst f) + (merid a ∙ sym (merid (ptSn (suc n))))) + (snd f) k j + ; (i = i1) → fst f (merid a j)}) + (fst f (compPath-filler (merid a) + (sym (merid (ptSn _))) (~ i) j)) +snd (∙Π-rUnit {A = A} {n = suc n} f i) j = snd f (~ i ∨ j) + +∙Π-lUnit : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ∙Π 1Π f ≡ f +fst (∙Π-lUnit {n = zero} f i) base = snd f (~ i) +fst (∙Π-lUnit {n = zero} f i) (loop j) = s i j + where + s : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) + ((refl ∙ refl) ∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) + (cong (fst f) loop) + s = (cong (_∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) + (sym (rUnit refl)) ∙ sym (lUnit _)) + ◁ λ i j → doubleCompPath-filler (sym (snd f)) + (cong (fst f) loop) (snd f) (~ i) j +snd (∙Π-lUnit {n = zero} f i) j = snd f (~ i ∨ j) +fst (∙Π-lUnit {n = suc n} f i) north = snd f (~ i) +fst (∙Π-lUnit {n = suc n} f i) south = + (sym (snd f) ∙ cong (fst f) (merid (ptSn _))) i +fst (∙Π-lUnit {n = suc n} f i) (merid a j) = help i j + where + help : PathP (λ i → snd f (~ i) + ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) + ((refl ∙ refl) ∙ ((sym (snd f)) + ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + ∙∙ snd f)) + (cong (fst f) (merid a)) + help = + (cong (_∙ ((sym (snd f)) + ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + ∙∙ snd f)) + (sym (rUnit refl)) + ∙ sym (lUnit _)) + ◁ λ i j → hcomp (λ k → + λ { (j = i0) → snd f (~ i ∧ k) + ; (j = i1) → compPath-filler' (sym (snd f)) + (cong (fst f) (merid (ptSn (suc n)))) k i + ; (i = i0) → doubleCompPath-filler (sym (snd f)) + (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + (snd f) k j + ; (i = i1) → fst f (merid a j)}) + (fst f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) +snd (∙Π-lUnit {n = suc n} f i) j = snd f (~ i ∨ j) + +∙Π-rCancel : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ∙Π f (-Π f) ≡ 1Π +fst (∙Π-rCancel {A = A} {n = zero} f i) base = pt A +fst (∙Π-rCancel {A = A} {n = zero} f i) (loop j) = + rCancel (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) i j +snd (∙Π-rCancel {A = A} {n = zero} f i) = refl +fst (∙Π-rCancel {A = A} {n = suc n} f i) north = pt A +fst (∙Π-rCancel {A = A} {n = suc n} f i) south = pt A +fst (∙Π-rCancel {A = A} {n = suc n} f i) (merid a i₁) = sl i i₁ + where + pl = (sym (snd f) + ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) + ∙∙ snd f) + sl : pl + ∙ ((sym (snd f) + ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) + ∙∙ snd f)) ≡ refl + sl = cong (pl ∙_) (cong (sym (snd f) ∙∙_∙∙ (snd f)) + (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) + ∙∙ cong₂ _∙_ refl + (cong (cong (fst f)) (rCancel (merid (ptSn _)))) + ∙∙ sym (rUnit _))) + ∙ rCancel pl +snd (∙Π-rCancel {A = A} {n = suc n} f i) = refl + +∙Π-lCancel : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ∙Π (-Π f) f ≡ 1Π +fst (∙Π-lCancel {A = A} {n = zero} f i) base = pt A +fst (∙Π-lCancel {A = A} {n = zero} f i) (loop j) = + rCancel (sym (snd f) ∙∙ cong (fst f) (sym loop) ∙∙ snd f) i j +fst (∙Π-lCancel {A = A} {n = suc n} f i) north = pt A +fst (∙Π-lCancel {A = A} {n = suc n} f i) south = pt A +fst (∙Π-lCancel {A = A} {n = suc n} f i) (merid a j) = sl i j + where + pl = (sym (snd f) + ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) + ∙∙ snd f) + + sl : (sym (snd f) + ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) + ∙∙ snd f) ∙ pl + ≡ refl + sl = cong (_∙ pl) (cong (sym (snd f) ∙∙_∙∙ (snd f)) + (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) + ∙∙ cong₂ _∙_ refl (cong (cong (fst f)) (rCancel (merid (ptSn _)))) + ∙∙ sym (rUnit _))) + ∙ lCancel pl +snd (∙Π-lCancel {A = A} {n = zero} f i) = refl +snd (∙Π-lCancel {A = A} {n = suc n} f i) = refl + +∙Π-assoc : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f g h : S₊∙ (suc n) →∙ A) + → ∙Π f (∙Π g h) ≡ ∙Π (∙Π f g) h +∙Π-assoc {n = n} f g h = + sym (Iso.leftInv (IsoSphereMapΩ (suc n)) (∙Π f (∙Π g h))) + ∙∙ cong (Ω→SphereMap n) (SphereMap→Ωpres∙ n f (∙Π g h) + ∙∙ cong (SphereMap→Ω n f ∙_) (SphereMap→Ωpres∙ n g h) + ∙∙ ∙assoc (SphereMap→Ω n f) (SphereMap→Ω n g) (SphereMap→Ω n h) + ∙∙ cong (_∙ SphereMap→Ω n h) (sym (SphereMap→Ωpres∙ n f g)) + ∙∙ sym (SphereMap→Ωpres∙ n (∙Π f g) h)) + ∙∙ Iso.leftInv (IsoSphereMapΩ (suc n)) (∙Π (∙Π f g) h) + +∙Π-comm : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f g : S₊∙ (suc (suc n)) →∙ A) + → ∙Π f g ≡ ∙Π g f +∙Π-comm {A = A} {n = n} f g = + sym (Iso.leftInv (IsoSphereMapΩ (suc (suc n))) (∙Π f g)) + ∙∙ cong (Ω→SphereMap (suc n)) (SphereMap→Ωpres∙ (suc n) f g + ∙∙ EH _ _ _ + ∙∙ sym (SphereMap→Ωpres∙ (suc n) g f)) + ∙∙ Iso.leftInv (IsoSphereMapΩ (suc (suc n))) (∙Π g f) + + +{- π'' as a group -} +1π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' n A +1π' n {A = A} = ∣ 1Π ∣₂ + +·π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' (suc n) A → π' (suc n) A → π' (suc n) A +·π' n = sRec2 squash₂ λ p q → ∣ ∙Π p q ∣₂ + +-π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' (suc n) A → π' (suc n) A +-π' n = sMap -Π + +π'-rUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) + → (·π' n x (1π' (suc n))) ≡ x +π'-rUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-rUnit p i ∣₂ + +π'-lUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) + → (·π' n (1π' (suc n)) x) ≡ x +π'-lUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-lUnit p i ∣₂ + +π'-rCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) + → (·π' n x (-π' n x)) ≡ 1π' (suc n) +π'-rCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-rCancel p i ∣₂ + +π'-lCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) + → (·π' n (-π' n x) x) ≡ 1π' (suc n) +π'-lCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-lCancel p i ∣₂ + +π'-assoc : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y z : π' (suc n) A) + → ·π' n x (·π' n y z) ≡ ·π' n (·π' n x y) z +π'-assoc n = sElim3 (λ _ _ _ → isSetPathImplicit) λ p q r i → ∣ ∙Π-assoc p q r i ∣₂ + +π'-comm : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y : π' (suc (suc n)) A) + → ·π' (suc n) x y ≡ ·π' (suc n) y x +π'-comm n = sElim2 (λ _ _ → isSetPathImplicit) λ p q i → ∣ ∙Π-comm p q i ∣₂ + +π'Gr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Group ℓ +fst (π'Gr n A) = π' (suc n) A +1g (snd (π'Gr n A)) = 1π' (suc n) +GroupStr._·_ (snd (π'Gr n A)) = ·π' n +inv (snd (π'Gr n A)) = -π' n +is-set (isSemigroup (isMonoid (isGroup (snd (π'Gr n A))))) = squash₂ +assoc (isSemigroup (isMonoid (isGroup (snd (π'Gr n A))))) = π'-assoc n +identity (isMonoid (isGroup (snd (π'Gr n A)))) x = (π'-rUnit n x) , (π'-lUnit n x) +inverse (isGroup (snd (π'Gr n A))) x = (π'-rCancel n x) , (π'-lCancel n x) + +-- Finally, the Iso +π'Gr≅πGr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → GroupIso (π'Gr n A) (πGr n A) +fst (π'Gr≅πGr n A) = setTruncIso (IsoSphereMapΩ (suc n)) +snd (π'Gr≅πGr n A) = + makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) + λ p q i → ∣ IsoSphereMapΩ-pres∙Π n p q i ∣₂) diff --git a/Cubical/Homotopy/Group/S3.agda b/Cubical/Homotopy/Group/S3.agda new file mode 100644 index 000000000..8a4eb3db8 --- /dev/null +++ b/Cubical/Homotopy/Group/S3.agda @@ -0,0 +1,686 @@ +{-# OPTIONS --safe #-} +{- +This file contains +1. The definition of πₙ as a truncated loop space +2. The definition of πₙ as a truncated function space (Sⁿ →∙ A) +3. A structure preserving equivalence Ωⁿ A ≃ (Sⁿ →∙ A) +4. A proof that the two constructions of homotopy groups are isomorphic +-} +module Cubical.Homotopy.Group.S3 where + +open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.Loopspace + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Path +open import Cubical.Foundations.Transport +open import Cubical.Functions.Embedding +open import Cubical.Foundations.Pointed.Homogeneous + +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim + ; elim2 to sElim2 ; elim3 to sElim3 ; map to sMap) +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; rec2 to pRec2 ; elim to pElim) +open import Cubical.HITs.Sn +open import Cubical.Data.Bool +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 +open import Cubical.Data.Sigma +open import Cubical.Data.Nat + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid + +open Iso +open IsGroup +open IsSemigroup +open IsMonoid +open GroupStr + +open import Cubical.Homotopy.Freudenthal +open import Cubical.Homotopy.Connected +open import Cubical.Foundations.Equiv +open import Cubical.HITs.Truncation + renaming (rec to trRec ; elim to trElim ; map to trMap) + + +open import Cubical.Data.Sum +open import Cubical.Relation.Nullary +open import Cubical.Data.Empty +open import Cubical.Data.Bool + +∙∙lCancel-fill : ∀ {ℓ} {A : Type ℓ} {x y : A} + → (p : x ≡ y) + → I → I → I → A +∙∙lCancel-fill p i j k = + hfill (λ k → λ { (i = i1) → p k + ; (j = i0) → p k + ; (j = i1) → p k}) + (inS (p i0)) k + +∙∙lCancel : ∀ {ℓ} {A : Type ℓ} {x y : A} + → (p : x ≡ y) + → sym p ∙∙ refl ∙∙ p ≡ refl +∙∙lCancel p i j = ∙∙lCancel-fill p i j i1 + +open import Cubical.Foundations.Equiv.HalfAdjoint + +suspMap : ∀ {ℓ} {A : Pointed ℓ}(n : ℕ) + → S₊∙ (suc n) →∙ A + → S₊∙ (suc (suc n)) →∙ ∙Susp (typ A) +fst (suspMap n f) north = north +fst (suspMap n f) south = north +fst (suspMap {A = A} n f) (merid a i) = (merid (fst f a) ∙ sym (merid (pt A))) i +snd (suspMap n f) = refl + +suspMap2 : ∀ {ℓ} {A : Pointed ℓ}(n : ℕ) + → ((Ω^ n) A) + →∙ ((Ω^ (suc n)) (∙Susp (typ A))) +fst (suspMap2 {A = A} zero) a = merid a ∙ sym (merid (pt A)) +snd (suspMap2 {A = A} zero) = rCancel (merid (pt A)) +fst (suspMap2 {A = A} (suc n)) p = sym (snd (suspMap2 n)) ∙∙ cong (fst (suspMap2 n)) p ∙∙ snd (suspMap2 n) +snd (suspMap2 {A = A} (suc n)) = ∙∙lCancel (snd (suspMap2 n)) + +lMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → typ ((Ω^ n) A) → (S₊∙ n →∙ A) +lMapId : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (a : _) → lMap n {A = A} (pt ((Ω^ n) A)) .fst a ≡ pt A +lMapId2 : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → lMap n {A = A} (pt ((Ω^ n) A)) ≡ ((λ _ → pt A) , refl) +fst (lMap zero {A = A} a) false = a +fst (lMap zero {A = A} a) true = pt A +snd (lMap zero {A = A} a) = refl +fst (lMap (suc zero) {A = A} p) base = pt A +fst (lMap (suc zero) {A = A} p) (loop i) = p i +snd (lMap (suc zero) {A = A} p) = refl +fst (lMap (suc (suc n)) {A = A} p) north = pt A +fst (lMap (suc (suc n)) {A = A} p) south = pt A +fst (lMap (suc (suc n)) {A = A} p) (merid a i) = + (sym (lMapId (suc n) a) ∙∙ (λ i → lMap (suc n) (p i) .fst a) ∙∙ lMapId (suc n) a) i +snd (lMap (suc (suc n)) {A = A} p) = refl +lMapId zero false = refl +lMapId zero true = refl +lMapId (suc zero) base = refl +lMapId (suc zero) (loop i) = refl +lMapId (suc (suc n)) north = refl +lMapId (suc (suc n)) south = refl +lMapId (suc (suc n)) {A = A} (merid a i) j = (∙∙lCancel (lMapId (suc n) {A = A} a)) j i +lMapId2 zero {A = A} = ΣPathP ((funExt (λ { false → refl ; true → refl})) , refl) +lMapId2 (suc zero) {A = A} = ΣPathP ((funExt (λ { base → refl ; (loop i) → refl})) , refl) +lMapId2 (suc (suc n)) {A = A} = + ΣPathP ((funExt (λ { north → refl ; south → refl ; (merid a i) j → ∙∙lCancel (lMapId (suc n) {A = A} a) j i})) , refl) + +{- +Ω (Ωⁿ A) -----------> Ω (Ωⁿ⁺¹ ΣA) + | | + | | + | | + v v + Ω (Sⁿ →∙ A) ----> Ω (Sⁿ⁺¹ →∙ ΣA) + | | + | | + | | + v v + (Sⁿ⁺¹ →∙ A) ----> Sⁿ⁺² →∙ Σ A +-} + +Ω-fun : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + → (A →∙ B) → (Ω A →∙ Ω B) +fst (Ω-fun {A = A} {B = B} (f , p)) q = sym p ∙∙ cong f q ∙∙ p +snd (Ω-fun {A = A} {B = B} (f , p)) = ∙∙lCancel p + +lMapSplit₁ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → typ ((Ω^ (suc n)) A) + → typ (Ω (S₊∙ n →∙ A ∙)) +lMapSplit₁ n = Ω-fun (lMap n , lMapId2 n) .fst + + +-- sphereMapEquiv +ΩSphereMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → typ (Ω (S₊∙ n →∙ A ∙)) + → (S₊∙ (suc n) →∙ A) +fst (ΩSphereMap {A = A} zero p) base = p i0 .fst false +fst (ΩSphereMap {A = A} zero p) (loop i) = p i .fst false +snd (ΩSphereMap {A = A} zero p) = refl +fst (ΩSphereMap {A = A} (suc n) p) north = pt A +fst (ΩSphereMap {A = A} (suc n) p) south = pt A +fst (ΩSphereMap {A = A} (suc n) p) (merid a i) = p i .fst a +snd (ΩSphereMap {A = A} (suc n) p) = refl + +open import Cubical.Foundations.Equiv.HalfAdjoint +coolBeans : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → (f : (A →∙ B)) → isEquiv (fst f) → isEquiv (Ω-fun f .fst) +coolBeans {B = (B , b)} = + uncurry λ f → + J (λ b y → isEquiv f → isEquiv (λ q → (λ i → y (~ i)) ∙∙ (λ i → f (q i)) ∙∙ y)) + λ eqf → subst isEquiv (funExt (rUnit ∘ cong f)) (isoToIsEquiv (congIso (equivToIso (f , eqf)))) + +SphereMapΩ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (S₊∙ (suc n) →∙ A) + → typ (Ω (S₊∙ n →∙ A ∙)) +SphereMapΩ {A = A} zero (f , p) = ΣPathP ((funExt λ { false → sym p ∙∙ cong f loop ∙∙ p ; true → refl}) , refl) +SphereMapΩ {A = A} (suc n) (f , p) = + ΣPathP (funExt (λ x → sym p ∙∙ cong f (merid x ∙ sym (merid (ptSn _))) ∙∙ p) + , flipSquare (cong (sym p ∙∙_∙∙ p) (cong (cong f) (rCancel (merid (ptSn _)))) ∙ ∙∙lCancel p)) + +SphereMapΩIso : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (S₊∙ (suc n) →∙ A) + (typ (Ω (S₊∙ n →∙ A ∙))) +fun (SphereMapΩIso {A = A} n) = SphereMapΩ n +inv (SphereMapΩIso {A = A} n) = ΩSphereMap n +fst (rightInv (SphereMapΩIso {A = A} zero) p k i) false = rUnit (funExt⁻ (cong fst p) false) (~ k) i +fst (rightInv (SphereMapΩIso {A = A} zero) p k i) true = snd (p i) (~ k) +snd (rightInv (SphereMapΩIso {A = A} zero) p k i) l = snd (p i) (~ k ∨ l) +fst (rightInv (SphereMapΩIso {A = A} (suc n)) p k i) x = help k i + where + help : refl ∙∙ cong (fst (ΩSphereMap {A = A} (suc n) p)) (merid x ∙ sym (merid (ptSn _))) ∙∙ refl + ≡ funExt⁻ (cong fst p) x + help = (cong (refl ∙∙_∙∙ refl) (cong-∙ (fst (ΩSphereMap {A = A} (suc n) p)) (merid x) (sym (merid (ptSn _)))) + ∙ cong (refl ∙∙_∙∙ refl) (cong (funExt⁻ (cong fst p) x ∙_) (cong sym (flipSquare (cong snd p))) ∙ sym (rUnit _))) + ∙ sym (rUnit _) +snd (rightInv (SphereMapΩIso {A = A} (suc n)) p k i) j = + hcomp (λ r → λ { (i = i0) → pt A + ; (i = i1) → pt A + ; (j = i0) → ((cong (refl ∙∙_∙∙ refl) (cong-∙ (fst (ΩSphereMap {A = A} (suc n) p)) (merid (ptSn (suc n))) (sym (merid (ptSn _)))) + ∙ cong (refl ∙∙_∙∙ refl) (cong (funExt⁻ (cong fst p) (ptSn (suc n)) ∙_) (cong sym (flipSquare (cong snd p))) ∙ sym (rUnit _))) + ∙ sym (rUnit _)) k i + ; (j = i1) → pt A + ; (k = i0) → (cong (sym refl ∙∙_∙∙ refl) (cong (cong (fst (ΩSphereMap {A = A} (suc n) p))) (rCancel (merid (ptSn _)))) ∙ sym (rUnit refl)) j i + ; (k = i1) → snd (p i) j}) + (hcomp (λ r → λ { (i = i0) → pt A + ; (i = i1) → pt A + ; (j = i0) → ((cong (refl ∙∙_∙∙ refl) (cong-∙ (fst (ΩSphereMap {A = A} (suc n) p)) (merid (ptSn (suc n))) (sym (merid (ptSn _)))) + ∙ cong (refl ∙∙_∙∙ refl) (cong (funExt⁻ (cong fst p) (ptSn (suc n)) ∙_) (cong sym (flipSquare (cong snd p))) ∙ sym (rUnit _))) + ∙ sym (rUnit _)) k i + ; (j = i1) → rUnit (λ _ → pt A) (~ r ∧ ~ k) i + ; (k = i0) → compPath-filler (cong (sym refl ∙∙_∙∙ refl) (cong (cong (fst (ΩSphereMap {A = A} (suc n) p))) (rCancel (merid (ptSn _))))) + (sym (rUnit refl)) r j i + ; (k = i1) → snd (p i) j}) + ((hcomp (λ r → λ { (i = i0) → rUnit (λ _ → pt A) (~ k ∨ r) i + ; (i = i1) → rUnit (λ _ → pt A) (~ k ∨ r) i + ; (j = i0) → ((cong (λ x → rUnit x r) (cong-∙ (fst (ΩSphereMap {A = A} (suc n) p)) (merid (ptSn (suc n))) (sym (merid (ptSn _)))) + ∙ cong (λ x → rUnit x r) (cong (funExt⁻ (cong fst p) (ptSn (suc n)) ∙_) (cong sym (flipSquare (cong snd p))) ∙ sym (rUnit _))) + ∙ λ i → rUnit (λ i₁ → funExt⁻ (λ i₂ → fst (p i₂)) (ptSn (suc n)) i₁) (~ i ∧ r)) k i + ; (j = i1) → rUnit (λ _ → pt A) (~ k ∧ r) i + ; (k = i0) → rUnit ((cong (fst (ΩSphereMap {A = A} (suc n) p))) (rCancel (merid (ptSn _)) j)) r i + ; (k = i1) → snd (p i) j})) + (hcomp (λ r → λ { (i = i0) → pt A + ; (i = i1) → pt A + ; (j = i0) → ((rUnit (compPath-filler' (cong-∙ (fst (ΩSphereMap {A = A} (suc n) p)) (merid (ptSn (suc n))) (sym (merid (ptSn _)))) + ((cong (funExt⁻ (cong fst p) (ptSn (suc n)) ∙_) (cong sym (flipSquare (cong snd p))) ∙ sym (rUnit _))) r) r) k i) + ; (j = i1) → pt A + ; (k = i0) → help (fst (ΩSphereMap {A = A} (suc n) p)) (merid (ptSn (suc n))) r j i + ; (k = i1) → snd (p i) j}) + (cool2 (funExt⁻ (cong fst p) (ptSn _)) (sym (flipSquare (cong snd p))) k j i)))) + where + cong-∙∙-filler : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y z w : A} (f : A → B) (p : w ≡ x) (q : x ≡ y) (r : y ≡ z) → I → I → I → B + cong-∙∙-filler f p q r k j i = hfill (λ k → λ { (j = i0) → f (doubleCompPath-filler p q r k i) + ; (i = i0) → f (p (~ k)) + ; (i = i1) → f (r k) }) + (inS (f (q i))) k + + cool2 : ∀ {ℓ} {A : Pointed ℓ} → (p : Path (typ A) (pt A) (pt A)) + → (q : refl ≡ p) + → PathP (λ i → (cong (p ∙_) (cong sym (sym q)) ∙ sym (rUnit p)) i ≡ refl) + (rCancel p) (sym q) + cool2 {A = A} p = J (λ p q → PathP (λ i → (cong (p ∙_) (cong sym (sym q)) ∙ sym (rUnit p)) i ≡ refl) + (rCancel p) (sym q)) + (flipSquare (sym (lUnit (sym (rUnit refl))) + ◁ λ k i → rCancel (λ _ → pt A) (k ∨ i))) + + help : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A} (f : A → B) (p : x ≡ y) + → PathP (λ k → cong-∙ f p (sym p) (~ k) ≡ refl) (rCancel (cong f p)) (cong (cong f) (rCancel p)) + help f p i j k = hcomp (λ r → λ { (i = i0) → rCancel-filler (cong f p) r j k + ; (i = i1) → f (rCancel-filler p r j k) + ; (j = i0) → cong-∙∙-filler f refl p (sym p) r (~ i) k + ; (j = i1) → f (p i0) + ; (k = i0) → f (p i0) + ; (k = i1) → f (p (~ r ∧ ~ j))}) + (f (p (~ j ∧ k))) + +leftInv (SphereMapΩIso {A = A} zero) (f , p) = ΣPathP ((funExt (λ { base → sym p + ; (loop i) j → doubleCompPath-filler (sym p) (cong f loop) p (~ j) i})) + , λ i j → p (j ∨ ~ i)) +leftInv (SphereMapΩIso {A = A} (suc n)) (f , p) = + ΣPathP (funExt (λ { north → sym p + ; south → sym p ∙ cong f (merid (ptSn (suc n))) + ; (merid a i) j → help a j i}) + , λ i j → p (j ∨ ~ i)) + where + help : (a : S₊ (suc n)) → PathP (λ i → p (~ i) ≡ (sym p ∙ cong f (merid (ptSn (suc n)))) i) + (sym p ∙∙ cong f (merid a ∙ sym (merid (ptSn _))) ∙∙ p) + (cong f (merid a)) + help a i j = hcomp (λ k → λ { (i = i0) → doubleCompPath-filler (sym p) (cong f (merid a ∙ sym (merid (ptSn _)))) p k j + ; (i = i1) → f (merid a j) + ; (j = i0) → p (~ i ∧ k) + ; (j = i1) → compPath-filler' (sym p) (cong f (merid (ptSn (suc n)))) k i}) + (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) + + +cool̂ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p : typ ((Ω^ (suc n)) A)) → lMap (suc n) p ≡ ΩSphereMap n (lMapSplit₁ n p) +cool̂ {A = A} zero p = ΣPathP ((funExt (λ { base → refl ; (loop i) j → h (~ j) i})) , refl) + where + h : funExt⁻ (cong fst (lMapSplit₁ zero p)) false ≡ p + h = (λ i → funExt⁻ (cong-∙∙ fst (sym (lMapId2 zero)) (cong (lMap zero) p) (lMapId2 zero) i) false) + ∙ sym (rUnit _) +cool̂ {A = A} (suc n) p = ΣPathP ((funExt (λ { north → refl ; south → refl ; (merid a i) j → h a j i})) , refl) + where + lem : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (a : S₊ (suc n)) → lMapId (suc n) {A = A} a ≡ (λ i → fst (lMapId2 (suc n) {A = A} i) a) + lem zero base = refl + lem zero (loop i) = refl + lem (suc n) north = refl + lem (suc n) south = refl + lem (suc n) (merid a i) = refl + + h : (a : S₊ (suc n)) → ((λ i₁ → lMapId (suc n) {A = A} a (~ i₁)) ∙∙ + (λ i₁ → lMap (suc n) (p i₁) .fst a) ∙∙ lMapId (suc n) a) ≡ (λ i → lMapSplit₁ (suc n) p i .fst a) + h a = cong (λ x → sym x ∙∙ funExt⁻ (cong fst (λ i → lMap (suc n) (p i))) a ∙∙ x) + (lem n a) + ∙∙ sym (cong-∙∙ (λ x → x a) (cong fst (λ i → lMapId2 (suc n) (~ i))) (cong fst (λ i → lMap (suc n) (p i))) (cong fst (lMapId2 (suc n)))) + ∙∙ (λ i → funExt⁻ (cong-∙∙ fst (sym (lMapId2 (suc n))) (cong (lMap (suc n)) p) (lMapId2 (suc n)) (~ i)) a) + +isEquiv-lMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → isEquiv (lMap (suc n) {A = A}) +isEquiv-lMap zero {A = A} = isoToIsEquiv (iso _ invFun sec λ p → sym (rUnit p)) + where + invFun : S₊∙ 1 →∙ A → typ (Ω A) + invFun (f , p) = sym p ∙∙ cong f loop ∙∙ p + + sec : section (lMap 1) invFun + sec (f , p) = ΣPathP ((funExt (λ { base → sym p ; (loop i) j → doubleCompPath-filler (sym p) (cong f loop) p (~ j) i})) + , λ i j → p (~ i ∨ j)) + +isEquiv-lMap (suc n) = + subst isEquiv (sym (funExt (cool̂ (suc n)))) + (snd (compEquiv ((lMapSplit₁ (suc n)) , (coolBeans (lMap (suc n) , lMapId2 (suc n)) (isEquiv-lMap n))) + (invEquiv (isoToEquiv (SphereMapΩIso (suc n)))))) + +isHom-lMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → (p q : _) → lMap (suc n) {A = A} (p ∙ q) ≡ ∙Π (lMap (suc n) p) (lMap (suc n) q) +isHom-lMap zero p q = ΣPathP ((funExt (λ { base → refl ; (loop i) j → (rUnit p j ∙ rUnit q j) i})) , refl) +isHom-lMap (suc n) {A = A} p q = ΣPathP ((funExt (λ { north → refl ; south → refl ; (merid a i) j → h a j i})) , refl) + where + doubleComp-lem : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) (q r : y ≡ y) + → (p ∙∙ q ∙∙ sym p) ∙ (p ∙∙ r ∙∙ sym p) ≡ (p ∙∙ (q ∙ r) ∙∙ sym p) + doubleComp-lem p q r i j = + hcomp (λ k → λ { (i = i0) → (doubleCompPath-filler p q (sym p) k + ∙ doubleCompPath-filler p r (sym p) k) j + ; (i = i1) → doubleCompPath-filler p (q ∙ r) (sym p) k j + ; (j = i0) → p (~ k) + ; (j = i1) → p (~ k)}) + ((q ∙ r) j) + + help2 : (p : typ ((Ω^ (suc (suc n))) A)) → cong (fst (lMap (suc (suc n)) p)) (merid (ptSn _)) ≡ refl + help2 p = cong (sym (lMapId (suc n) (ptSn _)) ∙∙_∙∙ lMapId (suc n) (ptSn _)) + (rUnit _ ∙ (λ j → (λ i → lMap (suc n) {A = A} refl .snd (i ∧ j)) + ∙∙ (λ i → lMap (suc n) {A = A} (p i) .snd j) + ∙∙ λ i → lMap (suc n) {A = A} refl .snd (~ i ∧ j)) + ∙ ∙∙lCancel _) + ∙ ∙∙lCancel _ + + h : (a : S₊ (suc n)) → sym (lMapId (suc n) a) ∙∙ funExt⁻ (cong fst (cong (lMap (suc n)) (p ∙ q))) a ∙∙ lMapId (suc n) a + ≡ cong (fst (∙Π (lMap (suc (suc n)) p) (lMap (suc (suc n)) q))) (merid a) + h a = (cong (sym (lMapId (suc n) a) ∙∙_∙∙ (lMapId (suc n) a)) + (cong-∙ (λ x → lMap (suc n) x .fst a) p q) + ∙ sym (doubleComp-lem (sym (lMapId (suc n) a)) _ _)) + ∙∙ cong₂ _∙_ (sym (cong (cong (fst (lMap (suc (suc n)) p)) (merid a) ∙_) (cong sym (help2 p)) ∙ sym (rUnit _))) + (sym (cong (cong (fst (lMap (suc (suc n)) q)) (merid a) ∙_) (cong sym (help2 q)) ∙ sym (rUnit _))) + ∙∙ λ i → (rUnit (cong-∙ (fst (lMap (suc (suc n)) p)) (merid a) (sym (merid (ptSn _))) (~ i)) i) + ∙ (rUnit (cong-∙ (fst (lMap (suc (suc n)) q)) (merid a) (sym (merid (ptSn _)))(~ i)) i) + +botMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} + → (S₊∙ n →∙ A) + → S₊∙ n →∙ Ω (∙Susp (typ A)) +fst (botMap n {A = A} f) x = merid (fst f x) ∙ sym (merid (pt A)) +snd (botMap n {A = A} f) = cong (_∙ merid (pt A) ⁻¹) (cong merid (snd f)) + ∙ rCancel (merid (pt A)) + +rMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → typ ((Ω^ (suc n)) (∙Susp (typ A))) → (S₊∙ n →∙ Ω (∙Susp (typ A))) +rMap n = lMap n ∘ Iso.fun (flipΩIso n) + +rMap1 : ∀ {ℓ} {A : Pointed ℓ} → typ ((Ω^ (suc 1)) (∙Susp (typ A))) → (S₊∙ 1 →∙ Ω (∙Susp (typ A))) +rMap1 = lMap 1 + +rMap≡rMap1 : ∀ {ℓ} {A : Pointed ℓ} → rMap 1 {A = A} ≡ rMap1 {A = A} +rMap≡rMap1 = funExt λ x → cong (lMap 1) (transportRefl x) + +flipΩrefl : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → fun (flipΩIso {A = A} (suc n)) refl ≡ refl +flipΩrefl {A = A} n j = transp (λ i₁ → fst (Ω (flipΩPath {A = A} n ((i₁ ∨ j))))) j (snd (Ω (flipΩPath n j))) + +cong-lMap-main' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p : _) + → cong (lMap (suc n) {A = Ω A}) (fun (flipΩIso (suc (suc n))) p) + ≡ (cong (lMap (suc n)) (sym (flipΩrefl n)) + ∙∙ cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) p + ∙∙ cong (lMap (suc n)) (flipΩrefl n)) +cong-lMap-main' {A = A} n p = (λ i → cong (lMap (suc n) {A = Ω A}) (fun (flipΩIso (suc (suc n))) (rUnit p i))) + ∙ (λ i j → lMap (suc n) {A = Ω A} + (transp (λ i₁ → fst (flipΩPath {A = A} (suc (suc n)) (i₁ ∨ i))) i + (((λ j → transp (λ i₁ → fst (Ω (flipΩPath {A = A} n ((i₁ ∨ ~ j) ∧ i)))) (~ i ∨ ~ j) (snd (Ω (flipΩPath n (~ j ∧ i))))) + ∙∙ (λ j → transp (λ i₁ → fst (flipΩPath {A = A} (suc n) (i₁ ∧ i))) (~ i) (p j)) + ∙∙ λ j → transp (λ i₁ → fst (Ω (flipΩPath {A = A} n ((i₁ ∨ j) ∧ i)))) (~ i ∨ j) (snd (Ω (flipΩPath n (j ∧ i)))))) j)) + ∙ cong-∙∙ (lMap (suc n)) (sym (flipΩrefl n)) (cong (fun (flipΩIso (suc n))) p) (flipΩrefl n) + ++nInd : ∀ {ℓ} {P : ℕ → Type ℓ} + → P 0 + → P 1 + → ((n : ℕ) → P (suc n) → P (suc (suc n))) + → (n : ℕ) → P n ++nInd 0c 1c indc zero = 0c ++nInd 0c 1c indc (suc zero) = 1c ++nInd {P = P} 0c 1c indc (suc (suc n)) = indc n (+nInd {P = P} 0c 1c indc (suc n)) + +{- suspMap +((Ω^ n) A) ------------> Ω^ (1 + n) (Susp A) + | | + | | + lMap | ≃ | rMap + v v +(S₊∙ n →∙ A) ----------> (S₊∙ n →∙ Ω (Susp A)) + \ botMap / + \ / + \ / + \ / + \ / + \ / + (Sⁿ⁺¹ →∙ Susp A) +-} + + +IsoΩSphereMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → Iso (typ ((Ω^ n) A)) ((S₊∙ n →∙ A)) +fun (IsoΩSphereMap zero) = lMap zero +inv (IsoΩSphereMap zero) f = fst f false +rightInv (IsoΩSphereMap zero) f = ΣPathP ((funExt (λ { false → refl ; true → sym (snd f)})) + , λ i j → snd f (~ i ∨ j)) +leftInv (IsoΩSphereMap zero) p = refl +IsoΩSphereMap (suc n) = equivToIso (_ , isEquiv-lMap n) + +Ω≡SphereMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → (typ ((Ω^ n) A)) ≡ ((S₊∙ n →∙ A)) +Ω≡SphereMap n = isoToPath (IsoΩSphereMap n) + +botₗ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → (S₊∙ n →∙ A) → S₊∙ (suc n) →∙ (∙Susp (typ A)) +fst (botₗ zero f) base = north +fst (botₗ {A = A} zero f) (loop i) = (merid (fst f false) ∙ sym (merid (pt A))) i +snd (botₗ zero f) = refl +botₗ (suc n) f = suspMap n f + + +botᵣ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → (S₊∙ n →∙ Ω (∙Susp (typ A))) → S₊∙ (suc n) →∙ ∙Susp (typ A) +fst (botᵣ zero (f , p)) base = north +fst (botᵣ zero (f , p)) (loop i) = f false i +snd (botᵣ zero (f , p)) = refl +fst (botᵣ (suc n) (f , p)) north = north +fst (botᵣ (suc n) (f , p)) south = north +fst (botᵣ (suc n) (f , p)) (merid a i) = f a i +snd (botᵣ (suc n) (f , p)) = refl + +finalΔ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (x : (S₊∙ n →∙ A)) + → botₗ n x ≡ botᵣ {A = A} n (botMap n x) +finalΔ zero (f , p) = + ΣPathP ((funExt (λ { base → refl ; (loop i) → refl})) , refl) +finalΔ (suc n) (f , p) = + ΣPathP ((funExt (λ { north → refl ; south → refl ; (merid a i) → refl})) , refl) + +botᵣ-inv' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → S₊∙ (suc n) →∙ ∙Susp (typ A) → (S₊ n → typ (Ω (∙Susp (typ A)))) +botᵣ-inv' zero f false = sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f +botᵣ-inv' zero f true = refl +botᵣ-inv' (suc n) f x = sym (snd f) ∙∙ cong (fst f) (merid x ∙ sym (merid (ptSn (suc n)))) ∙∙ snd f + +botᵣ-inv : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → S₊∙ (suc n) →∙ ∙Susp (typ A) → (S₊∙ n →∙ Ω (∙Susp (typ A))) +fst (botᵣ-inv {A = A} n f) = botᵣ-inv' {A = A} n f +snd (botᵣ-inv {A = A} zero f) = refl +snd (botᵣ-inv {A = A} (suc n) f) = + cong (sym (snd f) ∙∙_∙∙ snd f) (cong (cong (fst f)) (rCancel (merid (ptSn _)))) ∙ ∙∙lCancel (snd f) + + +isEquiv-rMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → isEquiv (rMap n {A = A}) +isEquiv-rMap zero = + compEquiv (isoToEquiv (flipΩIso zero)) (isoToEquiv (IsoΩSphereMap zero)) .snd +isEquiv-rMap (suc n) = + compEquiv (isoToEquiv (flipΩIso (suc n))) (isoToEquiv (IsoΩSphereMap (suc n))) .snd + + +IsoΩSphereMap' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (typ ((Ω^ (suc n)) (∙Susp (typ A)))) ((S₊∙ (suc n) →∙ ∙Susp (typ A))) +IsoΩSphereMap' {A = A} n = + compIso (equivToIso (_ , isEquiv-rMap {A = A} n)) + (iso (botᵣ {A = A} n) (botᵣ-inv {A = A} n) (h n) (retr n)) + where + h : (n : ℕ) → section (botᵣ {A = A} n) (botᵣ-inv {A = A} n) + h zero (f , p) = + ΣPathP (funExt (λ { base → sym p + ; (loop i) j → doubleCompPath-filler (sym p) (cong f loop) p (~ j) i}) + , λ i j → p (~ i ∨ j)) + h (suc n) (f , p) = + ΣPathP (funExt (λ { north → sym p + ; south → sym p ∙ cong f (merid (ptSn _)) + ; (merid a i) j → hcomp (λ k → λ { (i = i0) → p (~ j ∧ k) + ; (i = i1) → compPath-filler' (sym p) (cong f (merid (ptSn _))) k j + ; (j = i1) → f (merid a i)}) + (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ j) i))}) + , λ i j → p (~ i ∨ j)) + + retr : (n : ℕ) → retract (botᵣ {A = A} n) (botᵣ-inv {A = A} n) + retr zero (f , p) = ΣPathP ((funExt (λ { false → sym (rUnit _) ; true → sym p})) + , λ i j → p (~ i ∨ j)) + retr (suc n) (f , p) = + →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt λ x → (λ i → rUnit (cong-∙ (fst ((botᵣ {A = A}(suc n) (f , p)))) (merid x) (sym (merid (ptSn (suc n)))) i) (~ i)) + ∙∙ (λ i → f x ∙ sym (p i) ) + ∙∙ sym (rUnit (f x))) + +Ω≡SphereMap' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → (typ ((Ω^ (suc n)) (∙Susp (typ A)))) ≡ ((S₊∙ (suc n) →∙ ∙Susp (typ A))) +Ω≡SphereMap' {A = A} n = isoToPath (IsoΩSphereMap' {A = A} n) + +open import Cubical.Data.Bool +theComp' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} + → rMap n {A = A} ∘ ((suspMap2 {A = A} n) .fst) + ≡ botMap n ∘ lMap n {A = A} +theComp' {ℓ} = +nInd (λ {A} → funExt λ p → →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt λ { false → transportRefl _ + ; true → sym (rCancel _)})) + (λ {A} → funExt λ p → →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt λ x → (λ i → ((rMap≡rMap1 {A = A} i) ∘ suspMap2 1 .fst) p .fst x) ∙ sym (lem₁ p x))) + cool* + where + lem₁ : ∀ {ℓ} {A : Pointed ℓ} (p : _) (x : S¹) + → (botMap 1 ∘ lMap 1) p .fst x ≡ (rMap1 {A = A} ∘ ((suspMap2 {A = A} 1) .fst)) p .fst x + lem₁ {A = A} p base = rCancel (merid (pt A)) + lem₁ {A = A} p (loop i) j = + doubleCompPath-filler + (sym (rCancel (merid (pt A)))) + (cong (λ x → merid x ∙ sym (merid (pt A))) p) + (rCancel (merid (pt A))) j i + + cool* : (n : ℕ) → + ({A : Pointed ℓ} → + rMap (suc n) {A = A} ∘ suspMap2 (suc n) .fst ≡ + botMap (suc n) ∘ lMap (suc n)) → + {A : Pointed ℓ} → + rMap (suc (suc n)) {A = A} ∘ suspMap2 (suc (suc n)) .fst ≡ + botMap (suc (suc n)) ∘ lMap (suc (suc n)) + cool* n ind {A} = + funExt λ p → →∙Homogeneous≡ (isHomogeneousPath (Susp (typ A)) refl) + (funExt λ { north → sym (rCancel (merid (pt A))) + ; south → sym (rCancel (merid (pt A))) + ; (merid a i) j → hcomp (λ k → λ {(i = i0) → rCancel (merid (pt A)) (~ j) + ; (i = i1) → rCancel (merid (pt A)) (~ j) + ; (j = i0) → help2 p a (~ k) i + ; (j = i1) → (botMap (suc (suc n)) ∘ lMap (suc (suc n))) p .fst + (merid a i)}) + (doubleCompPath-filler (sym (rCancel (merid (pt A)))) + (λ i → (botMap (suc (suc n)) ∘ lMap (suc (suc n))) p .fst + (merid a i)) + (rCancel (merid (pt A))) (~ j) i)}) + where + module _ (p : typ (Ω _)) (a : S₊ (suc n)) where + abstract + indHyp : Path ((a₁ : fst (Ω ((Ω^ n) A))) → + Σ (fst (S₊∙ (suc n)) → fst (Ω (∙Susp (typ A)))) + (λ f → f (snd (S₊∙ (suc n))) ≡ snd (Ω (∙Susp (typ A))))) + ((rMap (suc n) {A = A} ∘ ((suspMap2 {A = A} (suc n)) .fst))) + (botMap (suc n) ∘ lMap (suc n) {A = A}) + indHyp = funExt λ a → →∙Homogeneous≡ (isHomogeneousPath _ _) (funExt (λ x → funExt⁻ (cong fst (funExt⁻ (ind {A = A}) a)) x)) + + pullPath : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) (r : z ≡ z) + → p ∙ (q ∙ r ∙ sym q) ∙ sym p ≡ (p ∙ q) ∙ r ∙ sym (p ∙ q) + pullPath p q r = + ∙assoc p (q ∙ r ∙ sym q) (sym p) + ∙ cong (_∙ sym p) (∙assoc p q (r ∙ sym q)) + ∙ sym (∙assoc (p ∙ q) (r ∙ sym q) (sym p)) + ∙ cong ((p ∙ q) ∙_) (sym (∙assoc r (sym q) (sym p)) + ∙ cong (r ∙_) (sym (symDistr p q))) + + appCoherence : ∀ {ℓ} {A : Type ℓ} {x y : A} (coh : x ≡ y) (p : x ≡ x) → p ≡ coh ∙ (sym coh ∙∙ p ∙∙ coh) ∙ sym coh + appCoherence {x = x} = J (λ y coh → (p : x ≡ x) → p ≡ coh ∙ (sym coh ∙∙ p ∙∙ coh) ∙ sym coh) + λ p → lUnit _ ∙ cong (refl ∙_) (rUnit _ ∙ cong (_∙ refl) (rUnit p)) + + coolLem : ∀ {ℓ} {A : Type ℓ} {x₀ x y z w : A} + (p : x₀ ≡ x) (q : x ≡ y) (r : y ≡ z) (s : z ≡ w) (mid : w ≡ w) + (coh : x₀ ≡ w) + → isComm∙ (A , x₀) + → (p ∙∙ q ∙∙ r ∙∙ s ∙∙ mid ∙∙ sym s ∙∙ sym r ∙∙ sym q ∙∙ sym p) + ≡ (coh ∙∙ mid ∙∙ sym coh) + coolLem p q r s mid coh comm = + doubleCompPath≡compPath p _ (sym p) + ∙∙ cong (λ x → p ∙ x ∙ (sym p)) + (doubleCompPath≡compPath q _ (sym q) + ∙ cong (λ x → q ∙ x ∙ (sym q)) + (doubleCompPath≡compPath r _ (sym r) + ∙ cong (λ x → r ∙ x ∙ (sym r)) + (doubleCompPath≡compPath s _ (sym s) + ∙ cong (λ x → s ∙ x ∙ sym s) (appCoherence (sym coh) mid) + ∙ pullPath s (sym coh) (coh ∙∙ mid ∙∙ sym coh)) + ∙ pullPath r (s ∙ sym coh) (coh ∙∙ mid ∙∙ sym coh)) + ∙ pullPath q (r ∙ s ∙ sym coh) (coh ∙∙ mid ∙∙ sym coh)) + ∙∙ pullPath p (q ∙ r ∙ s ∙ sym coh) (coh ∙∙ mid ∙∙ sym coh) + ∙∙ cong ((p ∙ q ∙ r ∙ s ∙ sym coh) ∙_) (comm (coh ∙∙ mid ∙∙ sym coh) (sym (p ∙ q ∙ r ∙ s ∙ sym coh))) + ∙∙ ∙assoc _ _ _ + ∙∙ cong (_∙ (coh ∙∙ mid ∙∙ sym coh)) (rCancel _) + ∙∙ sym (lUnit _) + + + cong-lMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (x : _) (p : refl ≡ x) (q : _ ≡ _) (a : S₊ (suc n)) + → (cong (lMap (suc n) {A = Ω A}) (fun (flipΩIso (suc (suc n))) (p ∙∙ q ∙∙ sym p))) + ≡ (cong (lMap (suc n)) (sym (flipΩrefl n)) + ∙∙ cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) p + ∙∙ cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) q + ∙∙ sym (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) p) + ∙∙ cong (lMap (suc n)) (flipΩrefl n)) + cong-lMap n x = J (λ x p → (q : x ≡ x) → + S₊ (suc n) → + cong (lMap (suc n)) + (fun (flipΩIso (suc (suc n))) (p ∙∙ q ∙∙ sym p)) + ≡ + (cong (lMap (suc n)) (sym (flipΩrefl n)) ∙∙ + cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) p ∙∙ + cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) q ∙∙ + sym (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) p) + ∙∙ cong (lMap (suc n)) (flipΩrefl n))) + λ q a → (λ j i → lMap (suc n) + (fun (flipΩIso (suc (suc n))) (rUnit q (~ j)) + i)) ∙ cong-lMap-main' n q + ∙ cong (cong (lMap (suc n)) (sym (flipΩrefl n)) ∙∙_∙∙ + cong (lMap (suc n)) (flipΩrefl n)) (rUnit (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) q)) + + help2 : cong ((rMap (suc (suc n)) {A = A} ∘ suspMap2 (suc (suc n)) .fst) p .fst) (merid a) + ≡ (sym (rCancel (merid (pt A))) ∙∙ (λ i → (botMap (suc (suc n)) ∘ lMap (suc (suc n)) {A = A}) p .fst (merid a i)) ∙∙ rCancel _) + help2 = ((λ _ → (sym (lMapId (suc n) a) ∙∙ cong (λ x → fst x a) + (cong (lMap (suc n)) + (fun (flipΩIso (suc (suc n))) + (sym (∙∙lCancel (snd (suspMap2 n))) + ∙∙ cong (sym (snd (suspMap2 n)) ∙∙_∙∙ snd (suspMap2 n)) (cong (cong (fst (suspMap2 n))) p) + ∙∙ ∙∙lCancel (snd (suspMap2 n))))) + ∙∙ lMapId (suc n) a))) + ∙∙ cong (sym (lMapId (suc n) a) ∙∙_∙∙ (lMapId (suc n) a)) + (cong (cong (λ x → fst x a)) (cong-lMap _ _ (sym (∙∙lCancel (snd (suspMap2 n)))) (cong ((suspMap2 (suc n)) .fst) p) a)) + ∙∙ cong (sym (lMapId (suc n) a) ∙∙_∙∙ (lMapId (suc n) a)) + ((cong-∙∙ (λ x → fst x a) (cong (lMap (suc n)) (sym (flipΩrefl n))) + (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) (sym (∙∙lCancel (snd (suspMap2 n)))) + ∙∙ cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) ((cong ((suspMap2 (suc n)) .fst) p)) + ∙∙ cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) (∙∙lCancel (snd (suspMap2 n)))) + (cong (lMap (suc n)) (flipΩrefl n)))) + ∙∙ cong (sym (lMapId (suc n) a) ∙∙_∙∙ (lMapId (suc n) a)) + (cong (cong (λ x → fst x a) (cong (lMap (suc n)) (sym (flipΩrefl n))) + ∙∙_∙∙ + cong (λ x → fst x a) (cong (lMap (suc n)) (flipΩrefl n))) + (cong-∙∙ (λ x → fst x a) + (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) (sym (∙∙lCancel (snd (suspMap2 n))))) + (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) (cong ((suspMap2 (suc n)) .fst) p)) + (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) (∙∙lCancel (snd (suspMap2 n)))))) + ∙∙ cong (sym (lMapId (suc n) a) ∙∙_∙∙ (lMapId (suc n) a)) + (cong (cong (λ x → fst x a) (cong (lMap (suc n)) (sym (flipΩrefl n))) + ∙∙_∙∙ + cong (λ x → fst x a) (cong (lMap (suc n)) (flipΩrefl n))) + (cong (cong (λ x → fst x a) (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) (sym (∙∙lCancel (snd (suspMap2 n))))) + ∙∙_∙∙ + cong (λ x → fst x a) (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) (∙∙lCancel (snd (suspMap2 n))))) + ((λ i j → rMap (suc n) {A = A} (suspMap2 (suc n) .fst (p j)) .fst a) + ∙ rUnit _ + ∙ λ i → (λ j → indHyp (i ∧ j) (snd (Ω ((Ω^ n) A))) .fst a) + ∙∙ (λ j → indHyp i (p j) .fst a) + ∙∙ λ j → indHyp (i ∧ ~ j) (snd (Ω ((Ω^ n) A))) .fst a))) + ∙∙ coolLem (sym (lMapId (suc n) a)) (λ i₁ → fst (lMap (suc n) (flipΩrefl n (~ i₁))) a) + ((λ i₁ → + fst + (lMap (suc n) + (fun (flipΩIso (suc n)) (∙∙lCancel (snd (suspMap2 {A = A} n)) (~ i₁)))) + a)) (λ j₁ → indHyp j₁ (snd (Ω ((Ω^ n) A))) .fst a) (λ j₁ → (botMap (suc n) ∘ lMap (suc n)) (p j₁) .fst a) + (sym (cong (_∙ sym (merid (pt A))) (cong merid (lMapId (suc n) a)) ∙ rCancel _)) + (EH 0) + ∙∙ helpLem ((rCancel (merid (pt A)))) + (cong (_∙ sym (merid (pt A))) (cong merid (lMapId (suc n) a))) + (λ j₁ → (botMap (suc n) ∘ lMap (suc n)) (p j₁) .fst a) + ∙∙ cong (sym (rCancel (merid (pt A))) ∙∙_∙∙ rCancel (merid (pt A))) + (sym (cong-∙∙ (λ x → merid x ∙ sym (merid (pt A))) (sym (lMapId (suc n) a)) (λ i → lMap (suc n) (p i) .fst a) (lMapId (suc n) a))) + ∙∙ λ _ → sym (rCancel (merid (pt A))) + ∙∙ (λ i → merid ((sym (lMapId (suc n) a) ∙∙ (λ i → lMap (suc n) (p i) .fst a) ∙∙ lMapId (suc n) a) i) ∙ sym (merid (pt A))) + ∙∙ rCancel (merid (pt A)) + where + helpLem : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : z ≡ x) (q : y ≡ z) (r : y ≡ y) + → (sym (q ∙ p) ∙∙ r ∙∙ (q ∙ p)) + ≡ (sym p ∙∙ sym q ∙∙ r ∙∙ q ∙∙ p) + helpLem p q r = + cong (λ x → x ∙∙ r ∙∙ (q ∙ p)) (symDistr q p) + ∙∙ doubleCompPath≡compPath (sym p ∙ sym q) r (q ∙ p) + ∙∙ (sym (∙assoc (sym p) (sym q) (r ∙ q ∙ p)) + ∙∙ cong (sym p ∙_) (∙assoc (sym q) r (q ∙ p) ∙ ∙assoc (sym q ∙ r) q p) + ∙∙ sym (doubleCompPath≡compPath (sym p) ((sym q ∙ r) ∙ q) p) + ∙ cong (sym p ∙∙_∙∙ p) (sym (∙assoc (sym q) r q) ∙ sym (doubleCompPath≡compPath (sym q) r q))) + +{- +Ω^ + + +-} + +suspMap→TranspType : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (typ (Ω ((Ω^ n) A)) → typ (Ω (Ω ((Ω^ n) (∙Susp (typ A)))))) + ≡ ((S₊∙ (suc n) →∙ A) → (S₊∙ (suc (suc n)) →∙ ∙Susp (typ A))) +suspMap→TranspType {A = A} n i = + Ω≡SphereMap {A = A} (suc n) i → Ω≡SphereMap' {A = A} (suc n) i + +suspMap→ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → PathP (λ i → suspMap→TranspType {A = A} n i) + (suspMap2 (suc n) .fst) + (suspMap n) +suspMap→ {A = A} n = + toPathP (funExt λ f → + (λ j → transportRefl (botᵣ {A = A} (suc n) + (rMap (suc n) {A = A} + (suspMap2 (suc n) .fst + ((invEq (_ , isEquiv-lMap n) (transportRefl f j)))))) j) + ∙∙ cong (botᵣ {A = A} (suc n)) + (funExt⁻ (theComp' (suc n)) (invEq (_ , isEquiv-lMap n) f)) + ∙∙ sym (finalΔ (suc n) (lMap (suc n) {A = A} (invEq (lMap (suc n) , isEquiv-lMap n) f))) + ∙ cong (suspMap n) (secEq ((lMap (suc n) , isEquiv-lMap n)) f)) + +isConnectedPres : ∀ {ℓ} {A : Pointed ℓ} (con n : ℕ) + → isConnectedFun con (suspMap2 {A = A} (suc n) .fst) + → isConnectedFun con (suspMap {A = A} n) +isConnectedPres {A = A} con n hyp = + transport (λ i → isConnectedFun con (suspMap→ {A = A} n i)) hyp + + +isConnectedSuspMap2 : ? +isConnectedSuspMap2 = ? diff --git a/Cubical/Homotopy/HomotopyGroup.agda b/Cubical/Homotopy/HomotopyGroup.agda index ba0cb6b8c..f9f0fc541 100644 --- a/Cubical/Homotopy/HomotopyGroup.agda +++ b/Cubical/Homotopy/HomotopyGroup.agda @@ -127,7 +127,6 @@ fst (-Π {A = A} {n = suc (suc n)} f) (merid a j) = fst f ((merid a ∙ sym (merid (ptSn _))) (~ j)) snd (-Π {A = A} {n = suc (suc n)} f) = snd f --- module _ {ℓ : Level} {A : Pointed ℓ} where flipΩPath : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) → ((Ω^ (suc n)) A) ≡ (Ω^ n) (Ω A) @@ -141,9 +140,11 @@ flipΩIso {A = A} n = pathToIso (cong fst (flipΩPath n)) flipΩIsopres· : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) → (f g : fst ((Ω^ (suc n)) (Ω A))) → Iso.inv (flipΩIso (suc n)) (f ∙ g) - ≡ (Iso.inv (flipΩIso (suc n)) f) ∙ (Iso.inv (flipΩIso (suc n)) g) + ≡ (Iso.inv (flipΩIso (suc n)) f) + ∙ (Iso.inv (flipΩIso (suc n)) g) flipΩIsopres· {A = A} n f g i = - transp (λ j → flipΩPath {A = A} n (~ i ∧ ~ j) .snd ≡ flipΩPath n (~ i ∧ ~ j) .snd) i + transp (λ j → flipΩPath {A = A} n (~ i ∧ ~ j) .snd + ≡ flipΩPath n (~ i ∧ ~ j) .snd) i (transp (λ j → flipΩPath {A = A} n (~ i ∨ ~ j) .snd ≡ flipΩPath n (~ i ∨ ~ j) .snd) (~ i) f ∙ transp (λ j → flipΩPath {A = A} n (~ i ∨ ~ j) .snd @@ -170,75 +171,92 @@ module _ {ℓ : Level} {A : Pointed ℓ} where (J (λ q Q → P refl q refl Q) b) J4 : ∀ {ℓ ℓ'} {A : Type ℓ} {x : A} - → (P : (p q r s : x ≡ x) → (refl ≡ p) → (p ≡ q) → (refl ≡ r) → (r ≡ s) → Type ℓ') + → (P : (p q r s : x ≡ x) + → (refl ≡ p) → (p ≡ q) → (refl ≡ r) → (r ≡ s) → Type ℓ') → P refl refl refl refl refl refl refl refl → (p q r s : x ≡ x) (P' : _) (Q : _) (R : _) (S : _) → P p q r s P' Q R S J4 P b p q r s = - J (λ p P' → (Q₁ : p ≡ q) (R : refl ≡ r) (S₁ : r ≡ s) → P p q r s P' Q₁ R S₁) - (J (λ q Q₁ → (R : refl ≡ r) (S₁ : r ≡ s) → P refl q r s refl Q₁ R S₁) - (J (λ r R → (S₁ : r ≡ s) → P refl refl r s refl refl R S₁) + J (λ p P' → (Q₁ : p ≡ q) (R : refl ≡ r) (S₁ : r ≡ s) + → P p q r s P' Q₁ R S₁) + (J (λ q Q₁ → (R : refl ≡ r) (S₁ : r ≡ s) + → P refl q r s refl Q₁ R S₁) + (J (λ r R → (S₁ : r ≡ s) + → P refl refl r s refl refl R S₁) (J (λ s S₁ → P refl refl refl s refl refl refl S₁) b))) looper-help : (n : ℕ) (f g : (S₊∙ (suc n) →∙ Ω A)) (a : _) → fst (∙Π f g) a ≡ fst f a ∙ fst g a - looper-help zero f g base = rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g)) + looper-help zero f g base = + rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g)) looper-help zero f g (loop i) k = - lazy-fill _ _ (sym (snd f)) (sym (snd g)) (cong (fst f) loop) (cong (fst g) loop) k i + lazy-fill _ _ + (sym (snd f)) (sym (snd g)) + (cong (fst f) loop) (cong (fst g) loop) k i where - lazy-fill : ∀ {ℓ} {A : Type ℓ} {x : A} (p q : x ≡ x) - (prefl : refl ≡ p) (qrefl : refl ≡ q) (fl : p ≡ p) (fr : q ≡ q) - → PathP (λ i → (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i - ≡ (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i) - ((prefl ∙∙ fl ∙∙ sym prefl) ∙ (qrefl ∙∙ fr ∙∙ sym qrefl)) - (cong₂ _∙_ fl fr) + lazy-fill : + ∀ {ℓ} {A : Type ℓ} {x : A} (p q : x ≡ x) + (prefl : refl ≡ p) (qrefl : refl ≡ q) (fl : p ≡ p) (fr : q ≡ q) + → PathP (λ i → (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i + ≡ (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i) + ((prefl ∙∙ fl ∙∙ sym prefl) ∙ (qrefl ∙∙ fr ∙∙ sym qrefl)) + (cong₂ _∙_ fl fr) lazy-fill {x = x} = J2 (λ p q prefl qrefl → (fl : p ≡ p) (fr : q ≡ q) - → PathP (λ i → (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i - ≡ (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i) - ((prefl ∙∙ fl ∙∙ sym prefl) ∙ (qrefl ∙∙ fr ∙∙ sym qrefl)) - (cong₂ _∙_ fl fr)) + → PathP (λ i → (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i + ≡ (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i) + ((prefl ∙∙ fl ∙∙ sym prefl) ∙ (qrefl ∙∙ fr ∙∙ sym qrefl)) + (cong₂ _∙_ fl fr)) λ fl fr → cong₂ _∙_ (sym (rUnit fl)) (sym (rUnit fr)) ◁ flipSquare (sym (rUnit (rUnit refl)) ◁ (flipSquare ((λ j → cong (λ x → rUnit x j) fl ∙ cong (λ x → lUnit x j) fr) ▷ sym (cong₂Funct _∙_ fl fr)) ▷ (rUnit (rUnit refl)))) - looper-help (suc n) f g north = rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g)) - looper-help (suc n) f g south = (rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) - ∙ cong₂ _∙_ (cong (fst f) (merid (ptSn _))) - (cong (fst g) (merid (ptSn _))) + looper-help (suc n) f g north = + rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g)) + looper-help (suc n) f g south = + (rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) + ∙ cong₂ _∙_ (cong (fst f) (merid (ptSn _))) + (cong (fst g) (merid (ptSn _))) looper-help (suc n) f g (merid a j) i = help i j where - help : PathP (λ i → (rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) i - ≡ ((rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) - ∙ cong₂ _∙_ (cong (fst f) (merid (ptSn _))) - (cong (fst g) (merid (ptSn _)))) i) - (cong (fst (∙Π f g)) (merid a)) - (cong₂ _∙_ (cong (fst f) (merid a)) (cong (fst g) (merid a))) + help : + PathP (λ i → (rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) i + ≡ ((rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) + ∙ cong₂ _∙_ (cong (fst f) (merid (ptSn _))) + (cong (fst g) (merid (ptSn _)))) i) + (cong (fst (∙Π f g)) (merid a)) + (cong₂ _∙_ (cong (fst f) (merid a)) + (cong (fst g) (merid a))) help = (cong₂ _∙_ (cong (sym (snd f) ∙∙_∙∙ snd f) - (cong-∙ (fst f) (merid a) (sym (merid (ptSn _))))) + (cong-∙ (fst f) + (merid a) (sym (merid (ptSn _))))) ((cong (sym (snd g) ∙∙_∙∙ snd g) - (cong-∙ (fst g) (merid a) (sym (merid (ptSn _)))))) + (cong-∙ (fst g) + (merid a) (sym (merid (ptSn _)))))) ◁ lazy-help _ _ _ _ (sym (snd f)) (cong (fst f) (merid (ptSn _))) (sym (snd g)) (cong (fst g) (merid (ptSn _))) (cong (fst f) (merid a)) (cong (fst g) (merid a))) where - lazy-help : ∀ {ℓ} {A : Type ℓ} {x : A} (p-north p-south q-north q-south : x ≡ x) - (prefl : refl ≡ p-north) (p-merid : p-north ≡ p-south) - (qrefl : refl ≡ q-north) (q-merid : q-north ≡ q-south) - (fl : p-north ≡ p-south) (fr : q-north ≡ q-south) - → PathP (λ i → (rUnit refl ∙ cong₂ _∙_ prefl qrefl) i - ≡ ((rUnit refl ∙ cong₂ _∙_ prefl qrefl) - ∙ cong₂ _∙_ p-merid q-merid) i) - ((prefl ∙∙ fl ∙ sym p-merid ∙∙ sym prefl) - ∙ (qrefl ∙∙ fr ∙ sym q-merid ∙∙ sym qrefl)) - (cong₂ _∙_ fl fr) + lazy-help : + ∀ {ℓ} {A : Type ℓ} {x : A} + (p-north p-south q-north q-south : x ≡ x) + (prefl : refl ≡ p-north) (p-merid : p-north ≡ p-south) + (qrefl : refl ≡ q-north) (q-merid : q-north ≡ q-south) + (fl : p-north ≡ p-south) (fr : q-north ≡ q-south) + → PathP (λ i → (rUnit refl ∙ cong₂ _∙_ prefl qrefl) i + ≡ ((rUnit refl ∙ cong₂ _∙_ prefl qrefl) + ∙ cong₂ _∙_ p-merid q-merid) i) + ((prefl ∙∙ fl ∙ sym p-merid ∙∙ sym prefl) + ∙ (qrefl ∙∙ fr ∙ sym q-merid ∙∙ sym qrefl)) + (cong₂ _∙_ fl fr) lazy-help {x = x} = J4 (λ p-north p-south q-north q-south prefl p-merid qrefl q-merid → (fl : p-north ≡ p-south) (fr : q-north ≡ q-south) → PathP (λ i₁ → (rUnit refl ∙ cong₂ _∙_ prefl qrefl) i₁ ≡ - ((rUnit refl ∙ cong₂ _∙_ prefl qrefl) ∙ cong₂ _∙_ p-merid q-merid) i₁) + ((rUnit refl ∙ cong₂ _∙_ prefl qrefl) + ∙ cong₂ _∙_ p-merid q-merid) i₁) ((prefl ∙∙ fl ∙ sym p-merid ∙∙ sym prefl) ∙ (qrefl ∙∙ fr ∙ sym q-merid ∙∙ sym qrefl)) (cong₂ _∙_ fl fr)) @@ -257,22 +275,31 @@ module _ {ℓ : Level} {A : Pointed ℓ} where main south = refl main (merid a i₁) j = help j i₁ where - help : fst (∙Π f g) a ≡ cong (fst (∙Π (suspension→ n f) (suspension→ n g))) (merid a) + help : fst (∙Π f g) a + ≡ cong (fst (∙Π (suspension→ n f) (suspension→ n g))) (merid a) help = looper-help n f g a - ∙ cong₂ _∙_ (sym (cong-∙ (fst (suspension→ n f)) (merid a) (sym (merid (ptSn _))) + ∙ cong₂ _∙_ (sym (cong-∙ (fst (suspension→ n f)) + (merid a) + (sym (merid (ptSn _))) ∙∙ cong (fst f a ∙_) (cong sym (snd f)) ∙∙ sym (rUnit _))) - (sym (cong-∙ (fst (suspension→ n g)) (merid a) (sym (merid (ptSn _))) + (sym (cong-∙ (fst (suspension→ n g)) + (merid a) + (sym (merid (ptSn _))) ∙∙ cong (fst g a ∙_) (cong sym (snd g)) ∙∙ sym (rUnit _))) - ∙ λ i → rUnit (cong (fst (suspension→ n f)) (merid a ∙ sym (merid (ptSn _)))) i - ∙ rUnit (cong (fst (suspension→ n g)) (merid a ∙ sym (merid (ptSn _)))) i + ∙ λ i → rUnit (cong (fst (suspension→ n f)) + (merid a ∙ sym (merid (ptSn _)))) i + ∙ rUnit (cong (fst (suspension→ n g)) + (merid a ∙ sym (merid (ptSn _)))) i - suspension←filler : (n : ℕ) → (S₊∙ (suc (suc n)) →∙ A) → I → I → I → fst A + suspension←filler : (n : ℕ) → (S₊∙ (suc (suc n)) →∙ A) + → I → I → I → fst A suspension←filler n f i j k = hfill (λ k → λ { (i = i0) → doubleCompPath-filler (sym (snd f)) - (cong (fst f) (merid (ptSn _) ∙ sym (merid (ptSn _)))) + (cong (fst f) (merid (ptSn _) + ∙ sym (merid (ptSn _)))) (snd f) k j ; (i = i1) → snd f k ; (j = i0) → snd f k @@ -297,7 +324,8 @@ module _ {ℓ : Level} {A : Pointed ℓ} where help (merid a j) i = hcomp (λ k → λ { (i = i1) → f (merid a j) ; (j = i0) → p (~ i ∧ k) - ; (j = i1) → compPath-filler' (sym p) (cong f (merid (ptSn _))) k i}) + ; (j = i1) → compPath-filler' (sym p) + (cong f (merid (ptSn _))) k i}) (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) Iso.leftInv (suspensionIso n) f = ΣPathP (funExt lem , PathP-filler) @@ -306,7 +334,8 @@ module _ {ℓ : Level} {A : Pointed ℓ} where side₁ : (x : S₊ (suc n)) → I → I → I → fst A side₁ x i j k = hfill (λ k → λ { (i = i0) → suspension→ n f .fst - (compPath-filler' (merid x) (sym (merid (ptSn _))) k j) + (compPath-filler' (merid x) + (sym (merid (ptSn _))) k j) ; (i = i1) → snd A ; (j = i0) → fst f x (i ∨ ~ k) ; (j = i1) → snd A}) @@ -339,16 +368,20 @@ module _ {ℓ : Level} {A : Pointed ℓ} where lem x i j = side₂ x i j i1 sideCube : Cube (λ j k → snd f j (~ k)) - (λ j k → fst (suspension→ n f) (rCancel' (merid (ptSn _)) j k)) + (λ j k → fst (suspension→ n f) + (rCancel' (merid (ptSn _)) j k)) (λ r k → suspension→ n f .fst - (compPath-filler' (merid (ptSn _)) (sym (merid (ptSn _))) r k)) + (compPath-filler' (merid (ptSn _)) + (sym (merid (ptSn _))) r k)) (λ _ _ → pt A) (λ r j → snd f j (~ r)) λ _ _ → pt A sideCube r j k = hcomp (λ i → λ {(r = i0) → snd f j (~ k ∨ ~ i) - ; (r = i1) → fst (suspension→ n f) (rCancel-filler' (merid (ptSn _)) (~ i) j k) + ; (r = i1) → fst (suspension→ n f) + (rCancel-filler' (merid (ptSn _)) (~ i) j k) ; (j = i0) → suspension→ n f .fst - (compPath-filler'' (merid (ptSn _)) (sym (merid (ptSn _))) i r k) + (compPath-filler'' (merid (ptSn _)) + (sym (merid (ptSn _))) i r k) ; (j = i1) → snd f (~ r) (~ i ∧ k) ; (k = i0) → snd f j (~ r) ; (k = i1) → snd f (~ r ∧ j) (~ i)}) @@ -392,7 +425,8 @@ SphereMap→Ω : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) → (S₊∙ (suc n) →∙ A) → (fst ((Ω^ (suc n)) A)) SphereMap→Ω zero (f , p) = sym p ∙∙ cong f loop ∙∙ p SphereMap→Ω {A = A} (suc n) (f , p) = - Iso.inv (flipΩIso _) (SphereMap→Ω {A = Ω A} n (Iso.inv (suspensionIso _) (f , p))) + Iso.inv (flipΩIso _) + (SphereMap→Ω {A = Ω A} n (Iso.inv (suspensionIso _) (f , p))) SphereMap→Ωpres∙ : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) → (p q : S₊∙ (suc n) →∙ A) @@ -401,14 +435,16 @@ SphereMap→Ωpres∙ zero f g = sym (rUnit _) SphereMap→Ωpres∙ (suc n) f g = cong (Iso.inv (flipΩIso (suc n))) (cong (SphereMap→Ω n) (suspension←-pres∙ n f g) - ∙ SphereMap→Ωpres∙ n (suspension← n f) ((suspension← n g))) + ∙ SphereMap→Ωpres∙ n (suspension← n f) (suspension← n g)) ∙ flipΩIsopres· n _ _ -Ω→SphereMap : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) → (fst ((Ω^ (suc n)) A)) → (S₊∙ (suc n) →∙ A) +Ω→SphereMap : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (fst ((Ω^ (suc n)) A)) → (S₊∙ (suc n) →∙ A) fst (Ω→SphereMap {A = A} zero f) base = pt A fst (Ω→SphereMap {A = A} zero f) (loop i₁) = f i₁ snd (Ω→SphereMap {A = A} zero f) = refl -Ω→SphereMap {A = A} (suc n) f = Iso.fun (suspensionIso _) (Ω→SphereMap n (Iso.fun (flipΩIso _) f)) +Ω→SphereMap {A = A} (suc n) f = + Iso.fun (suspensionIso _) (Ω→SphereMap n (Iso.fun (flipΩIso _) f)) Ω→SphereMap→Ω : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) → (x : fst ((Ω^ (suc n)) A)) @@ -425,15 +461,18 @@ snd (Ω→SphereMap {A = A} zero f) = refl SphereMap→Ω→SphereMap : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) → (x : (S₊∙ (suc n) →∙ A)) → Ω→SphereMap n (SphereMap→Ω n x) ≡ x -SphereMap→Ω→SphereMap zero (f , p) = ΣPathP (funExt fstp , λ i j → p (~ i ∨ j)) +SphereMap→Ω→SphereMap zero (f , p) = + ΣPathP (funExt fstp , λ i j → p (~ i ∨ j)) where fstp : (x : _) → _ ≡ f x fstp base = sym p fstp (loop i) j = doubleCompPath-filler (sym p) (cong f loop) p (~ j) i SphereMap→Ω→SphereMap (suc n) f = - cong (Iso.fun (suspensionIso n) ∘ Ω→SphereMap n) (Iso.rightInv (flipΩIso (suc n)) + cong (Iso.fun (suspensionIso n) ∘ Ω→SphereMap n) + (Iso.rightInv (flipΩIso (suc n)) (SphereMap→Ω n (Iso.inv (suspensionIso n) f))) - ∙∙ cong (Iso.fun (suspensionIso n)) (SphereMap→Ω→SphereMap n (Iso.inv (suspensionIso n) f)) + ∙∙ cong (Iso.fun (suspensionIso n)) + (SphereMap→Ω→SphereMap n (Iso.inv (suspensionIso n) f)) ∙∙ Iso.rightInv (suspensionIso _) _ @@ -446,7 +485,8 @@ fst (Iso.inv (IsoSphereMapΩ {A = A} zero) a) true = pt A snd (Iso.inv (IsoSphereMapΩ {A = A} zero) a) = refl Iso.rightInv (IsoSphereMapΩ {A = A} zero) a = refl Iso.leftInv (IsoSphereMapΩ {A = A} zero) (f , p) = - ΣPathP ((funExt (λ { false → refl ; true → sym p})) , λ i j → p (~ i ∨ j)) + ΣPathP ((funExt (λ { false → refl ; true → sym p})) + , λ i j → p (~ i ∨ j)) Iso.fun (IsoSphereMapΩ {A = A} (suc n)) = SphereMap→Ω n Iso.inv (IsoSphereMapΩ {A = A} (suc n)) = Ω→SphereMap n Iso.rightInv (IsoSphereMapΩ {A = A} (suc n)) = Ω→SphereMap→Ω n @@ -460,10 +500,12 @@ IsoSphereMapΩ-pres∙Π (suc n) f g = cong (Iso.inv (flipΩIso (suc n))) (cong (SphereMap→Ω n) (suspension←-pres∙ n f g) ∙ SphereMap→Ωpres∙ n (suspension← n f) (suspension← n g)) - ∙ (flipΩIsopres· n (SphereMap→Ω n (suspension← n f)) (SphereMap→Ω n (suspension← n g))) + ∙ (flipΩIsopres· n (SphereMap→Ω n (suspension← n f)) + (SphereMap→Ω n (suspension← n g))) --- It is useful to define the ``Group Structure'' on (S₊∙ n →∙ A) before doing it on π' --- These will be the equivalents of the usual groupoid laws on Ω A. +-- It is useful to define the ``Group Structure'' on (S₊∙ n →∙ A) +-- before doing it on π'. These will be the equivalents of the +-- usual groupoid laws on Ω A. 1Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → (S₊∙ n →∙ A) fst (1Π {A = A}) _ = pt A snd (1Π {A = A}) = refl @@ -478,28 +520,39 @@ fst (∙Π-rUnit {A = A} {n = zero} f i) (loop j) = help i j (((sym (snd f)) ∙∙ (cong (fst f) loop) ∙∙ snd f) ∙ (refl ∙ refl)) (cong (fst f) loop) - help = (cong ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) ∙_) (sym (rUnit refl)) ∙ sym (rUnit _)) - ◁ λ i j → doubleCompPath-filler (sym (snd f)) (cong (fst f) loop) (snd f) (~ i) j + help = (cong ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) ∙_) + (sym (rUnit refl)) ∙ sym (rUnit _)) + ◁ λ i j → doubleCompPath-filler (sym (snd f)) + (cong (fst f) loop) (snd f) (~ i) j snd (∙Π-rUnit {A = A} {n = zero} f i) j = snd f (~ i ∨ j) fst (∙Π-rUnit {A = A} {n = suc n} f i) north = snd f (~ i) -fst (∙Π-rUnit {A = A} {n = suc n} f i) south = (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i +fst (∙Π-rUnit {A = A} {n = suc n} f i) south = + (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i fst (∙Π-rUnit {A = A} {n = suc n} f i) (merid a j) = help i j where - help : PathP (λ i → snd f (~ i) ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) - (((sym (snd f)) ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) ∙∙ snd f) + help : PathP (λ i → snd f (~ i) + ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) + (((sym (snd f)) + ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + ∙∙ snd f) ∙ (refl ∙ refl)) (cong (fst f) (merid a)) - help = (cong (((sym (snd f)) ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) ∙∙ snd f) ∙_) + help = (cong (((sym (snd f)) + ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + ∙∙ snd f) ∙_) (sym (rUnit refl)) ∙ sym (rUnit _)) ◁ λ i j → hcomp (λ k → λ { (j = i0) → snd f (~ i ∧ k) - ; (j = i1) → compPath-filler' (sym (snd f)) (cong (fst f) (merid (ptSn (suc n)))) k i + ; (j = i1) → compPath-filler' (sym (snd f)) + (cong (fst f) (merid (ptSn (suc n)))) k i ; (i = i0) → doubleCompPath-filler (sym (snd f)) - (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + (cong (fst f) + (merid a ∙ sym (merid (ptSn (suc n))))) (snd f) k j ; (i = i1) → fst f (merid a j)}) - (fst f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) + (fst f (compPath-filler (merid a) + (sym (merid (ptSn _))) (~ i) j)) snd (∙Π-rUnit {A = A} {n = suc n} f i) j = snd f (~ i ∨ j) ∙Π-lUnit : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} @@ -511,42 +564,57 @@ fst (∙Π-lUnit {n = zero} f i) (loop j) = s i j s : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) ((refl ∙ refl) ∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) (cong (fst f) loop) - s = (cong (_∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) (sym (rUnit refl)) ∙ sym (lUnit _)) - ◁ λ i j → doubleCompPath-filler (sym (snd f)) (cong (fst f) loop) (snd f) (~ i) j + s = (cong (_∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) + (sym (rUnit refl)) ∙ sym (lUnit _)) + ◁ λ i j → doubleCompPath-filler (sym (snd f)) + (cong (fst f) loop) (snd f) (~ i) j snd (∙Π-lUnit {n = zero} f i) j = snd f (~ i ∨ j) fst (∙Π-lUnit {n = suc n} f i) north = snd f (~ i) -fst (∙Π-lUnit {n = suc n} f i) south = (sym (snd f) ∙ cong (fst f) (merid (ptSn _))) i +fst (∙Π-lUnit {n = suc n} f i) south = + (sym (snd f) ∙ cong (fst f) (merid (ptSn _))) i fst (∙Π-lUnit {n = suc n} f i) (merid a j) = help i j where - help : PathP (λ i → snd f (~ i) ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) - ((refl ∙ refl) ∙ ((sym (snd f)) ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) ∙∙ snd f)) + help : PathP (λ i → snd f (~ i) + ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) + ((refl ∙ refl) ∙ ((sym (snd f)) + ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + ∙∙ snd f)) (cong (fst f) (merid a)) help = - (cong (_∙ ((sym (snd f)) ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) ∙∙ snd f)) + (cong (_∙ ((sym (snd f)) + ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + ∙∙ snd f)) (sym (rUnit refl)) ∙ sym (lUnit _)) - ◁ λ i j → hcomp (λ k → λ { (j = i0) → snd f (~ i ∧ k) - ; (j = i1) → compPath-filler' (sym (snd f)) (cong (fst f) (merid (ptSn (suc n)))) k i - ; (i = i0) → doubleCompPath-filler (sym (snd f)) - (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) - (snd f) k j - ; (i = i1) → fst f (merid a j)}) - (fst f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) + ◁ λ i j → hcomp (λ k → + λ { (j = i0) → snd f (~ i ∧ k) + ; (j = i1) → compPath-filler' (sym (snd f)) + (cong (fst f) (merid (ptSn (suc n)))) k i + ; (i = i0) → doubleCompPath-filler (sym (snd f)) + (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + (snd f) k j + ; (i = i1) → fst f (merid a j)}) + (fst f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) snd (∙Π-lUnit {n = suc n} f i) j = snd f (~ i ∨ j) ∙Π-rCancel : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → (f : S₊∙ (suc n) →∙ A) → ∙Π f (-Π f) ≡ 1Π fst (∙Π-rCancel {A = A} {n = zero} f i) base = pt A -fst (∙Π-rCancel {A = A} {n = zero} f i) (loop j) = rCancel (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) i j +fst (∙Π-rCancel {A = A} {n = zero} f i) (loop j) = + rCancel (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) i j snd (∙Π-rCancel {A = A} {n = zero} f i) = refl fst (∙Π-rCancel {A = A} {n = suc n} f i) north = pt A fst (∙Π-rCancel {A = A} {n = suc n} f i) south = pt A fst (∙Π-rCancel {A = A} {n = suc n} f i) (merid a i₁) = sl i i₁ where - pl = (sym (snd f) ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f) + pl = (sym (snd f) + ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) + ∙∙ snd f) sl : pl - ∙ ((sym (snd f) ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f)) ≡ refl + ∙ ((sym (snd f) + ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) + ∙∙ snd f)) ≡ refl sl = cong (pl ∙_) (cong (sym (snd f) ∙∙_∙∙ (snd f)) (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) ∙∙ cong₂ _∙_ refl @@ -565,9 +633,14 @@ fst (∙Π-lCancel {A = A} {n = suc n} f i) north = pt A fst (∙Π-lCancel {A = A} {n = suc n} f i) south = pt A fst (∙Π-lCancel {A = A} {n = suc n} f i) (merid a j) = sl i j where - pl = (sym (snd f) ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f) - - sl : ((sym (snd f) ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f)) ∙ pl ≡ refl + pl = (sym (snd f) + ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) + ∙∙ snd f) + + sl : (sym (snd f) + ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) + ∙∙ snd f) ∙ pl + ≡ refl sl = cong (_∙ pl) (cong (sym (snd f) ∙∙_∙∙ (snd f)) (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) ∙∙ cong₂ _∙_ refl (cong (cong (fst f)) (rCancel (merid (ptSn _)))) diff --git a/Cubical/Homotopy/HopfInvariant/Base.agda b/Cubical/Homotopy/HopfInvariant/Base.agda index df81c6afd..d856f16c1 100644 --- a/Cubical/Homotopy/HopfInvariant/Base.agda +++ b/Cubical/Homotopy/HopfInvariant/Base.agda @@ -1,8 +1,12 @@ {-# OPTIONS --safe --experimental-lossy-unification #-} module Cubical.Homotopy.HopfInvariant.Base where +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Group.Base + +open import Cubical.Relation.Nullary + open import Cubical.ZCohomology.Base -open import Cubical.ZCohomology.Groups.Connected open import Cubical.ZCohomology.GroupStructure open import Cubical.ZCohomology.Properties open import Cubical.ZCohomology.MayerVietorisUnreduced @@ -10,68 +14,38 @@ open import Cubical.ZCohomology.Groups.Unit open import Cubical.ZCohomology.Groups.Wedge open import Cubical.ZCohomology.Groups.Sn open import Cubical.ZCohomology.RingStructure.CupProduct -open import Cubical.ZCohomology.RingStructure.RingLaws open import Cubical.ZCohomology.RingStructure.GradedCommutativity -open import Cubical.Relation.Nullary open import Cubical.ZCohomology.Gysin -open import Cubical.Functions.Embedding - open import Cubical.Foundations.Path -open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.HLevels open import Cubical.Foundations.Transport open import Cubical.Foundations.Function -open import Cubical.Foundations.Univalence open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Pointed.Homogeneous - -open import Cubical.Homotopy.HomotopyGroup - -open import Cubical.Foundations.Univalence - -open import Cubical.Relation.Nullary -open import Cubical.Data.Sum -open import Cubical.Data.Fin -open import Cubical.Data.Empty renaming (rec to ⊥-rec) open import Cubical.Data.Sigma open import Cubical.Data.Int hiding (_+'_) open import Cubical.Data.Nat renaming (_+_ to _+ℕ_ ; _·_ to _·ℕ_) -open import Cubical.Data.Nat.Order open import Cubical.Data.Unit -open import Cubical.Data.Bool + open import Cubical.Algebra.Group - renaming (ℤ to ℤGroup ; Unit to UnitGroup) hiding (Bool) + renaming (ℤ to ℤGroup ; Unit to UnitGroup) open import Cubical.Algebra.Group.ZModule -open import Cubical.HITs.Pushout.Flattening -open import Cubical.Homotopy.Connected -open import Cubical.Homotopy.EilenbergSteenrod open import Cubical.HITs.Pushout open import Cubical.HITs.Sn open import Cubical.HITs.Susp -open import Cubical.HITs.S1 renaming (_·_ to _*_) +open import Cubical.HITs.Wedge open import Cubical.HITs.Truncation - renaming (rec to trRec ; elim to trElim ; elim2 to trElim2) + renaming (rec to trRec) open import Cubical.HITs.SetTruncation - renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap ; elim3 to sElim3) + renaming (rec to sRec ; elim to sElim ; elim2 to sElim2 ; map to sMap) open import Cubical.HITs.PropositionalTruncation - renaming (rec to pRec ; elim to pElim) - -open import Cubical.Algebra.AbGroup - -open import Cubical.Homotopy.Loopspace - -open import Cubical.HITs.Join - -open import Cubical.Homotopy.Hopf - -open import Cubical.HITs.SetQuotients renaming (_/_ to _/'_) + renaming (rec to pRec) isContr→≡UnitGroup : {G : Group ℓ-zero} → isContr (fst G) → UnitGroup ≡ G isContr→≡UnitGroup c = @@ -79,7 +53,8 @@ isContr→≡UnitGroup c = (invGroupEquiv ((isContr→≃Unit c) , (makeIsGroupHom (λ _ _ → refl)))) -GroupIsoUnitGroup→isContr : {G : Group ℓ-zero} → GroupIso UnitGroup G → isContr (fst G) +GroupIsoUnitGroup→isContr : {G : Group ℓ-zero} + → GroupIso UnitGroup G → isContr (fst G) GroupIsoUnitGroup→isContr is = isOfHLevelRetractFromIso 0 (invIso (fst is)) isContrUnit @@ -138,14 +113,16 @@ Hopfα n f = ∣ (λ { (inl x) → 0ₖ _ where help : (a : S₊ (3 +ℕ n +ℕ n)) → ∣ fst f a ∣ ≡ 0ₖ (2 +ℕ n) help = sphereElim _ (λ _ → isOfHLevelPlus' {n = n} (3 +ℕ n) - (isOfHLevelPath' (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n)) _ _)) (cong ∣_∣ₕ (snd f)) + (isOfHLevelPath' (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n)) _ _)) + (cong ∣_∣ₕ (snd f)) Hopfβ : (n : ℕ) → (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) → coHom (4 +ℕ (n +ℕ n)) (HopfInvariantPush n (fst f)) Hopfβ n f = fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) ∣ ∣_∣ ∣₂ --- To define the Hopf invariant, we need to establish the non-trivial isos of Hⁿ(HopfInvariant). +-- To define the Hopf invariant, we need to establish the +-- non-trivial isos of Hⁿ(HopfInvariant). module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) where module M = MV _ _ _ (λ _ → tt) (fst f) ¬lemHopf : (n m : ℕ) → ¬ suc (suc (m +ℕ n)) ≡ suc n @@ -160,11 +137,13 @@ module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) whe GroupEquiv (coHomGr (3 +ℕ n +ℕ n) (S₊ (3 +ℕ n +ℕ n))) ((coHomGr (suc (3 +ℕ n +ℕ n)) (HopfInvariantPush _ (fst f)))) - fst (fst SphereHopfCohomIso) = fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) + fst (fst SphereHopfCohomIso) = + fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) snd (fst SphereHopfCohomIso) = help where abstract - help : isEquiv (fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n))) + help : isEquiv (fst (MV.d _ _ _ (λ _ → tt) (fst f) + (3 +ℕ n +ℕ n))) help = SES→isEquiv (isContr→≡UnitGroup @@ -188,10 +167,12 @@ module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) whe abstract d-inv : IsGroupHom (snd (coHomGr (3 +ℕ n +ℕ n) (S₊ (3 +ℕ n +ℕ n)))) (fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n))) - (snd (coHomGr (suc (3 +ℕ n +ℕ n)) (Pushout (λ _ → tt) (fst f)))) + (snd (coHomGr (suc (3 +ℕ n +ℕ n)) + (Pushout (λ _ → tt) (fst f)))) d-inv = snd (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) - Hopfβ-Iso : GroupIso (coHomGr (suc (3 +ℕ n +ℕ n)) (HopfInvariantPush _ (fst f))) + Hopfβ-Iso : GroupIso (coHomGr (suc (3 +ℕ n +ℕ n)) + (HopfInvariantPush _ (fst f))) ℤGroup Iso.fun (fst Hopfβ-Iso) x = Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) @@ -201,7 +182,8 @@ module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) whe (Iso.inv (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) f) Iso.rightInv (fst Hopfβ-Iso) f = cong (Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) - (retEq (fst SphereHopfCohomIso) (Iso.inv (fst (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) f)) + (retEq (fst SphereHopfCohomIso) + (Iso.inv (fst (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) f)) ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) f Iso.leftInv (fst Hopfβ-Iso) x = cong (fst (fst SphereHopfCohomIso)) @@ -211,10 +193,12 @@ module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) whe snd Hopfβ-Iso = grHom where abstract - grHom : IsGroupHom (coHomGr (suc (3 +ℕ n +ℕ n)) (Pushout (λ _ → tt) (fst f)) .snd) - (λ x → Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) - (invEq (fst SphereHopfCohomIso) x)) - (ℤGroup .snd) + grHom : + IsGroupHom (coHomGr (suc (3 +ℕ n +ℕ n)) + (Pushout (λ _ → tt) (fst f)) .snd) + (λ x → Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) + (invEq (fst SphereHopfCohomIso) x)) + (ℤGroup .snd) grHom = makeIsGroupHom λ x y → cong (Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) (IsGroupHom.pres· (isGroupHomInv SphereHopfCohomIso) x y) @@ -223,14 +207,19 @@ module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) whe Hⁿ-Sⁿ≅ℤ-nice-generator : (n : ℕ) → Iso.inv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) 1 ≡ ∣ ∣_∣ ∣₂ Hⁿ-Sⁿ≅ℤ-nice-generator zero = Iso.leftInv (fst (Hⁿ-Sⁿ≅ℤ (suc zero))) _ Hⁿ-Sⁿ≅ℤ-nice-generator (suc n) = - (λ i → Iso.inv (fst (suspensionAx-Sn (suc n) (suc n))) (Hⁿ-Sⁿ≅ℤ-nice-generator n i)) - ∙ cong ∣_∣₂ (funExt λ { north → refl - ; south → cong ∣_∣ₕ (merid north) - ; (merid a i) j → ∣ compPath-filler (merid a) (sym (merid north)) (~ j) i ∣ₕ}) + (λ i → Iso.inv (fst (suspensionAx-Sn (suc n) (suc n))) + (Hⁿ-Sⁿ≅ℤ-nice-generator n i)) + ∙ cong ∣_∣₂ + (funExt λ { north → refl + ; south → cong ∣_∣ₕ (merid north) + ; (merid a i) j → ∣ compPath-filler (merid a) + (sym (merid north)) (~ j) i ∣ₕ}) Hopfβ↦1 : (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) → Iso.fun (fst (Hopfβ-Iso n f)) (Hopfβ n f) ≡ 1 -Hopfβ↦1 n f = sym (cong (Iso.fun (fst (Hopfβ-Iso n f))) lem) ∙ Iso.rightInv (fst (Hopfβ-Iso n f)) (pos 1) +Hopfβ↦1 n f = + sym (cong (Iso.fun (fst (Hopfβ-Iso n f))) lem) + ∙ Iso.rightInv (fst (Hopfβ-Iso n f)) (pos 1) where lem : Iso.inv (fst (Hopfβ-Iso n f)) (pos 1) ≡ Hopfβ n f lem = (λ i → fst (MV.d _ _ _ (λ _ → tt) (fst f) (3 +ℕ n +ℕ n)) @@ -247,26 +236,34 @@ module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) whe H→Sphere : coHom 2+n H → coHom 2+n (S₊ (suc (suc n))) H→Sphere = sMap (_∘ inr) - grHom : IsGroupHom (snd (coHomGr 2+n H)) H→Sphere (snd (coHomGr 2+n (S₊ (suc (suc n))))) + grHom : IsGroupHom (snd (coHomGr 2+n H)) + H→Sphere (snd (coHomGr 2+n (S₊ (suc (suc n))))) grHom = makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ f g → refl) - preSphere→H : (g : (S₊ (suc (suc n)) → coHomK 2+n)) → H → coHomK (2 +ℕ n) + preSphere→H : (g : (S₊ (suc (suc n)) → coHomK 2+n)) + → H → coHomK (2 +ℕ n) preSphere→H g (inl x) = 0ₖ _ preSphere→H g (inr x) = g x -ₖ g north - preSphere→H g (push a i) = h23 a i + preSphere→H g (push a i) = lem a i where - h23 : (a : S₊ (suc (suc (suc (n +ℕ n))))) → 0ₖ (suc (suc n)) ≡ g (fst f a) -ₖ g north - h23 = sphereElim _ (λ x → isOfHLevelPlus' {n = n} (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n) _ _)) - (sym (rCancelₖ _ (g north)) ∙ cong (λ x → g x -ₖ g north) (sym (snd f))) + lem : (a : S₊ (suc (suc (suc (n +ℕ n))))) + → 0ₖ (suc (suc n)) ≡ g (fst f a) -ₖ g north + lem = + sphereElim _ + (λ x → isOfHLevelPlus' {n = n} (3 +ℕ n) + (isOfHLevelTrunc (4 +ℕ n) _ _)) + (sym (rCancelₖ _ (g north)) + ∙ cong (λ x → g x -ₖ g north) (sym (snd f))) Sphere→H : coHom 2+n (S₊ (suc (suc n))) → coHom 2+n H Sphere→H = sMap preSphere→H conCohom2+n : (x : _) → ∥ x ≡ 0ₖ (suc (suc n)) ∥ - conCohom2+n = coHomK-elim _ (λ _ → isProp→isOfHLevelSuc (suc n) squash) ∣ refl ∣ + conCohom2+n = + coHomK-elim _ (λ _ → isProp→isOfHLevelSuc (suc n) squash) ∣ refl ∣ HIPSphereCohomIso : Iso (coHom (2 +ℕ n) (HopfInvariantPush n (fst f))) (coHom (2 +ℕ n) ((S₊ (suc (suc n))))) @@ -275,37 +272,49 @@ module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) whe Iso.rightInv HIPSphereCohomIso = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ g → pRec (squash₂ _ _) - (λ p → cong ∣_∣₂ (funExt λ x → cong (g x +ₖ_) (cong (-ₖ_) p) ∙ rUnitₖ _ (g x))) + (λ p → cong ∣_∣₂ (funExt λ x → cong (g x +ₖ_) (cong (-ₖ_) p) + ∙ rUnitₖ _ (g x))) (conCohom2+n (g north)) Iso.leftInv HIPSphereCohomIso = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - λ g → pRec (squash₂ _ _) - (pRec (isPropΠ (λ _ → squash₂ _ _)) - (λ gn gtt → - trRec (isProp→isOfHLevelSuc n (squash₂ _ _)) - (λ r → cong ∣_∣₂ (funExt λ { - (inl x) → sym gtt - ; (inr x) → (λ i → g (inr x) -ₖ gn i) ∙ rUnitₖ _ (g (inr x)) - ; (push a i) - → sphereElim _ - {A = λ a → PathP (λ i → preSphere→H (λ x → g (inr x)) (push a i) ≡ g (push a i)) - (sym gtt) - ((λ i → g (inr (fst f a)) -ₖ gn i) ∙ rUnitₖ _ (g (inr (fst f a))))} - (λ _ → isOfHLevelPathP' (suc (suc (suc (n +ℕ n)))) - (isOfHLevelPath (suc (suc (suc (suc (n +ℕ n))))) - (isOfHLevelPlus' {n = n} (suc (suc (suc (suc n)))) - (isOfHLevelTrunc (suc (suc (suc (suc n)))))) _ _) _ _) - r a i})) - (push-helper g gtt gn)) - (conCohom2+n (g (inr north)))) - (conCohom2+n (g (inl tt))) + λ g → + pRec (squash₂ _ _) + (pRec (isPropΠ (λ _ → squash₂ _ _)) + (λ gn gtt → + trRec (isProp→isOfHLevelSuc n (squash₂ _ _)) + (λ r → cong ∣_∣₂ (funExt λ { + (inl x) → sym gtt + ; (inr x) → (λ i → g (inr x) -ₖ gn i) + ∙ rUnitₖ _ (g (inr x)) + ; (push a i) + → sphereElim _ + {A = λ a → + PathP (λ i → preSphere→H (λ x → g (inr x)) + (push a i) + ≡ g (push a i)) + (sym gtt) + ((λ i → g (inr (fst f a)) -ₖ gn i) + ∙ rUnitₖ _ (g (inr (fst f a))))} + (λ _ → isOfHLevelPathP' (suc (suc (suc (n +ℕ n)))) + (isOfHLevelPath (suc (suc (suc (suc (n +ℕ n))))) + (isOfHLevelPlus' {n = n} (suc (suc (suc (suc n)))) + (isOfHLevelTrunc (suc (suc (suc (suc n)))))) _ _) _ _) + r a i})) + (push-helper g gtt gn)) + (conCohom2+n (g (inr north)))) + (conCohom2+n (g (inl tt))) where push-helper : (g : HopfInvariantPush n (fst f) → coHomK (suc (suc n))) → (gtt : g (inl tt) ≡ 0ₖ (suc (suc n))) → (gn : g (inr north) ≡ 0ₖ (suc (suc n))) - → hLevelTrunc (suc n) (PathP (λ i → preSphere→H (λ x → g (inr x)) (push north i) ≡ g (push north i)) - (sym gtt) ((λ i → g (inr (fst f north)) -ₖ gn i) ∙ rUnitₖ _ (g (inr (fst f north))))) - push-helper g gtt gn = isConnectedPathP _ (isConnectedPath _ (isConnectedKn _) _ _) _ _ .fst + → hLevelTrunc (suc n) + (PathP (λ i → preSphere→H (λ x → g (inr x)) (push north i) + ≡ g (push north i)) + (sym gtt) + ((λ i → g (inr (fst f north)) -ₖ gn i) + ∙ rUnitₖ _ (g (inr (fst f north))))) + push-helper g gtt gn = + isConnectedPathP _ (isConnectedPath _ (isConnectedKn _) _ _) _ _ .fst Hopfα-Iso : GroupIso (coHomGr (2 +ℕ n) (HopfInvariantPush n (fst f))) ℤGroup Hopfα-Iso = @@ -338,3 +347,29 @@ HopfInvariant-π' n = sRec isSetℤ (HopfInvariant n) HopfInvariant-π : (n : ℕ) → π (3 +ℕ (n +ℕ n)) (S₊∙ (2 +ℕ n)) → ℤ HopfInvariant-π n = sRec isSetℤ λ x → HopfInvariant-π' n ∣ Ω→SphereMap _ x ∣₂ + + +-- Elimination principle for the pushout defining the Hopf Invariant +HopfInvariantPushElim : + ∀ {ℓ} n → (f : _) + → {P : HopfInvariantPush n f → Type ℓ} + → (isOfHLevel (suc (suc (suc (suc (n +ℕ n))))) (P (inl tt))) + → (e : P (inl tt)) + (g : (x : _) → P (inr x)) + (r : PathP (λ i → P (push north i)) e (g (f north))) + → (x : _) → P x +HopfInvariantPushElim n f {P = P} hlev e g r (inl x) = e +HopfInvariantPushElim n f {P = P} hlev e g r (inr x) = g x +HopfInvariantPushElim n f {P = P} hlev e g r (push a i₁) = help a i₁ + where + help : (a : Susp (Susp (S₊ (suc (n +ℕ n))))) + → PathP (λ i → P (push a i)) e (g (f a)) + help = + sphereElim _ + (sphereElim _ + (λ _ → isProp→isOfHLevelSuc (suc (suc (n +ℕ n))) + (isPropIsOfHLevel _)) + (isOfHLevelPathP' (suc (suc (suc (n +ℕ n)))) + (subst (isOfHLevel (suc (suc (suc (suc (n +ℕ n)))))) + (cong P (push north)) + hlev) _ _)) r diff --git a/Cubical/Homotopy/HopfInvariant/Homomorphism.agda b/Cubical/Homotopy/HopfInvariant/Homomorphism.agda index 6a5255fd1..c19962a98 100644 --- a/Cubical/Homotopy/HopfInvariant/Homomorphism.agda +++ b/Cubical/Homotopy/HopfInvariant/Homomorphism.agda @@ -1,10 +1,12 @@ {-# OPTIONS --safe --experimental-lossy-unification #-} module Cubical.Homotopy.HopfInvariant.Homomorphism where -open import Cubical.Homotopy.HopfInvariant.HopfMap open import Cubical.Homotopy.HopfInvariant.Base +open import Cubical.Homotopy.Group.Base + +open import Cubical.Relation.Nullary + open import Cubical.ZCohomology.Base -open import Cubical.ZCohomology.Groups.Connected open import Cubical.ZCohomology.GroupStructure open import Cubical.ZCohomology.Properties open import Cubical.ZCohomology.MayerVietorisUnreduced @@ -12,90 +14,49 @@ open import Cubical.ZCohomology.Groups.Unit open import Cubical.ZCohomology.Groups.Wedge open import Cubical.ZCohomology.Groups.Sn open import Cubical.ZCohomology.RingStructure.CupProduct -open import Cubical.ZCohomology.RingStructure.RingLaws open import Cubical.ZCohomology.RingStructure.GradedCommutativity -open import Cubical.Relation.Nullary open import Cubical.ZCohomology.Gysin -open import Cubical.Functions.Embedding - open import Cubical.Foundations.Path -open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.HLevels open import Cubical.Foundations.Transport open import Cubical.Foundations.Function -open import Cubical.Foundations.Univalence open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Pointed.Homogeneous - -open import Cubical.Homotopy.HomotopyGroup - -open import Cubical.Foundations.Univalence - -open import Cubical.Relation.Nullary -open import Cubical.Data.Sum -open import Cubical.Data.Fin -open import Cubical.Data.Empty renaming (rec to ⊥-rec) open import Cubical.Data.Sigma open import Cubical.Data.Int hiding (_+'_) open import Cubical.Data.Nat renaming (_+_ to _+ℕ_ ; _·_ to _·ℕ_) -open import Cubical.Data.Nat.Order open import Cubical.Data.Unit -open import Cubical.Data.Bool + open import Cubical.Algebra.Group - renaming (ℤ to ℤGroup ; Unit to UnitGroup) hiding (Bool) + renaming (ℤ to ℤGroup ; Unit to UnitGroup) open import Cubical.Algebra.Group.ZModule -open import Cubical.HITs.Pushout.Flattening -open import Cubical.Homotopy.Connected -open import Cubical.Homotopy.EilenbergSteenrod open import Cubical.HITs.Pushout open import Cubical.HITs.Sn open import Cubical.HITs.Susp -open import Cubical.HITs.S1 renaming (_·_ to _*_) +open import Cubical.HITs.Wedge open import Cubical.HITs.Truncation - renaming (rec to trRec ; elim to trElim ; elim2 to trElim2) open import Cubical.HITs.SetTruncation - renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap ; elim3 to sElim3) + renaming (elim to sElim ; elim2 to sElim2 ; map to sMap) open import Cubical.HITs.PropositionalTruncation - renaming (rec to pRec ; elim to pElim) - -open import Cubical.Algebra.AbGroup - -open import Cubical.Homotopy.Loopspace - -open import Cubical.HITs.Join - -open import Cubical.Homotopy.Hopf - -open import Cubical.HITs.SetQuotients renaming (_/_ to _/'_) - - -open import Cubical.Experiments.Brunerie -open import Cubical.HITs.Hopf -open import Cubical.HITs.Join - -open import Cubical.HITs.Wedge - --- probably to be moved to Wedge - -∨→∙ : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f g : S₊∙ n →∙ A) → ((S₊∙ n) ⋁ (S₊∙ n) , inl (ptSn _)) →∙ A -fst (∨→∙ {A = A} f g) = f ∨→ g -snd (∨→∙ {A = A} f g) = snd f +-- The pushout describing the hopf invariant of the multiplication (∙Π) of +-- two maps (S³⁺²ⁿ →∙ S²⁺ⁿ) +C·Π : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Type _ +C·Π n f g = Pushout (λ _ → tt) (∙Π f g .fst) +-- Another pushout, essentially the same, but starting with +-- S³⁺²ⁿ ∨ S³⁺²ⁿ. This will be used to break up the hopf invariant +-- of ∙Π f g. C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Type _ C* n f g = Pushout (λ _ → tt) (fst (∨→∙ f g)) -C·Π : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Type _ -C·Π n f g = Pushout (λ _ → tt) (∙Π f g .fst) - +-- The coequaliser of ∙Π and ∨→∙ θ : ∀ {ℓ} {A : Pointed ℓ} → Susp (fst A) → (Susp (fst A) , north) ⋁ (Susp (fst A) , north) θ north = inl north @@ -105,25 +66,15 @@ C·Π n f g = Pushout (λ _ → tt) (∙Π f g .fst) ∙∙ push tt ∙∙ λ i → inr ((merid a ∙ sym (merid (pt A))) i)) i₁ --- move to path - -doubleCompPath≡compPath : ∀ {ℓ} {A : Type ℓ} {x y z w : A} (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) - → p ∙∙ q ∙∙ r ≡ p ∙ q ∙ r -doubleCompPath≡compPath {x = x} {y = y} {z = z} {w = w} = - J (λ y p → (q : y ≡ z) (r : z ≡ w) → - (p ∙∙ q ∙∙ r) ≡ p ∙ q ∙ r) - λ q r → lUnit (q ∙ r) - - ∙Π≡∨→∘θ : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → (x : _) → ∙Π f g .fst x ≡ (f ∨→ g) (θ {A = S₊∙ _} x) + → (x : _) → ∙Π f g .fst x ≡ (f ∨→ g) (θ {A = S₊∙ _} x) ∙Π≡∨→∘θ n f g north = sym (snd f) ∙Π≡∨→∘θ n f g south = sym (snd g) ∙Π≡∨→∘θ n f g (merid a i₁) j = main j i₁ where help : cong (f ∨→ g) (cong (θ {A = S₊∙ _}) (merid a)) ≡ (refl ∙∙ cong (fst f) (merid a ∙ sym (merid north)) ∙∙ snd f) - ∙ (sym (snd g) ∙∙ cong (fst g) (merid a ∙ sym (merid north)) ∙∙ refl) + ∙ (sym (snd g) ∙∙ cong (fst g) (merid a ∙ sym (merid north)) ∙∙ refl) help = cong-∙∙ (f ∨→ g) ((λ i → inl ((merid a ∙ sym (merid north)) i))) (push tt) (λ i → inr ((merid a ∙ sym (merid north)) i)) @@ -134,10 +85,15 @@ doubleCompPath≡compPath {x = x} {y = y} {z = z} {w = w} = ∙∙ assoc _ _ _ ∙∙ cong₂ _∙_ refl (compPath≡compPath' _ _) - main : PathP (λ i → snd f (~ i) ≡ snd g (~ i)) (λ i₁ → ∙Π f g .fst (merid a i₁)) + main : PathP (λ i → snd f (~ i) + ≡ snd g (~ i)) (λ i₁ → ∙Π f g .fst (merid a i₁)) λ i₁ → (f ∨→ g) (θ {A = S₊∙ _} (merid a i₁)) - main = (λ i → ((λ j → snd f (~ i ∧ ~ j)) ∙∙ cong (fst f) (merid a ∙ sym (merid north)) ∙∙ snd f) - ∙ (sym (snd g) ∙∙ cong (fst g) (merid a ∙ sym (merid north)) ∙∙ λ j → snd g (~ i ∧ j))) + main = (λ i → ((λ j → snd f (~ i ∧ ~ j)) + ∙∙ cong (fst f) (merid a ∙ sym (merid north)) + ∙∙ snd f) + ∙ (sym (snd g) + ∙∙ cong (fst g) (merid a ∙ sym (merid north)) + ∙∙ λ j → snd g (~ i ∧ j))) ▷ sym help private @@ -156,59 +112,67 @@ private (subst (isOfHLevel ((3 +ℕ n) +ℕ n)) (cong P (push tt)) (isOfHLevelPlus' {n = n} (3 +ℕ n) (x _)))) (subst P (push tt) s) x₁ - WedgeElim n {P = P} x s (push a j) = transp (λ i → P (push tt (i ∧ j))) (~ j) s + WedgeElim n {P = P} x s (push a j) = + transp (λ i → P (push tt (i ∧ j))) (~ j) s H²C*≅ℤ : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → GroupIso (coHomGr (2 +ℕ n) (C* n f g)) ℤGroup -H²C*≅ℤ n f g = compGroupIso groupIso1 (Hⁿ-Sⁿ≅ℤ (suc n)) +H²C*≅ℤ n f g = compGroupIso is (Hⁿ-Sⁿ≅ℤ (suc n)) where - help : (coHom (2 +ℕ n) (C* n f g)) → coHom (2 +ℕ n) (S₊ (2 +ℕ n)) - help = sMap λ f → f ∘ inr - + ∘inr : (coHom (2 +ℕ n) (C* n f g)) → coHom (2 +ℕ n) (S₊ (2 +ℕ n)) + ∘inr = sMap λ f → f ∘ inr - invMapPrim : (S₊ (2 +ℕ n) → coHomK (2 +ℕ n)) → C* n f g → coHomK (2 +ℕ n) + invMapPrim : (S₊ (2 +ℕ n) → coHomK (2 +ℕ n)) + → C* n f g → coHomK (2 +ℕ n) invMapPrim h (inl x) = h (ptSn _) invMapPrim h (inr x) = h x - invMapPrim h (push a i₁) = WedgeElim n {P = λ a → h north ≡ h (fst (∨→∙ f g) a)} - (λ _ → isOfHLevelTrunc (4 +ℕ n) _ _) - (cong h (sym (snd f))) a i₁ + invMapPrim h (push a i₁) = + WedgeElim n {P = λ a → h north ≡ h (fst (∨→∙ f g) a)} + (λ _ → isOfHLevelTrunc (4 +ℕ n) _ _) + (cong h (sym (snd f))) a i₁ invMap : coHom (2 +ℕ n) (S₊ (2 +ℕ n)) → (coHom (2 +ℕ n) (C* n f g)) invMap = sMap invMapPrim - ret1 : (x : coHom (2 +ℕ n) (S₊ (2 +ℕ n))) → help (invMap x) ≡ x - ret1 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + ∘inrSect : (x : coHom (2 +ℕ n) (S₊ (2 +ℕ n))) → ∘inr (invMap x) ≡ x + ∘inrSect = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ a → refl - ret2 : (x : (coHom (2 +ℕ n) (C* n f g))) → invMap (help x) ≡ x - ret2 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - λ h → cong ∣_∣₂ (funExt λ { (inl x) → cong h ((λ i → inr ((snd f) (~ i))) - ∙ sym (push (inl north))) - ; (inr x) → refl - ; (push a i₁) → help2 h a i₁}) + ∘inrRetr : (x : (coHom (2 +ℕ n) (C* n f g))) → invMap (∘inr x) ≡ x + ∘inrRetr = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ h → cong ∣_∣₂ (funExt λ { (inl x) → cong h ((λ i → inr ((snd f) (~ i))) + ∙ sym (push (inl north))) + ; (inr x) → refl + ; (push a i) → lem₁ h a i}) where - help2 : (h : C* n f g → coHomK (2 +ℕ n)) - → (a : _) → PathP (λ i → invMapPrim (h ∘ inr) (push a i) ≡ h (push a i)) - (cong h ((λ i → inr ((snd f) (~ i))) ∙ sym (push (inl north)))) + lem₁ : (h : C* n f g → coHomK (2 +ℕ n)) (a : _) + → PathP (λ i → invMapPrim (h ∘ inr) (push a i) ≡ h (push a i)) + (cong h ((λ i → inr ((snd f) (~ i))) + ∙ sym (push (inl north)))) refl - help2 h = WedgeElim n (λ _ → isOfHLevelPathP (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n) _ _) _ _) - help4 + lem₁ h = + WedgeElim n (λ _ → isOfHLevelPathP (3 +ℕ n) + (isOfHLevelTrunc (4 +ℕ n) _ _) _ _) + lem₂ where - help4 : PathP (λ i → invMapPrim (h ∘ inr) (push (inl north) i) ≡ h (push (inl north) i)) - (cong h ((λ i → inr ((snd f) (~ i))) ∙ sym (push (inl north)))) + lem₂ : PathP (λ i → invMapPrim (h ∘ inr) (push (inl north) i) + ≡ h (push (inl north) i)) + (cong h ((λ i → inr ((snd f) (~ i))) + ∙ sym (push (inl north)))) refl - help4 i j = h (hcomp (λ k → λ { (i = i1) → inr (fst f north) + lem₂ i j = h (hcomp (λ k → λ { (i = i1) → inr (fst f north) ; (j = i0) → inr (snd f (~ i)) ; (j = i1) → push (inl north) (i ∨ ~ k)}) (inr (snd f (~ i ∧ ~ j)))) - groupIso1 : GroupIso ((coHomGr (2 +ℕ n) (C* n f g))) (coHomGr (2 +ℕ n) (S₊ (2 +ℕ n))) - Iso.fun (fst groupIso1) = help - Iso.inv (fst groupIso1) = invMap - Iso.rightInv (fst groupIso1) = ret1 - Iso.leftInv (fst groupIso1) = ret2 - snd groupIso1 = makeIsGroupHom + is : GroupIso (coHomGr (2 +ℕ n) (C* n f g)) (coHomGr (2 +ℕ n) (S₊ (2 +ℕ n))) + Iso.fun (fst is) = ∘inr + Iso.inv (fst is) = invMap + Iso.rightInv (fst is) = ∘inrSect + Iso.leftInv (fst is) = ∘inrRetr + snd is = makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ f g → refl) @@ -230,27 +194,37 @@ H⁴-C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n) H⁴-C* n f g = compGroupIso (GroupEquiv→GroupIso (invGroupEquiv fstIso)) (compGroupIso (transportCohomIso (cong (2 +ℕ_) (+-suc n n))) - (compGroupIso (Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) _) + (compGroupIso (Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) + (S₊∙ (suc (suc (suc (n +ℕ n))))) _) (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ _) (Hⁿ-Sⁿ≅ℤ _)))) where module Ms = MV _ _ _ (λ _ → tt) (fst (∨→∙ f g)) - fstIso : GroupEquiv (coHomGr ((suc (suc (n +ℕ suc n)))) (S₊∙ (3 +ℕ (n +ℕ n)) ⋁ S₊∙ (3 +ℕ (n +ℕ n)))) + fstIso : GroupEquiv (coHomGr ((suc (suc (n +ℕ suc n)))) + (S₊∙ (3 +ℕ (n +ℕ n)) ⋁ S₊∙ (3 +ℕ (n +ℕ n)))) (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) fst (fst fstIso) = Ms.d (suc (suc (n +ℕ suc n))) .fst snd (fst fstIso) = SES→isEquiv - (GroupPath _ _ .fst (compGroupEquiv (invEquiv (isContr→≃Unit ((tt , tt) , (λ _ → refl))) , makeIsGroupHom λ _ _ → refl) - (GroupEquivDirProd - (GroupIso→GroupEquiv (invGroupIso (Hⁿ-Unit≅0 _))) - (GroupIso→GroupEquiv - (invGroupIso - (Hⁿ-Sᵐ≅0 _ _ λ p → ¬lemHopf 0 ((λ _ → north) , refl) n n (cong suc (sym (+-suc n n)) ∙ p))))))) + (sym (fst (GroupPath _ _) + (isContr→≃Unit + (isContrΣ (subst isContr (isoToPath (invIso (fst (Hⁿ-Unit≅0 _)))) + isContrUnit) + (λ _ → subst isContr (isoToPath (invIso + (fst (Hⁿ-Sᵐ≅0 _ _ λ p → + ¬lemHopf 0 ((λ _ → north) , refl) + n n (cong suc (sym (+-suc n n)) ∙ p))))) + isContrUnit)) + , makeIsGroupHom λ _ _ → refl))) (GroupPath _ _ .fst - (compGroupEquiv ((invEquiv (isContr→≃Unit ((tt , tt) , (λ _ → refl))) , makeIsGroupHom λ _ _ → refl)) - ((GroupEquivDirProd - (GroupIso→GroupEquiv (invGroupIso (Hⁿ-Unit≅0 _))) - (GroupIso→GroupEquiv - (invGroupIso (Hⁿ-Sᵐ≅0 _ _ λ p → ¬lemHopf 0 ((λ _ → north) , refl) n (suc n) (cong (2 +ℕ_) (sym (+-suc n n)) ∙ p)))))))) + (compGroupEquiv (invEquiv (isContr→≃Unit ((tt , tt) , (λ _ → refl))) + , makeIsGroupHom λ _ _ → refl) + (GroupEquivDirProd + (GroupIso→GroupEquiv (invGroupIso (Hⁿ-Unit≅0 _))) + (GroupIso→GroupEquiv + (invGroupIso (Hⁿ-Sᵐ≅0 _ _ λ p → + ¬lemHopf 0 ((λ _ → north) , refl) + n (suc n) (cong (2 +ℕ_) (sym (+-suc n n)) + ∙ p))))))) (Ms.Δ (suc (suc (n +ℕ suc n)))) (Ms.d (suc (suc (n +ℕ suc n)))) (Ms.i (suc (suc (suc (n +ℕ suc n))))) @@ -293,27 +267,32 @@ module _ (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) βₗ'-fun (push (push a i₂) i₁) = Kn→ΩKn+10ₖ _ i₂ i₁ βₗ''-fun : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) - βₗ''-fun = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + βₗ''-fun = subst (λ m → coHom m (C* n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) ∣ βₗ'-fun ∣₂ βᵣ''-fun : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) - βᵣ''-fun = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + βᵣ''-fun = subst (λ m → coHom m (C* n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) ∣ βᵣ'-fun ∣₂ private - pp : (a : _) → 0ₖ (suc (suc n)) ≡ ∣ (f ∨→ g) a ∣ - pp = WedgeElim _ (λ _ → isOfHLevelTrunc (4 +ℕ n) _ _) + 0ₖ≡∨→ : (a : _) → 0ₖ (suc (suc n)) ≡ ∣ (f ∨→ g) a ∣ + 0ₖ≡∨→ = WedgeElim _ (λ _ → isOfHLevelTrunc (4 +ℕ n) _ _) (cong ∣_∣ₕ (sym (snd f))) - pp-inr : pp (inr north) ≡ cong ∣_∣ₕ (sym (snd g)) - pp-inr = (λ j → transport (λ i → 0ₖ (2 +ℕ n) ≡ ∣ compPath-filler' (snd f) (sym (snd g)) (~ j) i ∣ₕ) - λ i → ∣ snd f (~ i ∨ j) ∣ₕ) - ∙ λ j → transp (λ i → 0ₖ (2 +ℕ n) ≡ ∣ snd g (~ i ∧ ~ j) ∣ₕ) j λ i → ∣ snd g (~ i ∨ ~ j) ∣ₕ + 0ₖ≡∨→-inr : 0ₖ≡∨→ (inr north) ≡ cong ∣_∣ₕ (sym (snd g)) + 0ₖ≡∨→-inr = + (λ j → transport (λ i → 0ₖ (2 +ℕ n) + ≡ ∣ compPath-filler' (snd f) (sym (snd g)) (~ j) i ∣ₕ) + λ i → ∣ snd f (~ i ∨ j) ∣ₕ) + ∙ λ j → transp (λ i → 0ₖ (2 +ℕ n) ≡ ∣ snd g (~ i ∧ ~ j) ∣ₕ) j + λ i → ∣ snd g (~ i ∨ ~ j) ∣ₕ α*'-fun : C* n f g → coHomK (2 +ℕ n) α*'-fun (inl x) = 0ₖ _ α*'-fun (inr x) = ∣ x ∣ - α*'-fun (push a i₁) = pp a i₁ + α*'-fun (push a i₁) = 0ₖ≡∨→ a i₁ α*' : coHom (2 +ℕ n) (C* n f g) α*' = ∣ α*'-fun ∣₂ @@ -334,139 +313,165 @@ module _ (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) q (inr x) = inr x q (push a i₁) = (push (θ a) ∙ λ i → inr (∙Π≡∨→∘θ n f g a (~ i))) i₁ -α*↦1 : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +α↦1 : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Iso.fun (fst (H²C*≅ℤ n f g)) (α*' n f g) ≡ 1 -α*↦1 n f g = sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (Hⁿ-Sⁿ≅ℤ-nice-generator n)) - ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) +α↦1 n f g = + sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (Hⁿ-Sⁿ≅ℤ-nice-generator n)) + ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) -α*-lem : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) +α≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → α* n f g ≡ α*' n f g -α*-lem n f g = sym (Iso.leftInv (fst (H²C*≅ℤ n f g)) _) - ∙∙ cong (Iso.inv (fst (H²C*≅ℤ n f g))) help +α≡ n f g = sym (Iso.leftInv (fst (H²C*≅ℤ n f g)) _) + ∙∙ cong (Iso.inv (fst (H²C*≅ℤ n f g))) lem ∙∙ Iso.leftInv (fst (H²C*≅ℤ n f g)) _ where - help : Iso.fun (fst (H²C*≅ℤ n f g)) (α* n f g) ≡ Iso.fun (fst (H²C*≅ℤ n f g)) (α*' n f g) - help = (Iso.rightInv (fst (H²C*≅ℤ n f g)) (pos 1)) ∙ sym (α*↦1 n f g) + lem : Iso.fun (fst (H²C*≅ℤ n f g)) (α* n f g) + ≡ Iso.fun (fst (H²C*≅ℤ n f g)) (α*' n f g) + lem = (Iso.rightInv (fst (H²C*≅ℤ n f g)) (pos 1)) ∙ sym (α↦1 n f g) q-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → coHomFun _ (q n f g) (α* n f g) ≡ Hopfα n (∙Π f g) -q-α n f g = (λ i → coHomFun _ (q n f g) (α*-lem n f g i)) +q-α n f g = (λ i → coHomFun _ (q n f g) (α≡ n f g i)) ∙ sym (Iso.leftInv is _) - ∙∙ cong (Iso.inv is) help + ∙∙ cong (Iso.inv is) lem ∙∙ Iso.leftInv is _ where is = fst (Hopfα-Iso n (∙Π f g)) - help : Iso.fun is (coHomFun _ (q n f g) (α*' n f g)) + lem : Iso.fun is (coHomFun _ (q n f g) (α*' n f g)) ≡ Iso.fun is (Hopfα n (∙Π f g)) - help = sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (Hⁿ-Sⁿ≅ℤ-nice-generator n)) + lem = sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (Hⁿ-Sⁿ≅ℤ-nice-generator n)) ∙∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) ∙∙ sym (Hopfα-Iso-α n (∙Π f g)) - βₗ≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → βₗ n f g ≡ βₗ''-fun n f g -βₗ≡ n f g = cong (FF ∘ subber2) +βₗ≡ n f g = cong (∨-d ∘ subber₂) (cong (Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) (((suc (suc (n +ℕ n))))))))) help - ∙ help2) - ∙ funExt⁻ FF∘subber ∣ wedgeGen ∣₂ - ∙ cong subber3 (sym βₗ'-fun≡) + ∙ lem) + ∙ funExt⁻ ∨-d∘subber ∣ wedgeGen ∣₂ + ∙ cong subber₃ (sym βₗ'-fun≡) where - FF = MV.d _ _ _ (λ _ → tt) (fst (∨→∙ f g)) (suc (suc (n +ℕ suc n))) .fst - FF' = MV.d _ _ _ (λ _ → tt) (fst (∨→∙ f g)) (suc (suc (suc (n +ℕ n)))) .fst + ∨-d = MV.d _ _ _ (λ _ → tt) (fst (∨→∙ f g)) (suc (suc (n +ℕ suc n))) .fst + ∨-d' = MV.d _ _ _ (λ _ → tt) (fst (∨→∙ f g)) (suc (suc (suc (n +ℕ n)))) .fst - help : Iso.inv (fst (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (1 , 0) + help : Iso.inv (fst (GroupIsoDirProd + (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) + (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (1 , 0) ≡ (∣ ∣_∣ ∣₂ , 0ₕ _) - help = ΣPathP ((Hⁿ-Sⁿ≅ℤ-nice-generator _) , IsGroupHom.pres1 (snd (invGroupIso (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) - - subber3 = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - - subber2 = (subst (λ m → coHom m (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) - (sym (cong (2 +ℕ_) (+-suc n n)))) - - FF∘subber : FF ∘ subber2 ≡ subber3 ∘ FF' - FF∘subber = - funExt (λ a → (λ j → transp (λ i → coHom ((suc ∘ suc ∘ suc) (+-suc n n (~ i ∧ j))) (C* n f g)) (~ j) - (MV.d _ _ _ (λ _ → tt) (fst (∨→∙ f g)) ((suc (suc (+-suc n n j)))) .fst - (transp (λ i → coHom (2 +ℕ (+-suc n n (j ∨ ~ i))) - (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) j - a)))) - - wedgeGen : (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) → - coHomK (suc (suc (suc (n +ℕ n))))) + help = ΣPathP ((Hⁿ-Sⁿ≅ℤ-nice-generator _) + , IsGroupHom.pres1 (snd (invGroupIso (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) + + subber₃ = subst (λ m → coHom m (C* n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + + subber₂ = (subst (λ m → coHom m (S₊∙ (suc (suc (suc (n +ℕ n)))) + ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) + (sym (cong (2 +ℕ_) (+-suc n n)))) + + ∨-d∘subber : ∨-d ∘ subber₂ ≡ subber₃ ∘ ∨-d' + ∨-d∘subber = + funExt (λ a → + (λ j → transp (λ i → coHom ((suc ∘ suc ∘ suc) (+-suc n n (~ i ∧ j))) + (C* n f g)) (~ j) + (MV.d _ _ _ (λ _ → tt) (fst (∨→∙ f g)) + (suc (suc (+-suc n n j))) .fst + (transp (λ i → coHom (2 +ℕ (+-suc n n (j ∨ ~ i))) + (S₊∙ (suc (suc (suc (n +ℕ n)))) + ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) j + a)))) + + wedgeGen : (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) + → coHomK (suc (suc (suc (n +ℕ n))))) wedgeGen (inl x) = ∣ x ∣ wedgeGen (inr x) = 0ₖ _ wedgeGen (push a i₁) = 0ₖ _ - βₗ'-fun≡ : ∣ βₗ'-fun n f g ∣₂ ≡ FF' ∣ wedgeGen ∣₂ - βₗ'-fun≡ = cong ∣_∣₂ (funExt λ { (inl x) → refl - ; (inr x) → refl - ; (push (inl x) i₁) → refl - ; (push (inr x) i₁) j → Kn→ΩKn+10ₖ _ (~ j) i₁ - ; (push (push a i₂) i₁) j → Kn→ΩKn+10ₖ _ (~ j ∧ i₂) i₁}) - - help2 : Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) (((suc (suc (n +ℕ n)))))))) + βₗ'-fun≡ : ∣ βₗ'-fun n f g ∣₂ ≡ ∨-d' ∣ wedgeGen ∣₂ + βₗ'-fun≡ = + cong ∣_∣₂ (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push (inl x) i₁) → refl + ; (push (inr x) i₁) j → Kn→ΩKn+10ₖ _ (~ j) i₁ + ; (push (push a i₂) i₁) j → Kn→ΩKn+10ₖ _ (~ j ∧ i₂) i₁}) + + lem : Iso.inv (fst (Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) + (S₊∙ (suc (suc (suc (n +ℕ n))))) + (((suc (suc (n +ℕ n))))))) (∣ ∣_∣ ∣₂ , 0ₕ _) ≡ ∣ wedgeGen ∣₂ - help2 = cong ∣_∣₂ (funExt λ { (inl x) → rUnitₖ (suc (suc (suc (n +ℕ n)))) ∣ x ∣ - ; (inr x) → refl - ; (push a i₁) → refl}) + lem = cong ∣_∣₂ (funExt λ { (inl x) → rUnitₖ (suc (suc (suc (n +ℕ n)))) ∣ x ∣ + ; (inr x) → refl + ; (push a i₁) → refl}) βᵣ≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → βᵣ n f g ≡ βᵣ''-fun n f g βᵣ≡ n f g = - cong (FF ∘ subber2) - (cong (Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) - (S₊∙ (suc (suc (suc (n +ℕ n))))) - (((suc (suc (n +ℕ n))))))))) + cong (∨-d ∘ subber₂) + (cong (Iso.inv (fst (Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) + (S₊∙ (suc (suc (suc (n +ℕ n))))) + (suc (suc (n +ℕ n)))))) help - ∙ help2) - ∙ funExt⁻ FF∘subber ∣ wedgeGen ∣₂ - ∙ cong (subber3) (sym βᵣ'-fun≡) + ∙ lem) + ∙ funExt⁻ ∨-d∘subber ∣ wedgeGen ∣₂ + ∙ cong (subber₃) (sym βᵣ'-fun≡) where - FF = MV.d _ _ _ (λ _ → tt) (fst (∨→∙ f g)) (suc (suc (n +ℕ suc n))) .fst - FF' = MV.d _ _ _ (λ _ → tt) (fst (∨→∙ f g)) (suc (suc (suc (n +ℕ n)))) .fst + ∨-d = MV.d _ _ _ (λ _ → tt) (fst (∨→∙ f g)) (suc (suc (n +ℕ suc n))) .fst + ∨-d' = MV.d _ _ _ (λ _ → tt) (fst (∨→∙ f g)) (suc (suc (suc (n +ℕ n)))) .fst - help : Iso.inv (fst (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (0 , 1) + help : Iso.inv (fst (GroupIsoDirProd + (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) + (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (0 , 1) ≡ (0ₕ _ , ∣ ∣_∣ ∣₂) - help = ΣPathP (IsGroupHom.pres1 (snd (invGroupIso (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) , (Hⁿ-Sⁿ≅ℤ-nice-generator _)) - - subber3 = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - - subber2 = (subst (λ m → coHom m (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) - (sym (cong (2 +ℕ_) (+-suc n n)))) - - FF∘subber : FF ∘ subber2 ≡ subber3 ∘ FF' - FF∘subber = - funExt (λ a → (λ j → transp (λ i → coHom ((suc ∘ suc ∘ suc) (+-suc n n (~ i ∧ j))) (C* n f g)) (~ j) - (MV.d _ _ _ (λ _ → tt) (fst (∨→∙ f g)) ((suc (suc (+-suc n n j)))) .fst - (transp (λ i → coHom (2 +ℕ (+-suc n n (j ∨ ~ i))) - (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) j - a)))) - - wedgeGen : (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) → - coHomK (suc (suc (suc (n +ℕ n))))) + help = + ΣPathP (IsGroupHom.pres1 (snd (invGroupIso (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) + , (Hⁿ-Sⁿ≅ℤ-nice-generator _)) + + subber₃ = subst (λ m → coHom m (C* n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + + subber₂ = (subst (λ m → coHom m (S₊∙ (suc (suc (suc (n +ℕ n)))) + ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) + (sym (cong (2 +ℕ_) (+-suc n n)))) + + ∨-d∘subber : ∨-d ∘ subber₂ ≡ subber₃ ∘ ∨-d' + ∨-d∘subber = + funExt (λ a → + (λ j → transp (λ i → coHom ((suc ∘ suc ∘ suc) (+-suc n n (~ i ∧ j))) + (C* n f g)) (~ j) + (MV.d _ _ _ (λ _ → tt) + (fst (∨→∙ f g)) (suc (suc (+-suc n n j))) .fst + (transp (λ i → coHom (2 +ℕ (+-suc n n (j ∨ ~ i))) + (S₊∙ (suc (suc (suc (n +ℕ n)))) + ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) j + a)))) + + wedgeGen : (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) + → coHomK (suc (suc (suc (n +ℕ n))))) wedgeGen (inl x) = 0ₖ _ wedgeGen (inr x) = ∣ x ∣ wedgeGen (push a i₁) = 0ₖ _ - help2 : Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) - (S₊∙ (suc (suc (suc (n +ℕ n))))) (((suc (suc (n +ℕ n)))))))) + lem : Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) + (S₊∙ (suc (suc (suc (n +ℕ n))))) + (suc (suc (n +ℕ n)))))) (0ₕ _ , ∣ ∣_∣ ∣₂) ≡ ∣ wedgeGen ∣₂ - help2 = cong ∣_∣₂ (funExt λ { (inl x) → refl - ; (inr x) → lUnitₖ (suc (suc (suc (n +ℕ n)))) ∣ x ∣ - ; (push a i₁) → refl}) - - βᵣ'-fun≡ : ∣ βᵣ'-fun n f g ∣₂ ≡ FF' ∣ wedgeGen ∣₂ - βᵣ'-fun≡ = cong ∣_∣₂ (funExt λ { (inl x) → refl - ; (inr x) → refl - ; (push (inl x) i₁) j → Kn→ΩKn+10ₖ _ (~ j) i₁ - ; (push (inr x) i₁) → refl - ; (push (push a i₂) i₁) j → Kn→ΩKn+10ₖ _ (~ j ∧ ~ i₂) i₁}) + lem = cong ∣_∣₂ (funExt λ { (inl x) → refl + ; (inr x) → lUnitₖ (suc (suc (suc (n +ℕ n)))) ∣ x ∣ + ; (push a i₁) → refl}) + + βᵣ'-fun≡ : ∣ βᵣ'-fun n f g ∣₂ ≡ ∨-d' ∣ wedgeGen ∣₂ + βᵣ'-fun≡ = + cong ∣_∣₂ (funExt + λ { (inl x) → refl + ; (inr x) → refl + ; (push (inl x) i₁) j → Kn→ΩKn+10ₖ _ (~ j) i₁ + ; (push (inr x) i₁) → refl + ; (push (push a i₂) i₁) j → Kn→ΩKn+10ₖ _ (~ j ∧ ~ i₂) i₁}) q-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → coHomFun _ (q n f g) (βₗ n f g) @@ -475,58 +480,76 @@ q-βₗ n f g = cong (coHomFun _ (q n f g)) (βₗ≡ n f g) ∙ transportLem ∙ cong (subst (λ m → coHom m (C·Π' n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) - (sym (Iso.leftInv (fst (Hopfβ-Iso n (∙Π f g))) (Hopfβ n (∙Π f g))) - ∙ cong (Iso.inv ((fst (Hopfβ-Iso n (∙Π f g))))) (Hopfβ↦1 n (∙Π f g))) + (sym (Iso.leftInv (fst (Hopfβ-Iso n (∙Π f g))) + (Hopfβ n (∙Π f g))) + ∙ cong (Iso.inv ((fst (Hopfβ-Iso n (∙Π f g))))) + (Hopfβ↦1 n (∙Π f g))) where transportLem : coHomFun (suc (suc (suc (n +ℕ suc n)))) (q n f g) (βₗ''-fun n f g) - ≡ subst (λ m → coHom m (C·Π' n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ≡ subst (λ m → coHom m (C·Π' n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) (Hopfβ n (∙Π f g)) transportLem = - natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (q n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (q n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) ∙ cong (subst (λ m → coHom m (C·Π' n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) - (cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → loll a i₁})) + (cong ∣_∣₂ (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a i₁) → push-lem a i₁})) where - open import Cubical.Foundations.Path - loll : (x : fst (S₊∙ (3 +ℕ (n +ℕ n)))) → - PathP - (λ i₁ → - βₗ'-fun n f g ((push (θ x) ∙ λ i → inr (∙Π≡∨→∘θ n f g x (~ i))) i₁) ≡ - MV.d-pre Unit (fst (S₊∙ (2 +ℕ n))) (fst (S₊∙ (3 +ℕ n +ℕ n))) - (λ _ → tt) (fst (∙Π f g)) (3 +ℕ n +ℕ n) ∣_∣ (push x i₁)) + push-lem : (x : fst (S₊∙ (3 +ℕ (n +ℕ n)))) → + PathP (λ i₁ → βₗ'-fun n f g ((push (θ x) + ∙ λ i → inr (∙Π≡∨→∘θ n f g x (~ i))) i₁) + ≡ MV.d-pre Unit (fst (S₊∙ (2 +ℕ n))) (fst (S₊∙ (3 +ℕ n +ℕ n))) + (λ _ → tt) (fst (∙Π f g)) (3 +ℕ n +ℕ n) ∣_∣ (push x i₁)) (λ _ → βₗ'-fun n f g (q n f g (inl tt))) (λ _ → βₗ'-fun n f g (q n f g (inr (∙Π f g .fst x)))) - loll x = - flipSquare (cong-∙ (βₗ'-fun n f g) (push (θ x)) (λ i → inr (∙Π≡∨→∘θ n f g x (~ i))) + push-lem x = + flipSquare (cong-∙ (βₗ'-fun n f g) + (push (θ x)) (λ i → inr (∙Π≡∨→∘θ n f g x (~ i))) ∙∙ sym (rUnit _) ∙∙ lem₁ x) where - lem₁ : (x : _) → cong (βₗ'-fun n f g) (push (θ {A = S₊∙ _} x)) ≡ Kn→ΩKn+1 _ ∣ x ∣ + lem₁ : (x : _) → cong (βₗ'-fun n f g) (push (θ {A = S₊∙ _} x)) + ≡ Kn→ΩKn+1 _ ∣ x ∣ lem₁ north = refl - lem₁ south = sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)) - lem₁ (merid a j) k = lala k j + lem₁ south = + sym (Kn→ΩKn+10ₖ _) + ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)) + lem₁ (merid a j) k = merid-lem k j where - lala : PathP (λ k → Kn→ΩKn+1 _ ∣ north ∣ₕ ≡ (sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north))) k) - (λ j → cong (βₗ'-fun n f g) (push (θ {A = S₊∙ _} (merid a j)))) - (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) - lala = (cong-∙∙ (cong (βₗ'-fun n f g) ∘ push) - ((λ i₁ → inl ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁))) - (push tt) - ((λ i₁ → inr ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁))) - ∙ sym (compPath≡compPath' _ _) - ∙ (λ _ → (λ j → Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∣ (merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) j ∣ₕ) - ∙ Kn→ΩKn+10ₖ _) - ∙ cong (_∙ Kn→ΩKn+10ₖ _) (cong-∙ ((Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ)) - (merid a) (sym (merid north))) - ∙ sym (assoc _ _ _) - ∙ cong (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a) ∙_) - (sym (symDistr (sym ((Kn→ΩKn+10ₖ _))) (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)))))) - ◁ λ i j → compPath-filler (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) - (sym (sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) - (cong ∣_∣ₕ (merid north)))) - (~ i) j + p = pt (S₊∙ (suc (suc (n +ℕ n)))) + + merid-lem : + PathP (λ k → Kn→ΩKn+1 _ ∣ north ∣ₕ + ≡ (sym (Kn→ΩKn+10ₖ _) + ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) + (cong ∣_∣ₕ (merid north))) k) + (λ j → cong (βₗ'-fun n f g) (push (θ {A = S₊∙ _} (merid a j)))) + (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) + merid-lem = + (cong-∙∙ (cong (βₗ'-fun n f g) ∘ push) + (λ i₁ → inl ((merid a ∙ (sym (merid p))) i₁)) + (push tt) + ((λ i₁ → inr ((merid a ∙ (sym (merid p))) i₁))) + ∙ sym (compPath≡compPath' _ _) + ∙ cong (_∙ Kn→ΩKn+10ₖ _) + (cong-∙ ((Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ)) + (merid a) (sym (merid north))) + ∙ sym (assoc _ _ _) + ∙ cong (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a) ∙_) + (sym (symDistr (sym ((Kn→ΩKn+10ₖ _))) + (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) + (cong ∣_∣ₕ (merid north)))))) + ◁ λ i j → compPath-filler + (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) + (sym (sym (Kn→ΩKn+10ₖ _) + ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) + (cong ∣_∣ₕ (merid north)))) + (~ i) j q-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → coHomFun _ (q n f g) (βᵣ n f g) @@ -535,23 +558,26 @@ q-βᵣ n f g = cong (coHomFun _ (q n f g)) (βᵣ≡ n f g) ∙ transportLem ∙ cong (subst (λ m → coHom m (C·Π' n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) - (sym (Iso.leftInv (fst (Hopfβ-Iso n (∙Π f g))) (Hopfβ n (∙Π f g))) - ∙ cong (Iso.inv ((fst (Hopfβ-Iso n (∙Π f g))))) (Hopfβ↦1 n (∙Π f g))) + (sym (Iso.leftInv (fst (Hopfβ-Iso n (∙Π f g))) + (Hopfβ n (∙Π f g))) + ∙ cong (Iso.inv ((fst (Hopfβ-Iso n (∙Π f g))))) + (Hopfβ↦1 n (∙Π f g))) where transportLem : coHomFun (suc (suc (suc (n +ℕ suc n)))) (q n f g) (βᵣ''-fun n f g) - ≡ subst (λ m → coHom m (C·Π' n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ≡ subst (λ m → coHom m (C·Π' n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) (Hopfβ n (∙Π f g)) transportLem = - natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (q n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (q n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) ∙ cong (subst (λ m → coHom m (C·Π' n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) (cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl - ; (push a i₁) → loll a i₁})) + ; (push a i₁) → push-lem a i₁})) where - open import Cubical.Foundations.Path - loll : (x : fst (S₊∙ (3 +ℕ (n +ℕ n)))) → + push-lem : (x : fst (S₊∙ (3 +ℕ (n +ℕ n)))) → PathP (λ i₁ → βᵣ'-fun n f g ((push (θ x) ∙ λ i → inr (∙Π≡∨→∘θ n f g x (~ i))) i₁) ≡ @@ -559,37 +585,53 @@ q-βᵣ n f g = cong (coHomFun _ (q n f g)) (βᵣ≡ n f g) (λ _ → tt) (fst (∙Π f g)) (3 +ℕ n +ℕ n) ∣_∣ (push x i₁)) (λ _ → βᵣ'-fun n f g (q n f g (inl tt))) (λ _ → βᵣ'-fun n f g (q n f g (inr (∙Π f g .fst x)))) - loll x = flipSquare (cong-∙ (βᵣ'-fun n f g) (push (θ x)) (λ i → inr (∙Π≡∨→∘θ n f g x (~ i))) - ∙∙ sym (rUnit _) - ∙∙ lem₁ x) + push-lem x = flipSquare (cong-∙ (βᵣ'-fun n f g) (push (θ x)) + (λ i → inr (∙Π≡∨→∘θ n f g x (~ i))) + ∙∙ sym (rUnit _) + ∙∙ lem₁ x) where - lem₁ : (x : _) → cong (βᵣ'-fun n f g) (push (θ {A = S₊∙ _} x)) ≡ Kn→ΩKn+1 _ ∣ x ∣ + lem₁ : (x : _) → cong (βᵣ'-fun n f g) (push (θ {A = S₊∙ _} x)) + ≡ Kn→ΩKn+1 _ ∣ x ∣ lem₁ north = sym (Kn→ΩKn+10ₖ _) - lem₁ south = cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)) - lem₁ (merid a j) k = lala k j -- lala k j + lem₁ south = cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) + (cong ∣_∣ₕ (merid north)) + lem₁ (merid a j) k = merid-lem k j where - lala : PathP (λ k → (Kn→ΩKn+10ₖ _) (~ k) ≡ (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north))) k) - (λ j → cong (βᵣ'-fun n f g) (push (θ {A = S₊∙ _} (merid a j)))) - (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) - lala = (cong-∙∙ (cong (βᵣ'-fun n f g) ∘ push) - (λ i₁ → inl ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁)) + p = pt (S₊∙ (suc (suc (n +ℕ n)))) + + merid-lem : + PathP (λ k → (Kn→ΩKn+10ₖ _) (~ k) + ≡ (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) + (cong ∣_∣ₕ (merid north))) k) + (λ j → cong (βᵣ'-fun n f g) (push (θ {A = S₊∙ _} (merid a j)))) + (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) + merid-lem = + (cong-∙∙ (cong (βᵣ'-fun n f g) ∘ push) + (λ i₁ → inl ((merid a + ∙ (sym (merid p))) i₁)) (push tt) - (λ i₁ → inr ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁)) - ∙∙ (cong (sym (Kn→ΩKn+10ₖ _) ∙_) (cong-∙ (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a) (sym (merid (ptSn _))))) + (λ i₁ → inr ((merid a + ∙ (sym (merid p))) i₁)) + ∙∙ (cong (sym (Kn→ΩKn+10ₖ _) ∙_) + (cong-∙ (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) + (merid a) (sym (merid (ptSn _))))) ∙∙ sym (doubleCompPath≡compPath _ _ _)) ◁ symP (doubleCompPath-filler (sym (Kn→ΩKn+10ₖ _)) - (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) - (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (sym (merid north)))) -open import Cubical.Foundations.Path + (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) + (merid a)) + (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) + (sym (merid north)))) + jₗ-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → coHomFun _ (jₗ n f g) (α* n f g) ≡ Hopfα n f jₗ-α n f g = - cong (coHomFun _ (jₗ n f g)) (α*-lem n f g) + cong (coHomFun _ (jₗ n f g)) (α≡ n f g) ∙ cong ∣_∣₂ (funExt (HopfInvariantPushElim n (fst f) (isOfHLevelPath ((suc (suc (suc (suc (n +ℕ n)))))) - (isOfHLevelPlus' {n = n} (4 +ℕ n) (isOfHLevelTrunc (4 +ℕ n))) _ _) + (isOfHLevelPlus' {n = n} (4 +ℕ n) + (isOfHLevelTrunc (4 +ℕ n))) _ _) refl (λ _ → refl) λ j → refl)) @@ -601,8 +643,9 @@ jₗ-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ (Hopfβ n f) jₗ-βₗ n f g = cong (coHomFun _ (jₗ n f g)) (βₗ≡ n f g) - ∙∙ natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (jₗ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) + ∙∙ natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (jₗ n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) (cong ∣_∣₂ (funExt λ { (inl x) → refl @@ -614,38 +657,45 @@ jₗ-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ ≡ 0ₕ _ jₗ-βᵣ n f g = cong (coHomFun _ (jₗ n f g)) (βᵣ≡ n f g) - ∙∙ natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (jₗ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∙∙ natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (jₗ n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) cool where cool : coHomFun (suc (suc (suc (suc (n +ℕ n))))) (jₗ n f g) ∣ βᵣ'-fun n f g ∣₂ ≡ 0ₕ _ - cool = cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → refl}) + cool = cong ∣_∣₂ (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a i₁) → refl}) jᵣ-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → coHomFun _ (jᵣ n f g) (α* n f g) ≡ Hopfα n g -jᵣ-α n f g = cong (coHomFun _ (jᵣ n f g)) (α*-lem n f g) +jᵣ-α n f g = cong (coHomFun _ (jᵣ n f g)) (α≡ n f g) ∙ cong ∣_∣₂ (funExt (HopfInvariantPushElim n (fst g) (isOfHLevelPath ((suc (suc (suc (suc (n +ℕ n)))))) - (isOfHLevelPlus' {n = n} (4 +ℕ n) (isOfHLevelTrunc (4 +ℕ n))) _ _) + (isOfHLevelPlus' {n = n} (4 +ℕ n) + (isOfHLevelTrunc (4 +ℕ n))) _ _) refl (λ _ → refl) - (flipSquare (pp-inr n f g)))) + (flipSquare (0ₖ≡∨→-inr n f g)))) jᵣ-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → coHomFun _ (jᵣ n f g) (βₗ n f g) ≡ 0ₕ _ jᵣ-βₗ n f g = cong (coHomFun _ (jᵣ n f g)) (βₗ≡ n f g) - ∙∙ natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (jᵣ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∙∙ natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (jᵣ n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) cool where cool : coHomFun (suc (suc (suc (suc (n +ℕ n))))) (jᵣ n f g) ∣ βₗ'-fun n f g ∣₂ ≡ 0ₕ _ - cool = cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → refl}) + cool = cong ∣_∣₂ (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a i₁) → refl}) jᵣ-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) @@ -655,7 +705,8 @@ jᵣ-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ (Hopfβ n g) jᵣ-βᵣ n f g = cong (coHomFun _ (jᵣ n f g)) (βᵣ≡ n f g) - ∙∙ natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (jᵣ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + ∙∙ natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (jᵣ n f g)) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst g))) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) (cong ∣_∣₂ @@ -666,30 +717,33 @@ jᵣ-βᵣ n f g = gen₂ℤ×ℤ : gen₂-by (DirProd ℤGroup ℤGroup) (1 , 0) (0 , 1) fst (gen₂ℤ×ℤ (x , y)) = x , y snd (gen₂ℤ×ℤ (x , y)) = - ΣPathP ((cong₂ _+_ ((·Comm 1 x) ∙ cong fst (sym (distrLem 1 0 x))) ((·Comm 0 y) ∙ cong fst (sym (distrLem 0 1 y)))) - , +Comm y 0 ∙ cong₂ _+_ (·Comm 0 x ∙ cong snd (sym (distrLem 1 0 x))) (·Comm 1 y ∙ cong snd (sym (distrLem 0 1 y)))) + ΣPathP ((cong₂ _+_ ((·Comm 1 x) ∙ cong fst (sym (distrLem 1 0 x))) + ((·Comm 0 y) ∙ cong fst (sym (distrLem 0 1 y)))) + , +Comm y 0 + ∙ cong₂ _+_ (·Comm 0 x ∙ cong snd (sym (distrLem 1 0 x))) + (·Comm 1 y ∙ cong snd (sym (distrLem 0 1 y)))) where ℤ×ℤ = DirProd ℤGroup ℤGroup _+''_ = GroupStr._·_ (snd ℤ×ℤ) - ll3 : (x : ℤ) → - x ≡ 0 - x - ll3 (pos zero) = refl - ll3 (pos (suc zero)) = refl - ll3 (pos (suc (suc n))) = - cong predℤ (ll3 (pos (suc n))) - ll3 (negsuc zero) = refl - ll3 (negsuc (suc n)) = cong sucℤ (ll3 (negsuc n)) + -lem : (x : ℤ) → - x ≡ 0 - x + -lem (pos zero) = refl + -lem (pos (suc zero)) = refl + -lem (pos (suc (suc n))) = + cong predℤ (-lem (pos (suc n))) + -lem (negsuc zero) = refl + -lem (negsuc (suc n)) = cong sucℤ (-lem (negsuc n)) distrLem : (x y : ℤ) (z : ℤ) → Path (ℤ × ℤ) (z ℤ[ ℤ×ℤ ]· (x , y)) (z · x , z · y) distrLem x y (pos zero) = refl distrLem x y (pos (suc n)) = (cong ((x , y) +''_) (distrLem x y (pos n))) - distrLem x y (negsuc zero) = ΣPathP (sym (ll3 x) , sym (ll3 y)) + distrLem x y (negsuc zero) = ΣPathP (sym (-lem x) , sym (-lem y)) distrLem x y (negsuc (suc n)) = - cong₂ _+''_ (ΣPathP (sym (ll3 x) , sym (ll3 y))) + cong₂ _+''_ (ΣPathP (sym (-lem x) , sym (-lem y))) (distrLem x y (negsuc n)) - + genH²ⁿC* : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → gen₂-by (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) (βₗ n f g) @@ -703,7 +757,6 @@ genH²ⁿC* n f g = (invGroupIso (H⁴-C* n f g)) private - H⁴C* : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Group _ H⁴C* n f g = coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) @@ -714,122 +767,124 @@ private → ℤ Y n f g = (genH²ⁿC* n f g) (α* n f g ⌣ α* n f g) .fst .snd - α*≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) + α*≡⌣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → α* n f g ⌣ α* n f g ≡ ((X n f g) ℤ[ H⁴C* n f g ]· βₗ n f g) +ₕ ((Y n f g) ℤ[ H⁴C* n f g ]· βᵣ n f g) - α*≡ n f g = (genH²ⁿC* n f g) (α* n f g ⌣ α* n f g) .snd - - -eq₁ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → (Hopfα n (∙Π f g)) ⌣ (Hopfα n (∙Π f g)) - ≡ ((X n f g + Y n f g) ℤ[ coHomGr _ _ ]· (β·Π n f g)) -eq₁ n f g = - cong (λ x → x ⌣ x) (sym (q-α n f g) ∙ cong (coHomFun (suc (suc n)) (q n f g)) (α*-lem n f g)) - ∙ cong ((coHomFun _) (q _ f g)) (cong (λ x → x ⌣ x) (sym (α*-lem n f g)) - ∙ α*≡ n f g) - ∙ IsGroupHom.pres· (coHomMorph _ (q n f g) .snd) _ _ - ∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (q n f g)) (βₗ n f g) (X n f g) - ∙ cong (λ z → (X n f g) ℤ[ coHomGr _ _ ]· z) - (q-βₗ n f g)) - (homPresℤ· (coHomMorph _ (q n f g)) (βᵣ n f g) (Y n f g) - ∙ cong (λ z → (Y n f g) ℤ[ coHomGr _ _ ]· z) - (q-βᵣ n f g)) - ∙ sym (distrℤ· (coHomGr _ _) (β·Π n f g) (X n f g) (Y n f g)) + α*≡⌣ n f g = (genH²ⁿC* n f g) (α* n f g ⌣ α* n f g) .snd coHomFun⌣ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) → (n : ℕ) (x y : coHom n B) → coHomFun _ f (x ⌣ y) ≡ coHomFun _ f x ⌣ coHomFun _ f y coHomFun⌣ f n = sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl -eq₂ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → Hopfα n f ⌣ Hopfα n f - ≡ (X n f g ℤ[ coHomGr _ _ ]· subst (λ m → coHom m (HopfInvariantPush n (fst f))) - (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - (Hopfβ n f)) -eq₂ n f g = - cong (λ x → x ⌣ x) (sym (jₗ-α n f g)) - ∙∙ sym (coHomFun⌣ (jₗ n f g) _ _ _) - ∙∙ cong (coHomFun _ (jₗ n f g)) (α*≡ n f g) - ∙∙ IsGroupHom.pres· (snd (coHomMorph _ (jₗ n f g))) _ _ - ∙∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (jₗ n f g)) (βₗ n f g) (X n f g)) - (homPresℤ· (coHomMorph _ (jₗ n f g)) (βᵣ n f g) (Y n f g) - ∙ cong (λ z → (Y n f g) ℤ[ coHomGr _ _ ]· z) - (jₗ-βᵣ n f g)) - ∙∙ cong₂ _+ₕ_ refl (rUnitℤ· (coHomGr _ _) (Y n f g)) - ∙∙ (rUnitₕ _ _ - ∙ cong (X n f g ℤ[ coHomGr _ _ ]·_) (jₗ-βₗ n f g)) - -eq₃ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) - → Hopfα n g ⌣ Hopfα n g - ≡ (Y n f g ℤ[ coHomGr _ _ ]· subst (λ m → coHom m (HopfInvariantPush n (fst f))) - (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - (Hopfβ n g)) -eq₃ n f g = - cong (λ x → x ⌣ x) (sym (jᵣ-α n f g)) - ∙∙ sym (coHomFun⌣ (jᵣ n f g) _ _ _) - ∙∙ cong (coHomFun _ (jᵣ n f g)) (α*≡ n f g) - ∙∙ IsGroupHom.pres· (snd (coHomMorph _ (jᵣ n f g))) _ _ - ∙∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (jᵣ n f g)) (βₗ n f g) (X n f g) - ∙ cong (λ z → (X n f g) ℤ[ coHomGr _ _ ]· z) - (jᵣ-βₗ n f g)) - (homPresℤ· (coHomMorph _ (jᵣ n f g)) (βᵣ n f g) (Y n f g)) - ∙∙ cong₂ _+ₕ_ (rUnitℤ· (coHomGr _ _) (X n f g)) refl - ∙∙ ((lUnitₕ _ _) ∙ cong (Y n f g ℤ[ coHomGr _ _ ]·_) (jᵣ-βᵣ n f g)) - -eq₂-eq₂ : (n : ℕ) (f g : S₊∙ (suc (suc (suc (n +ℕ n)))) →∙ S₊∙ (suc (suc n))) +isHom-HopfInvariant : + (n : ℕ) (f g : S₊∙ (suc (suc (suc (n +ℕ n)))) →∙ S₊∙ (suc (suc n))) → HopfInvariant n (∙Π f g) ≡ HopfInvariant n f + HopfInvariant n g -eq₂-eq₂ n f g = +isHom-HopfInvariant n f g = mainL - ∙ sym (cong₂ _+_ (help n f g) (helpg n f g)) + ∙ sym (cong₂ _+_ f-id g-id) where - transpLem : ∀ {ℓ} {G : ℕ → Group ℓ} - → (n m : ℕ) - → (x : ℤ) - → (p : m ≡ n) - → (g : fst (G n)) - → subst (fst ∘ G) p (x ℤ[ G m ]· subst (fst ∘ G) (sym p) g) ≡ (x ℤ[ G n ]· g) + eq₁ : (Hopfα n (∙Π f g)) ⌣ (Hopfα n (∙Π f g)) + ≡ ((X n f g + Y n f g) ℤ[ coHomGr _ _ ]· (β·Π n f g)) + eq₁ = cong (λ x → x ⌣ x) (sym (q-α n f g) + ∙ cong (coHomFun (suc (suc n)) (q n f g)) (α≡ n f g)) + ∙ cong ((coHomFun _) (q _ f g)) (cong (λ x → x ⌣ x) (sym (α≡ n f g)) + ∙ α*≡⌣ n f g) + ∙ IsGroupHom.pres· (coHomMorph _ (q n f g) .snd) _ _ + ∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (q n f g)) (βₗ n f g) (X n f g) + ∙ cong (λ z → (X n f g) ℤ[ coHomGr _ _ ]· z) + (q-βₗ n f g)) + (homPresℤ· (coHomMorph _ (q n f g)) (βᵣ n f g) (Y n f g) + ∙ cong (λ z → (Y n f g) ℤ[ coHomGr _ _ ]· z) + (q-βᵣ n f g)) + ∙ sym (distrℤ· (coHomGr _ _) (β·Π n f g) (X n f g) (Y n f g)) + + eq₂ : Hopfα n f ⌣ Hopfα n f + ≡ (X n f g ℤ[ coHomGr _ _ ]· + subst (λ m → coHom m (HopfInvariantPush n (fst f))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + (Hopfβ n f)) + eq₂ = cong (λ x → x ⌣ x) (sym (jₗ-α n f g)) + ∙∙ sym (coHomFun⌣ (jₗ n f g) _ _ _) + ∙∙ cong (coHomFun _ (jₗ n f g)) (α*≡⌣ n f g) + ∙∙ IsGroupHom.pres· (snd (coHomMorph _ (jₗ n f g))) _ _ + ∙∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (jₗ n f g)) (βₗ n f g) (X n f g)) + (homPresℤ· (coHomMorph _ (jₗ n f g)) (βᵣ n f g) (Y n f g) + ∙ cong (λ z → (Y n f g) ℤ[ coHomGr _ _ ]· z) + (jₗ-βᵣ n f g)) + ∙∙ cong₂ _+ₕ_ refl (rUnitℤ· (coHomGr _ _) (Y n f g)) + ∙∙ (rUnitₕ _ _ + ∙ cong (X n f g ℤ[ coHomGr _ _ ]·_) (jₗ-βₗ n f g)) + + eq₃ : Hopfα n g ⌣ Hopfα n g + ≡ (Y n f g ℤ[ coHomGr _ _ ]· + subst (λ m → coHom m (HopfInvariantPush n (fst f))) + (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) + (Hopfβ n g)) + eq₃ = cong (λ x → x ⌣ x) (sym (jᵣ-α n f g)) + ∙∙ sym (coHomFun⌣ (jᵣ n f g) _ _ _) + ∙∙ cong (coHomFun _ (jᵣ n f g)) (α*≡⌣ n f g) + ∙∙ IsGroupHom.pres· (snd (coHomMorph _ (jᵣ n f g))) _ _ + ∙∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (jᵣ n f g)) (βₗ n f g) (X n f g) + ∙ cong (λ z → (X n f g) ℤ[ coHomGr _ _ ]· z) + (jᵣ-βₗ n f g)) + (homPresℤ· (coHomMorph _ (jᵣ n f g)) (βᵣ n f g) (Y n f g)) + ∙∙ cong₂ _+ₕ_ (rUnitℤ· (coHomGr _ _) (X n f g)) refl + ∙∙ ((lUnitₕ _ _) ∙ cong (Y n f g ℤ[ coHomGr _ _ ]·_) (jᵣ-βᵣ n f g)) + + transpLem : ∀ {ℓ} {G : ℕ → Group ℓ} (n m : ℕ) (x : ℤ) (p : m ≡ n) + (g : fst (G n)) + → subst (fst ∘ G) p (x ℤ[ G m ]· subst (fst ∘ G) (sym p) g) + ≡ (x ℤ[ G n ]· g) transpLem {G = G} n m x = - J (λ n p → (g : fst (G n)) → subst (fst ∘ G) p (x ℤ[ G m ]· subst (fst ∘ G) (sym p) g) - ≡ (x ℤ[ G n ]· g)) + J (λ n p → (g : fst (G n)) + → subst (fst ∘ G) p (x ℤ[ G m ]· subst (fst ∘ G) (sym p) g) + ≡ (x ℤ[ G n ]· g)) λ g → transportRefl _ ∙ cong (x ℤ[ G m ]·_) (transportRefl _) mainL : HopfInvariant n (∙Π f g) ≡ X n f g + Y n f g - mainL = cong (Iso.fun (fst (Hopfβ-Iso n (∙Π f g)))) - (cong (subst (λ x → coHom x (HopfInvariantPush n (fst (∙Π f g)))) - λ i₁ → suc (suc (suc (+-suc n n i₁)))) - (eq₁ n f g)) - ∙∙ cong (Iso.fun (fst (Hopfβ-Iso n (∙Π f g)))) - (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst (∙Π f g)))} _ _ - (X n f g + Y n f g) (λ i₁ → suc (suc (suc (+-suc n n i₁)))) - (Iso.inv (fst (Hopfβ-Iso n (∙Π f g))) 1)) - ∙∙ homPresℤ· (_ , snd (Hopfβ-Iso n (∙Π f g))) (Iso.inv (fst (Hopfβ-Iso n (∙Π f g))) 1) - (X n f g + Y n f g) - ∙∙ cong ((X n f g + Y n f g) ℤ[ ℤ , ℤGroup .snd ]·_) - (Iso.rightInv ((fst (Hopfβ-Iso n (∙Π f g)))) 1) - ∙∙ rUnitℤ·ℤ (X n f g + Y n f g) - - - help : (n : ℕ) (f g : _) → HopfInvariant n f ≡ X n f g - help n f g = cong (Iso.fun (fst (Hopfβ-Iso n f))) - (cong (subst (λ x → coHom x (HopfInvariantPush n (fst f))) - (λ i₁ → suc (suc (suc (+-suc n n i₁))))) (eq₂ n f g)) - ∙ cong (Iso.fun (fst (Hopfβ-Iso n f))) - (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst f))} _ _ - (X n f g) - ((cong (suc ∘ suc ∘ suc) (+-suc n n))) - (Hopfβ n f)) - ∙ homPresℤ· (_ , snd (Hopfβ-Iso n f)) (Hopfβ n f) (X n f g) - ∙ cong (X n f g ℤ[ ℤ , ℤGroup .snd ]·_) (Hopfβ↦1 n f) - ∙ rUnitℤ·ℤ (X n f g) - - helpg : (n : ℕ) (f g : _) → HopfInvariant n g ≡ Y n f g - helpg n f g = + mainL = + cong (Iso.fun (fst (Hopfβ-Iso n (∙Π f g)))) + (cong (subst (λ x → coHom x (HopfInvariantPush n (fst (∙Π f g)))) + λ i₁ → suc (suc (suc (+-suc n n i₁)))) + eq₁) + ∙∙ cong (Iso.fun (fst (Hopfβ-Iso n (∙Π f g)))) + (transpLem {G = λ x → coHomGr x + (HopfInvariantPush n (fst (∙Π f g)))} _ _ + (X n f g + Y n f g) + (λ i₁ → suc (suc (suc (+-suc n n i₁)))) + (Iso.inv (fst (Hopfβ-Iso n (∙Π f g))) 1)) + ∙∙ homPresℤ· (_ , snd (Hopfβ-Iso n (∙Π f g))) + (Iso.inv (fst (Hopfβ-Iso n (∙Π f g))) 1) + (X n f g + Y n f g) + ∙∙ cong ((X n f g + Y n f g) ℤ[ ℤ , ℤGroup .snd ]·_) + (Iso.rightInv ((fst (Hopfβ-Iso n (∙Π f g)))) 1) + ∙∙ rUnitℤ·ℤ (X n f g + Y n f g) + + f-id : HopfInvariant n f ≡ X n f g + f-id = + cong (Iso.fun (fst (Hopfβ-Iso n f))) + (cong (subst (λ x → coHom x (HopfInvariantPush n (fst f))) + (λ i₁ → suc (suc (suc (+-suc n n i₁))))) eq₂) + ∙ cong (Iso.fun (fst (Hopfβ-Iso n f))) + (transpLem {G = λ x → coHomGr x + (HopfInvariantPush n (fst f))} _ _ + (X n f g) + ((cong (suc ∘ suc ∘ suc) (+-suc n n))) + (Hopfβ n f)) + ∙ homPresℤ· (_ , snd (Hopfβ-Iso n f)) (Hopfβ n f) (X n f g) + ∙ cong (X n f g ℤ[ ℤ , ℤGroup .snd ]·_) (Hopfβ↦1 n f) + ∙ rUnitℤ·ℤ (X n f g) + + g-id : HopfInvariant n g ≡ Y n f g + g-id = cong (Iso.fun (fst (Hopfβ-Iso n g))) (cong (subst (λ x → coHom x (HopfInvariantPush n (fst g))) - (λ i₁ → suc (suc (suc (+-suc n n i₁))))) - (eq₃ n f g)) + (λ i₁ → suc (suc (suc (+-suc n n i₁))))) + eq₃) ∙∙ cong (Iso.fun (fst (Hopfβ-Iso n g))) - (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst g))} _ _ + (transpLem {G = λ x → coHomGr x + (HopfInvariantPush n (fst g))} _ _ (Y n f g) ((cong (suc ∘ suc ∘ suc) (+-suc n n))) (Hopfβ n g)) @@ -837,6 +892,22 @@ eq₂-eq₂ n f g = ∙∙ cong (Y n f g ℤ[ ℤ , ℤGroup .snd ]·_) (Hopfβ↦1 n g) ∙∙ rUnitℤ·ℤ (Y n f g) - --- GroupHom-HopfInvariant-π' : (n : ℕ) → GroupHom (π'Gr (suc (suc (n +ℕ n))) (S₊∙ (suc (suc n)))) ℤGroup --- GroupHom-HopfInvariant-π' n = HopfInvariant-π' n , makeIsGroupHom {!eq₂-eq₂ n!} +GroupHom-HopfInvariant-π' : (n : ℕ) + → GroupHom (π'Gr (suc (suc (n +ℕ n))) (S₊∙ (suc (suc n)))) ℤGroup +fst (GroupHom-HopfInvariant-π' n) = HopfInvariant-π' n +snd (GroupHom-HopfInvariant-π' n) = + makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 isSetℤ _ _) + (isHom-HopfInvariant n)) + +GroupHom-HopfInvariant-π : (n : ℕ) + → GroupHom (πGr (suc (suc (n +ℕ n))) (S₊∙ (suc (suc n)))) ℤGroup +fst (GroupHom-HopfInvariant-π n) = HopfInvariant-π n +snd (GroupHom-HopfInvariant-π n) = + makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 isSetℤ _ _) + λ p q → cong (HopfInvariant-π' n) + (IsGroupHom.pres· + (snd (invGroupIso (π'Gr≅πGr (suc (suc (n +ℕ n))) + (S₊∙ (suc (suc n)))))) + ∣ p ∣₂ ∣ q ∣₂) + ∙ IsGroupHom.pres· (GroupHom-HopfInvariant-π' n .snd) + ∣ Ω→SphereMap _ p ∣₂ ∣ Ω→SphereMap _ q ∣₂) diff --git a/Cubical/Homotopy/HopfInvariant/HopfMap.agda b/Cubical/Homotopy/HopfInvariant/HopfMap.agda index e554a9c85..c14c7979e 100644 --- a/Cubical/Homotopy/HopfInvariant/HopfMap.agda +++ b/Cubical/Homotopy/HopfInvariant/HopfMap.agda @@ -6,8 +6,12 @@ hopf invariant ±1. module Cubical.Homotopy.HopfInvariant.HopfMap where open import Cubical.Homotopy.HopfInvariant.Base +open import Cubical.Homotopy.Hopf +open import Cubical.Homotopy.Connected + +open import Cubical.Relation.Nullary + open import Cubical.ZCohomology.Base -open import Cubical.ZCohomology.Groups.Connected open import Cubical.ZCohomology.GroupStructure open import Cubical.ZCohomology.Properties open import Cubical.ZCohomology.MayerVietorisUnreduced @@ -17,80 +21,52 @@ open import Cubical.ZCohomology.Groups.Sn open import Cubical.ZCohomology.RingStructure.CupProduct open import Cubical.ZCohomology.RingStructure.RingLaws open import Cubical.ZCohomology.RingStructure.GradedCommutativity -open import Cubical.Relation.Nullary open import Cubical.ZCohomology.Gysin -open import Cubical.Functions.Embedding - open import Cubical.Foundations.Path -open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.HLevels open import Cubical.Foundations.Transport open import Cubical.Foundations.Function -open import Cubical.Foundations.Univalence open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Pointed.Homogeneous - -open import Cubical.Homotopy.HomotopyGroup - open import Cubical.Foundations.Univalence -open import Cubical.Relation.Nullary - open import Cubical.Data.Sum -open import Cubical.Data.Fin -open import Cubical.Data.Empty renaming (rec to ⊥-rec) open import Cubical.Data.Sigma open import Cubical.Data.Int hiding (_+'_) open import Cubical.Data.Nat renaming (_+_ to _+ℕ_ ; _·_ to _·ℕ_) -open import Cubical.Data.Nat.Order open import Cubical.Data.Unit -open import Cubical.Data.Bool + open import Cubical.Algebra.Group - renaming (ℤ to ℤGroup ; Unit to UnitGroup) hiding (Bool) + renaming (ℤ to ℤGroup ; Unit to UnitGroup) open import Cubical.Algebra.Group.ZModule -open import Cubical.HITs.Pushout.Flattening -open import Cubical.Homotopy.Connected -open import Cubical.Homotopy.EilenbergSteenrod open import Cubical.HITs.Pushout +open import Cubical.HITs.Hopf +open import Cubical.HITs.Join +open import Cubical.HITs.S1 renaming (_·_ to _*_) open import Cubical.HITs.Sn open import Cubical.HITs.Susp -open import Cubical.HITs.S1 renaming (_·_ to _*_) +open import Cubical.HITs.Wedge open import Cubical.HITs.Truncation - renaming (rec to trRec ; elim to trElim ; elim2 to trElim2) + renaming (rec to trRec ; elim to trElim) open import Cubical.HITs.SetTruncation - renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap ; elim3 to sElim3) + renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap) open import Cubical.HITs.PropositionalTruncation - renaming (rec to pRec ; elim to pElim) - -open import Cubical.Algebra.AbGroup - -open import Cubical.Homotopy.Loopspace - -open import Cubical.HITs.Join - -open import Cubical.Homotopy.Hopf - -open import Cubical.HITs.SetQuotients renaming (_/_ to _/'_) - - -open import Cubical.Experiments.Brunerie -open import Cubical.HITs.Hopf - -open import Cubical.HITs.Join - + renaming (rec to pRec) HopfMap : S₊∙ 3 →∙ S₊∙ 2 fst HopfMap x = JoinS¹S¹→TotalHopf (Iso.inv (IsoSphereJoin 1 1) x) .fst snd HopfMap = refl -- We use the Hopf fibration in order to connect it to the Gysin Sequence -module hopfS¹ = hopfBase S1-AssocHSpace (sphereElim2 _ (λ _ _ → squash) ∣ refl ∣) +module hopfS¹ = + hopfBase S1-AssocHSpace (sphereElim2 _ (λ _ _ → squash) ∣ refl ∣) + S¹Hopf = hopfS¹.Hopf E* = hopfS¹.totalSpaceMegaPush IsoE*join = hopfS¹.IsoJoin₁ @@ -114,12 +90,14 @@ conCP² x y = sRec2 squash₂ (λ p q → ∣ p ∙ sym q ∣₂) (conCP²' x) ( (λ _ → squash₂) ∣ sym (push (inl base)) ∣₂ x conCP²' (push a i₁) = main a i₁ where - indLem : ∀ {ℓ} {A : hopfS¹.TotalSpaceHopf' → Type ℓ} → ((a : _) → isProp (A a)) + indLem : ∀ {ℓ} {A : hopfS¹.TotalSpaceHopf' → Type ℓ} + → ((a : _) → isProp (A a)) → A (inl base) - → ((a : hopfS¹.TotalSpaceHopf') → A a) + → ((a : hopfS¹.TotalSpaceHopf') → A a) indLem {A = A} p b = - PushoutToProp p (sphereElim 0 (λ _ → p _) b) - (sphereElim 0 (λ _ → p _) (subst A (push (base , base)) b)) + PushoutToProp p + (sphereElim 0 (λ _ → p _) b) + (sphereElim 0 (λ _ → p _) (subst A (push (base , base)) b)) main : (a : hopfS¹.TotalSpaceHopf') → PathP (λ i → ∥ Path CP² (push a i) (inl tt) ∥₂) @@ -158,8 +136,8 @@ joinS¹S¹→Groupoid : ∀ {ℓ} (P : join S¹ S¹ → Type ℓ) → (x : _) → P x joinS¹S¹→Groupoid P grp b = transport (λ i → (x : (isoToPath (invIso (IsoSphereJoin 1 1))) i) - → P (transp (λ j → isoToPath (invIso (IsoSphereJoin 1 1)) (i ∨ j)) i x )) - (sphereElim _ (λ _ → grp _) b) + → P (transp (λ j → isoToPath (invIso (IsoSphereJoin 1 1)) (i ∨ j)) i x)) + (sphereElim _ (λ _ → grp _) b) TotalHopf→Gpd : ∀ {ℓ} (P : hopfS¹.TotalSpaceHopf' → Type ℓ) → ((x : _) → isGroupoid (P x)) @@ -176,7 +154,8 @@ TotalHopf→Gpd' : ∀ {ℓ} (P : Σ (S₊ 2) hopfS¹.Hopf → Type ℓ) → (x : _) → P x TotalHopf→Gpd' P grp b = transport (λ i → (x : ua (_ , hopfS¹.isEquivTotalSpaceHopf'→TotalSpace) i) - → P (transp (λ j → ua (_ , hopfS¹.isEquivTotalSpaceHopf'→TotalSpace) (i ∨ j)) i x)) + → P (transp (λ j → ua (_ , hopfS¹.isEquivTotalSpaceHopf'→TotalSpace) + (i ∨ j)) i x)) (TotalHopf→Gpd _ (λ _ → grp _) b) CP²→Groupoid : ∀ {ℓ} {P : CP² → Type ℓ} @@ -189,7 +168,7 @@ CP²→Groupoid {P = P} grp b s2 pp (inl x) = b CP²→Groupoid {P = P} grp b s2 pp (inr x) = s2 x CP²→Groupoid {P = P} grp b s2 pp (push a i₁) = lem a i₁ where - lem : (a : hopfS¹.TotalSpaceHopf') → PathP (λ i → P (push a i)) b (s2 _) + lem : (a : hopfS¹.TotalSpaceHopf') → PathP (λ i → P (push a i)) b (s2 _) lem = TotalHopf→Gpd _ (λ _ → isOfHLevelPathP' 3 (grp _) _ _) pp -- The function inducing the iso H²(S²) ≅ H²(CP²) @@ -197,15 +176,17 @@ H²S²→H²CP²-raw : (S₊ 2 → coHomK 2) → (CP² → coHomK 2) H²S²→H²CP²-raw f (inl x) = 0ₖ _ H²S²→H²CP²-raw f (inr x) = f x -ₖ f north H²S²→H²CP²-raw f (push a i) = - TotalHopf→Gpd (λ x → 0ₖ 2 ≡ f (hopfS¹.TotalSpaceHopf'→TotalSpace x .fst) -ₖ f north) - (λ _ → isOfHLevelTrunc 4 _ _) - (sym (rCancelₖ 2 (f north))) - a i + TotalHopf→Gpd (λ x → 0ₖ 2 + ≡ f (hopfS¹.TotalSpaceHopf'→TotalSpace x .fst) -ₖ f north) + (λ _ → isOfHLevelTrunc 4 _ _) + (sym (rCancelₖ 2 (f north))) + a i H²S²→H²CP² : coHomGr 2 (S₊ 2) .fst → coHomGr 2 CP² .fst H²S²→H²CP² = sMap H²S²→H²CP²-raw -cancel-H²S²→H²CP² : (f : CP² → coHomK 2) → ∥ H²S²→H²CP²-raw (f ∘ inr) ≡ f ∥ +cancel-H²S²→H²CP² : (f : CP² → coHomK 2) + → ∥ H²S²→H²CP²-raw (f ∘ inr) ≡ f ∥ cancel-H²S²→H²CP² f = pRec squash (λ p → pRec squash @@ -214,10 +195,11 @@ cancel-H²S²→H²CP² f = (connLem (f (inl tt))) where connLem : (x : coHomK 2) → ∥ x ≡ 0ₖ 2 ∥ - connLem = coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash) ∣ refl ∣ + connLem = coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash) ∣ refl ∣ - cancelLem : (p : f (inl tt) ≡ 0ₖ 2) → ∥ ((x : CP²) → H²S²→H²CP²-raw (f ∘ inr) x ≡ f x) ∥ - cancelLem p = trRec squash (λ pp → + cancelLem : (p : f (inl tt) ≡ 0ₖ 2) + → ∥ ((x : CP²) → H²S²→H²CP²-raw (f ∘ inr) x ≡ f x) ∥ + cancelLem p = trRec squash (λ pp → ∣ CP²→Groupoid (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) (sym p) (λ x → (λ i → f (inr x) -ₖ f (push (inl base) (~ i))) @@ -225,7 +207,10 @@ cancel-H²S²→H²CP² f = ∙∙ rUnitₖ 2 (f (inr x))) pp ∣) help where - help : hLevelTrunc 1 (PathP ((λ i₁ → H²S²→H²CP²-raw (f ∘ inr) (push (inl base) i₁) ≡ f (push (inl base) i₁))) + help : + hLevelTrunc 1 + (PathP (λ i₁ → H²S²→H²CP²-raw (f ∘ inr) (push (inl base) i₁) + ≡ f (push (inl base) i₁)) (sym p) (((λ i₁ → f (inr north) -ₖ f (push (inl base) (~ i₁))) ∙∙ (λ i₁ → f (inr north) -ₖ p i₁) ∙∙ rUnitₖ 2 (f (inr north))))) @@ -236,10 +221,12 @@ Iso.fun (fst H²CP²≅H²S²) = sMap (_∘ inr) Iso.inv (fst H²CP²≅H²S²) = H²S²→H²CP² Iso.rightInv (fst H²CP²≅H²S²) = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) - λ f → trRec {B = Iso.fun (fst H²CP²≅H²S²) (Iso.inv (fst H²CP²≅H²S²) ∣ f ∣₂) ≡ ∣ f ∣₂} - (isOfHLevelPath 2 squash₂ _ _) - (λ p → cong ∣_∣₂ (funExt λ x → cong (λ y → (f x) -ₖ y) p ∙ rUnitₖ 2 (f x))) - (Iso.fun (PathIdTruncIso _) (isContr→isProp (isConnectedKn 1) ∣ f north ∣ ∣ 0ₖ 2 ∣)) + λ f → trRec {B = Iso.fun (fst H²CP²≅H²S²) (Iso.inv (fst H²CP²≅H²S²) ∣ f ∣₂) + ≡ ∣ f ∣₂} + (isOfHLevelPath 2 squash₂ _ _) + (λ p → cong ∣_∣₂ (funExt λ x → cong (λ y → (f x) -ₖ y) p ∙ rUnitₖ 2 (f x))) + (Iso.fun (PathIdTruncIso _) + (isContr→isProp (isConnectedKn 1) ∣ f north ∣ ∣ 0ₖ 2 ∣)) Iso.leftInv (fst H²CP²≅H²S²) = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ f → pRec (squash₂ _ _) (cong ∣_∣₂) (cancel-H²S²→H²CP² f) @@ -253,10 +240,10 @@ H²CP²≅ℤ = compGroupIso H²CP²≅H²S² (Hⁿ-Sⁿ≅ℤ 1) -- ⌣ gives a groupEquiv H²(CP²) ≃ H⁴(CP²) ⌣Equiv : GroupEquiv (coHomGr 2 CP²) (coHomGr 4 CP²) fst (fst ⌣Equiv) = GysinS².⌣-hom 2 .fst -snd (fst ⌣Equiv) = subst isEquiv (cong fst (GysinS².ϕ∘j≡ 2)) eq' +snd (fst ⌣Equiv) = isEq⌣ where - eq' : isEquiv (GysinS².ϕ∘j 2 .fst) - eq' = + isEq⌣ : isEquiv (GysinS².⌣-hom 2 .fst) + isEq⌣ = SES→isEquiv (GroupPath _ _ .fst (invGroupEquiv (isContr→≃Unit isContrH³E @@ -265,10 +252,10 @@ snd (fst ⌣Equiv) = subst isEquiv (cong fst (GysinS².ϕ∘j≡ 2)) eq' (isContr→≃Unit isContrH⁴E , makeIsGroupHom λ _ _ → refl))) (GysinS².susp∘ϕ 1) - (GysinS².ϕ∘j 2) + (GysinS².⌣-hom 2) (GysinS².p-hom 4) - (GysinS².Ker-ϕ∘j⊂Im-Susp∘ϕ _) - (GysinS².Ker-p⊂Im-ϕ∘j _) + (GysinS².Ker-⌣e⊂Im-Susp∘ϕ _) + (GysinS².Ker-p⊂Im-⌣e _) snd ⌣Equiv = GysinS².⌣-hom 2 .snd -- The generator of H²(CP²) @@ -278,16 +265,19 @@ genCP² = ∣ CP²→Groupoid (λ _ → isOfHLevelTrunc 4) ∣_∣ refl ∣₂ -inrInjective : (f g : CP² → coHomK 2) → ∥ f ∘ inr ≡ g ∘ inr ∥ → Path (coHom 2 CP²) ∣ f ∣₂ ∣ g ∣₂ +inrInjective : (f g : CP² → coHomK 2) → ∥ f ∘ inr ≡ g ∘ inr ∥ + → Path (coHom 2 CP²) ∣ f ∣₂ ∣ g ∣₂ inrInjective f g = pRec (squash₂ _ _) (λ p → pRec (squash₂ _ _) (λ id → trRec (squash₂ _ _) (λ pp → cong ∣_∣₂ - (funExt (CP²→Groupoid (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) + (funExt (CP²→Groupoid + (λ _ → isOfHLevelPath 4 (isOfHLevelTrunc 4) _ _) id (funExt⁻ p) (compPathR→PathP pp)))) (conn2 (f (inl tt)) (g (inl tt)) id - (cong f (push (inl base)) ∙ (funExt⁻ p north) ∙ cong g (sym (push (inl base)))))) + (cong f (push (inl base)) + ∙ (funExt⁻ p north) ∙ cong g (sym (push (inl base)))))) (conn (f (inl tt)) (g (inl tt)))) where @@ -296,7 +286,9 @@ inrInjective f g = pRec (squash₂ _ _) (coHomK-elim _ (λ _ → isOfHLevelSuc 1 squash) ∣ refl ∣) conn2 : (x y : coHomK 2) (p q : x ≡ y) → hLevelTrunc 1 (p ≡ q) - conn2 x y = λ p q → Iso.fun (PathIdTruncIso _) (isContr→isProp (isConnectedPath _ (isConnectedKn 1) x y) ∣ p ∣ ∣ q ∣) + conn2 x y p q = + Iso.fun (PathIdTruncIso _) + (isContr→isProp (isConnectedPath _ (isConnectedKn 1) x y) ∣ p ∣ ∣ q ∣) -- A couple of basic lemma concerning the hSpace structure on S¹ private @@ -304,10 +296,15 @@ private rUnit* base = refl rUnit* (loop i₁) = refl - merid*-lem : (a x : S¹) → Path (Path (coHomK 2) _ _) (cong ∣_∣ₕ (merid (a * x) ∙ sym (merid base))) - ((cong ∣_∣ₕ (merid a ∙ sym (merid base))) ∙ (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) + merid*-lem : (a x : S¹) + → Path (Path (coHomK 2) _ _) + (cong ∣_∣ₕ (merid (a * x) ∙ sym (merid base))) + ((cong ∣_∣ₕ (merid a ∙ sym (merid base))) + ∙ (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) merid*-lem = wedgeconFun _ _ (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) - (λ x → lUnit _ ∙ cong (_∙ cong ∣_∣ₕ (merid x ∙ sym (merid base))) (cong (cong ∣_∣ₕ) (sym (rCancel (merid base))))) + (λ x → lUnit _ + ∙ cong (_∙ cong ∣_∣ₕ (merid x ∙ sym (merid base))) + (cong (cong ∣_∣ₕ) (sym (rCancel (merid base))))) (λ x → (λ i → cong ∣_∣ₕ (merid (rUnit* x i) ∙ sym (merid base))) ∙∙ rUnit _ ∙∙ cong (cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙_) @@ -331,15 +328,18 @@ private (cong ∣_∣ₕ (sym (merid x ∙ sym (merid base)))) meridInvLooperLem x = (lUnit _ ∙∙ cong (_∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base))) - (sym (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base))))) ∙∙ sym (assoc _ _ _)) + (sym (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base))))) + ∙∙ sym (assoc _ _ _)) ∙∙ cong (sym (cong ∣_∣ₕ (merid x ∙ sym (merid base))) ∙_) lem ∙∙ (assoc _ _ _ ∙∙ cong (_∙ (cong ∣_∣ₕ (sym (merid x ∙ sym (merid base))))) (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) ∙∙ sym (lUnit _)) where - lem : cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base)) - ≡ cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙ cong ∣_∣ₕ (sym (merid x ∙ sym (merid base))) + lem : cong ∣_∣ₕ (merid x ∙ sym (merid base)) + ∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base)) + ≡ cong ∣_∣ₕ (merid x ∙ sym (merid base)) + ∙ cong ∣_∣ₕ (sym (merid x ∙ sym (merid base))) lem = sym (merid*-lem x (invLooper x)) ∙ ((λ i → cong ∣_∣ₕ (merid (lemmie x (~ i)) ∙ sym (merid base))) ∙ cong (cong ∣_∣ₕ) (rCancel (merid base))) ∙ sym (rCancel _) @@ -350,7 +350,8 @@ private rUnit* refl - invLooperLem₁ : (a x : S¹) → (invEq (hopfS¹.μ-eq a) x) * a ≡ (invLooper a * x) * a + invLooperLem₁ : (a x : S¹) + → (invEq (hopfS¹.μ-eq a) x) * a ≡ (invLooper a * x) * a invLooperLem₁ a x = secEq (hopfS¹.μ-eq a) x ∙∙ cong (_* x) (lemmie a) @@ -374,9 +375,11 @@ private -- We prove that the generator of CP² given by Gysin is the same one -- as genCP², which is much easier to work with Gysin-e≡genCP² : GysinS².e ≡ genCP² -Gysin-e≡genCP² = inrInjective _ _ ∣ funExt (λ x → funExt⁻ (cong fst (main x)) south) ∣ +Gysin-e≡genCP² = + inrInjective _ _ ∣ funExt (λ x → funExt⁻ (cong fst (main x)) south) ∣ where - mainId : (x : Σ (S₊ 2) hopfS¹.Hopf) → Path (hLevelTrunc 4 _) ∣ fst x ∣ ∣ north ∣ + mainId : (x : Σ (S₊ 2) hopfS¹.Hopf) + → Path (hLevelTrunc 4 _) ∣ fst x ∣ ∣ north ∣ mainId = uncurry λ { north → λ y → cong ∣_∣ₕ (merid base ∙ sym (merid y)) ; south → λ y → cong ∣_∣ₕ (sym (merid y)) ; (merid a i) → main a i} @@ -388,27 +391,36 @@ Gysin-e≡genCP² = inrInjective _ _ ∣ funExt (λ x → funExt⁻ (cong fst (m main a = toPathP (funExt λ x → - cong (transport (λ i₁ → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a i₁ ∣ ∣ north ∣)) + cong (transport (λ i₁ + → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a i₁ ∣ ∣ north ∣)) ((λ i → (λ z → cong ∣_∣ₕ - (merid base ∙ sym (merid (transport (λ j → uaInvEquiv (hopfS¹.μ-eq a) (~ i) j) x))) z)) + (merid base + ∙ sym (merid (transport + (λ j → uaInvEquiv (hopfS¹.μ-eq a) (~ i) j) x))) z)) ∙ λ i → cong ∣_∣ₕ (merid base - ∙ sym (merid (transportRefl (invEq (hopfS¹.μ-eq a) x) i)))) - ∙∙ (λ i → transp (λ i₁ → Path (HubAndSpoke (Susp S¹) 3) ∣ merid a (i₁ ∨ i) ∣ ∣ north ∣) i + ∙ sym (merid (transportRefl (invEq (hopfS¹.μ-eq a) x) i)))) + ∙∙ (λ i → transp (λ i₁ → Path (HubAndSpoke (Susp S¹) 3) + ∣ merid a (i₁ ∨ i) ∣ ∣ north ∣) i (compPath-filler' (cong ∣_∣ₕ (sym (merid a))) - (cong ∣_∣ₕ (merid base ∙ sym (merid (invLooperLem₂ a x i)))) i)) + (cong ∣_∣ₕ (merid base + ∙ sym (merid (invLooperLem₂ a x i)))) i)) ∙∙ cong ((cong ∣_∣ₕ) (sym (merid a)) ∙_) - (cong (cong ∣_∣ₕ) (cong sym (symDistr (merid base) (sym (merid (invLooper a * x))))) - ∙ cong sym (merid*-lem (invLooper a) x) - ∙ symDistr ((cong ∣_∣ₕ) (merid (invLooper a) ∙ sym (merid base))) - ((cong ∣_∣ₕ) (merid x ∙ sym (merid base))) - ∙ isCommΩK 2 (sym (λ i₁ → ∣ (merid x ∙ (λ i₂ → merid base (~ i₂))) i₁ ∣)) - (sym (λ i₁ → ∣ (merid (invLooper a) ∙ (λ i₂ → merid base (~ i₂))) i₁ ∣)) - ∙ cong₂ _∙_ (cong sym (meridInvLooperLem a) - ∙ cong-∙ ∣_∣ₕ (merid a) (sym (merid base))) - (cong (cong ∣_∣ₕ) (symDistr (merid x) (sym (merid base))) - ∙ cong-∙ ∣_∣ₕ (merid base) (sym (merid x)))) + (cong (cong ∣_∣ₕ) (cong sym (symDistr (merid base) + (sym (merid (invLooper a * x))))) + ∙ cong sym (merid*-lem (invLooper a) x) + ∙ symDistr ((cong ∣_∣ₕ) (merid (invLooper a) ∙ sym (merid base))) + ((cong ∣_∣ₕ) (merid x ∙ sym (merid base))) + ∙ isCommΩK 2 (sym (λ i₁ → ∣ (merid x + ∙ (λ i₂ → merid base (~ i₂))) i₁ ∣)) + (sym (λ i₁ → ∣ (merid (invLooper a) + ∙ (λ i₂ → merid base (~ i₂))) i₁ ∣)) + ∙ cong₂ _∙_ (cong sym (meridInvLooperLem a) + ∙ cong-∙ ∣_∣ₕ (merid a) (sym (merid base))) + (cong (cong ∣_∣ₕ) (symDistr (merid x) (sym (merid base))) + ∙ cong-∙ ∣_∣ₕ (merid base) (sym (merid x)))) ∙∙ (λ j → (λ i₁ → ∣ merid a (~ i₁ ∨ j) ∣) ∙ ((λ i₁ → ∣ merid a (i₁ ∨ j) ∣) - ∙ (λ i₁ → ∣ merid base (~ i₁ ∨ j) ∣)) ∙ (λ i₁ → ∣ merid base (i₁ ∨ j) ∣) + ∙ (λ i₁ → ∣ merid base (~ i₁ ∨ j) ∣)) + ∙ (λ i₁ → ∣ merid base (i₁ ∨ j) ∣) ∙ (λ i₁ → ∣ merid x (~ i₁) ∣ₕ)) ∙∙ sym (lUnit _) ∙∙ sym (assoc _ _ _) @@ -422,25 +434,30 @@ Gysin-e≡genCP² = inrInjective _ _ ∣ funExt (λ x → funExt⁻ (cong fst (m gen'Id : GysinS².c (inr north) .fst ≡ gen' north .fst gen'Id = cong fst (GysinS².cEq (inr north) ∣ push (inl base) ∣₂) - ∙ (funExt l) + ∙ (funExt lem) where - l : (qb : _) → - ∣ (subst (fst ∘ preThom.Q (CP² , inl tt) fibr) (sym (push (inl base))) qb) ∣ + lem : (qb : _) → + ∣ (subst (fst ∘ preThom.Q (CP² , inl tt) fibr) + (sym (push (inl base))) qb) ∣ ≡ gen' north .fst qb - l north = refl - l south = cong ∣_∣ₕ (sym (merid base)) - l (merid a i) j = - hcomp (λ k → λ { (i = i0) → ∣ merid a (~ k ∧ j) ∣ - ; (i = i1) → ∣ merid base (~ j) ∣ - ; (j = i0) → ∣ transportRefl (merid a i) (~ k) ∣ - ; (j = i1) → ∣ compPath-filler (merid base) (sym (merid a)) k (~ i) ∣ₕ}) - (hcomp (λ k → λ { (i = i0) → ∣ merid a (j ∨ ~ k) ∣ - ; (i = i1) → ∣ merid base (~ j ∨ ~ k) ∣ - ; (j = i0) → ∣ merid a (~ k ∨ i) ∣ - ; (j = i1) → ∣ merid base (~ i ∨ ~ k) ∣ₕ}) - ∣ south ∣) - - setHelp : (x : S₊ 2) → isSet (preThom.Q (CP² , inl tt) fibr (inr x) →∙ coHomK-ptd 2) + lem north = refl + lem south = cong ∣_∣ₕ (sym (merid base)) + lem (merid a i) j = + hcomp (λ k → + λ { (i = i0) → ∣ merid a (~ k ∧ j) ∣ + ; (i = i1) → ∣ merid base (~ j) ∣ + ; (j = i0) → ∣ transportRefl (merid a i) (~ k) ∣ + ; (j = i1) → ∣ compPath-filler + (merid base) (sym (merid a)) k (~ i) ∣ₕ}) + (hcomp (λ k → + λ { (i = i0) → ∣ merid a (j ∨ ~ k) ∣ + ; (i = i1) → ∣ merid base (~ j ∨ ~ k) ∣ + ; (j = i0) → ∣ merid a (~ k ∨ i) ∣ + ; (j = i1) → ∣ merid base (~ i ∨ ~ k) ∣ₕ}) + ∣ south ∣) + + setHelp : (x : S₊ 2) + → isSet (preThom.Q (CP² , inl tt) fibr (inr x) →∙ coHomK-ptd 2) setHelp = sphereElim _ (λ _ → isProp→isOfHLevelSuc 1 (isPropIsOfHLevel 2)) (isOfHLevel↑∙' 0 1) @@ -466,9 +483,12 @@ HopfMap' x = hopfMap≡HopfMap' : HopfMap ≡ (HopfMap' , refl) hopfMap≡HopfMap' = - ΣPathP ((funExt (λ x → cong (λ x → JoinS¹S¹→TotalHopf x .fst) - (sym (Iso.rightInv IsoTotalHopf' (Iso.inv (IsoSphereJoin 1 1) x))) - ∙ sym (lem (Iso.inv IsoTotalHopf' (Iso.inv (IsoSphereJoin 1 1) x))))) + ΣPathP ((funExt (λ x → + cong (λ x → JoinS¹S¹→TotalHopf x .fst) + (sym (Iso.rightInv IsoTotalHopf' + (Iso.inv (IsoSphereJoin 1 1) x))) + ∙ sym (lem (Iso.inv IsoTotalHopf' + (Iso.inv (IsoSphereJoin 1 1) x))))) , flipSquare (sym (rUnit refl) ◁ λ _ _ → north)) where lem : (x : _) → hopfS¹.TotalSpaceHopf'→TotalSpace x .fst @@ -481,9 +501,9 @@ hopfMap≡HopfMap' = CP²' : Type _ CP²' = Pushout {A = S₊ 3} (λ _ → tt) HopfMap' -PushoutReplaceBase : - ∀ {ℓ ℓ' ℓ''} {A B : Type ℓ} {C : Type ℓ'} {D : Type ℓ''} {f : A → C} {g : A → D} - (e : B ≃ A) → Pushout (f ∘ fst e) (g ∘ fst e) ≡ Pushout f g +PushoutReplaceBase : ∀ {ℓ ℓ' ℓ''} {A B : Type ℓ} {C : Type ℓ'} {D : Type ℓ''} + {f : A → C} {g : A → D} (e : B ≃ A) + → Pushout (f ∘ fst e) (g ∘ fst e) ≡ Pushout f g PushoutReplaceBase {f = f} {g = g} = EquivJ (λ _ e → Pushout (f ∘ fst e) (g ∘ fst e) ≡ Pushout f g) refl @@ -493,54 +513,62 @@ CP2≡CP²' = (isoToEquiv (compIso (invIso (IsoSphereJoin 1 1)) (invIso IsoTotalHopf'))) -- packaging everything up: -⌣equiv→pres1 : ∀ {ℓ} {G H : Type ℓ} → (G ≡ H) → (g₁ : coHom 2 G) (h₁ : coHom 2 H) - → (fstEq : (Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] abs (fst (fst ϕ) g₁) ≡ 1)) - → (sndEq : ((Σ[ ϕ ∈ GroupEquiv (coHomGr 2 H) ℤGroup ] abs (fst (fst ϕ) h₁) ≡ 1))) +⌣equiv→pres1 : ∀ {ℓ} {G H : Type ℓ} → (G ≡ H) + → (g₁ : coHom 2 G) (h₁ : coHom 2 H) + → (fstEq : (Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] + abs (fst (fst ϕ) g₁) ≡ 1)) + → (sndEq : ((Σ[ ϕ ∈ GroupEquiv (coHomGr 2 H) ℤGroup ] + abs (fst (fst ϕ) h₁) ≡ 1))) → isEquiv {A = coHom 2 G} {B = coHom 4 G} (_⌣ g₁) → (3rdEq : GroupEquiv (coHomGr 4 H) ℤGroup) → abs (fst (fst 3rdEq) (h₁ ⌣ h₁)) ≡ 1 ⌣equiv→pres1 {G = G} = J (λ H _ → (g₁ : coHom 2 G) (h₁ : coHom 2 H) - → (fstEq : (Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] abs (fst (fst ϕ) g₁) ≡ 1)) - → (sndEq : ((Σ[ ϕ ∈ GroupEquiv (coHomGr 2 H) ℤGroup ] abs (fst (fst ϕ) h₁) ≡ 1))) + → (fstEq : (Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] + abs (fst (fst ϕ) g₁) ≡ 1)) + → (sndEq : ((Σ[ ϕ ∈ GroupEquiv (coHomGr 2 H) ℤGroup ] + abs (fst (fst ϕ) h₁) ≡ 1))) → isEquiv {A = coHom 2 G} {B = coHom 4 G} (_⌣ g₁) → (3rdEq : GroupEquiv (coHomGr 4 H) ℤGroup) → abs (fst (fst 3rdEq) (h₁ ⌣ h₁)) ≡ 1) help where help : (g₁ h₁ : coHom 2 G) - → (fstEq : (Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] abs (fst (fst ϕ) g₁) ≡ 1)) - → (sndEq : ((Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] abs (fst (fst ϕ) h₁) ≡ 1))) + → (fstEq : (Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] + abs (fst (fst ϕ) g₁) ≡ 1)) + → (sndEq : ((Σ[ ϕ ∈ GroupEquiv (coHomGr 2 G) ℤGroup ] + abs (fst (fst ϕ) h₁) ≡ 1))) → isEquiv {A = coHom 2 G} {B = coHom 4 G} (_⌣ g₁) → (3rdEq : GroupEquiv (coHomGr 4 G) ℤGroup) → abs (fst (fst 3rdEq) (h₁ ⌣ h₁)) ≡ 1 help g h (ϕ , idg) (ψ , idh) ⌣eq ξ = ⊎→abs _ _ (allisoPresGen _ - (compGroupEquiv main4 (compGroupEquiv (invGroupEquiv main4) ψ)) - h (abs→⊎ _ _ (cong abs (cong (fst (fst ψ)) (retEq (fst main4) h)) ∙ idh)) (compGroupEquiv main4 ξ)) + (compGroupEquiv main (compGroupEquiv (invGroupEquiv main) ψ)) + h (abs→⊎ _ _ (cong abs (cong (fst (fst ψ)) (retEq (fst main) h)) + ∙ idh)) (compGroupEquiv main ξ)) where - k1 : ((fst (fst ψ) h) ≡ 1) ⊎ (fst (fst ψ) h ≡ -1) + lem₁ : ((fst (fst ψ) h) ≡ 1) ⊎ (fst (fst ψ) h ≡ -1) → abs (fst (fst ϕ) h) ≡ 1 - k1 p = ⊎→abs _ _ (allisoPresGen _ ψ h p ϕ) + lem₁ p = ⊎→abs _ _ (allisoPresGen _ ψ h p ϕ) - help2 : ((fst (fst ϕ) h) ≡ 1) ⊎ (fst (fst ϕ) h ≡ -1) + lem₂ : ((fst (fst ϕ) h) ≡ 1) ⊎ (fst (fst ϕ) h ≡ -1) → ((fst (fst ϕ) g) ≡ 1) ⊎ (fst (fst ϕ) g ≡ -1) → (h ≡ g) ⊎ (h ≡ (-ₕ g)) - help2 (inl x) (inl x₁) = + lem₂ (inl x) (inl x₁) = inl (sym (retEq (fst ϕ) h) ∙∙ cong (invEq (fst ϕ)) (x ∙ sym x₁) ∙∙ retEq (fst ϕ) g) - help2 (inl x) (inr x₁) = + lem₂ (inl x) (inr x₁) = inr (sym (retEq (fst ϕ) h) ∙∙ cong (invEq (fst ϕ)) x ∙ IsGroupHom.presinv (snd (invGroupEquiv ϕ)) (negsuc zero) ∙∙ cong (-ₕ_) (cong (invEq (fst ϕ)) (sym x₁) ∙ (retEq (fst ϕ) g))) - help2 (inr x) (inl x₁) = + lem₂ (inr x) (inl x₁) = inr (sym (retEq (fst ϕ) h) ∙∙ cong (invEq (fst ϕ)) x ∙∙ (IsGroupHom.presinv (snd (invGroupEquiv ϕ)) 1 ∙ cong (-ₕ_) (cong (invEq (fst ϕ)) (sym x₁) ∙ (retEq (fst ϕ) g)))) - help2 (inr x) (inr x₁) = + lem₂ (inr x) (inr x₁) = inl (sym (retEq (fst ϕ) h) ∙∙ cong (invEq (fst ϕ)) (x ∙ sym x₁) ∙∙ retEq (fst ϕ) g) @@ -558,24 +586,26 @@ CP2≡CP²' = theEq : coHom 2 G ≃ coHom 4 G theEq = compEquiv (_ , ⌣eq) (isoToEquiv (-ₕeq 4)) - main3 : (h ≡ g) ⊎ (h ≡ (-ₕ g)) → isEquiv {A = coHom 2 G} (_⌣ h) - main3 (inl x) = subst isEquiv (λ i → _⌣ (x (~ i))) ⌣eq - main3 (inr x) = + lem₃ : (h ≡ g) ⊎ (h ≡ (-ₕ g)) → isEquiv {A = coHom 2 G} (_⌣ h) + lem₃ (inl x) = subst isEquiv (λ i → _⌣ (x (~ i))) ⌣eq + lem₃ (inr x) = subst isEquiv (funExt (λ x → -ₕDistᵣ 2 2 x g) ∙ (λ i → _⌣ (x (~ i)))) (theEq .snd) - main4 : GroupEquiv (coHomGr 2 G) (coHomGr 4 G) - fst main4 = _ , (main3 (help2 (abs→⊎ _ _ (k1 (abs→⊎ _ _ idh))) (abs→⊎ _ _ idg))) - snd main4 = + main : GroupEquiv (coHomGr 2 G) (coHomGr 4 G) + fst main = _ , + (lem₃ (lem₂ (abs→⊎ _ _ (lem₁ (abs→⊎ _ _ idh))) (abs→⊎ _ _ idg))) + snd main = makeIsGroupHom λ g1 g2 → rightDistr-⌣ _ _ g1 g2 h -- The hopf invariant is ±1 for both definitions of the hopf map -HopfInvariant-HopfMap' : abs (HopfInvariant zero (HopfMap' , λ _ → HopfMap' (snd (S₊∙ 3)))) - ≡ suc zero +HopfInvariant-HopfMap' : + abs (HopfInvariant zero (HopfMap' , λ _ → HopfMap' (snd (S₊∙ 3)))) ≡ suc zero HopfInvariant-HopfMap' = cong abs (cong (Iso.fun (fst (Hopfβ-Iso zero (HopfMap' , refl)))) (transportRefl (⌣-α 0 (HopfMap' , refl)))) - ∙ ⌣equiv→pres1 (sym CP2≡CP²') GysinS².e (Hopfα zero (HopfMap' , refl)) + ∙ ⌣equiv→pres1 (sym CP2≡CP²') + GysinS².e (Hopfα zero (HopfMap' , refl)) (l isGenerator≃ℤ-e) (GroupIso→GroupEquiv (Hopfα-Iso 0 (HopfMap' , refl)) , refl) (snd (fst ⌣Equiv)) @@ -590,790 +620,3 @@ HopfInvariant-HopfMap' = HopfInvariant-HopfMap : abs (HopfInvariant zero HopfMap) ≡ suc zero HopfInvariant-HopfMap = cong abs (cong (HopfInvariant zero) hopfMap≡HopfMap') ∙ HopfInvariant-HopfMap' - - - -HopfInvariantPushElim : ∀ {ℓ} n → (f : _) → {P : HopfInvariantPush n f → Type ℓ} - → (isOfHLevel (suc (suc (suc (suc (n +ℕ n))))) (P (inl tt))) - → (e : P (inl tt)) - (g : (x : _) → P (inr x)) - (r : PathP (λ i → P (push north i)) e (g (f north))) - → (x : _) → P x -HopfInvariantPushElim n f {P = P} hlev e g r (inl x) = e -HopfInvariantPushElim n f {P = P} hlev e g r (inr x) = g x -HopfInvariantPushElim n f {P = P} hlev e g r (push a i₁) = help a i₁ - where - help : (a : Susp (Susp (S₊ (suc (n +ℕ n))))) → PathP (λ i → P (push a i)) e (g (f a)) - help = sphereElim _ (sphereElim _ (λ _ → isProp→isOfHLevelSuc ((suc (suc (n +ℕ n)))) (isPropIsOfHLevel _)) - (isOfHLevelPathP' (suc (suc (suc (n +ℕ n)))) - (subst (isOfHLevel (suc (suc (suc (suc (n +ℕ n)))))) - (cong P (push north)) - hlev) _ _)) r - - --- open import Cubical.HITs.Wedge --- _∨→_ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} --- → (f : A →∙ C) (g : B →∙ C) --- → A ⋁ B → fst C --- (f ∨→ g) (inl x) = fst f x --- (f ∨→ g) (inr x) = fst g x --- (f ∨→ g) (push a i₁) = (snd f ∙ sym (snd g)) i₁ - - --- add2 : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → (f g : S₊∙ n →∙ A) → ((S₊∙ n) ⋁ (S₊∙ n) , inl (ptSn _)) →∙ A --- fst (add2 {A = A} f g) = f ∨→ g --- snd (add2 {A = A} f g) = snd f - --- C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Type _ --- C* n f g = Pushout (λ _ → tt) (fst (add2 f g)) - --- θ : ∀ {ℓ} {A : Pointed ℓ} → Susp (fst A) → (Susp (fst A) , north) ⋁ (Susp (fst A) , north) --- θ north = inl north --- θ south = inr north --- θ {A = A} (merid a i₁) = --- ((λ i → inl ((merid a ∙ sym (merid (pt A))) i)) --- ∙∙ push tt --- ∙∙ λ i → inr ((merid a ∙ sym (merid (pt A))) i)) i₁ - - --- help3 : ∀ {ℓ} {A : Type ℓ} {x y z w : A} (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) --- → p ∙∙ q ∙∙ r ≡ p ∙ q ∙ r --- help3 {x = x} {y = y} {z = z} {w = w} = --- J (λ y p → (q : y ≡ z) (r : z ≡ w) → --- (p ∙∙ q ∙∙ r) ≡ p ∙ q ∙ r) --- λ q r → lUnit (q ∙ r) - --- +≡ : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → (x : _) → ∙Π f g .fst x ≡ (f ∨→ g) (θ {A = S₊∙ _} x) --- +≡ n f g north = sym (snd f) --- +≡ n f g south = sym (snd g) --- +≡ n f g (merid a i₁) j = main j i₁ --- where - --- help : cong (f ∨→ g) (cong (θ {A = S₊∙ _}) (merid a)) --- ≡ (refl ∙∙ cong (fst f) (merid a ∙ sym (merid north)) ∙∙ snd f) --- ∙ (sym (snd g) ∙∙ cong (fst g) (merid a ∙ sym (merid north)) ∙∙ refl) --- help = cong-∙∙ (f ∨→ g) ((λ i → inl ((merid a ∙ sym (merid north)) i))) --- (push tt) --- (λ i → inr ((merid a ∙ sym (merid north)) i)) --- ∙∙ help3 _ _ _ --- ∙∙ cong (cong (f ∨→ g) --- (λ i₂ → inl ((merid a ∙ (λ i₃ → merid north (~ i₃))) i₂)) ∙_) --- (sym (assoc _ _ _)) --- ∙∙ assoc _ _ _ --- ∙∙ cong₂ _∙_ refl (compPath≡compPath' _ _) - --- main : PathP (λ i → snd f (~ i) ≡ snd g (~ i)) (λ i₁ → ∙Π f g .fst (merid a i₁)) --- λ i₁ → (f ∨→ g) (θ {A = S₊∙ _} (merid a i₁)) --- main = (λ i → ((λ j → snd f (~ i ∧ ~ j)) ∙∙ cong (fst f) (merid a ∙ sym (merid north)) ∙∙ snd f) --- ∙ (sym (snd g) ∙∙ cong (fst g) (merid a ∙ sym (merid north)) ∙∙ λ j → snd g (~ i ∧ j))) --- ▷ sym help - --- +≡' : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → ∙Π f g ≡ ((f ∨→ g) ∘ (θ {A = S₊∙ _}) , snd f) --- +≡' n f g = ΣPathP ((funExt (+≡ n f g)) , (λ j i → snd f (i ∨ ~ j))) - --- WedgeElim : (n : ℕ) → ∀ {ℓ} {P : S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) → Type ℓ} --- → ((x : _) → isOfHLevel (3 +ℕ n) (P x)) --- → P (inl north) --- → (x : _) → P x --- WedgeElim n {P = P} x s (inl x₁) = --- sphereElim _ {A = P ∘ inl} --- (λ _ → isOfHLevelPlus' {n = n} (3 +ℕ n) (x _)) s x₁ --- WedgeElim n {P = P} x s (inr x₁) = --- sphereElim _ {A = P ∘ inr} --- (sphereElim _ (λ _ → isProp→isOfHLevelSuc ((suc (suc (n +ℕ n)))) --- (isPropIsOfHLevel (suc (suc (suc (n +ℕ n)))))) --- (subst (isOfHLevel ((3 +ℕ n) +ℕ n)) (cong P (push tt)) --- (isOfHLevelPlus' {n = n} (3 +ℕ n) (x _)))) --- (subst P (push tt) s) x₁ --- WedgeElim n {P = P} x s (push a j) = transp (λ i → P (push tt (i ∧ j))) (~ j) s - - --- H²-C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → GroupIso (coHomGr (2 +ℕ n) (C* n f g)) ℤGroup --- H²-C* n f g = compGroupIso groupIso1 (Hⁿ-Sⁿ≅ℤ (suc n)) --- where --- help : (coHom (2 +ℕ n) (C* n f g)) → coHom (2 +ℕ n) (S₊ (2 +ℕ n)) --- help = sMap λ f → f ∘ inr - - --- invMapPrim : (S₊ (2 +ℕ n) → coHomK (2 +ℕ n)) → C* n f g → coHomK (2 +ℕ n) --- invMapPrim h (inl x) = h (ptSn _) --- invMapPrim h (inr x) = h x --- invMapPrim h (push a i₁) = WedgeElim n {P = λ a → h north ≡ h (fst (add2 f g) a)} --- (λ _ → isOfHLevelTrunc (4 +ℕ n) _ _) --- (cong h (sym (snd f))) a i₁ - --- invMap : coHom (2 +ℕ n) (S₊ (2 +ℕ n)) → (coHom (2 +ℕ n) (C* n f g)) --- invMap = sMap invMapPrim - --- ret1 : (x : coHom (2 +ℕ n) (S₊ (2 +ℕ n))) → help (invMap x) ≡ x --- ret1 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) --- λ a → refl - --- ret2 : (x : (coHom (2 +ℕ n) (C* n f g))) → invMap (help x) ≡ x --- ret2 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) --- λ h → cong ∣_∣₂ (funExt λ { (inl x) → cong h ((λ i → inr ((snd f) (~ i))) --- ∙ sym (push (inl north))) --- ; (inr x) → refl --- ; (push a i₁) → help2 h a i₁}) --- where --- help2 : (h : C* n f g → coHomK (2 +ℕ n)) --- → (a : _) → PathP (λ i → invMapPrim (h ∘ inr) (push a i) ≡ h (push a i)) --- (cong h ((λ i → inr ((snd f) (~ i))) ∙ sym (push (inl north)))) --- refl --- help2 h = WedgeElim n (λ _ → isOfHLevelPathP (3 +ℕ n) (isOfHLevelTrunc (4 +ℕ n) _ _) _ _) --- help4 - --- where --- help4 : PathP (λ i → invMapPrim (h ∘ inr) (push (inl north) i) ≡ h (push (inl north) i)) --- (cong h ((λ i → inr ((snd f) (~ i))) ∙ sym (push (inl north)))) --- refl --- help4 i j = h (hcomp (λ k → λ { (i = i1) → inr (fst f north) --- ; (j = i0) → inr (snd f (~ i)) --- ; (j = i1) → push (inl north) (i ∨ ~ k)}) --- (inr (snd f (~ i ∧ ~ j)))) - --- groupIso1 : GroupIso ((coHomGr (2 +ℕ n) (C* n f g))) (coHomGr (2 +ℕ n) (S₊ (2 +ℕ n))) --- Iso.fun (fst groupIso1) = help --- Iso.inv (fst groupIso1) = invMap --- Iso.rightInv (fst groupIso1) = ret1 --- Iso.leftInv (fst groupIso1) = ret2 --- snd groupIso1 = makeIsGroupHom --- (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) --- λ f g → refl) - - --- C+ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Type _ --- C+ n f g = Pushout (λ _ → tt) (∙Π f g .fst) - --- H⁴-C∨ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → GroupIso (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C+ n f g)) --- ℤGroup --- H⁴-C∨ n f g = compGroupIso --- (transportCohomIso (cong (3 +ℕ_) (+-suc n n))) (Hopfβ-Iso n (∙Π f g)) - --- lol : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) (x : ℤ) --- → Iso.inv (fst (H⁴-C∨ n f g)) x --- ≡ subst (λ m → coHom m (C+ n f g)) (sym (cong (3 +ℕ_) (+-suc n n))) (Iso.inv (fst (Hopfβ-Iso n (∙Π f g))) x) --- lol n f g = λ _ → refl - --- H²-C∨ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → GroupIso (coHomGr (2 +ℕ n) (C+ n f g)) --- ℤGroup --- H²-C∨ n f g = Hopfα-Iso n (∙Π f g) - - --- H⁴-C* : ∀ n → (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → GroupIso (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) --- (DirProd ℤGroup ℤGroup) --- H⁴-C* n f g = compGroupIso (GroupEquiv→GroupIso (invGroupEquiv fstIso)) --- (compGroupIso (transportCohomIso (cong (2 +ℕ_) (+-suc n n))) --- (compGroupIso (Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) _) --- (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ _) (Hⁿ-Sⁿ≅ℤ _)))) --- where --- module Ms = MV _ _ _ (λ _ → tt) (fst (add2 f g)) --- fstIso : GroupEquiv (coHomGr ((suc (suc (n +ℕ suc n)))) (S₊∙ (3 +ℕ (n +ℕ n)) ⋁ S₊∙ (3 +ℕ (n +ℕ n)))) --- (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) --- fst (fst fstIso) = Ms.d (suc (suc (n +ℕ suc n))) .fst --- snd (fst fstIso) = --- SES→isEquiv --- (GroupPath _ _ .fst (compGroupEquiv (invEquiv (isContr→≃Unit ((tt , tt) , (λ _ → refl))) , makeIsGroupHom λ _ _ → refl) --- (GroupEquivDirProd --- (GroupIso→GroupEquiv (invGroupIso (Hⁿ-Unit≅0 _))) --- (GroupIso→GroupEquiv --- (invGroupIso --- (Hⁿ-Sᵐ≅0 _ _ λ p → ¬lemHopf 0 ((λ _ → north) , refl) n n (cong suc (sym (+-suc n n)) ∙ p))))))) --- (GroupPath _ _ .fst --- (compGroupEquiv ((invEquiv (isContr→≃Unit ((tt , tt) , (λ _ → refl))) , makeIsGroupHom λ _ _ → refl)) --- ((GroupEquivDirProd --- (GroupIso→GroupEquiv (invGroupIso (Hⁿ-Unit≅0 _))) --- (GroupIso→GroupEquiv --- (invGroupIso (Hⁿ-Sᵐ≅0 _ _ λ p → ¬lemHopf 0 ((λ _ → north) , refl) n (suc n) (cong (2 +ℕ_) (sym (+-suc n n)) ∙ p)))))))) --- (Ms.Δ (suc (suc (n +ℕ suc n)))) --- (Ms.d (suc (suc (n +ℕ suc n)))) --- (Ms.i (suc (suc (suc (n +ℕ suc n))))) --- (Ms.Ker-d⊂Im-Δ _) --- (Ms.Ker-i⊂Im-d _) --- snd fstIso = Ms.d (suc (suc (n +ℕ suc n))) .snd - --- Hopfie : {n : ℕ} → ∥ S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n) ∥₂ → ℤ --- Hopfie = sRec isSetℤ (HopfInvariant _) - --- subber : (n : ℕ) (f : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → _ --- subber n f = subst (λ x → coHom x (HopfInvariantPush n (fst f))) --- (λ i → suc (suc (suc (+-suc n n i)))) - --- module _ (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) where - --- C+' = C+ n f g - --- βₗ : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) --- βₗ = Iso.inv (fst (H⁴-C* n f g)) (1 , 0) - --- βᵣ : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) --- βᵣ = Iso.inv (fst (H⁴-C* n f g)) (0 , 1) - --- βᵣ'-fun : C* n f g → coHomK ((4 +ℕ (n +ℕ n))) --- βᵣ'-fun (inl x) = 0ₖ _ --- βᵣ'-fun (inr x) = 0ₖ _ --- βᵣ'-fun (push (inl x) i₁) = 0ₖ _ --- βᵣ'-fun (push (inr x) i₁) = Kn→ΩKn+1 _ ∣ x ∣ i₁ --- βᵣ'-fun (push (push a i₂) i₁) = Kn→ΩKn+10ₖ _ (~ i₂) i₁ - - --- βₗ'-fun : C* n f g → coHomK (4 +ℕ (n +ℕ n)) --- βₗ'-fun (inl x) = 0ₖ _ --- βₗ'-fun (inr x) = 0ₖ _ --- βₗ'-fun (push (inl x) i₁) = Kn→ΩKn+1 _ ∣ x ∣ i₁ --- βₗ'-fun (push (inr x) i₁) = 0ₖ _ --- βₗ'-fun (push (push a i₂) i₁) = Kn→ΩKn+10ₖ _ i₂ i₁ - --- βₗ''-fun : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) --- βₗ''-fun = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) ∣ βₗ'-fun ∣₂ - --- βᵣ''-fun : coHom ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) --- βᵣ''-fun = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) ∣ βᵣ'-fun ∣₂ - --- α∨ : coHom (2 +ℕ n) (C* n f g) --- α∨ = Iso.inv (fst (H²-C* n f g)) 1 - --- private --- pp : (a : _) → 0ₖ (suc (suc n)) ≡ ∣ (f ∨→ g) a ∣ --- pp = WedgeElim _ (λ _ → isOfHLevelTrunc (4 +ℕ n) _ _) --- (cong ∣_∣ₕ (sym (snd f))) - --- pp-inr : pp (inr north) ≡ cong ∣_∣ₕ (sym (snd g)) --- pp-inr = (λ j → transport (λ i → 0ₖ (2 +ℕ n) ≡ ∣ compPath-filler' (snd f) (sym (snd g)) (~ j) i ∣ₕ) --- λ i → ∣ snd f (~ i ∨ j) ∣ₕ) --- ∙ λ j → transp (λ i → 0ₖ (2 +ℕ n) ≡ ∣ snd g (~ i ∧ ~ j) ∣ₕ) j λ i → ∣ snd g (~ i ∨ ~ j) ∣ₕ - --- α∨'-fun : C* n f g → coHomK (2 +ℕ n) --- α∨'-fun (inl x) = 0ₖ _ --- α∨'-fun (inr x) = ∣ x ∣ --- α∨'-fun (push a i₁) = pp a i₁ - --- α∨' : coHom (2 +ℕ n) (C* n f g) --- α∨' = ∣ α∨'-fun ∣₂ - - --- β+ : coHom ((2 +ℕ n) +' (2 +ℕ n)) C+' --- β+ = Iso.inv (fst (H⁴-C∨ n f g)) 1 - --- q : C+' → C* n f g --- q (inl x) = inl x --- q (inr x) = inr x --- q (push a i₁) = (push (θ a) ∙ λ i → inr (+≡ n f g a (~ i))) i₁ - --- jₗ : HopfInvariantPush n (fst f) → C* n f g --- jₗ (inl x) = inl x --- jₗ (inr x) = inr x --- jₗ (push a i₁) = push (inl a) i₁ - --- jᵣ : HopfInvariantPush n (fst g) → C* n f g --- jᵣ (inl x) = inl x --- jᵣ (inr x) = inr x --- jᵣ (push a i₁) = push (inr a) i₁ - --- α-∨→1 : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → Iso.fun (fst (H²-C* n f g)) (α∨' n f g) ≡ 1 --- α-∨→1 n f g = sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (Hⁿ-Sⁿ≅ℤ-nice-generator n)) --- ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) - --- α-∨-lem : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → α∨ n f g ≡ α∨' n f g --- α-∨-lem n f g = sym (Iso.leftInv (fst (H²-C* n f g)) _) --- ∙∙ cong (Iso.inv (fst (H²-C* n f g))) help --- ∙∙ Iso.leftInv (fst (H²-C* n f g)) _ --- where --- help : Iso.fun (fst (H²-C* n f g)) (α∨ n f g) ≡ Iso.fun (fst (H²-C* n f g)) (α∨' n f g) --- help = (Iso.rightInv (fst (H²-C* n f g)) (pos 1)) ∙ sym (α-∨→1 n f g) - --- q-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → coHomFun _ (q n f g) (α∨ n f g) ≡ Hopfα n (∙Π f g) --- q-α n f g = (λ i → coHomFun _ (q n f g) (α-∨-lem n f g i)) --- ∙ sym (Iso.leftInv is _) --- ∙∙ cong (Iso.inv is) help --- ∙∙ Iso.leftInv is _ --- where --- is = fst (Hopfα-Iso n (∙Π f g)) - --- help : Iso.fun is (coHomFun _ (q n f g) (α∨' n f g)) --- ≡ Iso.fun is (Hopfα n (∙Π f g)) --- help = sym (cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ (suc n)))) (Hⁿ-Sⁿ≅ℤ-nice-generator n)) --- ∙∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) (pos 1) --- ∙∙ sym (Hopfα-Iso-α n (∙Π f g)) - - --- βₗ≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → βₗ n f g ≡ βₗ''-fun n f g --- βₗ≡ n f g = cong (FF ∘ subber2) --- (cong (Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) --- (S₊∙ (suc (suc (suc (n +ℕ n))))) --- (((suc (suc (n +ℕ n))))))))) help --- ∙ help2) --- ∙ funExt⁻ FF∘subber ∣ wedgeGen ∣₂ --- ∙ cong subber3 (sym βₗ'-fun≡) --- where --- FF = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (n +ℕ suc n))) .fst --- FF' = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (suc (n +ℕ n)))) .fst - --- help : Iso.inv (fst (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (1 , 0) --- ≡ (∣ ∣_∣ ∣₂ , 0ₕ _) --- help = ΣPathP ((Hⁿ-Sⁿ≅ℤ-nice-generator _) , IsGroupHom.pres1 (snd (invGroupIso (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) - --- subber3 = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - --- subber2 = (subst (λ m → coHom m (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) --- (sym (cong (2 +ℕ_) (+-suc n n)))) - --- FF∘subber : FF ∘ subber2 ≡ subber3 ∘ FF' --- FF∘subber = --- funExt (λ a → (λ j → transp (λ i → coHom ((suc ∘ suc ∘ suc) (+-suc n n (~ i ∧ j))) (C* n f g)) (~ j) --- (MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) ((suc (suc (+-suc n n j)))) .fst --- (transp (λ i → coHom (2 +ℕ (+-suc n n (j ∨ ~ i))) --- (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) j --- a)))) - --- wedgeGen : (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) → --- coHomK (suc (suc (suc (n +ℕ n))))) --- wedgeGen (inl x) = ∣ x ∣ --- wedgeGen (inr x) = 0ₖ _ --- wedgeGen (push a i₁) = 0ₖ _ - --- βₗ'-fun≡ : ∣ βₗ'-fun n f g ∣₂ ≡ FF' ∣ wedgeGen ∣₂ --- βₗ'-fun≡ = cong ∣_∣₂ (funExt λ { (inl x) → refl --- ; (inr x) → refl --- ; (push (inl x) i₁) → refl --- ; (push (inr x) i₁) j → Kn→ΩKn+10ₖ _ (~ j) i₁ --- ; (push (push a i₂) i₁) j → Kn→ΩKn+10ₖ _ (~ j ∧ i₂) i₁}) - --- help2 : Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) (((suc (suc (n +ℕ n)))))))) --- (∣ ∣_∣ ∣₂ , 0ₕ _) --- ≡ ∣ wedgeGen ∣₂ --- help2 = cong ∣_∣₂ (funExt λ { (inl x) → rUnitₖ (suc (suc (suc (n +ℕ n)))) ∣ x ∣ --- ; (inr x) → refl --- ; (push a i₁) → refl}) - --- βᵣ≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → βᵣ n f g ≡ βᵣ''-fun n f g --- βᵣ≡ n f g = --- cong (FF ∘ subber2) --- (cong (Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) --- (S₊∙ (suc (suc (suc (n +ℕ n))))) --- (((suc (suc (n +ℕ n))))))))) --- help --- ∙ help2) --- ∙ funExt⁻ FF∘subber ∣ wedgeGen ∣₂ --- ∙ cong (subber3) (sym βᵣ'-fun≡) --- where --- FF = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (n +ℕ suc n))) .fst --- FF' = MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) (suc (suc (suc (n +ℕ n)))) .fst - --- help : Iso.inv (fst (GroupIsoDirProd (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))) (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (0 , 1) --- ≡ (0ₕ _ , ∣ ∣_∣ ∣₂) --- help = ΣPathP (IsGroupHom.pres1 (snd (invGroupIso (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) , (Hⁿ-Sⁿ≅ℤ-nice-generator _)) - --- subber3 = subst (λ m → coHom m (C* n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) - --- subber2 = (subst (λ m → coHom m (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) --- (sym (cong (2 +ℕ_) (+-suc n n)))) - --- FF∘subber : FF ∘ subber2 ≡ subber3 ∘ FF' --- FF∘subber = --- funExt (λ a → (λ j → transp (λ i → coHom ((suc ∘ suc ∘ suc) (+-suc n n (~ i ∧ j))) (C* n f g)) (~ j) --- (MV.d _ _ _ (λ _ → tt) (fst (add2 f g)) ((suc (suc (+-suc n n j)))) .fst --- (transp (λ i → coHom (2 +ℕ (+-suc n n (j ∨ ~ i))) --- (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))))) j --- a)))) - --- wedgeGen : (S₊∙ (suc (suc (suc (n +ℕ n)))) ⋁ S₊∙ (suc (suc (suc (n +ℕ n)))) → --- coHomK (suc (suc (suc (n +ℕ n))))) --- wedgeGen (inl x) = 0ₖ _ --- wedgeGen (inr x) = ∣ x ∣ --- wedgeGen (push a i₁) = 0ₖ _ - - --- help2 : Iso.inv (fst ((Hⁿ-⋁ (S₊∙ (suc (suc (suc (n +ℕ n))))) (S₊∙ (suc (suc (suc (n +ℕ n))))) (((suc (suc (n +ℕ n)))))))) --- (0ₕ _ , ∣ ∣_∣ ∣₂) --- ≡ ∣ wedgeGen ∣₂ --- help2 = cong ∣_∣₂ (funExt λ { (inl x) → refl --- ; (inr x) → lUnitₖ (suc (suc (suc (n +ℕ n)))) ∣ x ∣ --- ; (push a i₁) → refl}) - --- βᵣ'-fun≡ : ∣ βᵣ'-fun n f g ∣₂ ≡ FF' ∣ wedgeGen ∣₂ --- βᵣ'-fun≡ = cong ∣_∣₂ (funExt λ { (inl x) → refl --- ; (inr x) → refl --- ; (push (inl x) i₁) j → Kn→ΩKn+10ₖ _ (~ j) i₁ --- ; (push (inr x) i₁) → refl --- ; (push (push a i₂) i₁) j → Kn→ΩKn+10ₖ _ (~ j ∧ ~ i₂) i₁}) - --- q-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → coHomFun _ (q n f g) (βₗ n f g) --- ≡ β+ n f g --- q-βₗ n f g = cong (coHomFun _ (q n f g)) (βₗ≡ n f g) --- ∙ transportLem --- ∙ cong (subst (λ m → coHom m (C+' n f g)) --- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) --- (sym (Iso.leftInv (fst (Hopfβ-Iso n (∙Π f g))) (Hopfβ n (∙Π f g))) --- ∙ cong (Iso.inv ((fst (Hopfβ-Iso n (∙Π f g))))) (Hopfβ↦1 n (∙Π f g))) --- where --- transportLem : coHomFun (suc (suc (suc (n +ℕ suc n)))) (q n f g) --- (βₗ''-fun n f g) --- ≡ subst (λ m → coHom m (C+' n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- (Hopfβ n (∙Π f g)) --- transportLem = --- natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (q n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- ∙ cong (subst (λ m → coHom m (C+' n f g)) --- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) --- (cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → loll a i₁})) --- where --- open import Cubical.Foundations.Path --- loll : (x : fst (S₊∙ (3 +ℕ (n +ℕ n)))) → --- PathP --- (λ i₁ → --- βₗ'-fun n f g ((push (θ x) ∙ λ i → inr (+≡ n f g x (~ i))) i₁) ≡ --- MV.d-pre Unit (fst (S₊∙ (2 +ℕ n))) (fst (S₊∙ (3 +ℕ n +ℕ n))) --- (λ _ → tt) (fst (∙Π f g)) (3 +ℕ n +ℕ n) ∣_∣ (push x i₁)) --- (λ _ → βₗ'-fun n f g (q n f g (inl tt))) --- (λ _ → βₗ'-fun n f g (q n f g (inr (∙Π f g .fst x)))) --- loll x = --- flipSquare (cong-∙ (βₗ'-fun n f g) (push (θ x)) (λ i → inr (+≡ n f g x (~ i))) --- ∙∙ sym (rUnit _) --- ∙∙ lem₁ x) - --- where --- lem₁ : (x : _) → cong (βₗ'-fun n f g) (push (θ {A = S₊∙ _} x)) ≡ Kn→ΩKn+1 _ ∣ x ∣ --- lem₁ north = refl --- lem₁ south = sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)) --- lem₁ (merid a j) k = lala k j --- where --- lala : PathP (λ k → Kn→ΩKn+1 _ ∣ north ∣ₕ ≡ (sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north))) k) --- (λ j → cong (βₗ'-fun n f g) (push (θ {A = S₊∙ _} (merid a j)))) --- (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) --- lala = (cong-∙∙ (cong (βₗ'-fun n f g) ∘ push) --- ((λ i₁ → inl ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁))) --- (push tt) --- ((λ i₁ → inr ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁))) --- ∙ sym (compPath≡compPath' _ _) --- ∙ (λ _ → (λ j → Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∣ (merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) j ∣ₕ) --- ∙ Kn→ΩKn+10ₖ _) --- ∙ cong (_∙ Kn→ΩKn+10ₖ _) (cong-∙ ((Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ)) --- (merid a) (sym (merid north))) --- ∙ sym (assoc _ _ _) --- ∙ cong (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a) ∙_) --- (sym (symDistr (sym ((Kn→ΩKn+10ₖ _))) (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)))))) --- ◁ λ i j → compPath-filler (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) --- (sym (sym (Kn→ΩKn+10ₖ _) ∙ cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) --- (cong ∣_∣ₕ (merid north)))) --- (~ i) j - --- q-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → coHomFun _ (q n f g) (βᵣ n f g) --- ≡ β+ n f g --- q-βᵣ n f g = cong (coHomFun _ (q n f g)) (βᵣ≡ n f g) --- ∙ transportLem --- ∙ cong (subst (λ m → coHom m (C+' n f g)) --- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) --- (sym (Iso.leftInv (fst (Hopfβ-Iso n (∙Π f g))) (Hopfβ n (∙Π f g))) --- ∙ cong (Iso.inv ((fst (Hopfβ-Iso n (∙Π f g))))) (Hopfβ↦1 n (∙Π f g))) --- where --- transportLem : coHomFun (suc (suc (suc (n +ℕ suc n)))) (q n f g) --- (βᵣ''-fun n f g) --- ≡ subst (λ m → coHom m (C+' n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- (Hopfβ n (∙Π f g)) --- transportLem = --- natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (q n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- ∙ cong (subst (λ m → coHom m (C+' n f g)) --- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) --- (cong ∣_∣₂ (funExt λ { (inl x) → refl --- ; (inr x) → refl --- ; (push a i₁) → loll a i₁})) --- where --- open import Cubical.Foundations.Path --- loll : (x : fst (S₊∙ (3 +ℕ (n +ℕ n)))) → --- PathP --- (λ i₁ → --- βᵣ'-fun n f g ((push (θ x) ∙ λ i → inr (+≡ n f g x (~ i))) i₁) ≡ --- MV.d-pre Unit (fst (S₊∙ (2 +ℕ n))) (fst (S₊∙ (3 +ℕ n +ℕ n))) --- (λ _ → tt) (fst (∙Π f g)) (3 +ℕ n +ℕ n) ∣_∣ (push x i₁)) --- (λ _ → βᵣ'-fun n f g (q n f g (inl tt))) --- (λ _ → βᵣ'-fun n f g (q n f g (inr (∙Π f g .fst x)))) --- loll x = flipSquare (cong-∙ (βᵣ'-fun n f g) (push (θ x)) (λ i → inr (+≡ n f g x (~ i))) --- ∙∙ sym (rUnit _) --- ∙∙ lem₁ x) --- where --- lem₁ : (x : _) → cong (βᵣ'-fun n f g) (push (θ {A = S₊∙ _} x)) ≡ Kn→ΩKn+1 _ ∣ x ∣ --- lem₁ north = sym (Kn→ΩKn+10ₖ _) --- lem₁ south = cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north)) --- lem₁ (merid a j) k = lala k j -- lala k j --- where --- lala : PathP (λ k → (Kn→ΩKn+10ₖ _) (~ k) ≡ (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n))))) (cong ∣_∣ₕ (merid north))) k) --- (λ j → cong (βᵣ'-fun n f g) (push (θ {A = S₊∙ _} (merid a j)))) --- (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) --- lala = (cong-∙∙ (cong (βᵣ'-fun n f g) ∘ push) --- (λ i₁ → inl ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁)) --- (push tt) --- (λ i₁ → inr ((merid a ∙ (sym (merid (pt (S₊∙ (suc (suc (n +ℕ n)))))))) i₁)) --- ∙∙ (cong (sym (Kn→ΩKn+10ₖ _) ∙_) (cong-∙ (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a) (sym (merid (ptSn _))))) --- ∙∙ sym (help3 _ _ _)) --- ◁ symP (doubleCompPath-filler --- (sym (Kn→ΩKn+10ₖ _)) --- (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (merid a)) --- (cong (Kn→ΩKn+1 (suc (suc (suc (n +ℕ n)))) ∘ ∣_∣ₕ) (sym (merid north)))) --- open import Cubical.Foundations.Path --- jₗ-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → coHomFun _ (jₗ n f g) (α∨ n f g) --- ≡ Hopfα n f --- jₗ-α n f g = --- cong (coHomFun _ (jₗ n f g)) (α-∨-lem n f g) --- ∙ cong ∣_∣₂ (funExt (HopfInvariantPushElim n (fst f) --- (isOfHLevelPath ((suc (suc (suc (suc (n +ℕ n)))))) --- (isOfHLevelPlus' {n = n} (4 +ℕ n) (isOfHLevelTrunc (4 +ℕ n))) _ _) --- refl --- (λ _ → refl) --- λ j → refl)) - --- jₗ-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → coHomFun _ (jₗ n f g) (βₗ n f g) --- ≡ subst (λ m → coHom m (HopfInvariantPush n (fst f))) --- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- (Hopfβ n f) --- jₗ-βₗ n f g = --- cong (coHomFun _ (jₗ n f g)) (βₗ≡ n f g) --- ∙∙ natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (jₗ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) --- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) --- (cong ∣_∣₂ --- (funExt λ { (inl x) → refl --- ; (inr x) → refl --- ; (push a i₁) → refl})) - --- jₗ-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → coHomFun _ (jₗ n f g) (βᵣ n f g) --- ≡ 0ₕ _ --- jₗ-βᵣ n f g = --- cong (coHomFun _ (jₗ n f g)) (βᵣ≡ n f g) --- ∙∙ natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (jₗ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) --- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) --- cool --- where --- cool : coHomFun (suc (suc (suc (suc (n +ℕ n))))) (jₗ n f g) --- ∣ βᵣ'-fun n f g ∣₂ ≡ 0ₕ _ --- cool = cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → refl}) - --- jᵣ-α : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → coHomFun _ (jᵣ n f g) (α∨ n f g) --- ≡ Hopfα n g --- jᵣ-α n f g = cong (coHomFun _ (jᵣ n f g)) (α-∨-lem n f g) --- ∙ cong ∣_∣₂ (funExt (HopfInvariantPushElim n (fst g) --- (isOfHLevelPath ((suc (suc (suc (suc (n +ℕ n)))))) --- (isOfHLevelPlus' {n = n} (4 +ℕ n) (isOfHLevelTrunc (4 +ℕ n))) _ _) --- refl --- (λ _ → refl) --- (flipSquare (pp-inr n f g)))) - --- jᵣ-βₗ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → coHomFun _ (jᵣ n f g) (βₗ n f g) ≡ 0ₕ _ --- jᵣ-βₗ n f g = --- cong (coHomFun _ (jᵣ n f g)) (βₗ≡ n f g) --- ∙∙ natTranspLem ∣ βₗ'-fun n f g ∣₂ (λ m → coHomFun m (jᵣ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst f))) --- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) --- cool --- where --- cool : coHomFun (suc (suc (suc (suc (n +ℕ n))))) (jᵣ n f g) --- ∣ βₗ'-fun n f g ∣₂ ≡ 0ₕ _ --- cool = cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i₁) → refl}) - - --- jᵣ-βᵣ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → coHomFun _ (jᵣ n f g) (βᵣ n f g) --- ≡ subst (λ m → coHom m (HopfInvariantPush n (fst g))) --- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- (Hopfβ n g) --- jᵣ-βᵣ n f g = --- cong (coHomFun _ (jᵣ n f g)) (βᵣ≡ n f g) --- ∙∙ natTranspLem ∣ βᵣ'-fun n f g ∣₂ (λ m → coHomFun m (jᵣ n f g)) (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- ∙∙ cong (subst (λ m → coHom m (HopfInvariantPush n (fst g))) --- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n)))) --- (cong ∣_∣₂ --- (funExt λ { (inl x) → refl --- ; (inr x) → refl --- ; (push a i₁) → refl})) - --- gen₂ℤ×ℤ : gen₂-by (DirProd ℤGroup ℤGroup) (1 , 0) (0 , 1) --- fst (gen₂ℤ×ℤ (x , y)) = x , y --- snd (gen₂ℤ×ℤ (x , y)) = --- ΣPathP ((cong₂ _+_ ((·Comm 1 x) ∙ cong fst (sym (distrLem 1 0 x))) ((·Comm 0 y) ∙ cong fst (sym (distrLem 0 1 y)))) --- , +Comm y 0 ∙ cong₂ _+_ (·Comm 0 x ∙ cong snd (sym (distrLem 1 0 x))) (·Comm 1 y ∙ cong snd (sym (distrLem 0 1 y)))) --- where --- ℤ×ℤ = DirProd ℤGroup ℤGroup --- _+''_ = GroupStr._·_ (snd ℤ×ℤ) - --- ll3 : (x : ℤ) → - x ≡ 0 - x --- ll3 (pos zero) = refl --- ll3 (pos (suc zero)) = refl --- ll3 (pos (suc (suc n))) = --- cong predℤ (ll3 (pos (suc n))) --- ll3 (negsuc zero) = refl --- ll3 (negsuc (suc n)) = cong sucℤ (ll3 (negsuc n)) - --- distrLem : (x y : ℤ) (z : ℤ) --- → Path (ℤ × ℤ) (z ℤ[ ℤ×ℤ ]· (x , y)) (z · x , z · y) --- distrLem x y (pos zero) = refl --- distrLem x y (pos (suc n)) = --- (cong ((x , y) +''_) (distrLem x y (pos n))) --- distrLem x y (negsuc zero) = ΣPathP (sym (ll3 x) , sym (ll3 y)) --- distrLem x y (negsuc (suc n)) = --- cong₂ _+''_ (ΣPathP (sym (ll3 x) , sym (ll3 y))) --- (distrLem x y (negsuc n)) - --- genH²ⁿC* : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → gen₂-by (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) --- (βₗ n f g) --- (βᵣ n f g) --- genH²ⁿC* n f g = --- Iso-pres-gen₂ (DirProd ℤGroup ℤGroup) --- (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) --- (1 , 0) --- (0 , 1) --- gen₂ℤ×ℤ --- (invGroupIso (H⁴-C* n f g)) - --- private - --- H⁴C* : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → Group _ --- H⁴C* n f g = coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g) - --- X : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → ℤ --- X n f g = (genH²ⁿC* n f g) (α∨ n f g ⌣ α∨ n f g) .fst .fst --- Y : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → ℤ --- Y n f g = (genH²ⁿC* n f g) (α∨ n f g ⌣ α∨ n f g) .fst .snd - --- α∨≡ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → α∨ n f g ⌣ α∨ n f g ≡ ((X n f g) ℤ[ H⁴C* n f g ]· βₗ n f g) --- +ₕ ((Y n f g) ℤ[ H⁴C* n f g ]· βᵣ n f g) --- α∨≡ n f g = (genH²ⁿC* n f g) (α∨ n f g ⌣ α∨ n f g) .snd - - --- eq₁ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → (Hopfα n (∙Π f g)) ⌣ (Hopfα n (∙Π f g)) --- ≡ ((X n f g + Y n f g) ℤ[ coHomGr _ _ ]· (β+ n f g)) --- eq₁ n f g = --- cong (λ x → x ⌣ x) (sym (q-α n f g) ∙ cong (coHomFun (suc (suc n)) (q n f g)) (α-∨-lem n f g)) --- ∙ cong ((coHomFun _) (q _ f g)) (cong (λ x → x ⌣ x) (sym (α-∨-lem n f g)) --- ∙ α∨≡ n f g) --- ∙ IsGroupHom.pres· (coHomMorph _ (q n f g) .snd) _ _ --- ∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (q n f g)) (βₗ n f g) (X n f g) --- ∙ cong (λ z → (X n f g) ℤ[ coHomGr _ _ ]· z) --- (q-βₗ n f g)) --- (homPresℤ· (coHomMorph _ (q n f g)) (βᵣ n f g) (Y n f g) --- ∙ cong (λ z → (Y n f g) ℤ[ coHomGr _ _ ]· z) --- (q-βᵣ n f g)) --- ∙ sym (distrℤ· (coHomGr _ _) (β+ n f g) (X n f g) (Y n f g)) - --- coHomFun⌣ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) --- → (n : ℕ) (x y : coHom n B) --- → coHomFun _ f (x ⌣ y) ≡ coHomFun _ f x ⌣ coHomFun _ f y --- coHomFun⌣ f n = sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ _ _ → refl - --- eq₂ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → Hopfα n f ⌣ Hopfα n f --- ≡ (X n f g ℤ[ coHomGr _ _ ]· subst (λ m → coHom m (HopfInvariantPush n (fst f))) --- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- (Hopfβ n f)) --- eq₂ n f g = --- cong (λ x → x ⌣ x) (sym (jₗ-α n f g)) --- ∙∙ sym (coHomFun⌣ (jₗ n f g) _ _ _) --- ∙∙ cong (coHomFun _ (jₗ n f g)) (α∨≡ n f g) --- ∙∙ IsGroupHom.pres· (snd (coHomMorph _ (jₗ n f g))) _ _ --- ∙∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (jₗ n f g)) (βₗ n f g) (X n f g)) --- (homPresℤ· (coHomMorph _ (jₗ n f g)) (βᵣ n f g) (Y n f g) --- ∙ cong (λ z → (Y n f g) ℤ[ coHomGr _ _ ]· z) --- (jₗ-βᵣ n f g)) --- ∙∙ cong₂ _+ₕ_ refl (rUnitℤ· (coHomGr _ _) (Y n f g)) --- ∙∙ (rUnitₕ _ _ --- ∙ cong (X n f g ℤ[ coHomGr _ _ ]·_) (jₗ-βₗ n f g)) - --- eq₃ : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) --- → Hopfα n g ⌣ Hopfα n g --- ≡ (Y n f g ℤ[ coHomGr _ _ ]· subst (λ m → coHom m (HopfInvariantPush n (fst f))) --- (cong (suc ∘ suc ∘ suc) (sym (+-suc n n))) --- (Hopfβ n g)) --- eq₃ n f g = --- cong (λ x → x ⌣ x) (sym (jᵣ-α n f g)) --- ∙∙ sym (coHomFun⌣ (jᵣ n f g) _ _ _) --- ∙∙ cong (coHomFun _ (jᵣ n f g)) (α∨≡ n f g) --- ∙∙ IsGroupHom.pres· (snd (coHomMorph _ (jᵣ n f g))) _ _ --- ∙∙ cong₂ _+ₕ_ (homPresℤ· (coHomMorph _ (jᵣ n f g)) (βₗ n f g) (X n f g) --- ∙ cong (λ z → (X n f g) ℤ[ coHomGr _ _ ]· z) --- (jᵣ-βₗ n f g)) --- (homPresℤ· (coHomMorph _ (jᵣ n f g)) (βᵣ n f g) (Y n f g)) --- ∙∙ cong₂ _+ₕ_ (rUnitℤ· (coHomGr _ _) (X n f g)) refl --- ∙∙ ((lUnitₕ _ _) ∙ cong (Y n f g ℤ[ coHomGr _ _ ]·_) (jᵣ-βᵣ n f g)) - --- eq₂-eq₂ : (n : ℕ) (f g : S₊∙ (suc (suc (suc (n +ℕ n)))) →∙ S₊∙ (suc (suc n))) --- → HopfInvariant n (∙Π f g) ≡ HopfInvariant n f + HopfInvariant n g --- eq₂-eq₂ n f g = --- mainL --- ∙ sym (cong₂ _+_ (help n f g) (helpg n f g)) --- where --- transpLem : ∀ {ℓ} {G : ℕ → Group ℓ} --- → (n m : ℕ) --- → (x : ℤ) --- → (p : m ≡ n) --- → (g : fst (G n)) --- → subst (fst ∘ G) p (x ℤ[ G m ]· subst (fst ∘ G) (sym p) g) ≡ (x ℤ[ G n ]· g) --- transpLem {G = G} n m x = --- J (λ n p → (g : fst (G n)) → subst (fst ∘ G) p (x ℤ[ G m ]· subst (fst ∘ G) (sym p) g) --- ≡ (x ℤ[ G n ]· g)) --- λ g → transportRefl _ ∙ cong (x ℤ[ G m ]·_) (transportRefl _) - --- mainL : HopfInvariant n (∙Π f g) ≡ X n f g + Y n f g --- mainL = cong (Iso.fun (fst (Hopfβ-Iso n (∙Π f g)))) --- (cong (subst (λ x → coHom x (HopfInvariantPush n (fst (∙Π f g)))) --- λ i₁ → suc (suc (suc (+-suc n n i₁)))) --- (eq₁ n f g)) --- ∙∙ cong (Iso.fun (fst (Hopfβ-Iso n (∙Π f g)))) --- (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst (∙Π f g)))} _ _ --- (X n f g + Y n f g) (λ i₁ → suc (suc (suc (+-suc n n i₁)))) --- (Iso.inv (fst (Hopfβ-Iso n (∙Π f g))) 1)) --- ∙∙ homPresℤ· (_ , snd (Hopfβ-Iso n (∙Π f g))) (Iso.inv (fst (Hopfβ-Iso n (∙Π f g))) 1) --- (X n f g + Y n f g) --- ∙∙ cong ((X n f g + Y n f g) ℤ[ ℤ , ℤGroup .snd ]·_) --- (Iso.rightInv ((fst (Hopfβ-Iso n (∙Π f g)))) 1) --- ∙∙ rUnitℤ·ℤ (X n f g + Y n f g) - - --- help : (n : ℕ) (f g : _) → HopfInvariant n f ≡ X n f g --- help n f g = cong (Iso.fun (fst (Hopfβ-Iso n f))) --- (cong (subst (λ x → coHom x (HopfInvariantPush n (fst f))) --- (λ i₁ → suc (suc (suc (+-suc n n i₁))))) (eq₂ n f g)) --- ∙ cong (Iso.fun (fst (Hopfβ-Iso n f))) --- (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst f))} _ _ --- (X n f g) --- ((cong (suc ∘ suc ∘ suc) (+-suc n n))) --- (Hopfβ n f)) --- ∙ homPresℤ· (_ , snd (Hopfβ-Iso n f)) (Hopfβ n f) (X n f g) --- ∙ cong (X n f g ℤ[ ℤ , ℤGroup .snd ]·_) (Hopfβ↦1 n f) --- ∙ rUnitℤ·ℤ (X n f g) - --- helpg : (n : ℕ) (f g : _) → HopfInvariant n g ≡ Y n f g --- helpg n f g = --- cong (Iso.fun (fst (Hopfβ-Iso n g))) --- (cong (subst (λ x → coHom x (HopfInvariantPush n (fst g))) --- (λ i₁ → suc (suc (suc (+-suc n n i₁))))) --- (eq₃ n f g)) --- ∙∙ cong (Iso.fun (fst (Hopfβ-Iso n g))) --- (transpLem {G = λ x → coHomGr x (HopfInvariantPush n (fst g))} _ _ --- (Y n f g) --- ((cong (suc ∘ suc ∘ suc) (+-suc n n))) --- (Hopfβ n g)) --- ∙∙ homPresℤ· (_ , snd (Hopfβ-Iso n g)) (Hopfβ n g) (Y n f g) --- ∙∙ cong (Y n f g ℤ[ ℤ , ℤGroup .snd ]·_) (Hopfβ↦1 n g) --- ∙∙ rUnitℤ·ℤ (Y n f g) diff --git a/Cubical/ZCohomology/Gysin.agda b/Cubical/ZCohomology/Gysin.agda index fea5e95b2..df70c73e9 100644 --- a/Cubical/ZCohomology/Gysin.agda +++ b/Cubical/ZCohomology/Gysin.agda @@ -13,7 +13,7 @@ open import Cubical.ZCohomology.RingStructure.CupProduct open import Cubical.ZCohomology.RingStructure.RingLaws open import Cubical.ZCohomology.RingStructure.GradedCommutativity open import Cubical.Relation.Nullary -open import Cubical.Homotopy.HomotopyGroup +open import Cubical.Homotopy.Group.Base open import Cubical.Functions.Embedding @@ -97,62 +97,6 @@ private Int-ind P z one min ind (negsuc (suc n)) = ind (negsuc n) (negsuc zero) (Int-ind P z one min ind (negsuc n)) min - --- Move to embedding? -Whitehead : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} - → (f : A → B) - → isEmbedding f - → (∀ (b : B) → ∃[ a ∈ A ] f a ≡ b) - → isEquiv f -equiv-proof (Whitehead f emb p) b = - pRec isPropIsContr - (λ fib → fib , isEmbedding→hasPropFibers emb b fib) - (p b) - --- move to cohom? -transportCohomIso : ∀ {ℓ} {A : Type ℓ} {n m : ℕ} - → (p : n ≡ m) - → GroupIso (coHomGr n A) (coHomGr m A) -Iso.fun (fst (transportCohomIso {A = A} p)) = subst (λ n → coHom n A) p -Iso.inv (fst (transportCohomIso {A = A} p)) = subst (λ n → coHom n A) (sym p) -Iso.rightInv (fst (transportCohomIso p)) = transportTransport⁻ _ -Iso.leftInv (fst (transportCohomIso p)) = transportTransport⁻ _ -snd (transportCohomIso {A = A} {n = n} {m = m} p) = - makeIsGroupHom (λ x y → help x y p) - where - help : (x y : coHom n A) (p : n ≡ m) → subst (λ n → coHom n A) p (x +ₕ y) - ≡ subst (λ n → coHom n A) p x +ₕ subst (λ n → coHom n A) p y - help x y = J (λ m p → subst (λ n → coHom n A) p (x +ₕ y) - ≡ subst (λ n → coHom n A) p x +ₕ subst (λ n → coHom n A) p y) - (transportRefl (x +ₕ y) ∙ cong₂ _+ₕ_ (sym (transportRefl x)) (sym (transportRefl y))) - - -private - _·₀ₕ_ : ∀ {ℓ} {n : ℕ} {A : Type ℓ} → ℤ → coHom n A → coHom n A - _·₀ₕ_ {n = n} a b = a ℤ[ coHomGr n _ ]· b - - GroupHomℤ→ℤPres·₀ : ∀ {ℓ} {A : Type ℓ} (n : ℕ) - (e : GroupHom ℤGroup (coHomGr n A)) - → (a b : ℤ) - → fst e (a · b) - ≡ (a ·₀ₕ fst e b) - GroupHomℤ→ℤPres·₀ {A = A} n e a b = - cong (fst e) (ℤ·≡· a b) - ∙ homPresℤ· {H = coHomGr n A} (_ , snd e) b a - --- keep -genH : ∀ {ℓ} {A : Type ℓ} (n : ℕ) - → (e : GroupIso (coHomGr n A) - ℤGroup) - → Σ[ x ∈ coHom n A ] - gen₁-by (coHomGr n A) x -genH {A = A} n e = (Iso.inv (fst e) 1) - , λ h → (Iso.fun (fst e) h) , - (sym (Iso.leftInv (fst e) h) - ∙∙ sym (cong (Iso.inv (fst e)) (·Comm (Iso.fun (fst e) h) (pos 1))) - ∙∙ (cong (Iso.inv (fst e)) (ℤ·≡· _ _) - ∙ (homPresℤ· (_ , snd (invGroupIso e)) (pos 1) (Iso.fun (fst e) h)))) - -- The untruncated version (coHomRed n (S₊ n)), i.e. -- (S₊∙ n →∙ coHomK-ptd n) is in fact equal to -- (coHomRed n (S₊ n)). Its useful to formulate @@ -197,9 +141,11 @@ module _ where is-set (isSemigroup (isMonoid (isGroup (snd (πS zero))))) = isOfHLevelΣ 2 (isSetΠ (λ _ → isSetℤ)) λ _ → isOfHLevelPath 2 isSetℤ _ _ - is-set (isSemigroup (isMonoid (isGroup (snd (πS (suc n)))))) = isOfHLevel↑∙' 0 n + is-set (isSemigroup (isMonoid (isGroup (snd (πS (suc n)))))) = + isOfHLevel↑∙' 0 n IsSemigroup.assoc (isSemigroup (isMonoid (isGroup (snd (πS n))))) x y z = - →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ w → assocₖ n (fst x w) (fst y w) (fst z w)) + →∙Homogeneous≡ (isHomogeneousKn n) + (funExt λ w → assocₖ n (fst x w) (fst y w) (fst z w)) fst (identity (isMonoid (isGroup (snd (πS n)))) (f , p)) = →∙Homogeneous≡ (isHomogeneousKn n) (funExt λ x → rUnitₖ n (f x)) snd (identity (isMonoid (isGroup (snd (πS n)))) (f , p)) = @@ -264,19 +210,25 @@ Iso-πS-ℤ n = compIso (invIso (setTruncIdempotentIso (isOfHLevel↑∙' 0 n))) Iso-πS-ℤPres1 : (n : ℕ) → Iso.fun (fst (πS≅ℤ (suc n))) (∣_∣ , refl) ≡ pos 1 Iso-πS-ℤPres1 zero = refl -Iso-πS-ℤPres1 (suc n) = cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ n))) (lem n) ∙ Iso-πS-ℤPres1 n +Iso-πS-ℤPres1 (suc n) = + cong (Iso.fun (fst (Hⁿ-Sⁿ≅ℤ n))) (lem n) ∙ Iso-πS-ℤPres1 n where lem : (n : ℕ) → Iso.fun (fst (suspensionAx-Sn n n)) ∣ ∣_∣ ∣₂ ≡ ∣ ∣_∣ ∣₂ - lem zero = cong ∣_∣₂ (funExt λ x → transportRefl (∣ x ∣ +ₖ (0ₖ 1)) ∙ rUnitₖ 1 ∣ x ∣) + lem zero = + cong ∣_∣₂ (funExt λ x → transportRefl (∣ x ∣ +ₖ (0ₖ 1)) ∙ rUnitₖ 1 ∣ x ∣) lem (suc n) = cong ∣_∣₂ (funExt λ x → (λ i → transportRefl ((ΩKn+1→Kn (suc (suc n)) (transp (λ j → 0ₖ (suc (suc (suc n))) ≡ ∣ merid north (~ j ∧ ~ i) ∣) i - (λ z → ∣ compPath-filler (merid (transportRefl (transportRefl x i) i)) (sym (merid north)) i z + (λ z → ∣ compPath-filler + (merid (transportRefl (transportRefl x i) i)) + (sym (merid north)) i z ∣)))) i) ∙ Iso.leftInv (Iso-Kn-ΩKn+1 (suc (suc n))) ∣ x ∣) -Iso-πS-ℤPres1← : (n : ℕ) → Iso.inv (fst (πS≅ℤ (suc n))) (pos 1) ≡ (∣_∣ , refl) -Iso-πS-ℤPres1← n = sym (cong (Iso.inv (fst (πS≅ℤ (suc n)))) (Iso-πS-ℤPres1 n)) ∙ Iso.leftInv (fst (πS≅ℤ (suc n))) (∣_∣ , refl) +Iso-πS-ℤPres1← : (n : ℕ) + → Iso.inv (fst (πS≅ℤ (suc n))) (pos 1) ≡ (∣_∣ , refl) +Iso-πS-ℤPres1← n = sym (cong (Iso.inv (fst (πS≅ℤ (suc n)))) (Iso-πS-ℤPres1 n)) + ∙ Iso.leftInv (fst (πS≅ℤ (suc n))) (∣_∣ , refl) IsoFunSpace : (n : ℕ) → Iso (S₊∙ n →∙ coHomK-ptd n) ℤ IsoFunSpace n = fst (πS≅ℤ n) @@ -293,7 +245,8 @@ module g-base where fst (G n i x) y = (genFunSpace n) .fst y ⌣ₖ x snd (G n i x) = cong (_⌣ₖ x) ((genFunSpace n) .snd) ∙ 0ₖ-⌣ₖ n i x - -ₖ^-Iso : (n : ℕ) (i : ℕ) → (S₊∙ n →∙ coHomK-ptd (i +' n)) ≃ (S₊∙ n →∙ coHomK-ptd (i +' n)) + -ₖ^-Iso : (n : ℕ) (i : ℕ) + → (S₊∙ n →∙ coHomK-ptd (i +' n)) ≃ (S₊∙ n →∙ coHomK-ptd (i +' n)) -ₖ^-Iso n i = isoToEquiv (iso F F FF FF) where lem : (i n : ℕ) → (-ₖ^ i · n) (snd (coHomK-ptd (i +' n))) ≡ 0ₖ _ @@ -313,30 +266,38 @@ module g-base where →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ y → -ₖ-gen² i n _ _ (fst x y)) - rCancel'' : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) → sym p ∙∙ refl ∙∙ p ≡ refl - rCancel'' p = (λ j → (λ i → p (~ i ∨ j)) ∙∙ refl ∙∙ λ i → p (i ∨ j)) ∙ sym (rUnit refl) + rCancel'' : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) + → sym p ∙∙ refl ∙∙ p ≡ refl + rCancel'' p = (λ j → (λ i → p (~ i ∨ j)) ∙∙ refl ∙∙ λ i → p (i ∨ j)) + ∙ sym (rUnit refl) transpPres0ₖ : ∀ {k m : ℕ} (p : k ≡ m) → subst coHomK p (0ₖ k) ≡ 0ₖ m - transpPres0ₖ {k = k} = J (λ m p → subst coHomK p (0ₖ k) ≡ 0ₖ m) (transportRefl _) + transpPres0ₖ {k = k} = + J (λ m p → subst coHomK p (0ₖ k) ≡ 0ₖ m) (transportRefl _) -- There will be some index swapping going on. We statet this explicitly, since we will -- need to trace the maps later - indexSwap : (n : ℕ) (i : ℕ) → (S₊∙ n →∙ coHomK-ptd (n +' i)) ≃ (S₊∙ n →∙ coHomK-ptd (i +' n)) + indexSwap : (n : ℕ) (i : ℕ) + → (S₊∙ n →∙ coHomK-ptd (n +' i)) ≃ (S₊∙ n →∙ coHomK-ptd (i +' n)) indexSwap n i = isoToEquiv (iso (λ f → (λ x → subst coHomK (+'-comm n i) (fst f x)) , - cong (subst coHomK (+'-comm n i)) (snd f) ∙ transpPres0ₖ (+'-comm n i)) - (λ f → (λ x → subst coHomK (sym (+'-comm n i)) (fst f x)) - , (cong (subst coHomK (sym (+'-comm n i))) (snd f) ∙ transpPres0ₖ (sym (+'-comm n i)))) - (λ f → →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ x → transportTransport⁻ _ (f .fst x))) - λ f → →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ x → transportTransport⁻ _ (f .fst x))) + cong (subst coHomK (+'-comm n i)) (snd f) ∙ transpPres0ₖ (+'-comm n i)) + (λ f → (λ x → subst coHomK (sym (+'-comm n i)) (fst f x)) + , (cong (subst coHomK (sym (+'-comm n i))) (snd f) + ∙ transpPres0ₖ (sym (+'-comm n i)))) + (λ f → →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ x → transportTransport⁻ _ (f .fst x))) + λ f → →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ x → transportTransport⁻ _ (f .fst x))) -- g is a composition of G and our two previous equivs. - g≡ : (n : ℕ) (i : ℕ) → g n i ≡ λ x → fst (compEquiv (indexSwap n i) (-ₖ^-Iso n i)) ((G n i) x) + g≡ : (n : ℕ) (i : ℕ) → g n i ≡ λ x + → fst (compEquiv (indexSwap n i) (-ₖ^-Iso n i)) ((G n i) x) g≡ n i = funExt (λ f → →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ y → gradedComm-⌣ₖ _ _ f (genFunSpace n .fst y))) -- We need a third Iso. - + suspKn-Iso-fun : (n m : ℕ) → (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) → (S₊ n → coHomK m) @@ -348,11 +309,13 @@ module g-base where suspKn-Iso-fun∙ : (n m : ℕ) → (f : _) → suspKn-Iso-fun n m f (ptSn _) ≡ 0ₖ m suspKn-Iso-fun∙ zero m = λ _ → refl suspKn-Iso-fun∙ (suc n) m (f , p) = - cong (ΩKn+1→Kn m) (cong (sym p ∙∙_∙∙ p) (cong (cong f) (rCancel (merid (ptSn _))))) - ∙∙ cong (ΩKn+1→Kn m) (rCancel'' p) + cong (ΩKn+1→Kn m) + (cong (sym p ∙∙_∙∙ p) (cong (cong f) (rCancel (merid (ptSn _))))) + ∙∙ cong (ΩKn+1→Kn m) (rCancel'' p) ∙∙ ΩKn+1→Kn-refl m - suspKn-Iso-inv : (n m : ℕ) → (S₊∙ n →∙ coHomK-ptd m) → (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) + suspKn-Iso-inv : (n m : ℕ) → (S₊∙ n →∙ coHomK-ptd m) + → (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) fst (suspKn-Iso-inv zero m (f , p)) base = 0ₖ _ fst (suspKn-Iso-inv zero m (f , p)) (loop i) = Kn→ΩKn+1 m (f false) i snd (suspKn-Iso-inv zero m (f , p)) = refl @@ -364,38 +327,55 @@ module g-base where suspKn-Iso : (n m : ℕ) → Iso (S₊∙ (suc n) →∙ coHomK-ptd (suc m)) (S₊∙ n →∙ coHomK-ptd m) - Iso.fun (suspKn-Iso n m) f = (suspKn-Iso-fun n m f) , (suspKn-Iso-fun∙ n m f) + Iso.fun (suspKn-Iso n m) f = (suspKn-Iso-fun n m f) + , (suspKn-Iso-fun∙ n m f) Iso.inv (suspKn-Iso n m) = suspKn-Iso-inv n m Iso.rightInv (suspKn-Iso zero m) (f , p) = →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ { false → cong (ΩKn+1→Kn m) (sym (rUnit _)) ∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f false) + (funExt λ { false → cong (ΩKn+1→Kn m) (sym (rUnit _)) + ∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f false) ; true → sym p}) Iso.rightInv (suspKn-Iso (suc n) m) (f , p) = →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ x → (λ i → ΩKn+1→Kn m (sym (rUnit (cong-∙ (suspKn-Iso-inv (suc n) m (f , p) .fst) (merid x) (sym (merid (ptSn _))) i)) i)) - ∙∙ cong (ΩKn+1→Kn m) (cong (Kn→ΩKn+1 m (f x) ∙_) (cong sym (cong (Kn→ΩKn+1 m) p ∙ Kn→ΩKn+10ₖ m)) ∙ sym (rUnit _)) + (funExt λ x → + (λ i → ΩKn+1→Kn m + (sym (rUnit + (cong-∙ + (suspKn-Iso-inv (suc n) m (f , p) .fst) + (merid x) (sym (merid (ptSn _))) i)) i)) + ∙∙ cong (ΩKn+1→Kn m) + (cong (Kn→ΩKn+1 m (f x) ∙_) + (cong sym (cong (Kn→ΩKn+1 m) p ∙ Kn→ΩKn+10ₖ m)) + ∙ sym (rUnit _)) ∙∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f x)) Iso.leftInv (suspKn-Iso zero m) (f , p) = →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ { base → sym p - ; (loop i) j → h3 j i}) + ; (loop i) j → lem j i}) where - h3 : PathP (λ i → p (~ i) ≡ p (~ i)) (Kn→ΩKn+1 m (ΩKn+1→Kn m (sym p ∙∙ cong f loop ∙∙ p))) (cong f loop) - h3 = Iso.rightInv (Iso-Kn-ΩKn+1 _) _ + lem : PathP (λ i → p (~ i) ≡ p (~ i)) + (Kn→ΩKn+1 m (ΩKn+1→Kn m (sym p ∙∙ cong f loop ∙∙ p))) + (cong f loop) + lem = Iso.rightInv (Iso-Kn-ΩKn+1 _) _ ◁ λ i → doubleCompPath-filler (sym p) (cong f loop) p (~ i) Iso.leftInv (suspKn-Iso (suc n) m) (f , p) = →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ { north → sym p ; south → sym p ∙ cong f (merid (ptSn _)) - ; (merid a i) j → h3 a j i}) + ; (merid a i) j → lem a j i}) where - h3 : (a : S₊ (suc n)) → PathP (λ i → p (~ i) ≡ (sym p ∙ cong f (merid (ptSn (suc n)))) i) - (Kn→ΩKn+1 m (ΩKn+1→Kn m (sym p ∙∙ cong f (merid a ∙ sym (merid (ptSn _))) ∙∙ p))) - (cong f (merid a)) - h3 a = Iso.rightInv (Iso-Kn-ΩKn+1 _) _ - ◁ λ i j → hcomp (λ k → λ { (i = i1) → (f (merid a j)) - ; (j = i0) → p (k ∧ ~ i) - ; (j = i1) → compPath-filler' (sym p) (cong f (merid (ptSn _))) k i }) - (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) + lem : (a : S₊ (suc n)) + → PathP (λ i → p (~ i) ≡ (sym p ∙ cong f (merid (ptSn (suc n)))) i) + (Kn→ΩKn+1 m + (ΩKn+1→Kn m + (sym p ∙∙ cong f (merid a ∙ sym (merid (ptSn _))) ∙∙ p))) + (cong f (merid a)) + lem a = Iso.rightInv (Iso-Kn-ΩKn+1 _) _ + ◁ λ i j → hcomp (λ k → + λ { (i = i1) → (f (merid a j)) + ; (j = i0) → p (k ∧ ~ i) + ; (j = i1) → compPath-filler' + (sym p) (cong f (merid (ptSn _))) k i }) + (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) glIsoInvHom : (n m : ℕ) (x y : coHomK n) (z : S₊ (suc m)) → Iso.inv (suspKn-Iso _ _) (G m n (x +ₖ y)) .fst z @@ -403,41 +383,46 @@ module g-base where +ₖ Iso.inv (suspKn-Iso _ _) (G m n y) .fst z glIsoInvHom zero zero x y base = refl glIsoInvHom (suc n) zero x y base = refl - glIsoInvHom zero zero x y (loop i) j = h3 j i + glIsoInvHom zero zero x y (loop i) j = lem j i where - h3 : (cong (Iso.inv (suspKn-Iso _ _) (G zero zero (x + y)) .fst) loop) - ≡ cong₂ _+ₖ_ (cong (Iso.inv (suspKn-Iso _ _) (G zero zero x) .fst) loop) - (cong (Iso.inv (suspKn-Iso _ _) (G zero zero y) .fst) loop) - h3 = Kn→ΩKn+1-hom 0 x y + lem : (cong (Iso.inv (suspKn-Iso _ _) (G zero zero (x + y)) .fst) loop) + ≡ cong₂ _+ₖ_ (cong (Iso.inv (suspKn-Iso _ _) (G zero zero x) .fst) loop) + (cong (Iso.inv (suspKn-Iso _ _) (G zero zero y) .fst) loop) + lem = Kn→ΩKn+1-hom 0 x y ∙ ∙≡+₁ (cong (Iso.inv (suspKn-Iso _ _) (G zero zero x) .fst) loop) (cong (Iso.inv (suspKn-Iso _ _) (G zero zero y) .fst) loop) - glIsoInvHom (suc n) zero x y (loop i) j = h3 j i + glIsoInvHom (suc n) zero x y (loop i) j = lem j i where - h3 : Kn→ΩKn+1 (suc n) ((pos (suc zero)) ·₀ (x +ₖ y)) - ≡ cong₂ _+ₖ_ (cong (Iso.inv (suspKn-Iso zero (zero +' suc n)) (G zero (suc n) x) .fst) loop) - (cong (Iso.inv (suspKn-Iso zero (zero +' suc n)) (G zero (suc n) y) .fst) loop) - h3 = cong (Kn→ΩKn+1 (suc n)) (lUnit⌣ₖ _ (x +ₖ y)) + lem : Kn→ΩKn+1 (suc n) ((pos (suc zero)) ·₀ (x +ₖ y)) + ≡ cong₂ _+ₖ_ (cong (Iso.inv (suspKn-Iso zero (zero +' suc n)) + (G zero (suc n) x) .fst) loop) + (cong (Iso.inv (suspKn-Iso zero (zero +' suc n)) + (G zero (suc n) y) .fst) loop) + lem = cong (Kn→ΩKn+1 (suc n)) (lUnit⌣ₖ _ (x +ₖ y)) ∙∙ Kn→ΩKn+1-hom (suc n) x y ∙∙ (λ i → ∙≡+₂ n (Kn→ΩKn+1 (suc n) (lUnit⌣ₖ _ x (~ i))) (Kn→ΩKn+1 (suc n) (lUnit⌣ₖ _ y (~ i))) i) glIsoInvHom zero (suc m) x y north = refl glIsoInvHom zero (suc m) x y south = refl - glIsoInvHom zero (suc m) x y (merid a i) j = h3 j i + glIsoInvHom zero (suc m) x y (merid a i) j = lem j i where - h3 : Kn→ΩKn+1 (suc m) (_⌣ₖ_ {n = suc m} {m = zero} ∣ a ∣ (x + y)) + lem : Kn→ΩKn+1 (suc m) (_⌣ₖ_ {n = suc m} {m = zero} ∣ a ∣ (x + y)) ≡ cong₂ _+ₖ_ (Kn→ΩKn+1 (suc m) (_⌣ₖ_ {n = suc m} {m = zero} ∣ a ∣ x)) (Kn→ΩKn+1 (suc m) (_⌣ₖ_ {n = suc m} {m = zero} ∣ a ∣ y)) - h3 = cong (Kn→ΩKn+1 (suc m)) (leftDistr-⌣ₖ (suc m) 0 ∣ a ∣ x y) + lem = cong (Kn→ΩKn+1 (suc m)) (leftDistr-⌣ₖ (suc m) 0 ∣ a ∣ x y) ∙∙ Kn→ΩKn+1-hom (suc m) _ _ ∙∙ ∙≡+₂ _ _ _ glIsoInvHom (suc n) (suc m) x y north = refl glIsoInvHom (suc n) (suc m) x y south = refl - glIsoInvHom (suc n) (suc m) x y (merid a i) j = h3 j i + glIsoInvHom (suc n) (suc m) x y (merid a i) j = lem j i where - h3 : Kn→ΩKn+1 (suc (suc (m +ℕ n))) (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ (x +ₖ y)) - ≡ cong₂ _+ₖ_ (Kn→ΩKn+1 (suc (suc (m +ℕ n))) (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ x)) - (Kn→ΩKn+1 (suc (suc (m +ℕ n))) (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ y)) - h3 = cong (Kn→ΩKn+1 (suc (suc (m +ℕ n)))) + lem : Kn→ΩKn+1 (suc (suc (m +ℕ n))) + (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ (x +ₖ y)) + ≡ cong₂ _+ₖ_ (Kn→ΩKn+1 (suc (suc (m +ℕ n))) + (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ x)) + (Kn→ΩKn+1 (suc (suc (m +ℕ n))) + (_⌣ₖ_ {n = suc m} {m = suc n} ∣ a ∣ y)) + lem = cong (Kn→ΩKn+1 (suc (suc (m +ℕ n)))) (leftDistr-⌣ₖ (suc m) (suc n) ∣ a ∣ x y) ∙∙ Kn→ΩKn+1-hom _ _ _ ∙∙ ∙≡+₂ _ _ _ @@ -458,64 +443,83 @@ module g-base where →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ z → (λ i → x ·₀ ∣ z ∣) ∙ h3 x z ∙ sym (transportRefl _)) where - h3 : (x : ℤ) (z : S¹) → _·₀_ {n = 1} x ∣ z ∣ ≡ fst (Iso.inv (suspKn-Iso 0 zero) (G zero zero x)) z - h3 = Int-ind _ (λ { base → refl ; (loop i) j → Kn→ΩKn+10ₖ zero (~ j) i}) - (λ { base → refl ; (loop i) j → rUnit (cong ∣_∣ₕ (lUnit loop j)) j i}) - (λ { base → refl ; (loop i) j → rUnit (cong ∣_∣ₕ (sym loop)) j i}) - λ x y inx iny z - → rightDistr-⌣ₖ 0 1 x y ∣ z ∣ - ∙∙ cong₂ _+ₖ_ (inx z) (iny z) - ∙∙ sym (glIsoInvHom zero zero x y z) + h3 : (x : ℤ) (z : S¹) + → _·₀_ {n = 1} x ∣ z ∣ + ≡ fst (Iso.inv (suspKn-Iso 0 zero) (G zero zero x)) z + h3 = + Int-ind _ + (λ { base → refl ; (loop i) j → Kn→ΩKn+10ₖ zero (~ j) i}) + (λ { base → refl ; (loop i) j → rUnit (cong ∣_∣ₕ (lUnit loop j)) j i}) + (λ { base → refl ; (loop i) j → rUnit (cong ∣_∣ₕ (sym loop)) j i}) + λ x y inx iny z + → rightDistr-⌣ₖ 0 1 x y ∣ z ∣ + ∙∙ cong₂ _+ₖ_ (inx z) (iny z) + ∙∙ sym (glIsoInvHom zero zero x y z) decomposeG (suc i) zero x = →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ z → h3 z + (funExt λ z → lem z ∙ sym (transportRefl ((Iso.inv (suspKn-Iso zero (suc i)) (G zero (suc i) x)) .fst z))) where - h3 : (z : S₊ 1) → _ ≡ Iso.inv (suspKn-Iso zero (suc i)) (G zero (suc i) x) .fst z - h3 base = refl - h3 (loop k) j = Kn→ΩKn+1 (suc i) (lUnit⌣ₖ _ x (~ j)) k + lem : (z : S₊ 1) + → _ ≡ Iso.inv (suspKn-Iso zero (suc i)) (G zero (suc i) x) .fst z + lem base = refl + lem (loop k) j = Kn→ΩKn+1 (suc i) (lUnit⌣ₖ _ x (~ j)) k decomposeG zero (suc n) x = →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ z → h3 x z ∙ λ i → transportRefl (Iso.inv (suspKn-Iso (suc n) (suc n)) (G (suc n) zero x)) (~ i) .fst z) + (funExt λ z → main x z + ∙ λ i → transportRefl + (Iso.inv (suspKn-Iso (suc n) (suc n)) + (G (suc n) zero x)) (~ i) .fst z) where +merid : rUnitₖ (suc (suc n)) ∣ south ∣ ≡ cong ∣_∣ₕ (merid (ptSn _)) +merid = sym (lUnit _) ∙ cong (cong ∣_∣ₕ) - λ j i → transp (λ _ → S₊ (suc (suc n))) (i ∨ j) (merid (ptSn (suc n)) i) - - pp : (a : S₊ (suc n)) → PathP (λ i → ∣ merid a i ∣ₕ ≡ Kn→ΩKn+1 (suc n) (∣ a ∣ +ₖ (0ₖ _)) i) - refl (sym (rUnitₖ (suc (suc n)) ∣ south ∣)) - pp a = flipSquare ((λ i j → ∣ compPath-filler (merid a) (sym (merid (ptSn _))) i j ∣ₕ ) + λ j i → transp (λ _ → S₊ (suc (suc n))) (i ∨ j) + (merid (ptSn (suc n)) i) + + lem : (a : S₊ (suc n)) + → PathP (λ i → ∣ merid a i ∣ₕ + ≡ Kn→ΩKn+1 (suc n) (∣ a ∣ +ₖ (0ₖ _)) i) + refl (sym (rUnitₖ (suc (suc n)) ∣ south ∣)) + lem a = + flipSquare ((λ i j → ∣ compPath-filler (merid a) + (sym (merid (ptSn _))) i j ∣ₕ) ▷ cong (Kn→ΩKn+1 (suc n)) (sym (rUnitₖ (suc n) ∣ a ∣ₕ))) ▷ sym (cong sym +merid) - pp2 : (a : S₊ (suc n)) → (λ i → -ₖ ∣ merid a i ∣) + lem₂ : (a : S₊ (suc n)) → (λ i → -ₖ ∣ merid a i ∣) ≡ Kn→ΩKn+1 (suc n) (-ₖ ∣ a ∣) - pp2 a = cong (cong ∣_∣ₕ) (sym (symDistr (merid a) (sym (merid (ptSn (suc n)))))) + lem₂ a = + cong (cong ∣_∣ₕ) (sym (symDistr (merid a) + (sym (merid (ptSn (suc n)))))) ∙ sym (Kn→ΩKn+1-ₖ (suc n) ∣ a ∣) - h3 : (x : ℤ) (z : S₊ (suc (suc n))) + main : (x : ℤ) (z : S₊ (suc (suc n))) → _⌣ₖ_ {n = suc (suc n)} {m = 0} ∣ z ∣ x ≡ Iso.inv (suspKn-Iso (suc n) (suc n)) (G (suc n) zero x) .fst z - h3 = Int-ind _ - (λ { north → refl ; south → refl ; (merid a i) j → Kn→ΩKn+10ₖ (suc n) (~ j) i}) + main = Int-ind _ + (λ { north → refl ; south → refl + ; (merid a i) j → Kn→ΩKn+10ₖ (suc n) (~ j) i}) (λ { north → refl ; south → refl - ; (merid a i) j → hcomp (λ k → λ { (i = i0) → ∣ north ∣ - ; (i = i1) → rUnitₖ (suc (suc n)) ∣ south ∣ (~ k) - ; (j = i0) → rUnitₖ (suc (suc n)) ∣ merid a i ∣ (~ k) - ; (j = i1) → pp a i k}) - ∣ merid a i ∣ₕ}) + ; (merid a i) j → + hcomp (λ k → + λ { (i = i0) → ∣ north ∣ + ; (i = i1) → rUnitₖ (suc (suc n)) ∣ south ∣ (~ k) + ; (j = i0) → rUnitₖ (suc (suc n)) ∣ merid a i ∣ (~ k) + ; (j = i1) → lem a i k}) + ∣ merid a i ∣ₕ}) (λ { north → refl ; south → refl - ; (merid a i) j → pp2 a j i}) + ; (merid a i) j → lem₂ a j i}) λ x y indx indy z → leftDistr-⌣ₖ _ _ ∣ z ∣ x y ∙ cong₂ _+ₖ_ (indx z) (indy z) ∙ sym (glIsoInvHom _ _ _ _ _) decomposeG (suc i) (suc n) x = →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ z → lem z - ∙ λ j → transportRefl ((Iso.inv (suspKn-Iso (suc n) (suc (suc (n +ℕ i)))) + ∙ λ j → + transportRefl ((Iso.inv (suspKn-Iso (suc n) (suc (suc (n +ℕ i)))) (G (suc n) (suc i) x))) (~ j) .fst z) where lem : (z : S₊ (suc (suc n))) → _ @@ -547,8 +551,8 @@ module g-base where isEquiv-g : (n i : ℕ) → isEquiv (g n i) isEquiv-g n i = subst isEquiv (sym (g≡ n i)) - (compEquiv (G n i , isEquivG n i) (compEquiv (indexSwap n i) (-ₖ^-Iso n i)) .snd) - + (compEquiv (G n i , isEquivG n i) + (compEquiv (indexSwap n i) (-ₖ^-Iso n i)) .snd) -- We now generealise the equivalence g to also apply to arbitrary fibrations (Q : B → Type) -- satisfying (Q * ≃∙ Sⁿ) @@ -557,7 +561,8 @@ module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) (n : ℕ) (Q-is : Iso (typ (Q (pt B))) (S₊ n)) (Q-is-ptd : Iso.fun Q-is (pt (Q (pt B))) ≡ snd (S₊∙ n)) (c : (b : typ B) → (Q b →∙ coHomK-ptd n)) - (c-pt : c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) where + (c-pt : c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) + where g : (b : typ B) (i : ℕ) → coHomK i → (Q b →∙ coHomK-ptd (i +' n)) fst (g b i x) y = x ⌣ₖ c b .fst y @@ -565,12 +570,15 @@ module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) g-hom : (b : typ B) (i : ℕ) → (x y : coHomK i) → g b i (x +ₖ y) ≡ ((g b i x) ++ (g b i y)) - g-hom b i x y = →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ z → rightDistr-⌣ₖ i n x y (c b .fst z)) + g-hom b i x y = + →∙Homogeneous≡ (isHomogeneousKn _) + (funExt λ z → rightDistr-⌣ₖ i n x y (c b .fst z)) gPathP' : (i : ℕ) - → PathP (λ j → coHomK i → (isoToPath Q-is j , ua-gluePath (isoToEquiv Q-is) (Q-is-ptd) j) - →∙ coHomK-ptd (i +' n)) - (g (pt B) i) (g-base.g n i) + → PathP (λ j → coHomK i → + (isoToPath Q-is j , ua-gluePath (isoToEquiv Q-is) (Q-is-ptd) j) + →∙ coHomK-ptd (i +' n)) + (g (pt B) i) (g-base.g n i) gPathP' i = toPathP (funExt @@ -578,8 +586,9 @@ module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) (funExt λ y → (λ i → transportRefl (transportRefl x i ⌣ₖ c (pt B) .fst (Iso.inv Q-is (transportRefl y i))) i) - ∙ cong (x ⌣ₖ_) (funExt⁻ c-pt (Iso.inv Q-is y) - ∙ cong (genFunSpace n .fst) (Iso.rightInv Q-is y)))) + ∙ cong (x ⌣ₖ_) + (funExt⁻ c-pt (Iso.inv Q-is y) + ∙ cong (genFunSpace n .fst) (Iso.rightInv Q-is y)))) g-base : (i : ℕ) → isEquiv (g (pt B) i) g-base i = transport (λ j → isEquiv (gPathP' i (~ j))) (g-base.isEquiv-g n i) @@ -601,7 +610,8 @@ module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) sRec (isProp→isOfHLevelSuc 1 isPropIsSet) (J (λ b _ → isSet (Q b →∙ coHomK-ptd n)) (subst isSet (cong (_→∙ coHomK-ptd n) - (ua∙ (isoToEquiv (invIso Q-is)) (cong (Iso.inv Q-is) (sym Q-is-ptd) ∙ Iso.leftInv Q-is _))) + (ua∙ (isoToEquiv (invIso Q-is)) + (cong (Iso.inv Q-is) (sym Q-is-ptd) ∙ Iso.leftInv Q-is _))) (isOfHLevelRetractFromIso 2 (fst (πS≅ℤ n)) isSetℤ))) (conB (pt B) b) @@ -619,7 +629,7 @@ module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) -- We construct a term in c* : (b : B) → (Q b →∙ Kₙ) -- Which is equal to the generator of (Sⁿ →∙ Kₙ) over the base point. - c* : Σ[ c ∈ ((b : typ B) → (Q b →∙ coHomK-ptd n)) ] + c* : Σ[ c ∈ ((b : typ B) → (Q b →∙ coHomK-ptd n)) ] (c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) fst c* b = sRec (is-setQ→K b) @@ -631,20 +641,24 @@ module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) funExt λ x → (λ i → sRec (is-setQ→K (pt B)) (J (λ b _ → Q b →∙ coHomK-ptd n) ((λ x₁ → genFunSpace n .fst (Iso.fun Q-is x₁)) , - (λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) + (λ i → genFunSpace n .fst (Q-is-ptd i)) + ∙ genFunSpace n .snd)) (isPropPath (conB (pt B) (pt B)) ∣ refl ∣₂ i) .fst x) - ∙ (λ i → transportRefl (genFunSpace n .fst (Iso.fun Q-is (transportRefl x i))) i) + ∙ (λ i → transportRefl (genFunSpace n .fst + (Iso.fun Q-is (transportRefl x i))) i) p-help : {b : fst B} (p : pt B ≡ b) → (subst (fst ∘ Q) (sym p) (snd (Q b))) ≡ (snd (Q (pt B))) p-help {b = b} = - J (λ b p → subst (fst ∘ Q) (sym p) (snd (Q b)) ≡ snd (Q (pt B))) (transportRefl _) + J (λ b p → subst (fst ∘ Q) (sym p) (snd (Q b)) ≡ snd (Q (pt B))) + (transportRefl _) -- This form of c* will make things somewhat easier to work with later on. c≡ : (b : fst B) (p : ∥ pt B ≡ b ∥₂) → (c* .fst b) ≡ sRec (is-setQ→K b) - (λ pp → (λ qb → genFunSpace n .fst (Iso.fun Q-is (subst (fst ∘ Q) (sym pp) qb))) + (λ pp → (λ qb → + genFunSpace n .fst (Iso.fun Q-is (subst (fst ∘ Q) (sym pp) qb))) , cong (genFunSpace n .fst ∘ Iso.fun Q-is) (p-help pp) ∙ ((λ i → genFunSpace n .fst (Q-is-ptd i)) ∙ genFunSpace n .snd)) p c≡ b = @@ -663,7 +677,8 @@ module _ {ℓ} (B : Pointed ℓ) (Q : typ B → Pointed ℓ-zero) (isPropPath (conB (pt B) (pt B)) ∣ refl ∣₂ i)) ∙ →∙Homogeneous≡ (isHomogeneousKn n) (transportRefl ((λ x → genFunSpace n .fst (Iso.fun Q-is x))) - ∙ funExt λ x → cong (genFunSpace n .fst ∘ Iso.fun Q-is) (sym (transportRefl x))))) + ∙ funExt λ x → cong (genFunSpace n .fst ∘ Iso.fun Q-is) + (sym (transportRefl x))))) -- We are now almost ready to define the Thom isomorphism. -- The following module contains the types and functions occuring in it. @@ -708,11 +723,11 @@ module preThom {ℓ ℓ'} (B : Pointed ℓ) (P : typ B → Type ℓ') where Iso.inv IsoFE = invFE Iso.rightInv IsoFE (inl x) = refl Iso.rightInv IsoFE (inr x) = refl - Iso.rightInv IsoFE (push (x , a) i₁) k = h3 k i₁ + Iso.rightInv IsoFE (push (x , a) i₁) k = lem k i₁ where - h3 : cong funFE (((push x) ∙ λ i → inr (x , merid a i))) - ≡ push (x , a) - h3 = congFunct funFE (push x) (λ i → inr (x , merid a i)) + lem : cong funFE (((push x) ∙ λ i → inr (x , merid a i))) + ≡ push (x , a) + lem = congFunct funFE (push x) (λ i → inr (x , merid a i)) ∙ sym (lUnit (push (x , a))) Iso.leftInv IsoFE (inl x) = refl Iso.leftInv IsoFE (inr (x , north)) = push x @@ -721,7 +736,7 @@ module preThom {ℓ ℓ'} (B : Pointed ℓ) (P : typ B → Type ℓ') where compPath-filler' (push x) (λ i₁ → inr (x , merid a i₁)) (~ j) i Iso.leftInv IsoFE (push a i₁) k = push a (i₁ ∧ k) - + F̃→Q : ∀ {ℓ} {A : Pointed ℓ} → (F̃ , inl tt) →∙ A → (b : typ B) → Q b →∙ A fst (F̃→Q {A = A , a} (f , p) b) north = f (inr (b , north)) fst (F̃→Q {A = A , a} (f , p) b) south = f (inr (b , south)) @@ -740,15 +755,17 @@ module preThom {ℓ ℓ'} (B : Pointed ℓ) (P : typ B → Type ℓ') where → Q→F̃ (λ b → f b ++ g b) ≡ (Q→F̃ f ++ Q→F̃ g) Q→F̃-hom n f g = →∙Homogeneous≡ (isHomogeneousKn _) - (funExt λ { (inl x) → sym (rUnitₖ _ (0ₖ _)) - ; (inr (x , north)) → refl - ; (inr (x , south)) → refl - ; (inr (x , merid a i₁)) → refl - ; (push a j) i → compPath-filler (cong₂ _+ₖ_ (snd (f a)) (snd (g a))) - (rUnitₖ n (0ₖ n)) (~ i) (~ j)}) - - Q→F̃→Q : ∀ {ℓ} {A : Pointed ℓ} → (x : (b : typ B) → Q b →∙ A) (b : typ B) (q : Q b .fst) - → F̃→Q (Q→F̃ x) b .fst q ≡ x b .fst q + (funExt λ { (inl x) → sym (rUnitₖ _ (0ₖ _)) + ; (inr (x , north)) → refl + ; (inr (x , south)) → refl + ; (inr (x , merid a i₁)) → refl + ; (push a j) i → + compPath-filler (cong₂ _+ₖ_ (snd (f a)) (snd (g a))) + (rUnitₖ n (0ₖ n)) (~ i) (~ j)}) + + Q→F̃→Q : ∀ {ℓ} {A : Pointed ℓ} + → (x : (b : typ B) → Q b →∙ A) (b : typ B) (q : Q b .fst) + → F̃→Q (Q→F̃ x) b .fst q ≡ x b .fst q Q→F̃→Q f b north = refl Q→F̃→Q f b south = refl Q→F̃→Q f b (merid a i₁) = refl @@ -759,7 +776,8 @@ module preThom {ℓ ℓ'} (B : Pointed ℓ) (P : typ B → Type ℓ') where F̃→Q→F̃ f p (inr (x , north)) = refl F̃→Q→F̃ f p (inr (x , south)) = refl F̃→Q→F̃ f p (inr (x , merid a i₁)) = refl - F̃→Q→F̃ f p (push a i₁) j = compPath-filler (sym (cong f (push a))) p (~ j) (~ i₁) + F̃→Q→F̃ f p (push a i₁) j = + compPath-filler (sym (cong f (push a))) p (~ j) (~ i₁) IsoF̃Q : ∀ {ℓ} {A : Pointed ℓ} → Iso ((F̃ , inl tt) →∙ A) @@ -778,7 +796,7 @@ module preThom {ℓ ℓ'} (B : Pointed ℓ) (P : typ B → Type ℓ') where ((Ẽ , inl tt) →∙ coHomK-ptd k) ι k = compIso (invIso IsoF̃Q) IsoFE-extend where - + IsoFE-extend : Iso ((F̃ , inl tt) →∙ coHomK-ptd k) ((Ẽ , inl tt) →∙ coHomK-ptd k) Iso.fun IsoFE-extend G = (λ x → G .fst (Iso.inv IsoFE x)) @@ -800,37 +818,45 @@ module preThom {ℓ ℓ'} (B : Pointed ℓ) (P : typ B → Type ℓ') where (funExt λ x → funExt⁻ (cong fst (Q→F̃-hom _ f g)) (invFE x)) -- Packing everything up gives us the Thom Isomorphism between --- the nᵗʰ cohomology of B and the (n+i)ᵗʰ reduced cohomology of Ẽ, as defined above -module Thom {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) +-- the nᵗʰ cohomology of B and the (n+i)ᵗʰ reduced cohomology of Ẽ, +-- as defined above +module Thom {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) (conB : (x y : typ B) → ∥ x ≡ y ∥) (n : ℕ) (Q-is : Iso (typ (preThom.Q B P (pt B))) (S₊ n)) (Q-is-ptd : Iso.fun Q-is (pt (preThom.Q B P (pt B))) ≡ snd (S₊∙ n)) (c : (b : typ B) → (preThom.Q B P b →∙ coHomK-ptd n)) - (c-pt : c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) where + (c-pt : c (pt B) .fst ≡ ((λ x → genFunSpace n .fst (Iso.fun Q-is x)))) + where - ϕ : (i : ℕ) → GroupEquiv (coHomGr i (typ B)) (coHomRedGrDir (i +' n) (preThom.Ẽ B P , inl tt)) + ϕ : (i : ℕ) + → GroupEquiv (coHomGr i (typ B)) + (coHomRedGrDir (i +' n) (preThom.Ẽ B P , inl tt)) fst (ϕ i) = isoToEquiv (setTruncIso (compIso (codomainIsoDep - λ b → equivToIso ((g B (preThom.Q B P) conB n Q-is Q-is-ptd c c-pt b i) + λ b → + equivToIso ((g B (preThom.Q B P) conB n Q-is Q-is-ptd c c-pt b i) , g-equiv B (preThom.Q B P) conB n Q-is Q-is-ptd c c-pt b i)) (preThom.ι B P (i +' n)))) snd (ϕ i) = makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) - λ F G → cong ∣_∣₂ (cong (Iso.fun (preThom.ι B P (i +' n))) - (funExt (λ a → g-hom B (preThom.Q B P) - conB n Q-is Q-is-ptd c c-pt a i (F a) (G a))) - ∙ preThom.ι-hom B P (i +' n) _ _) - ∙ addAgree (i +' n) _ _) - --- We finally get the Gysin sequence -module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) + λ F G → + cong ∣_∣₂ (cong (Iso.fun (preThom.ι B P (i +' n))) + (funExt (λ a → + g-hom B (preThom.Q B P) + conB n Q-is Q-is-ptd c c-pt a i (F a) (G a))) + ∙ preThom.ι-hom B P (i +' n) _ _) + ∙ addAgree (i +' n) _ _) + +-- We finally get the Gysin sequence +module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) (conB : (x y : typ B) → ∥ x ≡ y ∥₂) (n : ℕ) (Q-is : Iso (typ (preThom.Q B P (pt B))) (S₊ n)) - (Q-is-ptd : Iso.fun Q-is (pt (preThom.Q B P (pt B))) ≡ snd (S₊∙ n)) where + (Q-is-ptd : Iso.fun Q-is (pt (preThom.Q B P (pt B))) ≡ snd (S₊∙ n)) + where 0-connB : (x y : typ B) → ∥ x ≡ y ∥ 0-connB x y = sRec (isProp→isSet squash) (∥_∥.∣_∣) (conB x y) @@ -852,8 +878,8 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) e = ∣ (λ b → c b .fst south) ∣₂ -- The maps of interest are ⌣, p, E-susp and j*. In reality, we are interested - -- in the composition of ϕ and j* (which is just the cup product), but it's easier - -- to first give an exact sequence involving p, E-susp and j* + -- in the composition of ϕ and j* (which is just the cup product), + -- but it's easier to first give an exact sequence involving p, E-susp and j* ⌣-hom : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr (i +' n) (typ B)) fst (⌣-hom i) x = x ⌣ e snd (⌣-hom i) = @@ -865,18 +891,21 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) p-hom : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr i E') p-hom i = coHomMorph _ p - E-susp : (i : ℕ) → GroupHom (coHomGr i E') (coHomRedGrDir (suc i) (E'̃ , inl tt)) - fst (E-susp i) = sMap λ f → (λ { (inl x) → 0ₖ _ - ; (inr x) → 0ₖ _ - ; (push a j) → Kn→ΩKn+1 _ (f a) j}) , refl + E-susp : (i : ℕ) → + GroupHom (coHomGr i E') (coHomRedGrDir (suc i) (E'̃ , inl tt)) + fst (E-susp i) = + sMap λ f → (λ { (inl x) → 0ₖ _ + ; (inr x) → 0ₖ _ + ; (push a j) → Kn→ΩKn+1 _ (f a) j}) , refl snd (E-susp zero) = makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ f g → cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn 1) (funExt λ { (inl x) → refl ; (inr x) → refl - ; (push a j) k → (Kn→ΩKn+1-hom zero (f a) (g a) - ∙ ∙≡+₁ (Kn→ΩKn+1 _ (f a)) (Kn→ΩKn+1 _ (g a))) k j}))) + ; (push a j) k → + (Kn→ΩKn+1-hom zero (f a) (g a) + ∙ ∙≡+₁ (Kn→ΩKn+1 _ (f a)) (Kn→ΩKn+1 _ (g a))) k j}))) snd (E-susp (suc i)) = makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ f g → @@ -884,10 +913,12 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a j) k → (Kn→ΩKn+1-hom _ (f a) (g a) - ∙ ∙≡+₂ _ (Kn→ΩKn+1 _ (f a)) (Kn→ΩKn+1 _ (g a))) k j}))) + ∙ ∙≡+₂ _ (Kn→ΩKn+1 _ (f a)) + (Kn→ΩKn+1 _ (g a))) k j}))) module cofibSeq where - j* : (i : ℕ) → GroupHom (coHomRedGrDir i (E'̃ , (inl tt))) (coHomGr i (typ B)) + j* : (i : ℕ) → + GroupHom (coHomRedGrDir i (E'̃ , (inl tt))) (coHomGr i (typ B)) fst (j* i) = sMap λ f → λ x → fst f (inr x) snd (j* zero) = makeIsGroupHom @@ -905,7 +936,8 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) λ f → pRec (squash₂ _ _) (uncurry (sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) λ g P → subst (isInKer (p-hom i)) P - (cong ∣_∣₂ (funExt λ x → cong (g .fst) (sym (push x)) ∙ g .snd)))) + (cong ∣_∣₂ (funExt λ x → + cong (g .fst) (sym (push x)) ∙ g .snd)))) Ker-p⊂Im-j : (i : ℕ) (x : _) → isInKer (p-hom i) x → isInIm (j* i) x Ker-p⊂Im-j i = @@ -913,10 +945,12 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) λ f ker → pRec squash (λ id → ∣ ∣ (λ { (inl x) → 0ₖ _ ; (inr x) → f x - ; (push a i₁) → funExt⁻ id a (~ i₁)}) , refl ∣₂ , refl ∣) + ; (push a i₁) → funExt⁻ id a (~ i₁)}) , refl ∣₂ + , refl ∣) (Iso.fun PathIdTrunc₀Iso ker) - Im-p⊂Ker-Susp : (i : ℕ) (x : _) → isInIm (p-hom i) x → isInKer (E-susp i) x + Im-p⊂Ker-Susp : (i : ℕ) (x : _) + → isInIm (p-hom i) x → isInKer (E-susp i) x Im-p⊂Ker-Susp i = sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) λ f → pRec (squash₂ _ _) @@ -927,15 +961,19 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) ; (inr x) → sym (Kn→ΩKn+1 _ (g x)) ; (push a j) k → Kn→ΩKn+1 i (g (fst a)) (~ k ∧ j)}))))) - Ker-Susp⊂Im-p : (i : ℕ) (x : _) → isInKer (E-susp i) x → isInIm (p-hom i) x + Ker-Susp⊂Im-p : (i : ℕ) (x : _) + → isInKer (E-susp i) x → isInIm (p-hom i) x Ker-Susp⊂Im-p i = sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) λ f ker → pRec squash (λ id → ∣ ∣ (λ x → ΩKn+1→Kn i (sym (funExt⁻ (cong fst id) (inr x)))) ∣₂ , cong ∣_∣₂ (funExt (λ { (a , b) → - cong (ΩKn+1→Kn i) (lUnit _ ∙ cong (_∙ sym (funExt⁻ (λ i₁ → fst (id i₁)) (inr a))) - (sym (flipSquare (cong snd id)) - ∙∙ (PathP→compPathR (cong (funExt⁻ (cong fst id)) (push (a , b)))) + cong (ΩKn+1→Kn i) + (lUnit _ + ∙ cong (_∙ sym (funExt⁻ (λ i₁ → fst (id i₁)) (inr a))) + (sym (flipSquare (cong snd id)) + ∙∙ (PathP→compPathR + (cong (funExt⁻ (cong fst id)) (push (a , b)))) ∙∙ assoc _ _ _ ∙ sym (rUnit _)) ∙ (sym (assoc _ _ _) @@ -944,7 +982,8 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) ∙ Iso.leftInv (Iso-Kn-ΩKn+1 _) (f (a , b))})) ∣) (Iso.fun PathIdTrunc₀Iso ker) - Im-Susp⊂Ker-j : (i : ℕ) (x : _) → isInIm (E-susp i) x → isInKer (cofibSeq.j* (suc i)) x + Im-Susp⊂Ker-j : (i : ℕ) (x : _) + → isInIm (E-susp i) x → isInKer (cofibSeq.j* (suc i)) x Im-Susp⊂Ker-j i = sElim (λ _ → isSetΠ (λ _ → isOfHLevelPath 2 squash₂ _ _)) λ g → pRec (squash₂ _ _) @@ -954,15 +993,18 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) (cong ∣_∣₂ refl)) (Iso.fun PathIdTrunc₀Iso id))) - Ker-j⊂Im-Susp : (i : ℕ) (x : _) → isInKer (cofibSeq.j* (suc i)) x → isInIm (E-susp i) x + Ker-j⊂Im-Susp : (i : ℕ) (x : _) + → isInKer (cofibSeq.j* (suc i)) x → isInIm (E-susp i) x Ker-j⊂Im-Susp i = sElim (λ _ → isSetΠ (λ _ → isProp→isSet squash)) λ f ker → pRec squash - (λ p → ∣ ∣ (λ x → ΩKn+1→Kn _ (sym (snd f) ∙∙ cong (fst f) (push x) ∙∙ funExt⁻ p (fst x))) ∣₂ + (λ p → ∣ ∣ (λ x → ΩKn+1→Kn _ (sym (snd f) + ∙∙ cong (fst f) (push x) + ∙∙ funExt⁻ p (fst x))) ∣₂ , cong ∣_∣₂ (→∙Homogeneous≡ (isHomogeneousKn _) (funExt (λ { (inl x) → sym (snd f) - ; (inr x) → sym (funExt⁻ p x) + ; (inr x) → sym (funExt⁻ p x) ; (push a j) k → h3 f p a k j}))) ∣) (Iso.fun PathIdTrunc₀Iso ker) where @@ -970,13 +1012,19 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) → (p : (fst f ∘ inr) ≡ (λ _ → 0ₖ (suc i))) → (a : preThom.E B P) → PathP (λ i → snd f (~ i) ≡ p (~ i) (fst a)) - (Kn→ΩKn+1 i (ΩKn+1→Kn i (sym (snd f) ∙∙ cong (fst f) (push a) ∙∙ funExt⁻ p (fst a)))) + (Kn→ΩKn+1 i (ΩKn+1→Kn i (sym (snd f) + ∙∙ cong (fst f) (push a) + ∙∙ funExt⁻ p (fst a)))) (cong (fst f) (push a)) - h3 f p a = Iso.rightInv (Iso-Kn-ΩKn+1 i) _ - ◁ λ i j → doubleCompPath-filler (sym (snd f)) (cong (fst f) (push a)) (funExt⁻ p (fst a)) (~ i) j - - -- We compose ϕ and j*. The above exact sequence will induce another one for this - -- composition. + h3 f p a = + Iso.rightInv (Iso-Kn-ΩKn+1 i) _ + ◁ λ i j → + doubleCompPath-filler (sym (snd f)) (cong (fst f) (push a)) + (funExt⁻ p (fst a)) (~ i) j + + -- We compose ϕ and j*. The above exact sequence will induce another one for + -- this composition. In fact, this group hom is definitionally equal to + -- (λ x → x ⌣ e) modulo truncation elimination and snd fields. ϕ∘j : (i : ℕ) → GroupHom (coHomGr i (typ B)) (coHomGr (i +' n) (typ B)) ϕ∘j i = compGroupHom (fst (fst (ϕ i)) , snd (ϕ i)) (cofibSeq.j* (i +' n)) @@ -987,7 +1035,8 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) +'-suc zero (suc n) = refl +'-suc (suc i₁) (suc n) = refl - Path→GroupPath : ∀ {ℓ ℓ'} {n m : ℕ} (G : ℕ → Group ℓ) (H : Group ℓ') (p : n ≡ m) + Path→GroupPath : ∀ {ℓ ℓ'} {n m : ℕ} (G : ℕ → Group ℓ) (H : Group ℓ') + (p : n ≡ m) → GroupEquiv (G n) H → GroupEquiv (G m) H Path→GroupPath {n = n} G H = @@ -997,18 +1046,24 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) h-ret : ∀ {ℓ ℓ'} {n m : ℕ} (G : ℕ → Group ℓ) (H : Group ℓ') → (e : GroupEquiv (G n) H) → (p : n ≡ m) - → (x : G m .fst) → invEq (fst e) (fst (fst (Path→GroupPath G H p e)) x) ≡ subst (λ x → G x .fst) (sym p) x + → (x : G m .fst) + → invEq (fst e) (fst (fst (Path→GroupPath G H p e)) x) + ≡ subst (λ x → G x .fst) (sym p) x h-ret G H e = - J (λ m p → ((x : G m .fst) → invEq (fst e) (fst (fst (Path→GroupPath G H p e)) x) ≡ subst (λ x → G x .fst) (sym p) x)) + J (λ m p → ((x : G m .fst) + → invEq (fst e) (fst (fst (Path→GroupPath G H p e)) x) + ≡ subst (λ x → G x .fst) (sym p) x)) λ x → cong (invEq (fst e) ) - (λ i → transportRefl (transportRefl (fst (fst e) (transportRefl (transportRefl x i) i)) i) i) + (λ i → transportRefl (transportRefl (fst (fst e) + (transportRefl (transportRefl x i) i)) i) i) ∙∙ retEq (fst e) x ∙∙ sym (transportRefl _) thomIso' : (i : ℕ) → GroupEquiv (coHomRedGrDir (suc (i +' n)) (E'̃ , inl tt)) (coHomGr (suc i) (typ B)) - thomIso' i = (Path→GroupPath (λ n → coHomRedGrDir n (E'̃ , inl tt)) _ (+'-suc i n) - (invGroupEquiv (ϕ (suc i)))) + thomIso' i = (Path→GroupPath + (λ n → coHomRedGrDir n (E'̃ , inl tt)) _ (+'-suc i n) + (invGroupEquiv (ϕ (suc i)))) ϕ' : (i : ℕ) → GroupHom (coHomRedGrDir (suc (i +' n)) (E'̃ , inl tt)) (coHomGr (suc i) (typ B)) @@ -1024,21 +1079,27 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) Ker-p⊂Im-ϕ∘j : (i : ℕ) (x : _) → isInKer (p-hom _) x → isInIm (ϕ∘j i) x Ker-p⊂Im-ϕ∘j i x p = - pRec squash (uncurry (λ f p → ∣ (invEq (fst (ϕ _)) f) - , (cong (fst (cofibSeq.j* _)) (secEq (fst (ϕ _)) f) ∙ p) ∣)) - (cofibSeq.Ker-p⊂Im-j _ x p) + pRec squash + (uncurry (λ f p → + ∣ (invEq (fst (ϕ _)) f) + , (cong (fst (cofibSeq.j* _)) (secEq (fst (ϕ _)) f) ∙ p) ∣)) + (cofibSeq.Ker-p⊂Im-j _ x p) - Im-p⊂KerSusp∘ϕ : (i : ℕ) (x : _) → isInIm (p-hom _) x → isInKer (susp∘ϕ i) x - Im-p⊂KerSusp∘ϕ i x p = cong (fst (ϕ' _)) (Im-p⊂Ker-Susp _ x p) ∙ IsGroupHom.pres1 (snd (ϕ' _)) + Im-p⊂KerSusp∘ϕ : (i : ℕ) (x : _) + → isInIm (p-hom _) x → isInKer (susp∘ϕ i) x + Im-p⊂KerSusp∘ϕ i x p = + cong (fst (ϕ' _)) (Im-p⊂Ker-Susp _ x p) ∙ IsGroupHom.pres1 (snd (ϕ' _)) - KerSusp∘ϕ⊂Im-p : (i : ℕ) (x : _) → isInKer (susp∘ϕ i) x → isInIm (p-hom _) x + KerSusp∘ϕ⊂Im-p : (i : ℕ) (x : _) + → isInKer (susp∘ϕ i) x → isInIm (p-hom _) x KerSusp∘ϕ⊂Im-p i x p = Ker-Susp⊂Im-p _ x (sym (retEq (fst (thomIso' _)) _) - ∙ (cong (invEq (fst (thomIso' _))) p - ∙ IsGroupHom.pres1 (snd (invGroupEquiv (thomIso' _))))) + ∙ (cong (invEq (fst (thomIso' _))) p + ∙ IsGroupHom.pres1 (snd (invGroupEquiv (thomIso' _))))) - Im-Susp∘ϕ⊂Ker-ϕ∘j : (i : ℕ) → (x : _) → isInIm (susp∘ϕ i) x → isInKer (ϕ∘j (suc i)) x + Im-Susp∘ϕ⊂Ker-ϕ∘j : (i : ℕ) → (x : _) + → isInIm (susp∘ϕ i) x → isInKer (ϕ∘j (suc i)) x Im-Susp∘ϕ⊂Ker-ϕ∘j i x = pRec (squash₂ _ _) (uncurry λ f → J (λ x p → isInKer (ϕ∘j (suc i)) x) @@ -1058,35 +1119,58 @@ module Gysin {ℓ} (B : Pointed ℓ) (P : typ B → Type ℓ-zero) tLem (suc i₁) zero = refl tLem (suc i₁) (suc n) = refl - Ker-ϕ∘j⊂Im-Susp∘ϕ : (i : ℕ) (x : _) → isInKer (ϕ∘j (suc i)) x → isInIm (susp∘ϕ i) x Ker-ϕ∘j⊂Im-Susp∘ϕ i x p = pRec squash - (uncurry (λ f p → ∣ f , cong (fst (fst (thomIso' i))) p ∙ secEq (fst (thomIso' _)) x ∣)) + (uncurry (λ f p → ∣ f , cong (fst (fst (thomIso' i))) p + ∙ secEq (fst (thomIso' _)) x ∣)) (Ker-j⊂Im-Susp _ (invEq (fst (thomIso' _)) x) - ((cong (cofibSeq.j* (suc (i +' n)) .fst ) lem2 + ((cong (cofibSeq.j* (suc (i +' n)) .fst ) lem₁ ∙ natTranspLem _ (λ n → cofibSeq.j* n .fst) (+'-suc i n)) ∙∙ cong (transport (λ j → (coHomGr (+'-suc i n j) (typ B) .fst))) p - ∙∙ h2 i n)) + ∙∙ lem₂ i n)) where - lem2 : invEq (fst (thomIso' i)) x - ≡ transport (λ j → coHomRed (+'-suc i n j) (E'̃ , inl tt)) (fst (fst (ϕ _)) x) - lem2 = cong (transport (λ j → coHomRed (+'-suc i n j) (E'̃ , inl tt))) + lem₁ : invEq (fst (thomIso' i)) x + ≡ transport (λ j → coHomRed (+'-suc i n j) + (E'̃ , inl tt)) (fst (fst (ϕ _)) x) + lem₁ = cong (transport (λ j → coHomRed (+'-suc i n j) (E'̃ , inl tt))) (transportRefl _ ∙ cong (fst (fst (ϕ _))) λ i → transportRefl (transportRefl x i) i) - h2 : (i n : ℕ) → transport (λ j → coHomGr (+'-suc i n j) (typ B) .fst) - (GroupStr.1g (coHomGr (suc i +' n) (typ B) .snd)) ≡ 0ₕ _ - h2 zero zero = refl - h2 zero (suc n) = refl - h2 (suc i₁) zero = refl - h2 (suc i₁) (suc n) = refl + lem₂ : (i n : ℕ) + → transport (λ j → coHomGr (+'-suc i n j) (typ B) .fst) + (GroupStr.1g (coHomGr (suc i +' n) (typ B) .snd)) ≡ 0ₕ _ + lem₂ zero zero = refl + lem₂ zero (suc n) = refl + lem₂ (suc i₁) zero = refl + lem₂ (suc i₁) (suc n) = refl - -- Finally, we have that ϕ∘j is just the cup product, and we have arrived at an exact - -- sequence involving it. + -- Finally, we have that ϕ∘j is just the cup product, and we have arrived + -- at an exact sequence involving it. ϕ∘j≡ : (i : ℕ) → ϕ∘j i ≡ ⌣-hom i ϕ∘j≡ i = Σ≡Prop (λ _ → isPropIsGroupHom _ _) (funExt (sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl)) + + -- We can now restate the previous resluts for (λ x → x ⌣ e) + Im-⌣e⊂Ker-p : (i : ℕ) (x : _) + → isInIm (⌣-hom i) x → isInKer (p-hom _) x + Im-⌣e⊂Ker-p i x p = + Im-ϕ∘j⊂Ker-p i x (subst (λ p → isInIm p x) (sym (ϕ∘j≡ i)) p) + + Ker-p⊂Im-⌣e : (i : ℕ) (x : _) + → isInKer (p-hom _) x → isInIm (⌣-hom i) x + Ker-p⊂Im-⌣e i x p = + subst (λ p → isInIm p x) (ϕ∘j≡ i) (Ker-p⊂Im-ϕ∘j i x p) + + Im-Susp∘ϕ⊂Ker-⌣e : (i : ℕ) (x : _) + → isInIm (susp∘ϕ i) x → isInKer (⌣-hom (suc i)) x + Im-Susp∘ϕ⊂Ker-⌣e i x p = + subst (λ p → isInKer p x) (ϕ∘j≡ (suc i)) (Im-Susp∘ϕ⊂Ker-ϕ∘j i x p) + + Ker-⌣e⊂Im-Susp∘ϕ : (i : ℕ) (x : _) + → isInKer (⌣-hom (suc i)) x → isInIm (susp∘ϕ i) x + Ker-⌣e⊂Im-Susp∘ϕ i x p = + Ker-ϕ∘j⊂Im-Susp∘ϕ i x (subst (λ p → isInKer p x) (sym (ϕ∘j≡ (suc i))) p) diff --git a/Cubical/ZCohomology/Properties.agda b/Cubical/ZCohomology/Properties.agda index e2392008f..0d42c109f 100644 --- a/Cubical/ZCohomology/Properties.agda +++ b/Cubical/ZCohomology/Properties.agda @@ -644,11 +644,31 @@ isOfHLevel↑∙∙ n m (suc l) = (λ i → coHomK-ptd (suc n) →∙ →∙KnPath (coHomK-ptd (suc m)) (suc (suc (l + n + m))) (~ i)) (isOfHLevel↑∙∙ n m l) --- Misc. -isoType→isoCohom : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (n : ℕ) +----------- Misc. ------------ +isoType→isoCohom : {A : Type ℓ} {B : Type ℓ'} (n : ℕ) → Iso A B → GroupIso (coHomGr n A) (coHomGr n B) fst (isoType→isoCohom n is) = setTruncIso (domIso is) snd (isoType→isoCohom n is) = makeIsGroupHom (sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) (λ _ _ → refl)) + +-- Explicit index swapping for cohomology groups +transportCohomIso : {A : Type ℓ} {n m : ℕ} + → (p : n ≡ m) + → GroupIso (coHomGr n A) (coHomGr m A) +Iso.fun (fst (transportCohomIso {A = A} p)) = subst (λ n → coHom n A) p +Iso.inv (fst (transportCohomIso {A = A} p)) = subst (λ n → coHom n A) (sym p) +Iso.rightInv (fst (transportCohomIso p)) = transportTransport⁻ _ +Iso.leftInv (fst (transportCohomIso p)) = transportTransport⁻ _ +snd (transportCohomIso {A = A} {n = n} {m = m} p) = + makeIsGroupHom (λ x y → help x y p) + where + help : (x y : coHom n A) (p : n ≡ m) + → subst (λ n → coHom n A) p (x +ₕ y) + ≡ subst (λ n → coHom n A) p x +ₕ subst (λ n → coHom n A) p y + help x y = + J (λ m p → subst (λ n → coHom n A) p (x +ₕ y) + ≡ subst (λ n → coHom n A) p x +ₕ subst (λ n → coHom n A) p y) + (transportRefl (x +ₕ y) + ∙ cong₂ _+ₕ_ (sym (transportRefl x)) (sym (transportRefl y))) From 452b27b5424a6f018381affece39f7effdac84c1 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Thu, 28 Oct 2021 16:24:17 +0200 Subject: [PATCH 73/89] stuff --- Cubical/Functions/Fibration.agda | 6 + Cubical/Homotopy/Connected.agda | 8 + Cubical/Homotopy/Freudenthal.agda | 76 ++ Cubical/Homotopy/Group/S3.agda | 1399 +++++++++++++++++++++-------- 4 files changed, 1105 insertions(+), 384 deletions(-) diff --git a/Cubical/Functions/Fibration.agda b/Cubical/Functions/Fibration.agda index 704da5cb9..3feb3b65f 100644 --- a/Cubical/Functions/Fibration.agda +++ b/Cubical/Functions/Fibration.agda @@ -101,5 +101,11 @@ fiber≡ {f = f} {b = b} h h' = FibrationStr : (B : Type ℓb) → Type ℓ → Type (ℓ-max ℓ ℓb) FibrationStr B A = A → B +fiberCong : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) {a₀ a₁ : A} (q : f a₀ ≡ f a₁) + → fiber (cong f) q ≡ Path (fiber f (f a₁)) (a₀ , q) (a₁ , refl) +fiberCong f q = + cong (fiber (cong f)) (cong sym (lUnit (sym q))) + ∙ sym (fiber≡ (_ , q) (_ , refl)) + Fibration : (B : Type ℓb) → (ℓ : Level) → Type (ℓ-max ℓb (ℓ-suc ℓ)) Fibration {ℓb = ℓb} B ℓ = Σ[ A ∈ Type ℓ ] FibrationStr B A diff --git a/Cubical/Homotopy/Connected.agda b/Cubical/Homotopy/Connected.agda index a53fd226d..492cfacbc 100644 --- a/Cubical/Homotopy/Connected.agda +++ b/Cubical/Homotopy/Connected.agda @@ -178,6 +178,14 @@ isConnectedPathP n con a₀ a₁ = subst (isConnected n) (sym (PathP≡Path _ _ _)) (isConnectedPath n con _ _) +isConnectedCong : ∀ {ℓ ℓ'} (n : HLevel) {A : Type ℓ} {B : Type ℓ'} (f : A → B) + → isConnectedFun (suc n) f + → ∀ {a₀ a₁} → isConnectedFun n {A = a₀ ≡ a₁} (cong f) +isConnectedCong n f cf {a₀} {a₁} q = + subst (isConnected n) + (sym (fiberCong f q)) + (isConnectedPath n (cf (f a₁)) (a₀ , q) (a₁ , refl)) + isConnectedRetract : ∀ {ℓ ℓ'} (n : HLevel) {A : Type ℓ} {B : Type ℓ'} (f : A → B) (g : B → A) diff --git a/Cubical/Homotopy/Freudenthal.agda b/Cubical/Homotopy/Freudenthal.agda index 108bada57..20a5284bb 100644 --- a/Cubical/Homotopy/Freudenthal.agda +++ b/Cubical/Homotopy/Freudenthal.agda @@ -8,6 +8,7 @@ module Cubical.Homotopy.Freudenthal where open import Cubical.Foundations.Everything open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order open import Cubical.Data.Sigma open import Cubical.HITs.Nullification open import Cubical.HITs.Susp renaming (toSusp to σ) @@ -16,6 +17,8 @@ open import Cubical.Homotopy.Connected open import Cubical.Homotopy.WedgeConnectivity open import Cubical.Homotopy.Loopspace open import Cubical.HITs.SmashProduct +open import Cubical.Relation.Nullary +open import Cubical.Data.Empty renaming (rec to ⊥-rec) open import Cubical.HITs.S1 hiding (encode) open import Cubical.HITs.Sn @@ -132,3 +135,76 @@ FreudenthalIso : ∀ {ℓ} (n : HLevel) (A : Pointed ℓ) → Iso (hLevelTrunc ((suc n) + (suc n)) (typ A)) (hLevelTrunc ((suc n) + (suc n)) (typ (Ω (Susp (typ A) , north)))) FreudenthalIso n A iscon = connectedTruncIso _ (σ A) (isConnectedσ _ iscon) + + +{- connectedness of congⁿ σ (called suspMapΩ here) -} +∙∙lCancel-fill : ∀ {ℓ} {A : Type ℓ} {x y : A} + → (p : x ≡ y) + → I → I → I → A +∙∙lCancel-fill p i j k = + hfill (λ k → λ { (i = i1) → p k + ; (j = i0) → p k + ; (j = i1) → p k}) + (inS (p i0)) k + +∙∙lCancel : ∀ {ℓ} {A : Type ℓ} {x y : A} + → (p : x ≡ y) + → sym p ∙∙ refl ∙∙ p ≡ refl +∙∙lCancel p i j = ∙∙lCancel-fill p i j i1 + +suspMapΩ∙ : ∀ {ℓ} {A : Pointed ℓ}(n : ℕ) + → ((Ω^ n) A) + →∙ ((Ω^ (suc n)) (Susp∙ (typ A))) +fst (suspMapΩ∙ {A = A} zero) a = merid a ∙ sym (merid (pt A)) +snd (suspMapΩ∙ {A = A} zero) = rCancel (merid (pt A)) +fst (suspMapΩ∙ {A = A} (suc n)) p = sym (snd (suspMapΩ∙ n)) ∙∙ cong (fst (suspMapΩ∙ n)) p ∙∙ snd (suspMapΩ∙ n) +snd (suspMapΩ∙ {A = A} (suc n)) = ∙∙lCancel (snd (suspMapΩ∙ n)) + +suspMapΩ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → typ ((Ω^ n) A) → typ ((Ω^ (suc n)) (Susp∙ (typ A))) +suspMapΩ {A = A} n = suspMapΩ∙ {A = A} n .fst + +isConnectedCong' : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x : A} {y : B} + (n : ℕ) (f : A → B) + → isConnectedFun (suc n) f + → (p : f x ≡ y) + → isConnectedFun n + λ (q : x ≡ x) → sym p ∙∙ cong f q ∙∙ p +isConnectedCong' {x = x} n f conf p = + transport (λ i → (isConnectedFun n + λ (q : x ≡ x) + → doubleCompPath-filler (sym p) (cong f q) p i)) + (isConnectedCong n f conf) + +suspMapΩ-connected : ∀ {ℓ} (n : HLevel) (m : ℕ) {A : Pointed ℓ} + (connA : isConnected (suc (suc n)) (typ A)) + → isConnectedFun ((suc n + suc n) ∸ m) (suspMapΩ {A = A} m) +suspMapΩ-connected n zero {A = A} connA = isConnectedσ n connA +suspMapΩ-connected n (suc m) {A = A} connA with ((n + suc n) ≟ m) +... | (lt p) = subst (λ x → isConnectedFun x (suspMapΩ {A = A} (suc m))) + (sym (h _ m p)) + λ b → tt* , (λ {tt* → refl}) + where + h : (n m : ℕ) → n < m → (n ∸ m) ≡ 0 + h zero zero p = refl + h (suc n) zero p = ⊥-rec (¬-<-zero p) + h zero (suc m) p = refl + h (suc n) (suc m) p = h n m (pred-≤-pred p) +... | (eq q) = subst (λ x → isConnectedFun x (suspMapΩ {A = A} (suc m))) + (sym (help m) ∙ cong (_∸ m) (sym q)) + λ b → tt* , (λ {tt* → refl}) + where + help : (n : ℕ) → n ∸ n ≡ 0 + help zero = refl + help (suc n) = help n +... | (gt p) = isConnectedCong' (n + suc n ∸ m) (suspMapΩ {A = A} m) + (subst (λ x → isConnectedFun x (suspMapΩ {A = A} m)) + (sym (help (n + suc n) m p)) + (suspMapΩ-connected n m connA)) + (snd (suspMapΩ∙ m)) + where + help : (n m : ℕ) → m < n → suc (n ∸ m) ≡ (suc n) ∸ m + help zero zero p = refl + help zero (suc m) p = ⊥-rec (¬-<-zero p) + help (suc n) zero p = refl + help (suc n) (suc m) p = (help n m (pred-≤-pred p)) diff --git a/Cubical/Homotopy/Group/S3.agda b/Cubical/Homotopy/Group/S3.agda index 8a4eb3db8..146d96a7e 100644 --- a/Cubical/Homotopy/Group/S3.agda +++ b/Cubical/Homotopy/Group/S3.agda @@ -20,6 +20,7 @@ open import Cubical.Foundations.Function open import Cubical.Foundations.Path open import Cubical.Foundations.Transport open import Cubical.Functions.Embedding +open import Cubical.Functions.Morphism open import Cubical.Foundations.Pointed.Homogeneous open import Cubical.HITs.SetTruncation @@ -28,7 +29,7 @@ open import Cubical.HITs.SetTruncation open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; rec2 to pRec2 ; elim to pElim) open import Cubical.HITs.Sn -open import Cubical.Data.Bool +open import Cubical.Data.Bool hiding (_≟_) open import Cubical.HITs.Susp open import Cubical.HITs.S1 open import Cubical.Data.Sigma @@ -54,43 +55,60 @@ open import Cubical.HITs.Truncation open import Cubical.Data.Sum open import Cubical.Relation.Nullary open import Cubical.Data.Empty -open import Cubical.Data.Bool - -∙∙lCancel-fill : ∀ {ℓ} {A : Type ℓ} {x y : A} - → (p : x ≡ y) - → I → I → I → A -∙∙lCancel-fill p i j k = - hfill (λ k → λ { (i = i1) → p k - ; (j = i0) → p k - ; (j = i1) → p k}) - (inS (p i0)) k - -∙∙lCancel : ∀ {ℓ} {A : Type ℓ} {x y : A} - → (p : x ≡ y) - → sym p ∙∙ refl ∙∙ p ≡ refl -∙∙lCancel p i j = ∙∙lCancel-fill p i j i1 open import Cubical.Foundations.Equiv.HalfAdjoint +-- Solves some termination issues +private + +nInd : ∀ {ℓ} {P : ℕ → Type ℓ} + → P 0 + → P 1 + → ((n : ℕ) → P (suc n) → P (suc (suc n))) + → (n : ℕ) → P n + +nInd 0c 1c indc zero = 0c + +nInd 0c 1c indc (suc zero) = 1c + +nInd {P = P} 0c 1c indc (suc (suc n)) = + indc n (+nInd {P = P} 0c 1c indc (suc n)) + +{- +The goal of this file is to prove that the function +suspMapΩ : Ωⁿ A → Ωⁿ⁺¹ (Susp A), induced by +the Freudenthal function σ : A → ΩΣA, gets taken to +the canonical suspension map +suspMap : (Sⁿ →∙ A) → (Sⁿ⁺¹ →∙ Susp A) +given some suitable structure preserving equivalences +Ωⁿ A ≃ (Sⁿ →∙ A). + +The idea is to fill the following diagram + + suspMapΩ +Ωⁿ A -------------------> Ωⁿ⁺¹ (Susp A) + | | + | | + | ≃ lMap | ≃ eq₂ + | | + v suspMap v + (Sⁿ →∙ A) -------------- > (Sⁿ⁺¹ →∙ Susp A) + +where lMap and eq₂ are structure preserving equivalences +(we choose them (intensionally) different for techinical reasons) +-} + suspMap : ∀ {ℓ} {A : Pointed ℓ}(n : ℕ) → S₊∙ (suc n) →∙ A - → S₊∙ (suc (suc n)) →∙ ∙Susp (typ A) + → S₊∙ (suc (suc n)) →∙ Susp∙ (typ A) fst (suspMap n f) north = north fst (suspMap n f) south = north -fst (suspMap {A = A} n f) (merid a i) = (merid (fst f a) ∙ sym (merid (pt A))) i +fst (suspMap {A = A} n f) (merid a i) = + (merid (fst f a) ∙ sym (merid (pt A))) i snd (suspMap n f) = refl -suspMap2 : ∀ {ℓ} {A : Pointed ℓ}(n : ℕ) - → ((Ω^ n) A) - →∙ ((Ω^ (suc n)) (∙Susp (typ A))) -fst (suspMap2 {A = A} zero) a = merid a ∙ sym (merid (pt A)) -snd (suspMap2 {A = A} zero) = rCancel (merid (pt A)) -fst (suspMap2 {A = A} (suc n)) p = sym (snd (suspMap2 n)) ∙∙ cong (fst (suspMap2 n)) p ∙∙ snd (suspMap2 n) -snd (suspMap2 {A = A} (suc n)) = ∙∙lCancel (snd (suspMap2 n)) - -lMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → typ ((Ω^ n) A) → (S₊∙ n →∙ A) -lMapId : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (a : _) → lMap n {A = A} (pt ((Ω^ n) A)) .fst a ≡ pt A -lMapId2 : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → lMap n {A = A} (pt ((Ω^ n) A)) ≡ ((λ _ → pt A) , refl) +lMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} + → typ ((Ω^ n) A) → (S₊∙ n →∙ A) +lMapId : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (a : _) + → lMap n {A = A} (pt ((Ω^ n) A)) .fst a ≡ pt A +lMapId2 : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} + → lMap n {A = A} (pt ((Ω^ n) A)) ≡ ((λ _ → pt A) , refl) fst (lMap zero {A = A} a) false = a fst (lMap zero {A = A} a) true = pt A snd (lMap zero {A = A} a) = refl @@ -100,7 +118,9 @@ snd (lMap (suc zero) {A = A} p) = refl fst (lMap (suc (suc n)) {A = A} p) north = pt A fst (lMap (suc (suc n)) {A = A} p) south = pt A fst (lMap (suc (suc n)) {A = A} p) (merid a i) = - (sym (lMapId (suc n) a) ∙∙ (λ i → lMap (suc n) (p i) .fst a) ∙∙ lMapId (suc n) a) i + (sym (lMapId (suc n) a) + ∙∙ (λ i → lMap (suc n) (p i) .fst a) + ∙∙ lMapId (suc n) a) i snd (lMap (suc (suc n)) {A = A} p) = refl lMapId zero false = refl lMapId zero true = refl @@ -108,12 +128,22 @@ lMapId (suc zero) base = refl lMapId (suc zero) (loop i) = refl lMapId (suc (suc n)) north = refl lMapId (suc (suc n)) south = refl -lMapId (suc (suc n)) {A = A} (merid a i) j = (∙∙lCancel (lMapId (suc n) {A = A} a)) j i -lMapId2 zero {A = A} = ΣPathP ((funExt (λ { false → refl ; true → refl})) , refl) -lMapId2 (suc zero) {A = A} = ΣPathP ((funExt (λ { base → refl ; (loop i) → refl})) , refl) +lMapId (suc (suc n)) {A = A} (merid a i) j = + ∙∙lCancel (lMapId (suc n) {A = A} a) j i +lMapId2 zero {A = A} = + ΣPathP ((funExt (λ { false → refl + ; true → refl})) + , refl) +lMapId2 (suc zero) {A = A} = + ΣPathP ((funExt (λ { base → refl + ; (loop i) → refl})) + , refl) lMapId2 (suc (suc n)) {A = A} = - ΣPathP ((funExt (λ { north → refl ; south → refl ; (merid a i) j → ∙∙lCancel (lMapId (suc n) {A = A} a) j i})) , refl) - + ΣPathP ((funExt (λ { north → refl + ; south → refl + ; (merid a i) j + → ∙∙lCancel (lMapId (suc n) {A = A} a) j i})) + , refl) {- Ω (Ωⁿ A) -----------> Ω (Ωⁿ⁺¹ ΣA) | | @@ -128,18 +158,31 @@ lMapId2 (suc (suc n)) {A = A} = (Sⁿ⁺¹ →∙ A) ----> Sⁿ⁺² →∙ Σ A -} +-- Move to loopspace Ω-fun : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → (A →∙ B) → (Ω A →∙ Ω B) fst (Ω-fun {A = A} {B = B} (f , p)) q = sym p ∙∙ cong f q ∙∙ p snd (Ω-fun {A = A} {B = B} (f , p)) = ∙∙lCancel p +-- move to loop space +open import Cubical.Foundations.Equiv.HalfAdjoint +isEquivΩfun : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + → (f : (A →∙ B)) + → isEquiv (fst f) → isEquiv (Ω-fun f .fst) +isEquivΩfun {B = (B , b)} = + uncurry λ f → + J (λ b y → isEquiv f → isEquiv (λ q → (λ i → y (~ i)) ∙∙ (λ i → f (q i)) ∙∙ y)) + λ eqf → subst isEquiv (funExt (rUnit ∘ cong f)) + (isoToIsEquiv (congIso (equivToIso (f , eqf)))) + + +-- We define the following maps which will be used to +-- show that lMap is an equivalence lMapSplit₁ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → typ ((Ω^ (suc n)) A) → typ (Ω (S₊∙ n →∙ A ∙)) lMapSplit₁ n = Ω-fun (lMap n , lMapId2 n) .fst - --- sphereMapEquiv ΩSphereMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → typ (Ω (S₊∙ n →∙ A ∙)) → (S₊∙ (suc n) →∙ A) @@ -151,99 +194,119 @@ fst (ΩSphereMap {A = A} (suc n) p) south = pt A fst (ΩSphereMap {A = A} (suc n) p) (merid a i) = p i .fst a snd (ΩSphereMap {A = A} (suc n) p) = refl -open import Cubical.Foundations.Equiv.HalfAdjoint -coolBeans : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → (f : (A →∙ B)) → isEquiv (fst f) → isEquiv (Ω-fun f .fst) -coolBeans {B = (B , b)} = - uncurry λ f → - J (λ b y → isEquiv f → isEquiv (λ q → (λ i → y (~ i)) ∙∙ (λ i → f (q i)) ∙∙ y)) - λ eqf → subst isEquiv (funExt (rUnit ∘ cong f)) (isoToIsEquiv (congIso (equivToIso (f , eqf)))) - SphereMapΩ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → (S₊∙ (suc n) →∙ A) → typ (Ω (S₊∙ n →∙ A ∙)) -SphereMapΩ {A = A} zero (f , p) = ΣPathP ((funExt λ { false → sym p ∙∙ cong f loop ∙∙ p ; true → refl}) , refl) +SphereMapΩ {A = A} zero (f , p) = + ΣPathP ((funExt λ { false → sym p ∙∙ cong f loop ∙∙ p + ; true → refl}) + , refl) SphereMapΩ {A = A} (suc n) (f , p) = ΣPathP (funExt (λ x → sym p ∙∙ cong f (merid x ∙ sym (merid (ptSn _))) ∙∙ p) - , flipSquare (cong (sym p ∙∙_∙∙ p) (cong (cong f) (rCancel (merid (ptSn _)))) ∙ ∙∙lCancel p)) + , flipSquare (cong (sym p ∙∙_∙∙ p) (cong (cong f) (rCancel (merid (ptSn _)))) + ∙ ∙∙lCancel p)) SphereMapΩIso : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → Iso (S₊∙ (suc n) →∙ A) (typ (Ω (S₊∙ n →∙ A ∙))) fun (SphereMapΩIso {A = A} n) = SphereMapΩ n inv (SphereMapΩIso {A = A} n) = ΩSphereMap n -fst (rightInv (SphereMapΩIso {A = A} zero) p k i) false = rUnit (funExt⁻ (cong fst p) false) (~ k) i +fst (rightInv (SphereMapΩIso {A = A} zero) p k i) false = + rUnit (funExt⁻ (cong fst p) false) (~ k) i fst (rightInv (SphereMapΩIso {A = A} zero) p k i) true = snd (p i) (~ k) snd (rightInv (SphereMapΩIso {A = A} zero) p k i) l = snd (p i) (~ k ∨ l) fst (rightInv (SphereMapΩIso {A = A} (suc n)) p k i) x = help k i where - help : refl ∙∙ cong (fst (ΩSphereMap {A = A} (suc n) p)) (merid x ∙ sym (merid (ptSn _))) ∙∙ refl + help : cong (fst (ΩSphereMap {A = A} (suc n) p)) (merid x ∙ sym (merid (ptSn _))) ∙ refl ≡ funExt⁻ (cong fst p) x - help = (cong (refl ∙∙_∙∙ refl) (cong-∙ (fst (ΩSphereMap {A = A} (suc n) p)) (merid x) (sym (merid (ptSn _)))) - ∙ cong (refl ∙∙_∙∙ refl) (cong (funExt⁻ (cong fst p) x ∙_) (cong sym (flipSquare (cong snd p))) ∙ sym (rUnit _))) - ∙ sym (rUnit _) + help = (cong (refl ∙∙_∙∙ refl) (cong-∙ (fst (ΩSphereMap {A = A} (suc n) p)) + (merid x) (sym (merid (ptSn _)))) + ∙ cong (refl ∙∙_∙∙ refl) (cong (funExt⁻ (cong fst p) x ∙_) + (cong sym (flipSquare (cong snd p))) + ∙ sym (rUnit _))) + ∙ sym (rUnit _) snd (rightInv (SphereMapΩIso {A = A} (suc n)) p k i) j = - hcomp (λ r → λ { (i = i0) → pt A - ; (i = i1) → pt A - ; (j = i0) → ((cong (refl ∙∙_∙∙ refl) (cong-∙ (fst (ΩSphereMap {A = A} (suc n) p)) (merid (ptSn (suc n))) (sym (merid (ptSn _)))) - ∙ cong (refl ∙∙_∙∙ refl) (cong (funExt⁻ (cong fst p) (ptSn (suc n)) ∙_) (cong sym (flipSquare (cong snd p))) ∙ sym (rUnit _))) - ∙ sym (rUnit _)) k i - ; (j = i1) → pt A - ; (k = i0) → (cong (sym refl ∙∙_∙∙ refl) (cong (cong (fst (ΩSphereMap {A = A} (suc n) p))) (rCancel (merid (ptSn _)))) ∙ sym (rUnit refl)) j i - ; (k = i1) → snd (p i) j}) - (hcomp (λ r → λ { (i = i0) → pt A - ; (i = i1) → pt A - ; (j = i0) → ((cong (refl ∙∙_∙∙ refl) (cong-∙ (fst (ΩSphereMap {A = A} (suc n) p)) (merid (ptSn (suc n))) (sym (merid (ptSn _)))) - ∙ cong (refl ∙∙_∙∙ refl) (cong (funExt⁻ (cong fst p) (ptSn (suc n)) ∙_) (cong sym (flipSquare (cong snd p))) ∙ sym (rUnit _))) - ∙ sym (rUnit _)) k i - ; (j = i1) → rUnit (λ _ → pt A) (~ r ∧ ~ k) i - ; (k = i0) → compPath-filler (cong (sym refl ∙∙_∙∙ refl) (cong (cong (fst (ΩSphereMap {A = A} (suc n) p))) (rCancel (merid (ptSn _))))) - (sym (rUnit refl)) r j i - ; (k = i1) → snd (p i) j}) - ((hcomp (λ r → λ { (i = i0) → rUnit (λ _ → pt A) (~ k ∨ r) i - ; (i = i1) → rUnit (λ _ → pt A) (~ k ∨ r) i - ; (j = i0) → ((cong (λ x → rUnit x r) (cong-∙ (fst (ΩSphereMap {A = A} (suc n) p)) (merid (ptSn (suc n))) (sym (merid (ptSn _)))) - ∙ cong (λ x → rUnit x r) (cong (funExt⁻ (cong fst p) (ptSn (suc n)) ∙_) (cong sym (flipSquare (cong snd p))) ∙ sym (rUnit _))) - ∙ λ i → rUnit (λ i₁ → funExt⁻ (λ i₂ → fst (p i₂)) (ptSn (suc n)) i₁) (~ i ∧ r)) k i - ; (j = i1) → rUnit (λ _ → pt A) (~ k ∧ r) i - ; (k = i0) → rUnit ((cong (fst (ΩSphereMap {A = A} (suc n) p))) (rCancel (merid (ptSn _)) j)) r i - ; (k = i1) → snd (p i) j})) - (hcomp (λ r → λ { (i = i0) → pt A - ; (i = i1) → pt A - ; (j = i0) → ((rUnit (compPath-filler' (cong-∙ (fst (ΩSphereMap {A = A} (suc n) p)) (merid (ptSn (suc n))) (sym (merid (ptSn _)))) - ((cong (funExt⁻ (cong fst p) (ptSn (suc n)) ∙_) (cong sym (flipSquare (cong snd p))) ∙ sym (rUnit _))) r) r) k i) - ; (j = i1) → pt A - ; (k = i0) → help (fst (ΩSphereMap {A = A} (suc n) p)) (merid (ptSn (suc n))) r j i - ; (k = i1) → snd (p i) j}) - (cool2 (funExt⁻ (cong fst p) (ptSn _)) (sym (flipSquare (cong snd p))) k j i)))) + hcomp (λ r → + λ { (i = i0) → pt A + ; (i = i1) → pt A + ; (j = i0) → ((wrap-refl (cong-∙ Ωp (merid (ptSn (suc n))) + (sym (merid (ptSn _)))) + ∙ wrap-refl (cong (p-refl ∙_) + (cong sym (flipSquare (cong snd p))) + ∙ sym (rUnit _))) ∙ sym (rUnit _)) k i + ; (j = i1) → rUnit (λ _ → pt A) (~ r ∧ ~ k) i + ; (k = i0) → compPath-filler (wrap-refl (cong (cong Ωp) + (rCancel (merid (ptSn _))))) + (sym (rUnit refl)) r j i + ; (k = i1) → snd (p i) j}) + (hcomp (λ r → + λ { (i = i0) → rUnit (λ _ → pt A) (~ k ∨ r) i + ; (i = i1) → rUnit (λ _ → pt A) (~ k ∨ r) i + ; (j = i0) → ((cong (λ x → rUnit x r) + (cong-∙ Ωp (merid (ptSn (suc n))) + (sym (merid (ptSn _)))) + ∙ cong (λ x → rUnit x r) + (cong (p-refl ∙_) + (cong sym (flipSquare (cong snd p))) + ∙ sym (rUnit _))) + ∙ λ i → rUnit p-refl (~ i ∧ r)) k i + ; (j = i1) → rUnit (λ _ → pt A) (~ k ∧ r) i + ; (k = i0) → rUnit (cong Ωp (rCancel (merid (ptSn _)) j)) r i + ; (k = i1) → snd (p i) j}) + (hcomp (λ r → + λ { (i = i0) → pt A + ; (i = i1) → pt A + ; (j = i0) → ((rUnit (compPath-filler' + (cong-∙ Ωp (merid (ptSn (suc n))) + (sym (merid (ptSn _)))) + ((cong (p-refl ∙_) + (cong sym (flipSquare (cong snd p))) + ∙ sym (rUnit _))) r) r) k i) + ; (j = i1) → pt A + ; (k = i0) → help Ωp (merid (ptSn (suc n))) r j i + ; (k = i1) → snd (p i) j}) + (cool2 p-refl (sym (flipSquare (cong snd p))) k j i))) where - cong-∙∙-filler : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y z w : A} (f : A → B) (p : w ≡ x) (q : x ≡ y) (r : y ≡ z) → I → I → I → B - cong-∙∙-filler f p q r k j i = hfill (λ k → λ { (j = i0) → f (doubleCompPath-filler p q r k i) - ; (i = i0) → f (p (~ k)) - ; (i = i1) → f (r k) }) - (inS (f (q i))) k - - cool2 : ∀ {ℓ} {A : Pointed ℓ} → (p : Path (typ A) (pt A) (pt A)) - → (q : refl ≡ p) - → PathP (λ i → (cong (p ∙_) (cong sym (sym q)) ∙ sym (rUnit p)) i ≡ refl) - (rCancel p) (sym q) - cool2 {A = A} p = J (λ p q → PathP (λ i → (cong (p ∙_) (cong sym (sym q)) ∙ sym (rUnit p)) i ≡ refl) - (rCancel p) (sym q)) - (flipSquare (sym (lUnit (sym (rUnit refl))) - ◁ λ k i → rCancel (λ _ → pt A) (k ∨ i))) + Ωp = (fst (ΩSphereMap {A = A} (suc n) p)) + wrap-refl : ∀ {ℓ} {A : Type ℓ} {x : A} {r s : x ≡ x} (p : r ≡ s) → _ + wrap-refl p = cong (refl ∙∙_∙∙ refl) p + p-refl = funExt⁻ (cong fst p) (ptSn (suc n)) + + cong-∙∙-filler : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y z w : A} + (f : A → B) (p : w ≡ x) (q : x ≡ y) (r : y ≡ z) → I → I → I → B + cong-∙∙-filler f p q r k j i = + hfill (λ k → λ { (j = i0) → f (doubleCompPath-filler p q r k i) + ; (i = i0) → f (p (~ k)) + ; (i = i1) → f (r k) }) + (inS (f (q i))) k + + cool2 : ∀ {ℓ} {A : Pointed ℓ} (p : Path (typ A) (pt A) (pt A)) (q : refl ≡ p) + → PathP (λ i → (cong (p ∙_) (cong sym (sym q)) ∙ sym (rUnit p)) i ≡ refl) + (rCancel p) (sym q) + cool2 {A = A} p = + J (λ p q → + PathP (λ i → (cong (p ∙_) (cong sym (sym q)) ∙ sym (rUnit p)) i ≡ refl) + (rCancel p) (sym q)) + (flipSquare (sym (lUnit (sym (rUnit refl))) + ◁ λ k i → rCancel (λ _ → pt A) (k ∨ i))) help : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A} (f : A → B) (p : x ≡ y) - → PathP (λ k → cong-∙ f p (sym p) (~ k) ≡ refl) (rCancel (cong f p)) (cong (cong f) (rCancel p)) - help f p i j k = hcomp (λ r → λ { (i = i0) → rCancel-filler (cong f p) r j k - ; (i = i1) → f (rCancel-filler p r j k) - ; (j = i0) → cong-∙∙-filler f refl p (sym p) r (~ i) k - ; (j = i1) → f (p i0) - ; (k = i0) → f (p i0) - ; (k = i1) → f (p (~ r ∧ ~ j))}) - (f (p (~ j ∧ k))) - -leftInv (SphereMapΩIso {A = A} zero) (f , p) = ΣPathP ((funExt (λ { base → sym p - ; (loop i) j → doubleCompPath-filler (sym p) (cong f loop) p (~ j) i})) - , λ i j → p (j ∨ ~ i)) + → PathP (λ k → cong-∙ f p (sym p) (~ k) ≡ refl) + (rCancel (cong f p)) (cong (cong f) (rCancel p)) + help f p i j k = + hcomp (λ r → λ { (i = i0) → rCancel-filler (cong f p) r j k + ; (i = i1) → f (rCancel-filler p r j k) + ; (j = i0) → cong-∙∙-filler f refl p (sym p) r (~ i) k + ; (j = i1) → f (p i0) + ; (k = i0) → f (p i0) + ; (k = i1) → f (p (~ r ∧ ~ j))}) + (f (p (~ j ∧ k))) + +leftInv (SphereMapΩIso {A = A} zero) (f , p) = + ΣPathP ((funExt (λ { base → sym p + ; (loop i) j + → doubleCompPath-filler (sym p) (cong f loop) p (~ j) i})) + , λ i j → p (j ∨ ~ i)) leftInv (SphereMapΩIso {A = A} (suc n)) (f , p) = ΣPathP (funExt (λ { north → sym p ; south → sym p ∙ cong f (merid (ptSn (suc n))) @@ -253,139 +316,187 @@ leftInv (SphereMapΩIso {A = A} (suc n)) (f , p) = help : (a : S₊ (suc n)) → PathP (λ i → p (~ i) ≡ (sym p ∙ cong f (merid (ptSn (suc n)))) i) (sym p ∙∙ cong f (merid a ∙ sym (merid (ptSn _))) ∙∙ p) (cong f (merid a)) - help a i j = hcomp (λ k → λ { (i = i0) → doubleCompPath-filler (sym p) (cong f (merid a ∙ sym (merid (ptSn _)))) p k j - ; (i = i1) → f (merid a j) - ; (j = i0) → p (~ i ∧ k) - ; (j = i1) → compPath-filler' (sym p) (cong f (merid (ptSn (suc n)))) k i}) - (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) + help a i j = + hcomp (λ k → + λ { (i = i0) → doubleCompPath-filler (sym p) + (cong f (merid a ∙ sym (merid (ptSn _)))) p k j + ; (i = i1) → f (merid a j) + ; (j = i0) → p (~ i ∧ k) + ; (j = i1) → compPath-filler' (sym p) + (cong f (merid (ptSn (suc n)))) k i}) + (f (compPath-filler (merid a) + (sym (merid (ptSn _))) (~ i) j)) + +{- +In order to show that lMap is an equivalence, we show that it factors + SphereMapΩ lMapSplit₁ +Ωⁿ⁺¹A ----------------> Ω (Sⁿ →∙ A) -----------> (Sⁿ⁺¹ →∙ A) +-} -cool̂ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p : typ ((Ω^ (suc n)) A)) → lMap (suc n) p ≡ ΩSphereMap n (lMapSplit₁ n p) -cool̂ {A = A} zero p = ΣPathP ((funExt (λ { base → refl ; (loop i) j → h (~ j) i})) , refl) + +lMap-split : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p : typ ((Ω^ (suc n)) A)) + → lMap (suc n) p ≡ ΩSphereMap n (lMapSplit₁ n p) +lMap-split {A = A} zero p = + ΣPathP ((funExt (λ { base → refl + ; (loop i) j → lem (~ j) i})) + , refl) where - h : funExt⁻ (cong fst (lMapSplit₁ zero p)) false ≡ p - h = (λ i → funExt⁻ (cong-∙∙ fst (sym (lMapId2 zero)) (cong (lMap zero) p) (lMapId2 zero) i) false) + lem : funExt⁻ (cong fst (lMapSplit₁ zero p)) false ≡ p + lem = (λ i → funExt⁻ (cong-∙∙ fst (sym (lMapId2 zero)) + (cong (lMap zero) p) + (lMapId2 zero) i) false) ∙ sym (rUnit _) -cool̂ {A = A} (suc n) p = ΣPathP ((funExt (λ { north → refl ; south → refl ; (merid a i) j → h a j i})) , refl) +lMap-split {A = A} (suc n) p = + ΣPathP ((funExt (λ { north → refl + ; south → refl + ; (merid a i) j → lem₂ a j i})) + , refl) where - lem : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (a : S₊ (suc n)) → lMapId (suc n) {A = A} a ≡ (λ i → fst (lMapId2 (suc n) {A = A} i) a) + lem : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (a : S₊ (suc n)) + → lMapId (suc n) {A = A} a + ≡ (λ i → fst (lMapId2 (suc n) {A = A} i) a) lem zero base = refl lem zero (loop i) = refl lem (suc n) north = refl lem (suc n) south = refl lem (suc n) (merid a i) = refl - h : (a : S₊ (suc n)) → ((λ i₁ → lMapId (suc n) {A = A} a (~ i₁)) ∙∙ - (λ i₁ → lMap (suc n) (p i₁) .fst a) ∙∙ lMapId (suc n) a) ≡ (λ i → lMapSplit₁ (suc n) p i .fst a) - h a = cong (λ x → sym x ∙∙ funExt⁻ (cong fst (λ i → lMap (suc n) (p i))) a ∙∙ x) + lem₂ : (a : S₊ (suc n)) + → ((λ i₁ → lMapId (suc n) {A = A} a (~ i₁)) + ∙∙ (λ i₁ → lMap (suc n) (p i₁) .fst a) + ∙∙ lMapId (suc n) a) + ≡ (λ i → lMapSplit₁ (suc n) p i .fst a) + lem₂ a = cong (λ x → sym x + ∙∙ funExt⁻ (cong fst (λ i → lMap (suc n) (p i))) a + ∙∙ x) (lem n a) - ∙∙ sym (cong-∙∙ (λ x → x a) (cong fst (λ i → lMapId2 (suc n) (~ i))) (cong fst (λ i → lMap (suc n) (p i))) (cong fst (lMapId2 (suc n)))) - ∙∙ (λ i → funExt⁻ (cong-∙∙ fst (sym (lMapId2 (suc n))) (cong (lMap (suc n)) p) (lMapId2 (suc n)) (~ i)) a) - -isEquiv-lMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → isEquiv (lMap (suc n) {A = A}) -isEquiv-lMap zero {A = A} = isoToIsEquiv (iso _ invFun sec λ p → sym (rUnit p)) + ∙∙ sym (cong-∙∙ (λ x → x a) + (cong fst (λ i → lMapId2 (suc n) (~ i))) + (cong fst (λ i → lMap (suc n) (p i))) + (cong fst (lMapId2 (suc n)))) + ∙∙ (λ i → funExt⁻ (cong-∙∙ fst (sym (lMapId2 (suc n))) + (cong (lMap (suc n)) p) + (lMapId2 (suc n)) (~ i)) a) + +isEquiv-lMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} + → isEquiv (lMap (suc n) {A = A}) +isEquiv-lMap zero {A = A} = + isoToIsEquiv (iso _ invFun sec λ p → sym (rUnit p)) where invFun : S₊∙ 1 →∙ A → typ (Ω A) invFun (f , p) = sym p ∙∙ cong f loop ∙∙ p sec : section (lMap 1) invFun - sec (f , p) = ΣPathP ((funExt (λ { base → sym p ; (loop i) j → doubleCompPath-filler (sym p) (cong f loop) p (~ j) i})) + sec (f , p) = + ΣPathP ((funExt (λ { base → sym p + ; (loop i) j → doubleCompPath-filler + (sym p) (cong f loop) p (~ j) i})) , λ i j → p (~ i ∨ j)) isEquiv-lMap (suc n) = - subst isEquiv (sym (funExt (cool̂ (suc n)))) - (snd (compEquiv ((lMapSplit₁ (suc n)) , (coolBeans (lMap (suc n) , lMapId2 (suc n)) (isEquiv-lMap n))) + subst isEquiv (sym (funExt (lMap-split (suc n)))) + (snd (compEquiv + ((lMapSplit₁ (suc n)) , + (isEquivΩfun (lMap (suc n) , lMapId2 (suc n)) + (isEquiv-lMap n))) (invEquiv (isoToEquiv (SphereMapΩIso (suc n)))))) -isHom-lMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → (p q : _) → lMap (suc n) {A = A} (p ∙ q) ≡ ∙Π (lMap (suc n) p) (lMap (suc n) q) -isHom-lMap zero p q = ΣPathP ((funExt (λ { base → refl ; (loop i) j → (rUnit p j ∙ rUnit q j) i})) , refl) -isHom-lMap (suc n) {A = A} p q = ΣPathP ((funExt (λ { north → refl ; south → refl ; (merid a i) j → h a j i})) , refl) - where - doubleComp-lem : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) (q r : y ≡ y) - → (p ∙∙ q ∙∙ sym p) ∙ (p ∙∙ r ∙∙ sym p) ≡ (p ∙∙ (q ∙ r) ∙∙ sym p) - doubleComp-lem p q r i j = - hcomp (λ k → λ { (i = i0) → (doubleCompPath-filler p q (sym p) k - ∙ doubleCompPath-filler p r (sym p) k) j - ; (i = i1) → doubleCompPath-filler p (q ∙ r) (sym p) k j - ; (j = i0) → p (~ k) - ; (j = i1) → p (~ k)}) - ((q ∙ r) j) - - help2 : (p : typ ((Ω^ (suc (suc n))) A)) → cong (fst (lMap (suc (suc n)) p)) (merid (ptSn _)) ≡ refl - help2 p = cong (sym (lMapId (suc n) (ptSn _)) ∙∙_∙∙ lMapId (suc n) (ptSn _)) - (rUnit _ ∙ (λ j → (λ i → lMap (suc n) {A = A} refl .snd (i ∧ j)) - ∙∙ (λ i → lMap (suc n) {A = A} (p i) .snd j) - ∙∙ λ i → lMap (suc n) {A = A} refl .snd (~ i ∧ j)) - ∙ ∙∙lCancel _) - ∙ ∙∙lCancel _ - - h : (a : S₊ (suc n)) → sym (lMapId (suc n) a) ∙∙ funExt⁻ (cong fst (cong (lMap (suc n)) (p ∙ q))) a ∙∙ lMapId (suc n) a - ≡ cong (fst (∙Π (lMap (suc (suc n)) p) (lMap (suc (suc n)) q))) (merid a) - h a = (cong (sym (lMapId (suc n) a) ∙∙_∙∙ (lMapId (suc n) a)) - (cong-∙ (λ x → lMap (suc n) x .fst a) p q) - ∙ sym (doubleComp-lem (sym (lMapId (suc n) a)) _ _)) - ∙∙ cong₂ _∙_ (sym (cong (cong (fst (lMap (suc (suc n)) p)) (merid a) ∙_) (cong sym (help2 p)) ∙ sym (rUnit _))) - (sym (cong (cong (fst (lMap (suc (suc n)) q)) (merid a) ∙_) (cong sym (help2 q)) ∙ sym (rUnit _))) - ∙∙ λ i → (rUnit (cong-∙ (fst (lMap (suc (suc n)) p)) (merid a) (sym (merid (ptSn _))) (~ i)) i) - ∙ (rUnit (cong-∙ (fst (lMap (suc (suc n)) q)) (merid a) (sym (merid (ptSn _)))(~ i)) i) - botMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → (S₊∙ n →∙ A) - → S₊∙ n →∙ Ω (∙Susp (typ A)) + → S₊∙ n →∙ Ω (Susp∙ (typ A)) fst (botMap n {A = A} f) x = merid (fst f x) ∙ sym (merid (pt A)) snd (botMap n {A = A} f) = cong (_∙ merid (pt A) ⁻¹) (cong merid (snd f)) ∙ rCancel (merid (pt A)) -rMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → typ ((Ω^ (suc n)) (∙Susp (typ A))) → (S₊∙ n →∙ Ω (∙Susp (typ A))) +rMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} + → typ ((Ω^ (suc n)) (Susp∙ (typ A))) + → (S₊∙ n →∙ Ω (Susp∙ (typ A))) rMap n = lMap n ∘ Iso.fun (flipΩIso n) -rMap1 : ∀ {ℓ} {A : Pointed ℓ} → typ ((Ω^ (suc 1)) (∙Susp (typ A))) → (S₊∙ 1 →∙ Ω (∙Susp (typ A))) -rMap1 = lMap 1 +private + rMap1 : ∀ {ℓ} {A : Pointed ℓ} + → typ ((Ω^ (suc 1)) (Susp∙ (typ A))) + → (S₊∙ 1 →∙ Ω (Susp∙ (typ A))) + rMap1 = lMap 1 + + rMap≡rMap1 : ∀ {ℓ} {A : Pointed ℓ} → rMap 1 {A = A} ≡ rMap1 {A = A} + rMap≡rMap1 = funExt λ x → cong (lMap 1) (transportRefl x) + +flipΩrefl : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → fun (flipΩIso {A = A} (suc n)) refl ≡ refl +flipΩrefl {A = A} n j = + transp (λ i₁ → fst (Ω (flipΩPath {A = A} n ((i₁ ∨ j))))) j (snd (Ω (flipΩPath n j))) + +cong-lMap-lem : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p : _) + → cong (lMap (suc n) {A = Ω A}) (fun (flipΩIso (suc (suc n))) p) + ≡ (cong (lMap (suc n)) (sym (flipΩrefl n)) + ∙∙ cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) p + ∙∙ cong (lMap (suc n)) (flipΩrefl n)) +cong-lMap-lem {A = A} n p = + (λ i → cong (lMap (suc n) {A = Ω A}) + (fun (flipΩIso (suc (suc n))) (rUnit p i))) + ∙∙ (λ i j → lMap (suc n) {A = Ω A} + (transp (λ k → fst (flipΩPath {A = A} (suc (suc n)) (k ∨ i))) i + (((λ j → transp (λ k → fst (Ω (flipΩPath {A = A} n ((k ∨ ~ j) ∧ i)))) + (~ i ∨ ~ j) + (snd (Ω (flipΩPath n (~ j ∧ i))))) + ∙∙ (λ j → transp (λ k → fst (flipΩPath {A = A} (suc n) + (k ∧ i))) (~ i) (p j)) + ∙∙ λ j → transp (λ k → fst (Ω (flipΩPath {A = A} n ((k ∨ j) ∧ i)))) + (~ i ∨ j) (snd (Ω (flipΩPath n (j ∧ i)))))) j)) + ∙∙ cong-∙∙ (lMap (suc n)) + (sym (flipΩrefl n)) + (cong (fun (flipΩIso (suc n))) p) + (flipΩrefl n) + +botₗ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (S₊∙ n →∙ A) → S₊∙ (suc n) →∙ (Susp∙ (typ A)) +fst (botₗ zero f) base = north +fst (botₗ {A = A} zero f) (loop i) = (merid (fst f false) ∙ sym (merid (pt A))) i +snd (botₗ zero f) = refl +botₗ (suc n) f = suspMap n f -rMap≡rMap1 : ∀ {ℓ} {A : Pointed ℓ} → rMap 1 {A = A} ≡ rMap1 {A = A} -rMap≡rMap1 = funExt λ x → cong (lMap 1) (transportRefl x) +botᵣ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (S₊∙ n →∙ Ω (Susp∙ (typ A))) → S₊∙ (suc n) →∙ Susp∙ (typ A) +fst (botᵣ zero (f , p)) base = north +fst (botᵣ zero (f , p)) (loop i) = f false i +snd (botᵣ zero (f , p)) = refl +fst (botᵣ (suc n) (f , p)) north = north +fst (botᵣ (suc n) (f , p)) south = north +fst (botᵣ (suc n) (f , p)) (merid a i) = f a i +snd (botᵣ (suc n) (f , p)) = refl -flipΩrefl : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → fun (flipΩIso {A = A} (suc n)) refl ≡ refl -flipΩrefl {A = A} n j = transp (λ i₁ → fst (Ω (flipΩPath {A = A} n ((i₁ ∨ j))))) j (snd (Ω (flipΩPath n j))) +{- +The goal now is to fill the following diagram. -cong-lMap-main' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p : _) - → cong (lMap (suc n) {A = Ω A}) (fun (flipΩIso (suc (suc n))) p) - ≡ (cong (lMap (suc n)) (sym (flipΩrefl n)) - ∙∙ cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) p - ∙∙ cong (lMap (suc n)) (flipΩrefl n)) -cong-lMap-main' {A = A} n p = (λ i → cong (lMap (suc n) {A = Ω A}) (fun (flipΩIso (suc (suc n))) (rUnit p i))) - ∙ (λ i j → lMap (suc n) {A = Ω A} - (transp (λ i₁ → fst (flipΩPath {A = A} (suc (suc n)) (i₁ ∨ i))) i - (((λ j → transp (λ i₁ → fst (Ω (flipΩPath {A = A} n ((i₁ ∨ ~ j) ∧ i)))) (~ i ∨ ~ j) (snd (Ω (flipΩPath n (~ j ∧ i))))) - ∙∙ (λ j → transp (λ i₁ → fst (flipΩPath {A = A} (suc n) (i₁ ∧ i))) (~ i) (p j)) - ∙∙ λ j → transp (λ i₁ → fst (Ω (flipΩPath {A = A} n ((i₁ ∨ j) ∧ i)))) (~ i ∨ j) (snd (Ω (flipΩPath n (j ∧ i)))))) j)) - ∙ cong-∙∙ (lMap (suc n)) (sym (flipΩrefl n)) (cong (fun (flipΩIso (suc n))) p) (flipΩrefl n) - -+nInd : ∀ {ℓ} {P : ℕ → Type ℓ} - → P 0 - → P 1 - → ((n : ℕ) → P (suc n) → P (suc (suc n))) - → (n : ℕ) → P n -+nInd 0c 1c indc zero = 0c -+nInd 0c 1c indc (suc zero) = 1c -+nInd {P = P} 0c 1c indc (suc (suc n)) = indc n (+nInd {P = P} 0c 1c indc (suc n)) - -{- suspMap + suspMap ((Ω^ n) A) ------------> Ω^ (1 + n) (Susp A) | | | | - lMap | ≃ | rMap + lMap | ≃ ≃ | rMap v v (S₊∙ n →∙ A) ----------> (S₊∙ n →∙ Ω (Susp A)) - \ botMap / - \ / - \ / - \ / - \ / - \ / + \ botMap / + \ / + \ / + botₗ \ / botᵣ + (suspMap) \ / + \ / + \ / + v (Sⁿ⁺¹ →∙ Susp A) + -} +-- The bottom part is trival▿ +filler▿ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (x : (S₊∙ n →∙ A)) + → botₗ n x ≡ botᵣ {A = A} n (botMap n x) +filler▿ zero (f , p) = + ΣPathP ((funExt (λ { base → refl ; (loop i) → refl})) , refl) +filler▿ (suc n) (f , p) = + ΣPathP ((funExt (λ { north → refl ; south → refl ; (merid a i) → refl})) , refl) + IsoΩSphereMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → Iso (typ ((Ω^ n) A)) ((S₊∙ n →∙ A)) fun (IsoΩSphereMap zero) = lMap zero @@ -398,93 +509,96 @@ IsoΩSphereMap (suc n) = equivToIso (_ , isEquiv-lMap n) Ω≡SphereMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → (typ ((Ω^ n) A)) ≡ ((S₊∙ n →∙ A)) Ω≡SphereMap n = isoToPath (IsoΩSphereMap n) -botₗ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → (S₊∙ n →∙ A) → S₊∙ (suc n) →∙ (∙Susp (typ A)) -fst (botₗ zero f) base = north -fst (botₗ {A = A} zero f) (loop i) = (merid (fst f false) ∙ sym (merid (pt A))) i -snd (botₗ zero f) = refl -botₗ (suc n) f = suspMap n f - - -botᵣ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → (S₊∙ n →∙ Ω (∙Susp (typ A))) → S₊∙ (suc n) →∙ ∙Susp (typ A) -fst (botᵣ zero (f , p)) base = north -fst (botᵣ zero (f , p)) (loop i) = f false i -snd (botᵣ zero (f , p)) = refl -fst (botᵣ (suc n) (f , p)) north = north -fst (botᵣ (suc n) (f , p)) south = north -fst (botᵣ (suc n) (f , p)) (merid a i) = f a i -snd (botᵣ (suc n) (f , p)) = refl - -finalΔ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (x : (S₊∙ n →∙ A)) - → botₗ n x ≡ botᵣ {A = A} n (botMap n x) -finalΔ zero (f , p) = - ΣPathP ((funExt (λ { base → refl ; (loop i) → refl})) , refl) -finalΔ (suc n) (f , p) = - ΣPathP ((funExt (λ { north → refl ; south → refl ; (merid a i) → refl})) , refl) - -botᵣ-inv' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → S₊∙ (suc n) →∙ ∙Susp (typ A) → (S₊ n → typ (Ω (∙Susp (typ A)))) -botᵣ-inv' zero f false = sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f -botᵣ-inv' zero f true = refl -botᵣ-inv' (suc n) f x = sym (snd f) ∙∙ cong (fst f) (merid x ∙ sym (merid (ptSn (suc n)))) ∙∙ snd f - -botᵣ-inv : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → S₊∙ (suc n) →∙ ∙Susp (typ A) → (S₊∙ n →∙ Ω (∙Susp (typ A))) -fst (botᵣ-inv {A = A} n f) = botᵣ-inv' {A = A} n f -snd (botᵣ-inv {A = A} zero f) = refl -snd (botᵣ-inv {A = A} (suc n) f) = +botᵣ⁻' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → S₊∙ (suc n) →∙ Susp∙ (typ A) → (S₊ n → typ (Ω (Susp∙ (typ A)))) +botᵣ⁻' zero f false = + sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f +botᵣ⁻' zero f true = refl +botᵣ⁻' (suc n) f x = + sym (snd f) + ∙∙ cong (fst f) (merid x ∙ sym (merid (ptSn (suc n)))) + ∙∙ snd f + +botᵣ⁻ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → S₊∙ (suc n) →∙ Susp∙ (typ A) + → (S₊∙ n →∙ Ω (Susp∙ (typ A))) +fst (botᵣ⁻ {A = A} n f) = botᵣ⁻' {A = A} n f +snd (botᵣ⁻ {A = A} zero f) = refl +snd (botᵣ⁻ {A = A} (suc n) f) = cong (sym (snd f) ∙∙_∙∙ snd f) (cong (cong (fst f)) (rCancel (merid (ptSn _)))) ∙ ∙∙lCancel (snd f) - isEquiv-rMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → isEquiv (rMap n {A = A}) isEquiv-rMap zero = - compEquiv (isoToEquiv (flipΩIso zero)) (isoToEquiv (IsoΩSphereMap zero)) .snd + compEquiv (isoToEquiv (flipΩIso zero)) + (isoToEquiv (IsoΩSphereMap zero)) .snd isEquiv-rMap (suc n) = - compEquiv (isoToEquiv (flipΩIso (suc n))) (isoToEquiv (IsoΩSphereMap (suc n))) .snd - - -IsoΩSphereMap' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → Iso (typ ((Ω^ (suc n)) (∙Susp (typ A)))) ((S₊∙ (suc n) →∙ ∙Susp (typ A))) -IsoΩSphereMap' {A = A} n = - compIso (equivToIso (_ , isEquiv-rMap {A = A} n)) - (iso (botᵣ {A = A} n) (botᵣ-inv {A = A} n) (h n) (retr n)) + compEquiv (isoToEquiv (flipΩIso (suc n))) + (isoToEquiv (IsoΩSphereMap (suc n))) .snd + +-- botᵣ is an Iso +botᵣIso : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (S₊∙ n →∙ Ω (Susp∙ (typ A))) + (S₊∙ (suc n) →∙ Susp∙ (typ A)) +botᵣIso {A = A} n = (iso (botᵣ {A = A} n) (botᵣ⁻ {A = A} n) (h n) (retr n)) where - h : (n : ℕ) → section (botᵣ {A = A} n) (botᵣ-inv {A = A} n) + h : (n : ℕ) → section (botᵣ {A = A} n) (botᵣ⁻ {A = A} n) h zero (f , p) = ΣPathP (funExt (λ { base → sym p ; (loop i) j → doubleCompPath-filler (sym p) (cong f loop) p (~ j) i}) - , λ i j → p (~ i ∨ j)) + , λ i j → p (~ i ∨ j)) h (suc n) (f , p) = ΣPathP (funExt (λ { north → sym p ; south → sym p ∙ cong f (merid (ptSn _)) - ; (merid a i) j → hcomp (λ k → λ { (i = i0) → p (~ j ∧ k) - ; (i = i1) → compPath-filler' (sym p) (cong f (merid (ptSn _))) k j - ; (j = i1) → f (merid a i)}) - (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ j) i))}) + ; (merid a i) j + → hcomp (λ k + → λ { (i = i0) → p (~ j ∧ k) + ; (i = i1) → compPath-filler' + (sym p) (cong f (merid (ptSn _))) k j + ; (j = i1) → f (merid a i)}) + (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ j) i))}) , λ i j → p (~ i ∨ j)) - retr : (n : ℕ) → retract (botᵣ {A = A} n) (botᵣ-inv {A = A} n) - retr zero (f , p) = ΣPathP ((funExt (λ { false → sym (rUnit _) ; true → sym p})) + retr : (n : ℕ) → retract (botᵣ {A = A} n) (botᵣ⁻ {A = A} n) + retr zero (f , p) = + ΣPathP ((funExt (λ { false → sym (rUnit _) ; true → sym p})) , λ i j → p (~ i ∨ j)) retr (suc n) (f , p) = →∙Homogeneous≡ (isHomogeneousPath _ _) - (funExt λ x → (λ i → rUnit (cong-∙ (fst ((botᵣ {A = A}(suc n) (f , p)))) (merid x) (sym (merid (ptSn (suc n)))) i) (~ i)) + (funExt λ x → (λ i + → rUnit (cong-∙ (fst ((botᵣ {A = A}(suc n) (f , p)))) + (merid x) + (sym (merid (ptSn (suc n)))) i) (~ i)) ∙∙ (λ i → f x ∙ sym (p i) ) ∙∙ sym (rUnit (f x))) -Ω≡SphereMap' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → (typ ((Ω^ (suc n)) (∙Susp (typ A)))) ≡ ((S₊∙ (suc n) →∙ ∙Susp (typ A))) +IsoΩSphereMap' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (typ ((Ω^ (suc n)) (Susp∙ (typ A)))) + (S₊∙ (suc n) →∙ Susp∙ (typ A)) +IsoΩSphereMap' {A = A} n = + compIso (equivToIso (_ , isEquiv-rMap {A = A} n)) + (botᵣIso {A = A} n) + +Ω≡SphereMap' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (typ ((Ω^ (suc n)) (Susp∙ (typ A)))) + ≡ (S₊∙ (suc n) →∙ Susp∙ (typ A)) Ω≡SphereMap' {A = A} n = isoToPath (IsoΩSphereMap' {A = A} n) -open import Cubical.Data.Bool -theComp' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} - → rMap n {A = A} ∘ ((suspMap2 {A = A} n) .fst) +filler-top□ : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} + → rMap n {A = A} ∘ (suspMapΩ {A = A} n) ≡ botMap n ∘ lMap n {A = A} -theComp' {ℓ} = +nInd (λ {A} → funExt λ p → →∙Homogeneous≡ (isHomogeneousPath _ _) - (funExt λ { false → transportRefl _ - ; true → sym (rCancel _)})) - (λ {A} → funExt λ p → →∙Homogeneous≡ (isHomogeneousPath _ _) - (funExt λ x → (λ i → ((rMap≡rMap1 {A = A} i) ∘ suspMap2 1 .fst) p .fst x) ∙ sym (lem₁ p x))) +filler-top□ {ℓ} = + +nInd (λ {A} → funExt λ p → →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt λ { false → transportRefl _ + ; true → sym (rCancel _)})) + (λ {A} → funExt λ p → →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt λ x + → (λ i → ((rMap≡rMap1 {A = A} i) ∘ suspMapΩ 1) p .fst x) + ∙ sym (lem₁ p x))) cool* where lem₁ : ∀ {ℓ} {A : Pointed ℓ} (p : _) (x : S¹) - → (botMap 1 ∘ lMap 1) p .fst x ≡ (rMap1 {A = A} ∘ ((suspMap2 {A = A} 1) .fst)) p .fst x + → (botMap 1 ∘ lMap 1) p .fst x + ≡ (rMap1 {A = A} ∘ (suspMapΩ {A = A} 1)) p .fst x lem₁ {A = A} p base = rCancel (merid (pt A)) lem₁ {A = A} p (loop i) j = doubleCompPath-filler @@ -494,54 +608,64 @@ theComp' {ℓ} = +nInd (λ {A} → funExt λ p → →∙Homogeneous≡ (isHomog cool* : (n : ℕ) → ({A : Pointed ℓ} → - rMap (suc n) {A = A} ∘ suspMap2 (suc n) .fst ≡ + rMap (suc n) {A = A} ∘ suspMapΩ (suc n) ≡ botMap (suc n) ∘ lMap (suc n)) → {A : Pointed ℓ} → - rMap (suc (suc n)) {A = A} ∘ suspMap2 (suc (suc n)) .fst ≡ + rMap (suc (suc n)) {A = A} ∘ suspMapΩ (suc (suc n)) ≡ botMap (suc (suc n)) ∘ lMap (suc (suc n)) cool* n ind {A} = funExt λ p → →∙Homogeneous≡ (isHomogeneousPath (Susp (typ A)) refl) - (funExt λ { north → sym (rCancel (merid (pt A))) - ; south → sym (rCancel (merid (pt A))) - ; (merid a i) j → hcomp (λ k → λ {(i = i0) → rCancel (merid (pt A)) (~ j) - ; (i = i1) → rCancel (merid (pt A)) (~ j) - ; (j = i0) → help2 p a (~ k) i - ; (j = i1) → (botMap (suc (suc n)) ∘ lMap (suc (suc n))) p .fst - (merid a i)}) - (doubleCompPath-filler (sym (rCancel (merid (pt A)))) - (λ i → (botMap (suc (suc n)) ∘ lMap (suc (suc n))) p .fst - (merid a i)) - (rCancel (merid (pt A))) (~ j) i)}) + (funExt + λ { north → sym (rCancel (merid (pt A))) + ; south → sym (rCancel (merid (pt A))) + ; (merid a i) j + → hcomp (λ k + → λ {( i = i0) → rCancel (merid (pt A)) (~ j) + ; (i = i1) → rCancel (merid (pt A)) (~ j) + ; (j = i0) → main p a (~ k) i + ; (j = i1) → (botMap (suc (suc n)) ∘ lMap (suc (suc n))) p .fst + (merid a i)}) + (doubleCompPath-filler + (sym (rCancel (merid (pt A)))) + (cong ((botMap (suc (suc n)) ∘ lMap (suc (suc n))) p .fst) + (merid a)) + (rCancel (merid (pt A))) (~ j) i)}) where module _ (p : typ (Ω _)) (a : S₊ (suc n)) where abstract indHyp : Path ((a₁ : fst (Ω ((Ω^ n) A))) → - Σ (fst (S₊∙ (suc n)) → fst (Ω (∙Susp (typ A)))) - (λ f → f (snd (S₊∙ (suc n))) ≡ snd (Ω (∙Susp (typ A))))) - ((rMap (suc n) {A = A} ∘ ((suspMap2 {A = A} (suc n)) .fst))) + Σ (fst (S₊∙ (suc n)) → fst (Ω (Susp∙ (typ A)))) + (λ f → f (snd (S₊∙ (suc n))) ≡ snd (Ω (Susp∙ (typ A))))) + ((rMap (suc n) {A = A} ∘ (suspMapΩ {A = A} (suc n)))) (botMap (suc n) ∘ lMap (suc n) {A = A}) - indHyp = funExt λ a → →∙Homogeneous≡ (isHomogeneousPath _ _) (funExt (λ x → funExt⁻ (cong fst (funExt⁻ (ind {A = A}) a)) x)) + indHyp = + funExt λ a → →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt (λ x → funExt⁻ (cong fst (funExt⁻ (ind {A = A}) a)) x)) - pullPath : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) (r : z ≡ z) + pathLem₁ : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) (r : z ≡ z) → p ∙ (q ∙ r ∙ sym q) ∙ sym p ≡ (p ∙ q) ∙ r ∙ sym (p ∙ q) - pullPath p q r = + pathLem₁ p q r = ∙assoc p (q ∙ r ∙ sym q) (sym p) ∙ cong (_∙ sym p) (∙assoc p q (r ∙ sym q)) ∙ sym (∙assoc (p ∙ q) (r ∙ sym q) (sym p)) ∙ cong ((p ∙ q) ∙_) (sym (∙assoc r (sym q) (sym p)) ∙ cong (r ∙_) (sym (symDistr p q))) - appCoherence : ∀ {ℓ} {A : Type ℓ} {x y : A} (coh : x ≡ y) (p : x ≡ x) → p ≡ coh ∙ (sym coh ∙∙ p ∙∙ coh) ∙ sym coh - appCoherence {x = x} = J (λ y coh → (p : x ≡ x) → p ≡ coh ∙ (sym coh ∙∙ p ∙∙ coh) ∙ sym coh) - λ p → lUnit _ ∙ cong (refl ∙_) (rUnit _ ∙ cong (_∙ refl) (rUnit p)) + pathLem₂ : ∀ {ℓ} {A : Type ℓ} {x y : A} (coh : x ≡ y) (p : x ≡ x) + → p ≡ coh ∙ (sym coh ∙∙ p ∙∙ coh) ∙ sym coh + pathLem₂ {x = x} = + J (λ y coh → (p : x ≡ x) → p ≡ coh ∙ (sym coh ∙∙ p ∙∙ coh) ∙ sym coh) + λ p → lUnit _ ∙ cong (refl ∙_) (rUnit _ ∙ cong (_∙ refl) (rUnit p)) - coolLem : ∀ {ℓ} {A : Type ℓ} {x₀ x y z w : A} - (p : x₀ ≡ x) (q : x ≡ y) (r : y ≡ z) (s : z ≡ w) (mid : w ≡ w) - (coh : x₀ ≡ w) + pathLem₃ : ∀ {ℓ} {A : Type ℓ} {x₀ x y z w : A} + (p : x₀ ≡ x) (q : x ≡ y) (r : y ≡ z) (s : z ≡ w) (mid : w ≡ w) + (coh : x₀ ≡ w) → isComm∙ (A , x₀) - → (p ∙∙ q ∙∙ r ∙∙ s ∙∙ mid ∙∙ sym s ∙∙ sym r ∙∙ sym q ∙∙ sym p) + → (p ∙∙ q ∙∙ r + ∙∙ s ∙∙ mid ∙∙ sym s + ∙∙ sym r ∙∙ sym q ∙∙ sym p) ≡ (coh ∙∙ mid ∙∙ sym coh) - coolLem p q r s mid coh comm = + pathLem₃ p q r s mid coh comm = doubleCompPath≡compPath p _ (sym p) ∙∙ cong (λ x → p ∙ x ∙ (sym p)) (doubleCompPath≡compPath q _ (sym q) @@ -549,27 +673,45 @@ theComp' {ℓ} = +nInd (λ {A} → funExt λ p → →∙Homogeneous≡ (isHomog (doubleCompPath≡compPath r _ (sym r) ∙ cong (λ x → r ∙ x ∙ (sym r)) (doubleCompPath≡compPath s _ (sym s) - ∙ cong (λ x → s ∙ x ∙ sym s) (appCoherence (sym coh) mid) - ∙ pullPath s (sym coh) (coh ∙∙ mid ∙∙ sym coh)) - ∙ pullPath r (s ∙ sym coh) (coh ∙∙ mid ∙∙ sym coh)) - ∙ pullPath q (r ∙ s ∙ sym coh) (coh ∙∙ mid ∙∙ sym coh)) - ∙∙ pullPath p (q ∙ r ∙ s ∙ sym coh) (coh ∙∙ mid ∙∙ sym coh) - ∙∙ cong ((p ∙ q ∙ r ∙ s ∙ sym coh) ∙_) (comm (coh ∙∙ mid ∙∙ sym coh) (sym (p ∙ q ∙ r ∙ s ∙ sym coh))) + ∙ cong (λ x → s ∙ x ∙ sym s) (pathLem₂ (sym coh) mid) + ∙ pathLem₁ s (sym coh) (coh ∙∙ mid ∙∙ sym coh)) + ∙ pathLem₁ r (s ∙ sym coh) (coh ∙∙ mid ∙∙ sym coh)) + ∙ pathLem₁ q (r ∙ s ∙ sym coh) (coh ∙∙ mid ∙∙ sym coh)) + ∙∙ pathLem₁ p (q ∙ r ∙ s ∙ sym coh) (coh ∙∙ mid ∙∙ sym coh) + ∙∙ cong ((p ∙ q ∙ r ∙ s ∙ sym coh) ∙_) + (comm (coh ∙∙ mid ∙∙ sym coh) (sym (p ∙ q ∙ r ∙ s ∙ sym coh))) ∙∙ ∙assoc _ _ _ ∙∙ cong (_∙ (coh ∙∙ mid ∙∙ sym coh)) (rCancel _) ∙∙ sym (lUnit _) - - cong-lMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (x : _) (p : refl ≡ x) (q : _ ≡ _) (a : S₊ (suc n)) - → (cong (lMap (suc n) {A = Ω A}) (fun (flipΩIso (suc (suc n))) (p ∙∙ q ∙∙ sym p))) + pathLem₄ : ∀ {ℓ} {A : Type ℓ} {x y z : A} + (p : z ≡ x) (q : y ≡ z) (r : y ≡ y) + → (sym (q ∙ p) ∙∙ r ∙∙ (q ∙ p)) + ≡ (sym p ∙∙ sym q ∙∙ r ∙∙ q ∙∙ p) + pathLem₄ p q r = + cong (λ x → x ∙∙ r ∙∙ (q ∙ p)) (symDistr q p) + ∙∙ doubleCompPath≡compPath (sym p ∙ sym q) r (q ∙ p) + ∙∙ (sym (∙assoc (sym p) (sym q) (r ∙ q ∙ p)) + ∙∙ cong (sym p ∙_) (∙assoc (sym q) r (q ∙ p) + ∙ ∙assoc (sym q ∙ r) q p) + ∙∙ sym (doubleCompPath≡compPath (sym p) ((sym q ∙ r) ∙ q) p) + ∙ cong (sym p ∙∙_∙∙ p) + (sym (∙assoc (sym q) r q) + ∙ sym (doubleCompPath≡compPath (sym q) r q))) + + cong-lMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (x : _) + (p : refl ≡ x) (q : _ ≡ _) (a : S₊ (suc n)) + → (cong (lMap (suc n) {A = Ω A}) + (fun (flipΩIso (suc (suc n))) (p ∙∙ q ∙∙ sym p))) ≡ (cong (lMap (suc n)) (sym (flipΩrefl n)) ∙∙ cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) p ∙∙ cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) q ∙∙ sym (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) p) ∙∙ cong (lMap (suc n)) (flipΩrefl n)) - cong-lMap n x = J (λ x p → (q : x ≡ x) → - S₊ (suc n) → - cong (lMap (suc n)) + cong-lMap n x = + J (λ x p → (q : x ≡ x) + → S₊ (suc n) + → cong (lMap (suc n)) (fun (flipΩIso (suc (suc n))) (p ∙∙ q ∙∙ sym p)) ≡ (cong (lMap (suc n)) (sym (flipΩrefl n)) ∙∙ @@ -577,110 +719,599 @@ theComp' {ℓ} = +nInd (λ {A} → funExt λ p → →∙Homogeneous≡ (isHomog cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) q ∙∙ sym (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) p) ∙∙ cong (lMap (suc n)) (flipΩrefl n))) - λ q a → (λ j i → lMap (suc n) - (fun (flipΩIso (suc (suc n))) (rUnit q (~ j)) - i)) ∙ cong-lMap-main' n q - ∙ cong (cong (lMap (suc n)) (sym (flipΩrefl n)) ∙∙_∙∙ - cong (lMap (suc n)) (flipΩrefl n)) (rUnit (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) q)) - - help2 : cong ((rMap (suc (suc n)) {A = A} ∘ suspMap2 (suc (suc n)) .fst) p .fst) (merid a) - ≡ (sym (rCancel (merid (pt A))) ∙∙ (λ i → (botMap (suc (suc n)) ∘ lMap (suc (suc n)) {A = A}) p .fst (merid a i)) ∙∙ rCancel _) - help2 = ((λ _ → (sym (lMapId (suc n) a) ∙∙ cong (λ x → fst x a) - (cong (lMap (suc n)) - (fun (flipΩIso (suc (suc n))) - (sym (∙∙lCancel (snd (suspMap2 n))) - ∙∙ cong (sym (snd (suspMap2 n)) ∙∙_∙∙ snd (suspMap2 n)) (cong (cong (fst (suspMap2 n))) p) - ∙∙ ∙∙lCancel (snd (suspMap2 n))))) - ∙∙ lMapId (suc n) a))) + λ q a → (λ j i → lMap (suc n) + (fun (flipΩIso (suc (suc n))) (rUnit q (~ j)) i)) + ∙∙ cong-lMap-lem n q + ∙∙ cong (cong (lMap (suc n)) (sym (flipΩrefl n)) + ∙∙_∙∙ + cong (lMap (suc n)) (flipΩrefl n)) + (rUnit (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) q)) + + main : cong ((rMap (suc (suc n)) {A = A} + ∘ suspMapΩ (suc (suc n))) p .fst) (merid a) + ≡ (sym (rCancel (merid (pt A))) + ∙∙ (λ i → (botMap (suc (suc n)) ∘ lMap (suc (suc n)) {A = A}) + p .fst (merid a i)) + ∙∙ rCancel _) + main = cong (sym (lMapId (suc n) a) ∙∙_∙∙ (lMapId (suc n) a)) + (cong (cong (λ x → fst x a)) + (cong-lMap _ _ + (sym (∙∙lCancel (snd (suspMapΩ∙ n)))) + (cong (suspMapΩ (suc n)) p) a)) ∙∙ cong (sym (lMapId (suc n) a) ∙∙_∙∙ (lMapId (suc n) a)) - (cong (cong (λ x → fst x a)) (cong-lMap _ _ (sym (∙∙lCancel (snd (suspMap2 n)))) (cong ((suspMap2 (suc n)) .fst) p) a)) + ((cong-∙∙ (λ x → fst x a) + (cong (lMap (suc n)) (sym (flipΩrefl n))) + (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) + (sym (∙∙lCancel (snd (suspMapΩ∙ n)))) + ∙∙ cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) + (cong ((suspMapΩ∙ (suc n)) .fst) p) + ∙∙ cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) + (∙∙lCancel (snd (suspMapΩ∙ n)))) + (cong (lMap (suc n)) (flipΩrefl n)))) ∙∙ cong (sym (lMapId (suc n) a) ∙∙_∙∙ (lMapId (suc n) a)) - ((cong-∙∙ (λ x → fst x a) (cong (lMap (suc n)) (sym (flipΩrefl n))) - (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) (sym (∙∙lCancel (snd (suspMap2 n)))) - ∙∙ cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) ((cong ((suspMap2 (suc n)) .fst) p)) - ∙∙ cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) (∙∙lCancel (snd (suspMap2 n)))) - (cong (lMap (suc n)) (flipΩrefl n)))) - ∙∙ cong (sym (lMapId (suc n) a) ∙∙_∙∙ (lMapId (suc n) a)) - (cong (cong (λ x → fst x a) (cong (lMap (suc n)) (sym (flipΩrefl n))) + (cong (cong (λ x → fst x a) + (cong (lMap (suc n)) (sym (flipΩrefl n))) ∙∙_∙∙ - cong (λ x → fst x a) (cong (lMap (suc n)) (flipΩrefl n))) + cong (λ x → fst x a) + (cong (lMap (suc n)) (flipΩrefl n))) (cong-∙∙ (λ x → fst x a) - (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) (sym (∙∙lCancel (snd (suspMap2 n))))) - (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) (cong ((suspMap2 (suc n)) .fst) p)) - (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) (∙∙lCancel (snd (suspMap2 n)))))) + (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) + (sym (∙∙lCancel (snd (suspMapΩ∙ n))))) + (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) + (cong ((suspMapΩ∙ (suc n)) .fst) p)) + (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) + (∙∙lCancel (snd (suspMapΩ∙ n)))))) ∙∙ cong (sym (lMapId (suc n) a) ∙∙_∙∙ (lMapId (suc n) a)) - (cong (cong (λ x → fst x a) (cong (lMap (suc n)) (sym (flipΩrefl n))) - ∙∙_∙∙ - cong (λ x → fst x a) (cong (lMap (suc n)) (flipΩrefl n))) - (cong (cong (λ x → fst x a) (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) (sym (∙∙lCancel (snd (suspMap2 n))))) - ∙∙_∙∙ - cong (λ x → fst x a) (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) (∙∙lCancel (snd (suspMap2 n))))) - ((λ i j → rMap (suc n) {A = A} (suspMap2 (suc n) .fst (p j)) .fst a) - ∙ rUnit _ - ∙ λ i → (λ j → indHyp (i ∧ j) (snd (Ω ((Ω^ n) A))) .fst a) - ∙∙ (λ j → indHyp i (p j) .fst a) - ∙∙ λ j → indHyp (i ∧ ~ j) (snd (Ω ((Ω^ n) A))) .fst a))) - ∙∙ coolLem (sym (lMapId (suc n) a)) (λ i₁ → fst (lMap (suc n) (flipΩrefl n (~ i₁))) a) - ((λ i₁ → - fst - (lMap (suc n) - (fun (flipΩIso (suc n)) (∙∙lCancel (snd (suspMap2 {A = A} n)) (~ i₁)))) - a)) (λ j₁ → indHyp j₁ (snd (Ω ((Ω^ n) A))) .fst a) (λ j₁ → (botMap (suc n) ∘ lMap (suc n)) (p j₁) .fst a) - (sym (cong (_∙ sym (merid (pt A))) (cong merid (lMapId (suc n) a)) ∙ rCancel _)) - (EH 0) - ∙∙ helpLem ((rCancel (merid (pt A)))) - (cong (_∙ sym (merid (pt A))) (cong merid (lMapId (suc n) a))) + (cong (cong (λ x → fst x a) + (cong (lMap (suc n)) (sym (flipΩrefl n))) + ∙∙_∙∙ + cong (λ x → fst x a) + (cong (lMap (suc n)) (flipΩrefl n))) + (cong (cong (λ x → fst x a) + (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) + (sym (∙∙lCancel (snd (suspMapΩ∙ n))))) + ∙∙_∙∙ + cong (λ x → fst x a) + (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) + (∙∙lCancel (snd (suspMapΩ∙ n))))) + ((λ i j → rMap (suc n) {A = A} + (suspMapΩ∙ (suc n) .fst (p j)) .fst a) + ∙ rUnit _ + ∙ λ i → (λ j → indHyp (i ∧ j) + (snd (Ω ((Ω^ n) A))) .fst a) + ∙∙ (λ j → indHyp i (p j) .fst a) + ∙∙ λ j → indHyp (i ∧ ~ j) + (snd (Ω ((Ω^ n) A))) .fst a))) + ∙∙ pathLem₃ + (sym (lMapId (suc n) a)) + (λ i₁ → fst (lMap (suc n) (flipΩrefl n (~ i₁))) a) + (λ i₁ → fst (lMap (suc n) (fun (flipΩIso (suc n)) + (∙∙lCancel (snd (suspMapΩ∙ {A = A} n)) (~ i₁)))) a) + (λ j₁ → indHyp j₁ (snd (Ω ((Ω^ n) A))) .fst a) + (λ j₁ → (botMap (suc n) ∘ lMap (suc n)) (p j₁) .fst a) + (sym (cong (_∙ sym (merid (pt A))) + (cong merid (lMapId (suc n) a)) ∙ rCancel _)) + (EH 0) + ∙∙ pathLem₄ (rCancel (merid (pt A))) + (cong (_∙ sym (merid (pt A))) + (cong merid (lMapId (suc n) a))) (λ j₁ → (botMap (suc n) ∘ lMap (suc n)) (p j₁) .fst a) ∙∙ cong (sym (rCancel (merid (pt A))) ∙∙_∙∙ rCancel (merid (pt A))) - (sym (cong-∙∙ (λ x → merid x ∙ sym (merid (pt A))) (sym (lMapId (suc n) a)) (λ i → lMap (suc n) (p i) .fst a) (lMapId (suc n) a))) - ∙∙ λ _ → sym (rCancel (merid (pt A))) - ∙∙ (λ i → merid ((sym (lMapId (suc n) a) ∙∙ (λ i → lMap (suc n) (p i) .fst a) ∙∙ lMapId (suc n) a) i) ∙ sym (merid (pt A))) - ∙∙ rCancel (merid (pt A)) - where - helpLem : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : z ≡ x) (q : y ≡ z) (r : y ≡ y) - → (sym (q ∙ p) ∙∙ r ∙∙ (q ∙ p)) - ≡ (sym p ∙∙ sym q ∙∙ r ∙∙ q ∙∙ p) - helpLem p q r = - cong (λ x → x ∙∙ r ∙∙ (q ∙ p)) (symDistr q p) - ∙∙ doubleCompPath≡compPath (sym p ∙ sym q) r (q ∙ p) - ∙∙ (sym (∙assoc (sym p) (sym q) (r ∙ q ∙ p)) - ∙∙ cong (sym p ∙_) (∙assoc (sym q) r (q ∙ p) ∙ ∙assoc (sym q ∙ r) q p) - ∙∙ sym (doubleCompPath≡compPath (sym p) ((sym q ∙ r) ∙ q) p) - ∙ cong (sym p ∙∙_∙∙ p) (sym (∙assoc (sym q) r q) ∙ sym (doubleCompPath≡compPath (sym q) r q))) - -{- -Ω^ - - --} + (sym (cong-∙∙ (λ x → merid x ∙ sym (merid (pt A))) + (sym (lMapId (suc n) a)) + (λ i → lMap (suc n) (p i) .fst a) + (lMapId (suc n) a))) suspMap→TranspType : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → (typ (Ω ((Ω^ n) A)) → typ (Ω (Ω ((Ω^ n) (∙Susp (typ A)))))) - ≡ ((S₊∙ (suc n) →∙ A) → (S₊∙ (suc (suc n)) →∙ ∙Susp (typ A))) + → (typ (Ω ((Ω^ n) A)) → typ (Ω (Ω ((Ω^ n) (Susp∙ (typ A)))))) + ≡ ((S₊∙ (suc n) →∙ A) → (S₊∙ (suc (suc n)) →∙ Susp∙ (typ A))) suspMap→TranspType {A = A} n i = Ω≡SphereMap {A = A} (suc n) i → Ω≡SphereMap' {A = A} (suc n) i suspMap→ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → PathP (λ i → suspMap→TranspType {A = A} n i) - (suspMap2 (suc n) .fst) + (suspMapΩ∙ (suc n) .fst) (suspMap n) suspMap→ {A = A} n = toPathP (funExt λ f → (λ j → transportRefl (botᵣ {A = A} (suc n) (rMap (suc n) {A = A} - (suspMap2 (suc n) .fst + (suspMapΩ∙ (suc n) .fst ((invEq (_ , isEquiv-lMap n) (transportRefl f j)))))) j) ∙∙ cong (botᵣ {A = A} (suc n)) - (funExt⁻ (theComp' (suc n)) (invEq (_ , isEquiv-lMap n) f)) - ∙∙ sym (finalΔ (suc n) (lMap (suc n) {A = A} (invEq (lMap (suc n) , isEquiv-lMap n) f))) + (funExt⁻ (filler-top□ (suc n)) (invEq (_ , isEquiv-lMap n) f)) + ∙∙ sym (filler▿ (suc n) (lMap (suc n) {A = A} (invEq (lMap (suc n) , isEquiv-lMap n) f))) ∙ cong (suspMap n) (secEq ((lMap (suc n) , isEquiv-lMap n)) f)) -isConnectedPres : ∀ {ℓ} {A : Pointed ℓ} (con n : ℕ) - → isConnectedFun con (suspMap2 {A = A} (suc n) .fst) - → isConnectedFun con (suspMap {A = A} n) -isConnectedPres {A = A} con n hyp = - transport (λ i → isConnectedFun con (suspMap→ {A = A} n i)) hyp +{- +We would like for the above dependent path to preserve composition +(in order to deduce properties of suspMap from the corresponding +ones of suspMapΩ). For this, we need that both Ω≡SphereMap and +Ω≡SphereMap' are structure preserving +-} +-- We define an alternative notion of composition (to ∙Π) on +-- Sⁿ →∙ Ω (Susp∙ (typ A)). It will be easier to work with +private + invComp : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → S₊∙ n →∙ Ω (Susp∙ (typ A)) + → S₊∙ n →∙ Ω (Susp∙ (typ A)) + → S₊∙ n →∙ Ω (Susp∙ (typ A)) + fst (invComp n f g) x = (fst f x) ∙ (fst g x) + snd (invComp n f g) = cong₂ _∙_ (snd f) (snd g) ∙ sym (rUnit refl) + + -- We prove that it agrees with ∙Π + ∙Π≡invComp : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (f g : S₊∙ (suc n) →∙ Ω (Susp∙ (typ A))) + → ∙Π f g ≡ invComp {A = A} (suc n) f g + ∙Π≡invComp zero f g = + →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt (λ { base → rUnit refl + ∙ sym (cong (_∙ fst g (snd (S₊∙ 1))) (snd f) + ∙ cong (refl ∙_) (snd g)) + ; (loop i) j → + hcomp (λ k → + λ { (i = i0) → (rUnit refl + ∙ sym (cong (_∙ fst g (snd (S₊∙ 1))) (snd f) + ∙ cong (refl ∙_) (snd g))) j + ; (i = i1) → (rUnit refl + ∙ sym (cong (_∙ fst g (snd (S₊∙ 1))) (snd f) + ∙ cong (refl ∙_) (snd g))) j + ; (j = i0) → ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) + ∙ (sym (snd g) ∙∙ cong (fst g) loop ∙∙ snd g)) i + ; (j = i1) → cong₂Funct _∙_ + (cong (fst f) loop) (cong (fst g) loop) (~ k) i}) + (hcomp (λ k → + λ { (i = i0) → (rUnit refl + ∙ sym (cong (_∙ snd g (~ k)) (λ j → snd f (j ∨ ~ k)) + ∙ cong (refl ∙_) (λ j → snd g (j ∨ ~ k)))) j + ; (i = i1) → (rUnit refl + ∙ sym (cong (_∙ snd g (~ k)) (λ j → snd f (j ∨ ~ k)) + ∙ cong (refl ∙_) (λ j → snd g (j ∨ ~ k)))) j + ; (j = i0) → ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) + ∙ (sym (snd g) ∙∙ cong (fst g) loop ∙∙ snd g)) i + ; (j = i1) → (cong (_∙ snd g (~ k)) + (doubleCompPath-filler (sym (snd f)) + (cong (fst f) loop) + (snd f) (~ k)) ∙ + cong ((snd f (~ k)) ∙_) + (doubleCompPath-filler (sym (snd g)) + (cong (fst g) loop) (snd g) (~ k))) i}) + (hcomp (λ k → + λ { (i = i0) → (rUnit (rUnit refl) + ∙ cong (rUnit refl ∙_) (cong sym (rUnit refl))) k j + ; (i = i1) → (rUnit (rUnit refl) + ∙ cong (rUnit refl ∙_) (cong sym (rUnit refl))) k j + ; (j = i0) → ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) + ∙ (sym (snd g) ∙∙ cong (fst g) loop ∙∙ snd g)) i + ; (j = i1) → (cong (_∙ refl) + ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) + ∙ cong (refl ∙_) + (sym (snd g) ∙∙ cong (fst g) loop ∙∙ snd g)) i}) + ((cong (λ x → rUnit x j) + (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) + ∙ cong (λ x → lUnit x j) + (sym (snd g) ∙∙ cong (fst g) loop ∙∙ snd g)) i)))})) + ∙Π≡invComp {A = A} (suc n) f g = + →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt λ { north → rUnit refl + ∙ sym (cong (fst f north ∙_) (snd g) + ∙ cong (_∙ refl) (snd f)) + ; south → rUnit refl + ∙ sym (cong₂ _∙_ + (cong (fst f) (sym (merid (ptSn _))) ∙ snd f) + (cong (fst g) (sym (merid (ptSn _))) ∙ snd g)) + ; (merid a i) j → p a i j}) + where + module _ (a : S₊ (suc n)) where + f-n = fst f north + g-n = fst g north + cong-f = (cong (fst f) (merid a ∙ sym (merid (ptSn _)))) + cong-g = (cong (fst g) (merid a ∙ sym (merid (ptSn _)))) + c-f = sym (snd f) ∙∙ cong-f ∙∙ snd f + c-g = sym (snd g) ∙∙ cong-g ∙∙ snd g + + p : I → I → fst (Ω (Susp∙ (typ A))) + p i j = + hcomp (λ k → + λ { (i = i0) → (rUnit (λ _ → snd (Susp∙ (typ A))) + ∙ sym ((cong (fst f north ∙_) (snd g) + ∙ cong (_∙ refl) (snd f)))) j + ; (i = i1) → (rUnit refl + ∙ sym (cong₂ _∙_ + (compPath-filler' + (cong (fst f) (sym (merid (ptSn _)))) (snd f) k) + (compPath-filler' + (cong (fst g) (sym (merid (ptSn _)))) (snd g) k))) j + ; (j = i0) → (c-f ∙ c-g) i + ; (j = i1) → fst f (compPath-filler + (merid a) + (sym (merid (ptSn _))) (~ k) i) + ∙ fst g (compPath-filler + (merid a) + (sym (merid (ptSn _))) (~ k) i)}) + (hcomp (λ k → + λ {(i = i0) → (rUnit (λ _ → snd (Susp∙ (typ A))) + ∙ sym ((cong (fst f north ∙_) (snd g) + ∙ cong (_∙ refl) (snd f)))) j + ; (i = i1) → (rUnit refl ∙ sym (cong₂ _∙_ (snd f) (snd g))) j + ; (j = i0) → (c-f ∙ c-g) i + ; (j = i1) → cong₂Funct _∙_ cong-f cong-g (~ k) i}) + (hcomp (λ k → + λ {(i = i0) → (rUnit refl + ∙ sym (compPath-filler' + ((cong (fst f north ∙_) (snd g))) + (cong (_∙ refl) (snd f)) k)) j + ; (i = i1) → (rUnit refl + ∙ sym (cong₂ _∙_ (λ i → snd f (i ∨ ~ k)) (snd g))) j + ; (j = i0) → (c-f ∙ c-g) i + ; (j = i1) → (cong (λ x → x ∙ snd g (~ k)) + (doubleCompPath-filler refl + cong-f (snd f) (~ k)) + ∙ cong ((snd f (~ k)) ∙_) + (doubleCompPath-filler + (sym (snd g)) cong-g refl (~ k))) i}) + (hcomp (λ k → + λ {(i = i0) → compPath-filler + (rUnit (λ _ → snd (Susp∙ (typ A)))) + (sym ((cong (_∙ refl) (snd f)))) k j + ; (i = i1) → compPath-filler + (rUnit refl) + (sym (cong (refl ∙_) (snd g))) k j + ; (j = i0) → (c-f ∙ c-g) i + ; (j = i1) → (cong (_∙ refl) + ((λ j → snd f (~ j ∧ ~ k)) ∙∙ cong-f ∙∙ snd f) + ∙ cong (refl ∙_) + (sym (snd g) ∙∙ cong-g ∙∙ (λ j → snd g (j ∧ ~ k)))) i}) + (((cong (λ x → rUnit x j) c-f) ∙ (cong (λ x → lUnit x j) c-g)) i)))) + + hom-botᵣ⁻ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (f g : S₊∙ (suc n) →∙ Susp∙ (typ A)) + → botᵣ⁻ {A = A} n (∙Π f g) + ≡ invComp {A = A} n (botᵣ⁻ {A = A} n f) (botᵣ⁻ {A = A} n g) + +-- We ge that +hom-botᵣ⁻ zero f g = + ΣPathP ((funExt (λ { false → sym (rUnit _) + ; true → (rUnit _)})) + , ((λ i j → rUnit refl (i ∧ ~ j)) + ▷ lUnit (sym (rUnit refl)))) +hom-botᵣ⁻ (suc n) f g = + →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt (λ x → (sym (rUnit (cong (fst (∙Π f g)) (merid x ∙ sym (merid (ptSn _)))))) + ∙∙ cong-∙ (fst (∙Π f g)) (merid x) (sym (merid (ptSn _))) + ∙∙ cong (cong (fst (∙Π f g)) (merid x) ∙_) (cong sym lem) + ∙ sym (rUnit (cong (fst (∙Π f g)) (merid x))))) + where + lem : cong (fst (∙Π f g)) (merid (ptSn (suc n))) ≡ refl + lem = (λ i → (sym (snd f) ∙∙ cong (fst f) (rCancel (merid (ptSn _)) i) ∙∙ snd f) + ∙ (sym (snd g) ∙∙ cong (fst g) (rCancel (merid (ptSn _)) i) ∙∙ snd g)) + ∙ (λ i → ∙∙lCancel (snd f) i ∙ ∙∙lCancel (snd g) i) + ∙ sym (rUnit refl) + +-- We get that botᵣ⁻ (and hence botᵣ) is homomorphism +hom-botᵣ⁻' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (f g : S₊∙ (suc (suc n)) →∙ Susp∙ (typ A)) + → botᵣ⁻ {A = A} (suc n) (∙Π f g) + ≡ ∙Π (botᵣ⁻ {A = A} (suc n) f) (botᵣ⁻ {A = A} (suc n) g) +hom-botᵣ⁻' {A = A} n f g = + hom-botᵣ⁻ {A = A} (suc n) f g + ∙ sym (∙Π≡invComp {A = A} _ (botᵣ⁻ {A = A} _ f) (botᵣ⁻ {A = A} _ g)) + +isHom-lMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (p q : _) + → lMap (suc n) {A = A} (p ∙ q) + ≡ ∙Π (lMap (suc n) p) (lMap (suc n) q) +isHom-lMap zero p q = + ΣPathP ((funExt (λ { base → refl + ; (loop i) j → (rUnit p j ∙ rUnit q j) i})) + , refl) +isHom-lMap (suc n) {A = A} p q = + ΣPathP ((funExt (λ { north → refl + ; south → refl + ; (merid a i) j → h a j i})) + , refl) + where + doubleComp-lem : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) (q r : y ≡ y) + → (p ∙∙ q ∙∙ sym p) ∙ (p ∙∙ r ∙∙ sym p) + ≡ (p ∙∙ (q ∙ r) ∙∙ sym p) + doubleComp-lem p q r i j = + hcomp (λ k → λ { (i = i0) → (doubleCompPath-filler p q (sym p) k + ∙ doubleCompPath-filler p r (sym p) k) j + ; (i = i1) → doubleCompPath-filler p (q ∙ r) (sym p) k j + ; (j = i0) → p (~ k) + ; (j = i1) → p (~ k)}) + ((q ∙ r) j) -isConnectedSuspMap2 : ? -isConnectedSuspMap2 = ? + help2 : (p : typ ((Ω^ (suc (suc n))) A)) + → cong (fst (lMap (suc (suc n)) p)) (merid (ptSn _)) ≡ refl + help2 p = + cong (sym (lMapId (suc n) (ptSn _)) ∙∙_∙∙ lMapId (suc n) (ptSn _)) + (rUnit _ ∙ (λ j → (λ i → lMap (suc n) {A = A} refl .snd (i ∧ j)) + ∙∙ (λ i → lMap (suc n) {A = A} (p i) .snd j) + ∙∙ λ i → lMap (suc n) {A = A} refl .snd (~ i ∧ j)) + ∙ ∙∙lCancel _) + ∙ ∙∙lCancel _ + + h : (a : S₊ (suc n)) + → sym (lMapId (suc n) a) + ∙∙ funExt⁻ (cong fst (cong (lMap (suc n)) (p ∙ q))) a + ∙∙ lMapId (suc n) a + ≡ cong (fst (∙Π (lMap (suc (suc n)) p) (lMap (suc (suc n)) q))) (merid a) + h a = (cong (sym (lMapId (suc n) a) ∙∙_∙∙ (lMapId (suc n) a)) + (cong-∙ (λ x → lMap (suc n) x .fst a) p q) + ∙ sym (doubleComp-lem (sym (lMapId (suc n) a)) _ _)) + ∙∙ cong₂ _∙_ (sym (cong (cong (fst (lMap (suc (suc n)) p)) (merid a) ∙_) + (cong sym (help2 p)) ∙ sym (rUnit _))) + (sym (cong (cong (fst (lMap (suc (suc n)) q)) (merid a) ∙_) + (cong sym (help2 q)) ∙ sym (rUnit _))) + ∙∙ λ i → (rUnit (cong-∙ (fst (lMap (suc (suc n)) p)) + (merid a) (sym (merid (ptSn _))) (~ i)) i) + ∙ (rUnit (cong-∙ (fst (lMap (suc (suc n)) q)) + (merid a) (sym (merid (ptSn _)))(~ i)) i) + +hom-rMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (p q : typ ((Ω^ (suc (suc n))) (Susp∙ (typ A)))) + → rMap (suc n) {A = A} (p ∙ q) + ≡ ∙Π (rMap (suc n) {A = A} p) (rMap (suc n) {A = A} q) +hom-rMap {A = A} n p q = + cong (lMap (suc n) {A = Ω (Susp∙ (typ A))}) + (morphLemmas.isMorphInv _∙_ _∙_ + (Iso.inv (flipΩIso (suc n))) (flipΩIsopres· n) + (Iso.fun (flipΩIso (suc n))) + (Iso.leftInv (flipΩIso (suc n))) (Iso.rightInv (flipΩIso (suc n))) + p q) + ∙ isHom-lMap _ (fun (flipΩIso (suc n)) p) (fun (flipΩIso (suc n)) q) + +isHom-IsoΩSphereMap' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + (p q : typ (Ω ((Ω^ n) (Susp∙ (typ A))))) + → Iso.fun (IsoΩSphereMap' {A = A} n) (p ∙ q) + ≡ ∙Π {n = (suc n)} (Iso.fun (IsoΩSphereMap' {A = A} n) p) + (Iso.fun (IsoΩSphereMap' {A = A} n) q) +isHom-IsoΩSphereMap' {A = A} zero p q = + (cong (botᵣ {A = A} 0) + (cong ((lMap 0 {A = Ω (Susp∙ (typ A))})) + (transportRefl (p ∙ q)))) + ∙ ΣPathP (funExt (λ { base → refl + ; (loop i) j + → lem j i}) + , refl) + where + lem : p ∙ q ≡ (cong (fst (fun (IsoΩSphereMap' {A = A} zero) p)) loop ∙ refl) + ∙ (cong (fst (fun (IsoΩSphereMap' {A = A} zero) q)) loop ∙ refl) + lem = cong₂ _∙_ + (sym (transportRefl p) + ∙ rUnit (cong (fst (fun (IsoΩSphereMap' {A = A} zero) p)) loop)) + (sym (transportRefl q) + ∙ rUnit (cong (fst (fun (IsoΩSphereMap' {A = A} zero) q)) loop)) +isHom-IsoΩSphereMap' {A = A} (suc n) p q = + cong (botᵣ {A = A} (suc n) ∘ lMap (suc n) {A = Ω (Susp∙ (typ A))}) + (morphLemmas.isMorphInv + _∙_ _∙_ + (Iso.inv (flipΩIso (suc n))) + (flipΩIsopres· n) + (Iso.fun (flipΩIso (suc n))) + (Iso.leftInv (flipΩIso (suc n))) + (Iso.rightInv (flipΩIso (suc n))) p q) + ∙ (cong (botᵣ {A = A} (suc n)) + (isHom-lMap _ {A = Ω (Susp∙ (typ A))} + (fun (flipΩIso (suc n)) p) (fun (flipΩIso (suc n)) q)) + ∙ morphLemmas.isMorphInv ∙Π ∙Π (botᵣ⁻ {A = A} (suc n)) + (hom-botᵣ⁻' {A = A} n) + (botᵣ {A = A} (suc n)) + (Iso.leftInv (botᵣIso {A = A} (suc n))) + (Iso.rightInv (botᵣIso {A = A} (suc n))) + (lMap (suc n) (fun (flipΩIso (suc n)) p)) + (lMap (suc n) (fun (flipΩIso (suc n)) q))) + + +suspMapΩ→hom : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p q : typ ((Ω^ (suc n)) A)) + → suspMapΩ∙ (suc n) .fst (p ∙ q) + ≡ suspMapΩ∙ (suc n) .fst p ∙ suspMapΩ∙ (suc n) .fst q +suspMapΩ→hom {A = A} n p q = + cong (sym (snd (suspMapΩ∙ {A = A} n)) ∙∙_∙∙ snd (suspMapΩ∙ {A = A} n)) + (cong-∙ (fst (suspMapΩ∙ {A = A} n)) p q) + ∙ help (snd (suspMapΩ∙ {A = A} n)) _ _ + where + help : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) (q s : x ≡ x) + → sym p ∙∙ (q ∙ s) ∙∙ p + ≡ (sym p ∙∙ q ∙∙ p) ∙ (sym p ∙∙ s ∙∙ p) + help {x = x} = + J (λ y p → (q s : x ≡ x) + → sym p ∙∙ (q ∙ s) ∙∙ p + ≡ (sym p ∙∙ q ∙∙ p ) ∙ (sym p ∙∙ s ∙∙ p)) + λ q s → sym (rUnit (q ∙ s)) + ∙ cong₂ _∙_ (rUnit q) (rUnit s) + +private + transpLemTyp : ∀ {ℓ} → {A1 A2 B1 B2 : Type ℓ} + (AB1 : A1 ≡ B1) (AB2 : A2 ≡ B2) + → (f : A1 → A2) (g : B1 → B2) + → PathP (λ i → AB1 i → AB2 i) f g + → Type ℓ + transpLemTyp {ℓ} {A1} {A2} {B1} {B2} AB1 AB2 f g p = + (+A1 : A1 → A1 → A1) + → (+A2 : A2 → A2 → A2) + → (+B1 : B1 → B1 → B1) + → (+B2 : B2 → B2 → B2) + → ((x y : A1) → transport AB1 (+A1 x y) + ≡ +B1 (transport AB1 x) (transport AB1 y)) + → ((x y : A2) → transport AB2 (+A2 x y) + ≡ +B2 (transport AB2 x) (transport AB2 y)) + → ((x y : A1) → f (+A1 x y) ≡ +A2 (f x) (f y)) + → ((x y : B1) → g (+B1 x y) ≡ +B2 (g x) (g y)) + + transpLem : ∀ {ℓ} → {A1 A2 B1 B2 : Type ℓ} (AB1 : A1 ≡ B1) (AB2 : A2 ≡ B2) + → (f : A1 → A2) (g : B1 → B2) + → (eq : PathP (λ i → AB1 i → AB2 i) f g) + → transpLemTyp AB1 AB2 f g eq + transpLem {ℓ} {A1} {A2} {B1} {B2} = + J (λ B1 AB1 → (AB2 : A2 ≡ B2) + → (f : A1 → A2) (g : B1 → B2) + → (eq : PathP (λ i → AB1 i → AB2 i) f g) + → transpLemTyp AB1 AB2 f g eq) + (J (λ B2 AB2 → (f : A1 → A2) (g : A1 → B2) + → (eq : PathP (λ i → A1 → AB2 i) f g) + → transpLemTyp refl AB2 f g eq) + λ f g → J (λ g p → transpLemTyp refl refl f g p) + λ +A1 +A2 +B1 +B2 trId1 trId2 fhom1 + → λ x y → cong f (cong₂ +B1 (sym (transportRefl x)) (sym (transportRefl y)) + ∙∙ sym (trId1 x y) + ∙∙ transportRefl (+A1 x y)) + ∙∙ fhom1 x y + ∙∙ ((sym (transportRefl _) ∙ (trId2 (f x) (f y))) + ∙ cong₂ +B2 (transportRefl (f x)) (transportRefl (f y)))) + +open import Cubical.Foundations.Univalence +suspMap→hom : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f g : S₊∙ (suc n) →∙ A) + → suspMap n (∙Π f g) ≡ ∙Π (suspMap n f) (suspMap n g) +suspMap→hom {A = A} n = + transpLem (Ω≡SphereMap {A = A} (suc n)) (Ω≡SphereMap' {A = A} (suc n)) + (suspMapΩ {A = A} (suc n)) (suspMap {A = A} n) + (suspMap→ {A = A} n) + _∙_ _∙_ ∙Π ∙Π + (λ x y → uaβ (isoToEquiv (IsoΩSphereMap {A = A} (suc n))) (x ∙ y) + ∙∙ isHom-lMap n x y + ∙∙ cong₂ ∙Π + (sym (uaβ (isoToEquiv (IsoΩSphereMap {A = A} (suc n))) x)) + (sym (uaβ (isoToEquiv (IsoΩSphereMap {A = A} (suc n))) y))) + (λ x y → uaβ (isoToEquiv (IsoΩSphereMap' {A = A} (suc n))) (x ∙ y) + ∙∙ isHom-IsoΩSphereMap' {A = A} (suc n) x y + ∙∙ cong₂ ∙Π + (sym (uaβ (isoToEquiv (IsoΩSphereMap' {A = A} (suc n))) x)) + (sym (uaβ (isoToEquiv (IsoΩSphereMap' {A = A} (suc n))) y))) + (suspMapΩ→hom {A = A} n) + +private + isConnectedPres : ∀ {ℓ} {A : Pointed ℓ} (con n : ℕ) + → isConnectedFun con (suspMapΩ∙ {A = A} (suc n) .fst) + → isConnectedFun con (suspMap {A = A} n) + isConnectedPres {A = A} con n hyp = + transport (λ i → isConnectedFun con (suspMap→ {A = A} n i)) hyp + +isConnectedSuspMap : (n m : ℕ) + → isConnectedFun ((m + suc m) ∸ n) (suspMap {A = S₊∙ (suc m)} n) +isConnectedSuspMap n m = + isConnectedPres _ _ (suspMapΩ-connected m (suc n) (sphereConnected (suc m))) + +suspMapπ' : (n m : ℕ) + → π' (suc n) (S₊∙ (suc m)) + → π' (suc (suc n)) (S₊∙ (suc (suc m))) +suspMapπ' n m = sMap (suspMap {A = S₊∙ (suc m)} n) + +suspMapHom : (n m : ℕ) + → GroupHom (π'Gr n (S₊∙ (suc m))) + (π'Gr (suc n) (S₊∙ (suc (suc m)))) +fst (suspMapHom n m) = suspMapπ' n m +snd (suspMapHom n m) = + makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) + λ f g → cong ∣_∣₂ (suspMap→hom n f g)) + +isSurjectiveSuspMap : (n : ℕ) → isSurjective (suspMapHom (2 + n) (suc n)) +isSurjectiveSuspMap n = + sElim (λ _ → isProp→isSet squash) + λ f → trRec (subst (λ x → isOfHLevel x (isInIm (suspMapHom (2 + n) (suc n)) ∣ f ∣₂)) + (sym (snd (lem n n))) + (isProp→isOfHLevelSuc {A = isInIm (suspMapHom (2 + n) (suc n)) ∣ f ∣₂} + (fst (lem n n)) squash)) + (λ p → ∣ ∣ fst p ∣₂ , (cong ∣_∣₂ (snd p)) ∣) + (fst (isConnectedSuspMap (2 + n) (suc n) f)) + where + lem : (m n : ℕ) → Σ[ x ∈ ℕ ] ((m + suc (suc n) ∸ suc n) ≡ suc x) + lem zero zero = 0 , refl + lem (suc m) zero = suc m , +-comm m 2 + lem zero (suc n) = lem zero n + lem (suc m) (suc n) = fst (lem (suc m) n) , (cong (_∸ (suc n)) (+-comm m (3 + n)) + ∙∙ cong (_∸ n) (+-comm (suc (suc n)) m) + ∙∙ snd (lem (suc m) n)) + + +-- open import Cubical.Data.Nat.Order +-- open import Cubical.Data.Empty renaming (rec to ⊥-rec) + +-- n∸n : (n : ℕ) → n ∸ n ≡ 0 +-- n∸n zero = refl +-- n∸n (suc n) = n∸n n + +-- ∸≤ : (n m : ℕ) → n < m → n ∸ m ≡ 0 +-- ∸≤ n zero p = ⊥-rec (¬-<-zero p) +-- ∸≤ zero (suc m) p = refl +-- ∸≤ (suc n) (suc m) p = ∸≤ n m (pred-≤-pred p) + +-- suc∸ : (n m : ℕ) → m < n → suc (n ∸ m) ≡ (suc n) ∸ m +-- suc∸ n zero p = refl +-- suc∸ zero (suc m) p = ⊥-rec (¬-<-zero p) +-- suc∸ (suc n) (suc m) p = suc∸ n m (pred-≤-pred p) + +-- isConnectedΩ : ∀ {ℓ} {A : Pointed ℓ} (n m : ℕ) +-- → isConnected n (typ A) +-- → isConnected (n ∸ m) (typ ((Ω^ m) A)) +-- isConnectedΩ {A = A} n zero conA = conA +-- isConnectedΩ {A = A} n (suc m) conA with (n ≟ (suc m)) +-- ... | (lt p) = subst (λ x → isConnected x (typ ((Ω^ (suc m)) A))) +-- (sym (∸≤ n (suc m) p)) +-- (tt* , (λ{tt* → refl})) +-- ... | (eq p) = subst (λ x → isConnected x (typ ((Ω^ (suc m)) A))) +-- (sym (n∸n n) ∙ (cong (n ∸_) p)) +-- (tt* , (λ{tt* → refl})) +-- ... | (gt p) = subst (λ x → isConnected x (typ ((Ω^ (suc m)) A))) +-- (sym (cong predℕ (snd (lem₂ n m p))) ∙ sym (lem₁ n m p)) +-- conΩ +-- where +-- lem₁ : (n m : ℕ) → suc m < n → n ∸ (suc m) ≡ predℕ (n ∸ m) +-- lem₁ zero m p = ⊥-rec (¬-<-zero p) +-- lem₁ (suc n) zero p = refl +-- lem₁ (suc n) (suc m) p = lem₁ n m (pred-≤-pred p) + +-- lem₂ : (n m : ℕ) → suc m < n → Σ[ x ∈ ℕ ] (n ∸ m ≡ suc x) +-- lem₂ n m = uncurry λ x +-- → J (λ n p → Σ[ x ∈ ℕ ] (n ∸ m ≡ suc x)) +-- (help m x) +-- where +-- help : (m x : ℕ) → Σ[ y ∈ ℕ ] (((x + suc (suc m)) ∸ m) ≡ suc y) +-- help zero zero = 1 , refl +-- help (suc m) zero = help m zero +-- help m (suc x) = _ , sym (suc∸ (x + (suc (suc m))) m (lem m x)) +-- where +-- lem : (m x : ℕ) → m < x + suc (suc m) +-- lem m zero = suc zero , refl +-- lem m (suc x) = <-trans (lem m x) (zero , refl) + +-- x = fst (lem₂ n m p) + +-- conΩ : isConnected x ((typ ((Ω^ (suc m)) A))) +-- conΩ = isConnectedPath x (subst (λ x → isConnected x (typ ((Ω^ m) A))) +-- (snd (lem₂ n m p)) +-- (isConnectedΩ n m conA)) _ _ + + +-- isConnectedΩ-S : (n m : ℕ) → isConnected (suc (m ∸ n)) (typ ((Ω^ n) (S₊∙ m))) +-- isConnectedΩ-S n m with (m ≟ n) +-- ... | (lt p) = subst (λ x → isConnected x (typ ((Ω^ n) (S₊∙ m)))) +-- (cong suc (sym (∸≤ m n p))) +-- (∣ pt ((Ω^ n) (S₊∙ m)) ∣ , (isOfHLevelTrunc 1 _)) +-- ... | (eq p) = subst (λ x → isConnected x (typ ((Ω^ n) (S₊∙ m)))) +-- (cong suc (sym (n∸n m) ∙ (λ i → m ∸ p i))) +-- (∣ pt ((Ω^ n) (S₊∙ m)) ∣ , (isOfHLevelTrunc 1 _)) +-- ... | (gt p) = subst (λ x → isConnected x ((typ ((Ω^ n) (S₊∙ m))))) +-- (sym (suc∸ m n p)) +-- (isConnectedΩ (suc m) n (sphereConnected m)) + +-- open import Cubical.Homotopy.WedgeConnectivity +-- isConnectedΩ'-S : (n m : ℕ) → isConnected (suc (m ∸ n)) (S₊∙ n →∙ S₊∙ m) +-- isConnectedΩ'-S n m = +-- subst (isConnected (suc (m ∸ n))) (Ω≡SphereMap n) (isConnectedΩ-S n m) + +-- suspMap-preComp' : ∀ {ℓ} {A : Pointed ℓ} (n m : ℕ) +-- (f g : S₊∙ (suc n) →∙ S₊∙ m) +-- → ∥ suspMap n (∙Π f g) ≡ ∙Π (suspMap n f) (suspMap n g) ∥ +-- suspMap-preComp' n m f g = {!WC!} +-- where +-- lem1 : (n m : ℕ) → 1 ≤ (m ∸ suc n + (m ∸ suc n)) +-- lem1 n m = {!!} + +-- module WC = WedgeConnectivity 1 1 +-- (S₊∙ (suc n) →∙ S₊∙ m ∙) {!!} +-- (S₊∙ (suc n) →∙ S₊∙ m ∙) {!!} +-- {!!} +-- {!!} +-- {!!} From ade3cd2ab3c22f3187c8c1e73609dd016be0b40c Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Fri, 29 Oct 2021 16:00:06 +0200 Subject: [PATCH 74/89] stuff --- Cubical/Homotopy/Group/S3.agda | 109 ++++----------------------------- 1 file changed, 11 insertions(+), 98 deletions(-) diff --git a/Cubical/Homotopy/Group/S3.agda b/Cubical/Homotopy/Group/S3.agda index 146d96a7e..3462594ca 100644 --- a/Cubical/Homotopy/Group/S3.agda +++ b/Cubical/Homotopy/Group/S3.agda @@ -85,7 +85,7 @@ The idea is to fill the following diagram Ωⁿ A -------------------> Ωⁿ⁺¹ (Susp A) | | | | - | ≃ lMap | ≃ eq₂ + | ≃ eq₁ | ≃ eq₂ | | v suspMap v (Sⁿ →∙ A) -------------- > (Sⁿ⁺¹ →∙ Susp A) @@ -495,14 +495,19 @@ filler▿ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (x : (S₊∙ n →∙ A)) filler▿ zero (f , p) = ΣPathP ((funExt (λ { base → refl ; (loop i) → refl})) , refl) filler▿ (suc n) (f , p) = - ΣPathP ((funExt (λ { north → refl ; south → refl ; (merid a i) → refl})) , refl) - + ΣPathP ((funExt (λ { north → refl + ; south → refl + ; (merid a i) → refl})) + , refl) -IsoΩSphereMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → Iso (typ ((Ω^ n) A)) ((S₊∙ n →∙ A)) +IsoΩSphereMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (typ ((Ω^ n) A)) ((S₊∙ n →∙ A)) fun (IsoΩSphereMap zero) = lMap zero inv (IsoΩSphereMap zero) f = fst f false -rightInv (IsoΩSphereMap zero) f = ΣPathP ((funExt (λ { false → refl ; true → sym (snd f)})) - , λ i j → snd f (~ i ∨ j)) +rightInv (IsoΩSphereMap zero) f = + ΣPathP ((funExt (λ { false → refl + ; true → sym (snd f)})) + , λ i j → snd f (~ i ∨ j)) leftInv (IsoΩSphereMap zero) p = refl IsoΩSphereMap (suc n) = equivToIso (_ , isEquiv-lMap n) @@ -1223,95 +1228,3 @@ isSurjectiveSuspMap n = lem (suc m) (suc n) = fst (lem (suc m) n) , (cong (_∸ (suc n)) (+-comm m (3 + n)) ∙∙ cong (_∸ n) (+-comm (suc (suc n)) m) ∙∙ snd (lem (suc m) n)) - - --- open import Cubical.Data.Nat.Order --- open import Cubical.Data.Empty renaming (rec to ⊥-rec) - --- n∸n : (n : ℕ) → n ∸ n ≡ 0 --- n∸n zero = refl --- n∸n (suc n) = n∸n n - --- ∸≤ : (n m : ℕ) → n < m → n ∸ m ≡ 0 --- ∸≤ n zero p = ⊥-rec (¬-<-zero p) --- ∸≤ zero (suc m) p = refl --- ∸≤ (suc n) (suc m) p = ∸≤ n m (pred-≤-pred p) - --- suc∸ : (n m : ℕ) → m < n → suc (n ∸ m) ≡ (suc n) ∸ m --- suc∸ n zero p = refl --- suc∸ zero (suc m) p = ⊥-rec (¬-<-zero p) --- suc∸ (suc n) (suc m) p = suc∸ n m (pred-≤-pred p) - --- isConnectedΩ : ∀ {ℓ} {A : Pointed ℓ} (n m : ℕ) --- → isConnected n (typ A) --- → isConnected (n ∸ m) (typ ((Ω^ m) A)) --- isConnectedΩ {A = A} n zero conA = conA --- isConnectedΩ {A = A} n (suc m) conA with (n ≟ (suc m)) --- ... | (lt p) = subst (λ x → isConnected x (typ ((Ω^ (suc m)) A))) --- (sym (∸≤ n (suc m) p)) --- (tt* , (λ{tt* → refl})) --- ... | (eq p) = subst (λ x → isConnected x (typ ((Ω^ (suc m)) A))) --- (sym (n∸n n) ∙ (cong (n ∸_) p)) --- (tt* , (λ{tt* → refl})) --- ... | (gt p) = subst (λ x → isConnected x (typ ((Ω^ (suc m)) A))) --- (sym (cong predℕ (snd (lem₂ n m p))) ∙ sym (lem₁ n m p)) --- conΩ --- where --- lem₁ : (n m : ℕ) → suc m < n → n ∸ (suc m) ≡ predℕ (n ∸ m) --- lem₁ zero m p = ⊥-rec (¬-<-zero p) --- lem₁ (suc n) zero p = refl --- lem₁ (suc n) (suc m) p = lem₁ n m (pred-≤-pred p) - --- lem₂ : (n m : ℕ) → suc m < n → Σ[ x ∈ ℕ ] (n ∸ m ≡ suc x) --- lem₂ n m = uncurry λ x --- → J (λ n p → Σ[ x ∈ ℕ ] (n ∸ m ≡ suc x)) --- (help m x) --- where --- help : (m x : ℕ) → Σ[ y ∈ ℕ ] (((x + suc (suc m)) ∸ m) ≡ suc y) --- help zero zero = 1 , refl --- help (suc m) zero = help m zero --- help m (suc x) = _ , sym (suc∸ (x + (suc (suc m))) m (lem m x)) --- where --- lem : (m x : ℕ) → m < x + suc (suc m) --- lem m zero = suc zero , refl --- lem m (suc x) = <-trans (lem m x) (zero , refl) - --- x = fst (lem₂ n m p) - --- conΩ : isConnected x ((typ ((Ω^ (suc m)) A))) --- conΩ = isConnectedPath x (subst (λ x → isConnected x (typ ((Ω^ m) A))) --- (snd (lem₂ n m p)) --- (isConnectedΩ n m conA)) _ _ - - --- isConnectedΩ-S : (n m : ℕ) → isConnected (suc (m ∸ n)) (typ ((Ω^ n) (S₊∙ m))) --- isConnectedΩ-S n m with (m ≟ n) --- ... | (lt p) = subst (λ x → isConnected x (typ ((Ω^ n) (S₊∙ m)))) --- (cong suc (sym (∸≤ m n p))) --- (∣ pt ((Ω^ n) (S₊∙ m)) ∣ , (isOfHLevelTrunc 1 _)) --- ... | (eq p) = subst (λ x → isConnected x (typ ((Ω^ n) (S₊∙ m)))) --- (cong suc (sym (n∸n m) ∙ (λ i → m ∸ p i))) --- (∣ pt ((Ω^ n) (S₊∙ m)) ∣ , (isOfHLevelTrunc 1 _)) --- ... | (gt p) = subst (λ x → isConnected x ((typ ((Ω^ n) (S₊∙ m))))) --- (sym (suc∸ m n p)) --- (isConnectedΩ (suc m) n (sphereConnected m)) - --- open import Cubical.Homotopy.WedgeConnectivity --- isConnectedΩ'-S : (n m : ℕ) → isConnected (suc (m ∸ n)) (S₊∙ n →∙ S₊∙ m) --- isConnectedΩ'-S n m = --- subst (isConnected (suc (m ∸ n))) (Ω≡SphereMap n) (isConnectedΩ-S n m) - --- suspMap-preComp' : ∀ {ℓ} {A : Pointed ℓ} (n m : ℕ) --- (f g : S₊∙ (suc n) →∙ S₊∙ m) --- → ∥ suspMap n (∙Π f g) ≡ ∙Π (suspMap n f) (suspMap n g) ∥ --- suspMap-preComp' n m f g = {!WC!} --- where --- lem1 : (n m : ℕ) → 1 ≤ (m ∸ suc n + (m ∸ suc n)) --- lem1 n m = {!!} - --- module WC = WedgeConnectivity 1 1 --- (S₊∙ (suc n) →∙ S₊∙ m ∙) {!!} --- (S₊∙ (suc n) →∙ S₊∙ m ∙) {!!} --- {!!} --- {!!} --- {!!} From 96d1f28d450674bec455abf2d983d149ff601acf Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 1 Nov 2021 09:17:18 +0100 Subject: [PATCH 75/89] stuf --- Cubical/Homotopy/Group/Base.agda | 927 +++++++++---------------- Cubical/Homotopy/Group/Properties.agda | 710 +++++++++++++++++++ Cubical/Homotopy/Group/S3.agda | 423 +---------- 3 files changed, 1062 insertions(+), 998 deletions(-) create mode 100644 Cubical/Homotopy/Group/Properties.agda diff --git a/Cubical/Homotopy/Group/Base.agda b/Cubical/Homotopy/Group/Base.agda index 5d619d011..a5f7a27cc 100644 --- a/Cubical/Homotopy/Group/Base.agda +++ b/Cubical/Homotopy/Group/Base.agda @@ -40,14 +40,15 @@ open IsMonoid open GroupStr -{- homotopy Group -} +{- Homotopy group -} π : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Type ℓ π n A = ∥ typ ((Ω^ n) A) ∥₂ +{- Alternative formulation. This will be given a group structure in + the Properties file -} π' : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Type ℓ π' n A = ∥ S₊∙ n →∙ A ∥₂ - {- π as a group -} 1π : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π n A 1π zero {A = A} = ∣ pt A ∣₂ @@ -98,6 +99,7 @@ inverse (isGroup (snd (πGr n A))) x = (π-rCancel n x) , (π-lCancel n x) -- (S₊∙ n →∙ A). We first give multiplication and inversion and use this -- to give a structure preserving equivalence (S₊∙ n →∙ A) ≃ Ωⁿ A. + ∙Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → (S₊∙ n →∙ A) → (S₊∙ n →∙ A) @@ -127,598 +129,329 @@ fst (-Π {A = A} {n = suc (suc n)} f) (merid a j) = fst f ((merid a ∙ sym (merid (ptSn _))) (~ j)) snd (-Π {A = A} {n = suc (suc n)} f) = snd f -module _ {ℓ : Level} {A : Pointed ℓ} where -flipΩPath : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → ((Ω^ (suc n)) A) ≡ (Ω^ n) (Ω A) -flipΩPath {A = A} zero = refl -flipΩPath {A = A} (suc n) = cong Ω (flipΩPath {A = A} n) - -flipΩIso : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → Iso (fst ((Ω^ (suc n)) A)) (fst ((Ω^ n) (Ω A))) -flipΩIso {A = A} n = pathToIso (cong fst (flipΩPath n)) - -flipΩIsopres· : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → (f g : fst ((Ω^ (suc n)) (Ω A))) - → Iso.inv (flipΩIso (suc n)) (f ∙ g) - ≡ (Iso.inv (flipΩIso (suc n)) f) - ∙ (Iso.inv (flipΩIso (suc n)) g) -flipΩIsopres· {A = A} n f g i = - transp (λ j → flipΩPath {A = A} n (~ i ∧ ~ j) .snd - ≡ flipΩPath n (~ i ∧ ~ j) .snd) i - (transp (λ j → flipΩPath {A = A} n (~ i ∨ ~ j) .snd - ≡ flipΩPath n (~ i ∨ ~ j) .snd) (~ i) f - ∙ transp (λ j → flipΩPath {A = A} n (~ i ∨ ~ j) .snd - ≡ flipΩPath n (~ i ∨ ~ j) .snd) (~ i) g) - -module _ {ℓ : Level} {A : Pointed ℓ} where - suspension→ : (n : ℕ) → (S₊∙ (suc n) →∙ Ω A) → (S₊∙ (suc (suc n)) →∙ A) - fst (suspension→ n f) north = pt A - fst (suspension→ n f) south = pt A - fst (suspension→ n f) (merid a i₁) = fst f a i₁ - snd (suspension→ n f) = refl - - suspension-pres∙ : - (n : ℕ) (f g : (S₊∙ (suc n) →∙ Ω A)) - → suspension→ n (∙Π f g) ≡ ∙Π (suspension→ n f) (suspension→ n g) - suspension-pres∙ n f g = ΣPathP ((funExt main) , refl) - where - J2 : ∀ {ℓ ℓ'} {A : Type ℓ} {x : A} - → (P : (p q : x ≡ x) → (refl ≡ p) → (refl ≡ q) → Type ℓ') - → P refl refl refl refl - → (p q : x ≡ x) (P' : _) (Q : _) → P p q P' Q - J2 P b p q = - J (λ p P' → (Q₁ : refl ≡ q) → P p q P' Q₁) - (J (λ q Q → P refl q refl Q) b) - - J4 : ∀ {ℓ ℓ'} {A : Type ℓ} {x : A} - → (P : (p q r s : x ≡ x) - → (refl ≡ p) → (p ≡ q) → (refl ≡ r) → (r ≡ s) → Type ℓ') - → P refl refl refl refl refl refl refl refl - → (p q r s : x ≡ x) (P' : _) (Q : _) (R : _) (S : _) → P p q r s P' Q R S - J4 P b p q r s = - J (λ p P' → (Q₁ : p ≡ q) (R : refl ≡ r) (S₁ : r ≡ s) - → P p q r s P' Q₁ R S₁) - (J (λ q Q₁ → (R : refl ≡ r) (S₁ : r ≡ s) - → P refl q r s refl Q₁ R S₁) - (J (λ r R → (S₁ : r ≡ s) - → P refl refl r s refl refl R S₁) - (J (λ s S₁ → P refl refl refl s refl refl refl S₁) - b))) - - looper-help : (n : ℕ) (f g : (S₊∙ (suc n) →∙ Ω A)) (a : _) - → fst (∙Π f g) a ≡ fst f a ∙ fst g a - looper-help zero f g base = - rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g)) - looper-help zero f g (loop i) k = - lazy-fill _ _ - (sym (snd f)) (sym (snd g)) - (cong (fst f) loop) (cong (fst g) loop) k i - where - lazy-fill : - ∀ {ℓ} {A : Type ℓ} {x : A} (p q : x ≡ x) - (prefl : refl ≡ p) (qrefl : refl ≡ q) (fl : p ≡ p) (fr : q ≡ q) - → PathP (λ i → (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i - ≡ (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i) - ((prefl ∙∙ fl ∙∙ sym prefl) ∙ (qrefl ∙∙ fr ∙∙ sym qrefl)) - (cong₂ _∙_ fl fr) - lazy-fill {x = x} = - J2 (λ p q prefl qrefl → (fl : p ≡ p) (fr : q ≡ q) - → PathP (λ i → (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i - ≡ (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i) - ((prefl ∙∙ fl ∙∙ sym prefl) ∙ (qrefl ∙∙ fr ∙∙ sym qrefl)) - (cong₂ _∙_ fl fr)) - λ fl fr → cong₂ _∙_ (sym (rUnit fl)) (sym (rUnit fr)) - ◁ flipSquare (sym (rUnit (rUnit refl)) - ◁ (flipSquare ((λ j → cong (λ x → rUnit x j) fl - ∙ cong (λ x → lUnit x j) fr) - ▷ sym (cong₂Funct _∙_ fl fr)) - ▷ (rUnit (rUnit refl)))) - looper-help (suc n) f g north = - rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g)) - looper-help (suc n) f g south = - (rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) - ∙ cong₂ _∙_ (cong (fst f) (merid (ptSn _))) - (cong (fst g) (merid (ptSn _))) - looper-help (suc n) f g (merid a j) i = help i j - where - help : - PathP (λ i → (rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) i - ≡ ((rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) - ∙ cong₂ _∙_ (cong (fst f) (merid (ptSn _))) - (cong (fst g) (merid (ptSn _)))) i) - (cong (fst (∙Π f g)) (merid a)) - (cong₂ _∙_ (cong (fst f) (merid a)) - (cong (fst g) (merid a))) - help = (cong₂ _∙_ (cong (sym (snd f) ∙∙_∙∙ snd f) - (cong-∙ (fst f) - (merid a) (sym (merid (ptSn _))))) - ((cong (sym (snd g) ∙∙_∙∙ snd g) - (cong-∙ (fst g) - (merid a) (sym (merid (ptSn _)))))) - ◁ lazy-help _ _ _ _ (sym (snd f)) (cong (fst f) (merid (ptSn _))) - (sym (snd g)) (cong (fst g) (merid (ptSn _))) - (cong (fst f) (merid a)) (cong (fst g) (merid a))) - where - lazy-help : - ∀ {ℓ} {A : Type ℓ} {x : A} - (p-north p-south q-north q-south : x ≡ x) - (prefl : refl ≡ p-north) (p-merid : p-north ≡ p-south) - (qrefl : refl ≡ q-north) (q-merid : q-north ≡ q-south) - (fl : p-north ≡ p-south) (fr : q-north ≡ q-south) - → PathP (λ i → (rUnit refl ∙ cong₂ _∙_ prefl qrefl) i - ≡ ((rUnit refl ∙ cong₂ _∙_ prefl qrefl) - ∙ cong₂ _∙_ p-merid q-merid) i) - ((prefl ∙∙ fl ∙ sym p-merid ∙∙ sym prefl) - ∙ (qrefl ∙∙ fr ∙ sym q-merid ∙∙ sym qrefl)) - (cong₂ _∙_ fl fr) - lazy-help {x = x} = - J4 (λ p-north p-south q-north q-south prefl p-merid qrefl q-merid - → (fl : p-north ≡ p-south) (fr : q-north ≡ q-south) → - PathP (λ i₁ → (rUnit refl ∙ cong₂ _∙_ prefl qrefl) i₁ ≡ - ((rUnit refl ∙ cong₂ _∙_ prefl qrefl) - ∙ cong₂ _∙_ p-merid q-merid) i₁) - ((prefl ∙∙ fl ∙ sym p-merid ∙∙ sym prefl) ∙ - (qrefl ∙∙ fr ∙ sym q-merid ∙∙ sym qrefl)) - (cong₂ _∙_ fl fr)) - λ fl fr → (cong₂ _∙_ (sym (rUnit (fl ∙ refl)) ∙ sym (rUnit fl)) - (sym (rUnit (fr ∙ refl)) ∙ sym (rUnit fr)) - ◁ flipSquare ((sym (rUnit _) - ◁ flipSquare - ((λ j → cong (λ x → rUnit x j) fl - ∙ cong (λ x → lUnit x j) fr) - ▷ sym (cong₂Funct _∙_ fl fr))) - ▷ (rUnit _ ∙ rUnit _))) - main : (x : fst (S₊∙ (suc (suc n)))) → - fst (suspension→ n (∙Π f g)) x ≡ - fst (∙Π (suspension→ n f) (suspension→ n g)) x - main north = refl - main south = refl - main (merid a i₁) j = help j i₁ - where - help : fst (∙Π f g) a - ≡ cong (fst (∙Π (suspension→ n f) (suspension→ n g))) (merid a) - help = looper-help n f g a - ∙ cong₂ _∙_ (sym (cong-∙ (fst (suspension→ n f)) - (merid a) - (sym (merid (ptSn _))) - ∙∙ cong (fst f a ∙_) (cong sym (snd f)) - ∙∙ sym (rUnit _))) - (sym (cong-∙ (fst (suspension→ n g)) - (merid a) - (sym (merid (ptSn _))) - ∙∙ cong (fst g a ∙_) (cong sym (snd g)) - ∙∙ sym (rUnit _))) - ∙ λ i → rUnit (cong (fst (suspension→ n f)) - (merid a ∙ sym (merid (ptSn _)))) i - ∙ rUnit (cong (fst (suspension→ n g)) - (merid a ∙ sym (merid (ptSn _)))) i - - suspension←filler : (n : ℕ) → (S₊∙ (suc (suc n)) →∙ A) - → I → I → I → fst A - suspension←filler n f i j k = - hfill (λ k → λ { (i = i0) → doubleCompPath-filler - (sym (snd f)) - (cong (fst f) (merid (ptSn _) - ∙ sym (merid (ptSn _)))) - (snd f) k j - ; (i = i1) → snd f k - ; (j = i0) → snd f k - ; (j = i1) → snd f k}) - (inS ((fst f) (rCancel' (merid (ptSn _)) i j))) - k - - suspension← : (n : ℕ) → (S₊∙ (suc (suc n)) →∙ A) → (S₊∙ (suc n) →∙ Ω A) - fst (suspension← n f) a = - sym (snd f) ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f - snd (suspension← n f) i j = suspension←filler n f i j i1 - - suspensionIso : (n : ℕ) → Iso (S₊∙ (suc n) →∙ Ω A) (S₊∙ (suc (suc n)) →∙ A) - Iso.fun (suspensionIso n) f = suspension→ n f - Iso.inv (suspensionIso n) f = suspension← n f - Iso.rightInv (suspensionIso n) (f , p) = - ΣPathP (funExt help , λ i j → p (~ i ∨ j)) - where - help : (x : _) → fst (suspension→ n (suspension← n (f , p))) x ≡ f x - help north = sym p - help south = sym p ∙ cong f (merid (ptSn _)) - help (merid a j) i = - hcomp (λ k → λ { (i = i1) → f (merid a j) - ; (j = i0) → p (~ i ∧ k) - ; (j = i1) → compPath-filler' (sym p) - (cong f (merid (ptSn _))) k i}) - (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) - Iso.leftInv (suspensionIso n) f = - ΣPathP (funExt lem , PathP-filler) - where - - side₁ : (x : S₊ (suc n)) → I → I → I → fst A - side₁ x i j k = - hfill (λ k → λ { (i = i0) → suspension→ n f .fst - (compPath-filler' (merid x) - (sym (merid (ptSn _))) k j) - ; (i = i1) → snd A - ; (j = i0) → fst f x (i ∨ ~ k) - ; (j = i1) → snd A}) - (inS (snd f i (~ j))) - k - - side₂ : (x : S₊ (suc n)) → I → I → I → fst A - side₂ x i j k = - hfill (λ k → λ { (i = i0) → doubleCompPath-filler - refl - (cong (suspension→ n f .fst) - (merid x ∙ sym (merid (ptSn _)))) - refl k j - ; (i = i1) → fst f x (j ∨ ~ k) - ; (j = i0) → fst f x (i ∧ ~ k) - ; (j = i1) → snd A}) - (inS (side₁ x i j i1)) - k - - compPath-filler'' : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) - → I → I → I → A - compPath-filler'' {z = z} p q k j i = - hfill (λ k → λ { (i = i0) → p (~ j) - ; (i = i1) → q k - ; (j = i0) → q (i ∧ k) }) - (inS (p (i ∨ ~ j))) - k - - lem : (x : _) → fst (suspension← n (suspension→ n f)) x ≡ fst f x - lem x i j = side₂ x i j i1 - - sideCube : Cube (λ j k → snd f j (~ k)) - (λ j k → fst (suspension→ n f) - (rCancel' (merid (ptSn _)) j k)) - (λ r k → suspension→ n f .fst - (compPath-filler' (merid (ptSn _)) - (sym (merid (ptSn _))) r k)) - (λ _ _ → pt A) - (λ r j → snd f j (~ r)) λ _ _ → pt A - sideCube r j k = - hcomp (λ i → λ {(r = i0) → snd f j (~ k ∨ ~ i) - ; (r = i1) → fst (suspension→ n f) - (rCancel-filler' (merid (ptSn _)) (~ i) j k) - ; (j = i0) → suspension→ n f .fst - (compPath-filler'' (merid (ptSn _)) - (sym (merid (ptSn _))) i r k) - ; (j = i1) → snd f (~ r) (~ i ∧ k) - ; (k = i0) → snd f j (~ r) - ; (k = i1) → snd f (~ r ∧ j) (~ i)}) - (hcomp (λ i → λ {(r = i0) → pt A - ; (r = i1) → snd f (~ i) k - ; (j = i0) → snd f (~ i) (k ∨ ~ r) - ; (j = i1) → snd f (~ r ∨ ~ i) k - ; (k = i0) → snd f (j ∨ ~ i) (~ r) - ; (k = i1) → pt A}) - (pt A)) - - PathP-filler : PathP _ _ _ - PathP-filler i j k = - hcomp (λ r → λ {(i = i0) → suspension←filler n (suspension→ n f) j k r - ; (i = i1) → snd f j (k ∨ ~ r) - ; (j = i0) → side₂ (ptSn _) i k r - ; (j = i1) → snd A - ; (k = i0) → snd f j (i ∧ ~ r) - ; (k = i1) → snd A}) - (hcomp ((λ r → λ {(i = i0) → sideCube r j k - ; (i = i1) → pt A - ; (j = i0) → side₁ (ptSn _) i k r - ; (j = i1) → pt A - ; (k = i0) → snd f j (i ∨ ~ r) - ; (k = i1) → pt A})) - (snd f (i ∨ j) (~ k))) - - suspension←-pres∙ : (n : ℕ) → (f g : S₊∙ (suc (suc n)) →∙ A) - → suspension← n (∙Π f g) - ≡ ∙Π (suspension← n f) (suspension← n g) - suspension←-pres∙ n f g = - sym (Iso.leftInv (suspensionIso n) (suspension← n (∙Π f g))) - ∙∙ cong (suspension← n) - (Iso.rightInv (suspensionIso n) (∙Π f g) - ∙∙ cong₂ ∙Π (sym (Iso.rightInv (suspensionIso n) f)) - (sym (Iso.rightInv (suspensionIso n) g)) - ∙∙ sym (suspension-pres∙ _ (suspension← n f) (suspension← n g))) - ∙∙ (Iso.leftInv (suspensionIso n) (∙Π (suspension← n f) (suspension← n g))) - -SphereMap→Ω : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → (S₊∙ (suc n) →∙ A) → (fst ((Ω^ (suc n)) A)) -SphereMap→Ω zero (f , p) = sym p ∙∙ cong f loop ∙∙ p -SphereMap→Ω {A = A} (suc n) (f , p) = - Iso.inv (flipΩIso _) - (SphereMap→Ω {A = Ω A} n (Iso.inv (suspensionIso _) (f , p))) - -SphereMap→Ωpres∙ : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → (p q : S₊∙ (suc n) →∙ A) - → SphereMap→Ω n (∙Π p q) ≡ SphereMap→Ω n p ∙ SphereMap→Ω n q -SphereMap→Ωpres∙ zero f g = sym (rUnit _) -SphereMap→Ωpres∙ (suc n) f g = - cong (Iso.inv (flipΩIso (suc n))) - (cong (SphereMap→Ω n) (suspension←-pres∙ n f g) - ∙ SphereMap→Ωpres∙ n (suspension← n f) (suspension← n g)) - ∙ flipΩIsopres· n _ _ - -Ω→SphereMap : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → (fst ((Ω^ (suc n)) A)) → (S₊∙ (suc n) →∙ A) -fst (Ω→SphereMap {A = A} zero f) base = pt A -fst (Ω→SphereMap {A = A} zero f) (loop i₁) = f i₁ -snd (Ω→SphereMap {A = A} zero f) = refl -Ω→SphereMap {A = A} (suc n) f = - Iso.fun (suspensionIso _) (Ω→SphereMap n (Iso.fun (flipΩIso _) f)) - -Ω→SphereMap→Ω : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → (x : fst ((Ω^ (suc n)) A)) - → SphereMap→Ω n (Ω→SphereMap n x) ≡ x -Ω→SphereMap→Ω zero x = sym (rUnit _) -Ω→SphereMap→Ω (suc n) x = - cong (Iso.inv (flipΩIso (suc n)) ∘ SphereMap→Ω n) - (Iso.leftInv (suspensionIso _) - (Ω→SphereMap n (Iso.fun (flipΩIso (suc n)) x))) - ∙∙ cong (Iso.inv (flipΩIso (suc n))) - (Ω→SphereMap→Ω n (Iso.fun (flipΩIso (suc n)) x)) - ∙∙ Iso.leftInv (flipΩIso (suc n)) x - -SphereMap→Ω→SphereMap : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → (x : (S₊∙ (suc n) →∙ A)) - → Ω→SphereMap n (SphereMap→Ω n x) ≡ x -SphereMap→Ω→SphereMap zero (f , p) = - ΣPathP (funExt fstp , λ i j → p (~ i ∨ j)) - where - fstp : (x : _) → _ ≡ f x - fstp base = sym p - fstp (loop i) j = doubleCompPath-filler (sym p) (cong f loop) p (~ j) i -SphereMap→Ω→SphereMap (suc n) f = - cong (Iso.fun (suspensionIso n) ∘ Ω→SphereMap n) - (Iso.rightInv (flipΩIso (suc n)) - (SphereMap→Ω n (Iso.inv (suspensionIso n) f))) - ∙∙ cong (Iso.fun (suspensionIso n)) - (SphereMap→Ω→SphereMap n (Iso.inv (suspensionIso n) f)) - ∙∙ Iso.rightInv (suspensionIso _) _ - - --- We finally arrive at the main result -IsoSphereMapΩ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → Iso (S₊∙ n →∙ A) (fst ((Ω^ n) A)) -Iso.fun (IsoSphereMapΩ {A = A} zero) (f , p) = f false -fst (Iso.inv (IsoSphereMapΩ {A = A} zero) a) false = a -fst (Iso.inv (IsoSphereMapΩ {A = A} zero) a) true = pt A -snd (Iso.inv (IsoSphereMapΩ {A = A} zero) a) = refl -Iso.rightInv (IsoSphereMapΩ {A = A} zero) a = refl -Iso.leftInv (IsoSphereMapΩ {A = A} zero) (f , p) = - ΣPathP ((funExt (λ { false → refl ; true → sym p})) - , λ i j → p (~ i ∨ j)) -Iso.fun (IsoSphereMapΩ {A = A} (suc n)) = SphereMap→Ω n -Iso.inv (IsoSphereMapΩ {A = A} (suc n)) = Ω→SphereMap n -Iso.rightInv (IsoSphereMapΩ {A = A} (suc n)) = Ω→SphereMap→Ω n -Iso.leftInv (IsoSphereMapΩ {A = A} (suc n)) = SphereMap→Ω→SphereMap n - --- The iso is structure preserving -IsoSphereMapΩ-pres∙Π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f g : S₊∙ (suc n) →∙ A) - → SphereMap→Ω n (∙Π f g) ≡ SphereMap→Ω n f ∙ SphereMap→Ω n g -IsoSphereMapΩ-pres∙Π zero f g = sym (rUnit _) -IsoSphereMapΩ-pres∙Π (suc n) f g = - cong (Iso.inv (flipΩIso (suc n))) - (cong (SphereMap→Ω n) (suspension←-pres∙ n f g) - ∙ SphereMap→Ωpres∙ n (suspension← n f) (suspension← n g)) - ∙ (flipΩIsopres· n (SphereMap→Ω n (suspension← n f)) - (SphereMap→Ω n (suspension← n g))) - --- It is useful to define the ``Group Structure'' on (S₊∙ n →∙ A) --- before doing it on π'. These will be the equivalents of the --- usual groupoid laws on Ω A. -1Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → (S₊∙ n →∙ A) -fst (1Π {A = A}) _ = pt A -snd (1Π {A = A}) = refl - -∙Π-rUnit : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f : S₊∙ (suc n) →∙ A) - → ∙Π f 1Π ≡ f -fst (∙Π-rUnit {A = A} {n = zero} f i) base = snd f (~ i) -fst (∙Π-rUnit {A = A} {n = zero} f i) (loop j) = help i j - where - help : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) - (((sym (snd f)) ∙∙ (cong (fst f) loop) ∙∙ snd f) - ∙ (refl ∙ refl)) - (cong (fst f) loop) - help = (cong ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) ∙_) - (sym (rUnit refl)) ∙ sym (rUnit _)) - ◁ λ i j → doubleCompPath-filler (sym (snd f)) - (cong (fst f) loop) (snd f) (~ i) j -snd (∙Π-rUnit {A = A} {n = zero} f i) j = snd f (~ i ∨ j) -fst (∙Π-rUnit {A = A} {n = suc n} f i) north = snd f (~ i) -fst (∙Π-rUnit {A = A} {n = suc n} f i) south = - (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i -fst (∙Π-rUnit {A = A} {n = suc n} f i) (merid a j) = help i j - where - help : PathP (λ i → snd f (~ i) - ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) - (((sym (snd f)) - ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) - ∙∙ snd f) - ∙ (refl ∙ refl)) - (cong (fst f) (merid a)) - help = (cong (((sym (snd f)) - ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) - ∙∙ snd f) ∙_) - (sym (rUnit refl)) - ∙ sym (rUnit _)) - ◁ λ i j → hcomp (λ k → - λ { (j = i0) → snd f (~ i ∧ k) - ; (j = i1) → compPath-filler' (sym (snd f)) - (cong (fst f) (merid (ptSn (suc n)))) k i - ; (i = i0) → doubleCompPath-filler (sym (snd f)) - (cong (fst f) - (merid a ∙ sym (merid (ptSn (suc n))))) - (snd f) k j - ; (i = i1) → fst f (merid a j)}) - (fst f (compPath-filler (merid a) - (sym (merid (ptSn _))) (~ i) j)) -snd (∙Π-rUnit {A = A} {n = suc n} f i) j = snd f (~ i ∨ j) - -∙Π-lUnit : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f : S₊∙ (suc n) →∙ A) - → ∙Π 1Π f ≡ f -fst (∙Π-lUnit {n = zero} f i) base = snd f (~ i) -fst (∙Π-lUnit {n = zero} f i) (loop j) = s i j - where - s : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) - ((refl ∙ refl) ∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) - (cong (fst f) loop) - s = (cong (_∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) - (sym (rUnit refl)) ∙ sym (lUnit _)) - ◁ λ i j → doubleCompPath-filler (sym (snd f)) - (cong (fst f) loop) (snd f) (~ i) j -snd (∙Π-lUnit {n = zero} f i) j = snd f (~ i ∨ j) -fst (∙Π-lUnit {n = suc n} f i) north = snd f (~ i) -fst (∙Π-lUnit {n = suc n} f i) south = - (sym (snd f) ∙ cong (fst f) (merid (ptSn _))) i -fst (∙Π-lUnit {n = suc n} f i) (merid a j) = help i j - where - help : PathP (λ i → snd f (~ i) - ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) - ((refl ∙ refl) ∙ ((sym (snd f)) - ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) - ∙∙ snd f)) - (cong (fst f) (merid a)) - help = - (cong (_∙ ((sym (snd f)) - ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) - ∙∙ snd f)) - (sym (rUnit refl)) - ∙ sym (lUnit _)) - ◁ λ i j → hcomp (λ k → - λ { (j = i0) → snd f (~ i ∧ k) - ; (j = i1) → compPath-filler' (sym (snd f)) - (cong (fst f) (merid (ptSn (suc n)))) k i - ; (i = i0) → doubleCompPath-filler (sym (snd f)) - (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) - (snd f) k j - ; (i = i1) → fst f (merid a j)}) - (fst f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) -snd (∙Π-lUnit {n = suc n} f i) j = snd f (~ i ∨ j) - -∙Π-rCancel : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f : S₊∙ (suc n) →∙ A) - → ∙Π f (-Π f) ≡ 1Π -fst (∙Π-rCancel {A = A} {n = zero} f i) base = pt A -fst (∙Π-rCancel {A = A} {n = zero} f i) (loop j) = - rCancel (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) i j -snd (∙Π-rCancel {A = A} {n = zero} f i) = refl -fst (∙Π-rCancel {A = A} {n = suc n} f i) north = pt A -fst (∙Π-rCancel {A = A} {n = suc n} f i) south = pt A -fst (∙Π-rCancel {A = A} {n = suc n} f i) (merid a i₁) = sl i i₁ - where - pl = (sym (snd f) - ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) - ∙∙ snd f) - sl : pl - ∙ ((sym (snd f) - ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) - ∙∙ snd f)) ≡ refl - sl = cong (pl ∙_) (cong (sym (snd f) ∙∙_∙∙ (snd f)) - (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) - ∙∙ cong₂ _∙_ refl - (cong (cong (fst f)) (rCancel (merid (ptSn _)))) - ∙∙ sym (rUnit _))) - ∙ rCancel pl -snd (∙Π-rCancel {A = A} {n = suc n} f i) = refl - -∙Π-lCancel : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f : S₊∙ (suc n) →∙ A) - → ∙Π (-Π f) f ≡ 1Π -fst (∙Π-lCancel {A = A} {n = zero} f i) base = pt A -fst (∙Π-lCancel {A = A} {n = zero} f i) (loop j) = - rCancel (sym (snd f) ∙∙ cong (fst f) (sym loop) ∙∙ snd f) i j -fst (∙Π-lCancel {A = A} {n = suc n} f i) north = pt A -fst (∙Π-lCancel {A = A} {n = suc n} f i) south = pt A -fst (∙Π-lCancel {A = A} {n = suc n} f i) (merid a j) = sl i j - where - pl = (sym (snd f) - ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) - ∙∙ snd f) - - sl : (sym (snd f) - ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) - ∙∙ snd f) ∙ pl - ≡ refl - sl = cong (_∙ pl) (cong (sym (snd f) ∙∙_∙∙ (snd f)) - (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) - ∙∙ cong₂ _∙_ refl (cong (cong (fst f)) (rCancel (merid (ptSn _)))) - ∙∙ sym (rUnit _))) - ∙ lCancel pl -snd (∙Π-lCancel {A = A} {n = zero} f i) = refl -snd (∙Π-lCancel {A = A} {n = suc n} f i) = refl - -∙Π-assoc : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f g h : S₊∙ (suc n) →∙ A) - → ∙Π f (∙Π g h) ≡ ∙Π (∙Π f g) h -∙Π-assoc {n = n} f g h = - sym (Iso.leftInv (IsoSphereMapΩ (suc n)) (∙Π f (∙Π g h))) - ∙∙ cong (Ω→SphereMap n) (SphereMap→Ωpres∙ n f (∙Π g h) - ∙∙ cong (SphereMap→Ω n f ∙_) (SphereMap→Ωpres∙ n g h) - ∙∙ ∙assoc (SphereMap→Ω n f) (SphereMap→Ω n g) (SphereMap→Ω n h) - ∙∙ cong (_∙ SphereMap→Ω n h) (sym (SphereMap→Ωpres∙ n f g)) - ∙∙ sym (SphereMap→Ωpres∙ n (∙Π f g) h)) - ∙∙ Iso.leftInv (IsoSphereMapΩ (suc n)) (∙Π (∙Π f g) h) - -∙Π-comm : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f g : S₊∙ (suc (suc n)) →∙ A) - → ∙Π f g ≡ ∙Π g f -∙Π-comm {A = A} {n = n} f g = - sym (Iso.leftInv (IsoSphereMapΩ (suc (suc n))) (∙Π f g)) - ∙∙ cong (Ω→SphereMap (suc n)) (SphereMap→Ωpres∙ (suc n) f g - ∙∙ EH _ _ _ - ∙∙ sym (SphereMap→Ωpres∙ (suc n) g f)) - ∙∙ Iso.leftInv (IsoSphereMapΩ (suc (suc n))) (∙Π g f) - - -{- π'' as a group -} -1π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' n A -1π' n {A = A} = ∣ 1Π ∣₂ - -·π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' (suc n) A → π' (suc n) A → π' (suc n) A -·π' n = sRec2 squash₂ λ p q → ∣ ∙Π p q ∣₂ - --π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' (suc n) A → π' (suc n) A --π' n = sMap -Π - -π'-rUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) - → (·π' n x (1π' (suc n))) ≡ x -π'-rUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-rUnit p i ∣₂ - -π'-lUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) - → (·π' n (1π' (suc n)) x) ≡ x -π'-lUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-lUnit p i ∣₂ - -π'-rCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) - → (·π' n x (-π' n x)) ≡ 1π' (suc n) -π'-rCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-rCancel p i ∣₂ - -π'-lCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) - → (·π' n (-π' n x) x) ≡ 1π' (suc n) -π'-lCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-lCancel p i ∣₂ - -π'-assoc : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y z : π' (suc n) A) - → ·π' n x (·π' n y z) ≡ ·π' n (·π' n x y) z -π'-assoc n = sElim3 (λ _ _ _ → isSetPathImplicit) λ p q r i → ∣ ∙Π-assoc p q r i ∣₂ - -π'-comm : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y : π' (suc (suc n)) A) - → ·π' (suc n) x y ≡ ·π' (suc n) y x -π'-comm n = sElim2 (λ _ _ → isSetPathImplicit) λ p q i → ∣ ∙Π-comm p q i ∣₂ - -π'Gr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Group ℓ -fst (π'Gr n A) = π' (suc n) A -1g (snd (π'Gr n A)) = 1π' (suc n) -GroupStr._·_ (snd (π'Gr n A)) = ·π' n -inv (snd (π'Gr n A)) = -π' n -is-set (isSemigroup (isMonoid (isGroup (snd (π'Gr n A))))) = squash₂ -assoc (isSemigroup (isMonoid (isGroup (snd (π'Gr n A))))) = π'-assoc n -identity (isMonoid (isGroup (snd (π'Gr n A)))) x = (π'-rUnit n x) , (π'-lUnit n x) -inverse (isGroup (snd (π'Gr n A))) x = (π'-rCancel n x) , (π'-lCancel n x) - --- Finally, the Iso -π'Gr≅πGr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → GroupIso (π'Gr n A) (πGr n A) -fst (π'Gr≅πGr n A) = setTruncIso (IsoSphereMapΩ (suc n)) -snd (π'Gr≅πGr n A) = - makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) - λ p q i → ∣ IsoSphereMapΩ-pres∙Π n p q i ∣₂) + + +-- module _ {ℓ : Level} {A : Pointed ℓ} where +-- suspension→ : (n : ℕ) → (S₊∙ (suc n) →∙ Ω A) → (S₊∙ (suc (suc n)) →∙ A) +-- fst (suspension→ n f) north = pt A +-- fst (suspension→ n f) south = pt A +-- fst (suspension→ n f) (merid a i₁) = fst f a i₁ +-- snd (suspension→ n f) = refl + +-- suspension-pres∙ : +-- (n : ℕ) (f g : (S₊∙ (suc n) →∙ Ω A)) +-- → suspension→ n (∙Π f g) ≡ ∙Π (suspension→ n f) (suspension→ n g) +-- suspension-pres∙ n f g = ΣPathP ((funExt main) , refl) +-- where +-- J2 : ∀ {ℓ ℓ'} {A : Type ℓ} {x : A} +-- → (P : (p q : x ≡ x) → (refl ≡ p) → (refl ≡ q) → Type ℓ') +-- → P refl refl refl refl +-- → (p q : x ≡ x) (P' : _) (Q : _) → P p q P' Q +-- J2 P b p q = +-- J (λ p P' → (Q₁ : refl ≡ q) → P p q P' Q₁) +-- (J (λ q Q → P refl q refl Q) b) + +-- J4 : ∀ {ℓ ℓ'} {A : Type ℓ} {x : A} +-- → (P : (p q r s : x ≡ x) +-- → (refl ≡ p) → (p ≡ q) → (refl ≡ r) → (r ≡ s) → Type ℓ') +-- → P refl refl refl refl refl refl refl refl +-- → (p q r s : x ≡ x) (P' : _) (Q : _) (R : _) (S : _) → P p q r s P' Q R S +-- J4 P b p q r s = +-- J (λ p P' → (Q₁ : p ≡ q) (R : refl ≡ r) (S₁ : r ≡ s) +-- → P p q r s P' Q₁ R S₁) +-- (J (λ q Q₁ → (R : refl ≡ r) (S₁ : r ≡ s) +-- → P refl q r s refl Q₁ R S₁) +-- (J (λ r R → (S₁ : r ≡ s) +-- → P refl refl r s refl refl R S₁) +-- (J (λ s S₁ → P refl refl refl s refl refl refl S₁) +-- b))) + +-- looper-help : (n : ℕ) (f g : (S₊∙ (suc n) →∙ Ω A)) (a : _) +-- → fst (∙Π f g) a ≡ fst f a ∙ fst g a +-- looper-help zero f g base = +-- rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g)) +-- looper-help zero f g (loop i) k = +-- lazy-fill _ _ +-- (sym (snd f)) (sym (snd g)) +-- (cong (fst f) loop) (cong (fst g) loop) k i +-- where +-- lazy-fill : +-- ∀ {ℓ} {A : Type ℓ} {x : A} (p q : x ≡ x) +-- (prefl : refl ≡ p) (qrefl : refl ≡ q) (fl : p ≡ p) (fr : q ≡ q) +-- → PathP (λ i → (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i +-- ≡ (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i) +-- ((prefl ∙∙ fl ∙∙ sym prefl) ∙ (qrefl ∙∙ fr ∙∙ sym qrefl)) +-- (cong₂ _∙_ fl fr) +-- lazy-fill {x = x} = +-- J2 (λ p q prefl qrefl → (fl : p ≡ p) (fr : q ≡ q) +-- → PathP (λ i → (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i +-- ≡ (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i) +-- ((prefl ∙∙ fl ∙∙ sym prefl) ∙ (qrefl ∙∙ fr ∙∙ sym qrefl)) +-- (cong₂ _∙_ fl fr)) +-- λ fl fr → cong₂ _∙_ (sym (rUnit fl)) (sym (rUnit fr)) +-- ◁ flipSquare (sym (rUnit (rUnit refl)) +-- ◁ (flipSquare ((λ j → cong (λ x → rUnit x j) fl +-- ∙ cong (λ x → lUnit x j) fr) +-- ▷ sym (cong₂Funct _∙_ fl fr)) +-- ▷ (rUnit (rUnit refl)))) +-- looper-help (suc n) f g north = +-- rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g)) +-- looper-help (suc n) f g south = +-- (rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) +-- ∙ cong₂ _∙_ (cong (fst f) (merid (ptSn _))) +-- (cong (fst g) (merid (ptSn _))) +-- looper-help (suc n) f g (merid a j) i = help i j +-- where +-- help : +-- PathP (λ i → (rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) i +-- ≡ ((rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) +-- ∙ cong₂ _∙_ (cong (fst f) (merid (ptSn _))) +-- (cong (fst g) (merid (ptSn _)))) i) +-- (cong (fst (∙Π f g)) (merid a)) +-- (cong₂ _∙_ (cong (fst f) (merid a)) +-- (cong (fst g) (merid a))) +-- help = (cong₂ _∙_ (cong (sym (snd f) ∙∙_∙∙ snd f) +-- (cong-∙ (fst f) +-- (merid a) (sym (merid (ptSn _))))) +-- ((cong (sym (snd g) ∙∙_∙∙ snd g) +-- (cong-∙ (fst g) +-- (merid a) (sym (merid (ptSn _)))))) +-- ◁ lazy-help _ _ _ _ (sym (snd f)) (cong (fst f) (merid (ptSn _))) +-- (sym (snd g)) (cong (fst g) (merid (ptSn _))) +-- (cong (fst f) (merid a)) (cong (fst g) (merid a))) +-- where +-- lazy-help : +-- ∀ {ℓ} {A : Type ℓ} {x : A} +-- (p-north p-south q-north q-south : x ≡ x) +-- (prefl : refl ≡ p-north) (p-merid : p-north ≡ p-south) +-- (qrefl : refl ≡ q-north) (q-merid : q-north ≡ q-south) +-- (fl : p-north ≡ p-south) (fr : q-north ≡ q-south) +-- → PathP (λ i → (rUnit refl ∙ cong₂ _∙_ prefl qrefl) i +-- ≡ ((rUnit refl ∙ cong₂ _∙_ prefl qrefl) +-- ∙ cong₂ _∙_ p-merid q-merid) i) +-- ((prefl ∙∙ fl ∙ sym p-merid ∙∙ sym prefl) +-- ∙ (qrefl ∙∙ fr ∙ sym q-merid ∙∙ sym qrefl)) +-- (cong₂ _∙_ fl fr) +-- lazy-help {x = x} = +-- J4 (λ p-north p-south q-north q-south prefl p-merid qrefl q-merid +-- → (fl : p-north ≡ p-south) (fr : q-north ≡ q-south) → +-- PathP (λ i₁ → (rUnit refl ∙ cong₂ _∙_ prefl qrefl) i₁ ≡ +-- ((rUnit refl ∙ cong₂ _∙_ prefl qrefl) +-- ∙ cong₂ _∙_ p-merid q-merid) i₁) +-- ((prefl ∙∙ fl ∙ sym p-merid ∙∙ sym prefl) ∙ +-- (qrefl ∙∙ fr ∙ sym q-merid ∙∙ sym qrefl)) +-- (cong₂ _∙_ fl fr)) +-- λ fl fr → (cong₂ _∙_ (sym (rUnit (fl ∙ refl)) ∙ sym (rUnit fl)) +-- (sym (rUnit (fr ∙ refl)) ∙ sym (rUnit fr)) +-- ◁ flipSquare ((sym (rUnit _) +-- ◁ flipSquare +-- ((λ j → cong (λ x → rUnit x j) fl +-- ∙ cong (λ x → lUnit x j) fr) +-- ▷ sym (cong₂Funct _∙_ fl fr))) +-- ▷ (rUnit _ ∙ rUnit _))) +-- main : (x : fst (S₊∙ (suc (suc n)))) → +-- fst (suspension→ n (∙Π f g)) x ≡ +-- fst (∙Π (suspension→ n f) (suspension→ n g)) x +-- main north = refl +-- main south = refl +-- main (merid a i₁) j = help j i₁ +-- where +-- help : fst (∙Π f g) a +-- ≡ cong (fst (∙Π (suspension→ n f) (suspension→ n g))) (merid a) +-- help = looper-help n f g a +-- ∙ cong₂ _∙_ (sym (cong-∙ (fst (suspension→ n f)) +-- (merid a) +-- (sym (merid (ptSn _))) +-- ∙∙ cong (fst f a ∙_) (cong sym (snd f)) +-- ∙∙ sym (rUnit _))) +-- (sym (cong-∙ (fst (suspension→ n g)) +-- (merid a) +-- (sym (merid (ptSn _))) +-- ∙∙ cong (fst g a ∙_) (cong sym (snd g)) +-- ∙∙ sym (rUnit _))) +-- ∙ λ i → rUnit (cong (fst (suspension→ n f)) +-- (merid a ∙ sym (merid (ptSn _)))) i +-- ∙ rUnit (cong (fst (suspension→ n g)) +-- (merid a ∙ sym (merid (ptSn _)))) i + +-- suspension←filler : (n : ℕ) → (S₊∙ (suc (suc n)) →∙ A) +-- → I → I → I → fst A +-- suspension←filler n f i j k = +-- hfill (λ k → λ { (i = i0) → doubleCompPath-filler +-- (sym (snd f)) +-- (cong (fst f) (merid (ptSn _) +-- ∙ sym (merid (ptSn _)))) +-- (snd f) k j +-- ; (i = i1) → snd f k +-- ; (j = i0) → snd f k +-- ; (j = i1) → snd f k}) +-- (inS ((fst f) (rCancel' (merid (ptSn _)) i j))) +-- k + +-- suspension← : (n : ℕ) → (S₊∙ (suc (suc n)) →∙ A) → (S₊∙ (suc n) →∙ Ω A) +-- fst (suspension← n f) a = +-- sym (snd f) ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f +-- snd (suspension← n f) i j = suspension←filler n f i j i1 + +-- suspensionIso : (n : ℕ) → Iso (S₊∙ (suc n) →∙ Ω A) (S₊∙ (suc (suc n)) →∙ A) +-- Iso.fun (suspensionIso n) f = suspension→ n f +-- Iso.inv (suspensionIso n) f = suspension← n f +-- Iso.rightInv (suspensionIso n) (f , p) = +-- ΣPathP (funExt help , λ i j → p (~ i ∨ j)) +-- where +-- help : (x : _) → fst (suspension→ n (suspension← n (f , p))) x ≡ f x +-- help north = sym p +-- help south = sym p ∙ cong f (merid (ptSn _)) +-- help (merid a j) i = +-- hcomp (λ k → λ { (i = i1) → f (merid a j) +-- ; (j = i0) → p (~ i ∧ k) +-- ; (j = i1) → compPath-filler' (sym p) +-- (cong f (merid (ptSn _))) k i}) +-- (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) +-- Iso.leftInv (suspensionIso n) f = +-- ΣPathP (funExt lem , PathP-filler) +-- where + +-- side₁ : (x : S₊ (suc n)) → I → I → I → fst A +-- side₁ x i j k = +-- hfill (λ k → λ { (i = i0) → suspension→ n f .fst +-- (compPath-filler' (merid x) +-- (sym (merid (ptSn _))) k j) +-- ; (i = i1) → snd A +-- ; (j = i0) → fst f x (i ∨ ~ k) +-- ; (j = i1) → snd A}) +-- (inS (snd f i (~ j))) +-- k + +-- side₂ : (x : S₊ (suc n)) → I → I → I → fst A +-- side₂ x i j k = +-- hfill (λ k → λ { (i = i0) → doubleCompPath-filler +-- refl +-- (cong (suspension→ n f .fst) +-- (merid x ∙ sym (merid (ptSn _)))) +-- refl k j +-- ; (i = i1) → fst f x (j ∨ ~ k) +-- ; (j = i0) → fst f x (i ∧ ~ k) +-- ; (j = i1) → snd A}) +-- (inS (side₁ x i j i1)) +-- k + +-- compPath-filler'' : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) +-- → I → I → I → A +-- compPath-filler'' {z = z} p q k j i = +-- hfill (λ k → λ { (i = i0) → p (~ j) +-- ; (i = i1) → q k +-- ; (j = i0) → q (i ∧ k) }) +-- (inS (p (i ∨ ~ j))) +-- k + +-- lem : (x : _) → fst (suspension← n (suspension→ n f)) x ≡ fst f x +-- lem x i j = side₂ x i j i1 + +-- sideCube : Cube (λ j k → snd f j (~ k)) +-- (λ j k → fst (suspension→ n f) +-- (rCancel' (merid (ptSn _)) j k)) +-- (λ r k → suspension→ n f .fst +-- (compPath-filler' (merid (ptSn _)) +-- (sym (merid (ptSn _))) r k)) +-- (λ _ _ → pt A) +-- (λ r j → snd f j (~ r)) λ _ _ → pt A +-- sideCube r j k = +-- hcomp (λ i → λ {(r = i0) → snd f j (~ k ∨ ~ i) +-- ; (r = i1) → fst (suspension→ n f) +-- (rCancel-filler' (merid (ptSn _)) (~ i) j k) +-- ; (j = i0) → suspension→ n f .fst +-- (compPath-filler'' (merid (ptSn _)) +-- (sym (merid (ptSn _))) i r k) +-- ; (j = i1) → snd f (~ r) (~ i ∧ k) +-- ; (k = i0) → snd f j (~ r) +-- ; (k = i1) → snd f (~ r ∧ j) (~ i)}) +-- (hcomp (λ i → λ {(r = i0) → pt A +-- ; (r = i1) → snd f (~ i) k +-- ; (j = i0) → snd f (~ i) (k ∨ ~ r) +-- ; (j = i1) → snd f (~ r ∨ ~ i) k +-- ; (k = i0) → snd f (j ∨ ~ i) (~ r) +-- ; (k = i1) → pt A}) +-- (pt A)) + +-- PathP-filler : PathP _ _ _ +-- PathP-filler i j k = +-- hcomp (λ r → λ {(i = i0) → suspension←filler n (suspension→ n f) j k r +-- ; (i = i1) → snd f j (k ∨ ~ r) +-- ; (j = i0) → side₂ (ptSn _) i k r +-- ; (j = i1) → snd A +-- ; (k = i0) → snd f j (i ∧ ~ r) +-- ; (k = i1) → snd A}) +-- (hcomp ((λ r → λ {(i = i0) → sideCube r j k +-- ; (i = i1) → pt A +-- ; (j = i0) → side₁ (ptSn _) i k r +-- ; (j = i1) → pt A +-- ; (k = i0) → snd f j (i ∨ ~ r) +-- ; (k = i1) → pt A})) +-- (snd f (i ∨ j) (~ k))) + +-- suspension←-pres∙ : (n : ℕ) → (f g : S₊∙ (suc (suc n)) →∙ A) +-- → suspension← n (∙Π f g) +-- ≡ ∙Π (suspension← n f) (suspension← n g) +-- suspension←-pres∙ n f g = +-- sym (Iso.leftInv (suspensionIso n) (suspension← n (∙Π f g))) +-- ∙∙ cong (suspension← n) +-- (Iso.rightInv (suspensionIso n) (∙Π f g) +-- ∙∙ cong₂ ∙Π (sym (Iso.rightInv (suspensionIso n) f)) +-- (sym (Iso.rightInv (suspensionIso n) g)) +-- ∙∙ sym (suspension-pres∙ _ (suspension← n f) (suspension← n g))) +-- ∙∙ (Iso.leftInv (suspensionIso n) (∙Π (suspension← n f) (suspension← n g))) + +-- SphereMap→Ω : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) +-- → (S₊∙ (suc n) →∙ A) → (fst ((Ω^ (suc n)) A)) +-- SphereMap→Ω zero (f , p) = sym p ∙∙ cong f loop ∙∙ p +-- SphereMap→Ω {A = A} (suc n) (f , p) = +-- Iso.inv (flipΩIso _) +-- (SphereMap→Ω {A = Ω A} n (Iso.inv (suspensionIso _) (f , p))) + +-- SphereMap→Ωpres∙ : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) +-- → (p q : S₊∙ (suc n) →∙ A) +-- → SphereMap→Ω n (∙Π p q) ≡ SphereMap→Ω n p ∙ SphereMap→Ω n q +-- SphereMap→Ωpres∙ zero f g = sym (rUnit _) +-- SphereMap→Ωpres∙ (suc n) f g = +-- cong (Iso.inv (flipΩIso (suc n))) +-- (cong (SphereMap→Ω n) (suspension←-pres∙ n f g) +-- ∙ SphereMap→Ωpres∙ n (suspension← n f) (suspension← n g)) +-- ∙ flipΩIsopres· n _ _ + +-- Ω→SphereMap : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) +-- → (fst ((Ω^ (suc n)) A)) → (S₊∙ (suc n) →∙ A) +-- fst (Ω→SphereMap {A = A} zero f) base = pt A +-- fst (Ω→SphereMap {A = A} zero f) (loop i₁) = f i₁ +-- snd (Ω→SphereMap {A = A} zero f) = refl +-- Ω→SphereMap {A = A} (suc n) f = +-- Iso.fun (suspensionIso _) (Ω→SphereMap n (Iso.fun (flipΩIso _) f)) + +-- Ω→SphereMap→Ω : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) +-- → (x : fst ((Ω^ (suc n)) A)) +-- → SphereMap→Ω n (Ω→SphereMap n x) ≡ x +-- Ω→SphereMap→Ω zero x = sym (rUnit _) +-- Ω→SphereMap→Ω (suc n) x = +-- cong (Iso.inv (flipΩIso (suc n)) ∘ SphereMap→Ω n) +-- (Iso.leftInv (suspensionIso _) +-- (Ω→SphereMap n (Iso.fun (flipΩIso (suc n)) x))) +-- ∙∙ cong (Iso.inv (flipΩIso (suc n))) +-- (Ω→SphereMap→Ω n (Iso.fun (flipΩIso (suc n)) x)) +-- ∙∙ Iso.leftInv (flipΩIso (suc n)) x + +-- SphereMap→Ω→SphereMap : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) +-- → (x : (S₊∙ (suc n) →∙ A)) +-- → Ω→SphereMap n (SphereMap→Ω n x) ≡ x +-- SphereMap→Ω→SphereMap zero (f , p) = +-- ΣPathP (funExt fstp , λ i j → p (~ i ∨ j)) +-- where +-- fstp : (x : _) → _ ≡ f x +-- fstp base = sym p +-- fstp (loop i) j = doubleCompPath-filler (sym p) (cong f loop) p (~ j) i +-- SphereMap→Ω→SphereMap (suc n) f = +-- cong (Iso.fun (suspensionIso n) ∘ Ω→SphereMap n) +-- (Iso.rightInv (flipΩIso (suc n)) +-- (SphereMap→Ω n (Iso.inv (suspensionIso n) f))) +-- ∙∙ cong (Iso.fun (suspensionIso n)) +-- (SphereMap→Ω→SphereMap n (Iso.inv (suspensionIso n) f)) +-- ∙∙ Iso.rightInv (suspensionIso _) _ diff --git a/Cubical/Homotopy/Group/Properties.agda b/Cubical/Homotopy/Group/Properties.agda new file mode 100644 index 000000000..fa1033dde --- /dev/null +++ b/Cubical/Homotopy/Group/Properties.agda @@ -0,0 +1,710 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Homotopy.Group.Properties where + +open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.Group.S3 +open import Cubical.Homotopy.Group.S3 + using (IsoΩSphereMap) public +open import Cubical.Homotopy.Loopspace + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Path +open import Cubical.Foundations.Transport +open import Cubical.Functions.Embedding +open import Cubical.Functions.Morphism +open import Cubical.Foundations.Pointed.Homogeneous + +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim + ; elim2 to sElim2 ; elim3 to sElim3 ; map to sMap) +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; rec2 to pRec2 ; elim to pElim) +open import Cubical.HITs.Sn +open import Cubical.Data.Bool +open import Cubical.HITs.Susp +open import Cubical.HITs.S1 +open import Cubical.Data.Sigma +open import Cubical.Data.Nat + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Semigroup +open import Cubical.Algebra.Monoid + +open Iso +open IsGroup +open IsSemigroup +open IsMonoid +open GroupStr + +open import Cubical.Homotopy.Freudenthal +open import Cubical.Homotopy.Connected +open import Cubical.HITs.Truncation + renaming (rec to trRec) + + +open import Cubical.Data.Sum +open import Cubical.Relation.Nullary +open import Cubical.Data.Empty + +open import Cubical.Foundations.Equiv.HalfAdjoint + +-- We finally arrive at the main result + +IsoSphereMapΩ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (S₊∙ n →∙ A) (fst ((Ω^ n) A)) +IsoSphereMapΩ n = invIso (IsoΩSphereMap n) + +SphereMap→Ω : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (S₊∙ n →∙ A) → (fst ((Ω^ n) A)) +SphereMap→Ω n = fun (IsoSphereMapΩ n) + +Ω→SphereMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (fst ((Ω^ n) A)) → (S₊∙ n →∙ A) +Ω→SphereMap n = inv (IsoSphereMapΩ n) + +isHom-lMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (p q : _) + → lMap (suc n) {A = A} (p ∙ q) + ≡ ∙Π (lMap (suc n) p) (lMap (suc n) q) +isHom-lMap zero p q = + ΣPathP ((funExt (λ { base → refl + ; (loop i) j → (rUnit p j ∙ rUnit q j) i})) + , refl) +isHom-lMap (suc n) {A = A} p q = + ΣPathP ((funExt (λ { north → refl + ; south → refl + ; (merid a i) j → h a j i})) + , refl) + where + doubleComp-lem : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) (q r : y ≡ y) + → (p ∙∙ q ∙∙ sym p) ∙ (p ∙∙ r ∙∙ sym p) + ≡ (p ∙∙ (q ∙ r) ∙∙ sym p) + doubleComp-lem p q r i j = + hcomp (λ k → λ { (i = i0) → (doubleCompPath-filler p q (sym p) k + ∙ doubleCompPath-filler p r (sym p) k) j + ; (i = i1) → doubleCompPath-filler p (q ∙ r) (sym p) k j + ; (j = i0) → p (~ k) + ; (j = i1) → p (~ k)}) + ((q ∙ r) j) + + help2 : (p : typ ((Ω^ (suc (suc n))) A)) + → cong (fst (lMap (suc (suc n)) p)) (merid (ptSn _)) ≡ refl + help2 p = + cong (sym (lMapId (suc n) (ptSn _)) ∙∙_∙∙ lMapId (suc n) (ptSn _)) + (rUnit _ ∙ (λ j → (λ i → lMap (suc n) {A = A} refl .snd (i ∧ j)) + ∙∙ (λ i → lMap (suc n) {A = A} (p i) .snd j) + ∙∙ λ i → lMap (suc n) {A = A} refl .snd (~ i ∧ j)) + ∙ ∙∙lCancel _) + ∙ ∙∙lCancel _ + + h : (a : S₊ (suc n)) + → sym (lMapId (suc n) a) + ∙∙ funExt⁻ (cong fst (cong (lMap (suc n)) (p ∙ q))) a + ∙∙ lMapId (suc n) a + ≡ cong (fst (∙Π (lMap (suc (suc n)) p) (lMap (suc (suc n)) q))) (merid a) + h a = (cong (sym (lMapId (suc n) a) ∙∙_∙∙ (lMapId (suc n) a)) + (cong-∙ (λ x → lMap (suc n) x .fst a) p q) + ∙ sym (doubleComp-lem (sym (lMapId (suc n) a)) _ _)) + ∙∙ cong₂ _∙_ (sym (cong (cong (fst (lMap (suc (suc n)) p)) (merid a) ∙_) + (cong sym (help2 p)) ∙ sym (rUnit _))) + (sym (cong (cong (fst (lMap (suc (suc n)) q)) (merid a) ∙_) + (cong sym (help2 q)) ∙ sym (rUnit _))) + ∙∙ λ i → (rUnit (cong-∙ (fst (lMap (suc (suc n)) p)) + (merid a) (sym (merid (ptSn _))) (~ i)) i) + ∙ (rUnit (cong-∙ (fst (lMap (suc (suc n)) q)) + (merid a) (sym (merid (ptSn _)))(~ i)) i) + +-- The iso is structure preserving +IsoSphereMapΩ-pres∙Π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f g : S₊∙ (suc n) →∙ A) + → SphereMap→Ω (suc n) (∙Π f g) + ≡ SphereMap→Ω (suc n) f ∙ SphereMap→Ω (suc n) g +IsoSphereMapΩ-pres∙Π n = + morphLemmas.isMorphInv _∙_ ∙Π (Ω→SphereMap (suc n)) + (isHom-lMap n) + (SphereMap→Ω (suc n)) + (Iso.rightInv (IsoΩSphereMap (suc n))) + (Iso.leftInv (IsoΩSphereMap (suc n))) + +-- It is useful to define the ``Group Structure'' on (S₊∙ n →∙ A) +-- before doing it on π'. These will be the equivalents of the +-- usual groupoid laws on Ω A. +1Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → (S₊∙ n →∙ A) +fst (1Π {A = A}) _ = pt A +snd (1Π {A = A}) = refl + +∙Π-rUnit : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ∙Π f 1Π ≡ f +fst (∙Π-rUnit {A = A} {n = zero} f i) base = snd f (~ i) +fst (∙Π-rUnit {A = A} {n = zero} f i) (loop j) = help i j + where + help : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) + (((sym (snd f)) ∙∙ (cong (fst f) loop) ∙∙ snd f) + ∙ (refl ∙ refl)) + (cong (fst f) loop) + help = (cong ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) ∙_) + (sym (rUnit refl)) ∙ sym (rUnit _)) + ◁ λ i j → doubleCompPath-filler (sym (snd f)) + (cong (fst f) loop) (snd f) (~ i) j +snd (∙Π-rUnit {A = A} {n = zero} f i) j = snd f (~ i ∨ j) +fst (∙Π-rUnit {A = A} {n = suc n} f i) north = snd f (~ i) +fst (∙Π-rUnit {A = A} {n = suc n} f i) south = + (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i +fst (∙Π-rUnit {A = A} {n = suc n} f i) (merid a j) = help i j + where + help : PathP (λ i → snd f (~ i) + ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) + (((sym (snd f)) + ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + ∙∙ snd f) + ∙ (refl ∙ refl)) + (cong (fst f) (merid a)) + help = (cong (((sym (snd f)) + ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + ∙∙ snd f) ∙_) + (sym (rUnit refl)) + ∙ sym (rUnit _)) + ◁ λ i j → hcomp (λ k → + λ { (j = i0) → snd f (~ i ∧ k) + ; (j = i1) → compPath-filler' (sym (snd f)) + (cong (fst f) (merid (ptSn (suc n)))) k i + ; (i = i0) → doubleCompPath-filler (sym (snd f)) + (cong (fst f) + (merid a ∙ sym (merid (ptSn (suc n))))) + (snd f) k j + ; (i = i1) → fst f (merid a j)}) + (fst f (compPath-filler (merid a) + (sym (merid (ptSn _))) (~ i) j)) +snd (∙Π-rUnit {A = A} {n = suc n} f i) j = snd f (~ i ∨ j) + +∙Π-lUnit : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ∙Π 1Π f ≡ f +fst (∙Π-lUnit {n = zero} f i) base = snd f (~ i) +fst (∙Π-lUnit {n = zero} f i) (loop j) = s i j + where + s : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) + ((refl ∙ refl) ∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) + (cong (fst f) loop) + s = (cong (_∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) + (sym (rUnit refl)) ∙ sym (lUnit _)) + ◁ λ i j → doubleCompPath-filler (sym (snd f)) + (cong (fst f) loop) (snd f) (~ i) j +snd (∙Π-lUnit {n = zero} f i) j = snd f (~ i ∨ j) +fst (∙Π-lUnit {n = suc n} f i) north = snd f (~ i) +fst (∙Π-lUnit {n = suc n} f i) south = + (sym (snd f) ∙ cong (fst f) (merid (ptSn _))) i +fst (∙Π-lUnit {n = suc n} f i) (merid a j) = help i j + where + help : PathP (λ i → snd f (~ i) + ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) + ((refl ∙ refl) ∙ ((sym (snd f)) + ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + ∙∙ snd f)) + (cong (fst f) (merid a)) + help = + (cong (_∙ ((sym (snd f)) + ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + ∙∙ snd f)) + (sym (rUnit refl)) + ∙ sym (lUnit _)) + ◁ λ i j → hcomp (λ k → + λ { (j = i0) → snd f (~ i ∧ k) + ; (j = i1) → compPath-filler' (sym (snd f)) + (cong (fst f) (merid (ptSn (suc n)))) k i + ; (i = i0) → doubleCompPath-filler (sym (snd f)) + (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + (snd f) k j + ; (i = i1) → fst f (merid a j)}) + (fst f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) +snd (∙Π-lUnit {n = suc n} f i) j = snd f (~ i ∨ j) + +∙Π-rCancel : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ∙Π f (-Π f) ≡ 1Π +fst (∙Π-rCancel {A = A} {n = zero} f i) base = pt A +fst (∙Π-rCancel {A = A} {n = zero} f i) (loop j) = + rCancel (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) i j +snd (∙Π-rCancel {A = A} {n = zero} f i) = refl +fst (∙Π-rCancel {A = A} {n = suc n} f i) north = pt A +fst (∙Π-rCancel {A = A} {n = suc n} f i) south = pt A +fst (∙Π-rCancel {A = A} {n = suc n} f i) (merid a i₁) = sl i i₁ + where + pl = (sym (snd f) + ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) + ∙∙ snd f) + sl : pl + ∙ ((sym (snd f) + ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) + ∙∙ snd f)) ≡ refl + sl = cong (pl ∙_) (cong (sym (snd f) ∙∙_∙∙ (snd f)) + (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) + ∙∙ cong₂ _∙_ refl + (cong (cong (fst f)) (rCancel (merid (ptSn _)))) + ∙∙ sym (rUnit _))) + ∙ rCancel pl +snd (∙Π-rCancel {A = A} {n = suc n} f i) = refl + +∙Π-lCancel : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ∙Π (-Π f) f ≡ 1Π +fst (∙Π-lCancel {A = A} {n = zero} f i) base = pt A +fst (∙Π-lCancel {A = A} {n = zero} f i) (loop j) = + rCancel (sym (snd f) ∙∙ cong (fst f) (sym loop) ∙∙ snd f) i j +fst (∙Π-lCancel {A = A} {n = suc n} f i) north = pt A +fst (∙Π-lCancel {A = A} {n = suc n} f i) south = pt A +fst (∙Π-lCancel {A = A} {n = suc n} f i) (merid a j) = sl i j + where + pl = (sym (snd f) + ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) + ∙∙ snd f) + + sl : (sym (snd f) + ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) + ∙∙ snd f) ∙ pl + ≡ refl + sl = cong (_∙ pl) (cong (sym (snd f) ∙∙_∙∙ (snd f)) + (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) + ∙∙ cong₂ _∙_ refl (cong (cong (fst f)) (rCancel (merid (ptSn _)))) + ∙∙ sym (rUnit _))) + ∙ lCancel pl +snd (∙Π-lCancel {A = A} {n = zero} f i) = refl +snd (∙Π-lCancel {A = A} {n = suc n} f i) = refl + +∙Π-assoc : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f g h : S₊∙ (suc n) →∙ A) + → ∙Π f (∙Π g h) ≡ ∙Π (∙Π f g) h +∙Π-assoc {n = n} f g h = + sym (Iso.leftInv (IsoSphereMapΩ (suc n)) (∙Π f (∙Π g h))) + ∙∙ cong (Ω→SphereMap (suc n)) (IsoSphereMapΩ-pres∙Π n f (∙Π g h) + ∙∙ cong (SphereMap→Ω (suc n) f ∙_) (IsoSphereMapΩ-pres∙Π n g h) + ∙∙ ∙assoc (SphereMap→Ω (suc n) f) (SphereMap→Ω (suc n) g) (SphereMap→Ω (suc n) h) + ∙∙ cong (_∙ SphereMap→Ω (suc n) h) (sym (IsoSphereMapΩ-pres∙Π n f g)) + ∙∙ sym (IsoSphereMapΩ-pres∙Π n (∙Π f g) h)) + ∙∙ Iso.leftInv (IsoSphereMapΩ (suc n)) (∙Π (∙Π f g) h) + +∙Π-comm : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f g : S₊∙ (suc (suc n)) →∙ A) + → ∙Π f g ≡ ∙Π g f +∙Π-comm {A = A} {n = n} f g = + sym (Iso.leftInv (IsoSphereMapΩ (suc (suc n))) (∙Π f g)) + ∙∙ cong (Ω→SphereMap (suc (suc n))) (IsoSphereMapΩ-pres∙Π (suc n) f g + ∙∙ EH _ _ _ + ∙∙ sym (IsoSphereMapΩ-pres∙Π (suc n) g f)) + ∙∙ Iso.leftInv (IsoSphereMapΩ (suc (suc n))) (∙Π g f) + +{- π'' as a group -} +1π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' n A +1π' n {A = A} = ∣ 1Π ∣₂ + +·π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' (suc n) A → π' (suc n) A → π' (suc n) A +·π' n = sRec2 squash₂ λ p q → ∣ ∙Π p q ∣₂ + +-π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' (suc n) A → π' (suc n) A +-π' n = sMap -Π + +π'-rUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) + → (·π' n x (1π' (suc n))) ≡ x +π'-rUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-rUnit p i ∣₂ + +π'-lUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) + → (·π' n (1π' (suc n)) x) ≡ x +π'-lUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-lUnit p i ∣₂ + +π'-rCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) + → (·π' n x (-π' n x)) ≡ 1π' (suc n) +π'-rCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-rCancel p i ∣₂ + +π'-lCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) + → (·π' n (-π' n x) x) ≡ 1π' (suc n) +π'-lCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-lCancel p i ∣₂ + +π'-assoc : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y z : π' (suc n) A) + → ·π' n x (·π' n y z) ≡ ·π' n (·π' n x y) z +π'-assoc n = sElim3 (λ _ _ _ → isSetPathImplicit) λ p q r i → ∣ ∙Π-assoc p q r i ∣₂ + +π'-comm : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y : π' (suc (suc n)) A) + → ·π' (suc n) x y ≡ ·π' (suc n) y x +π'-comm n = sElim2 (λ _ _ → isSetPathImplicit) λ p q i → ∣ ∙Π-comm p q i ∣₂ + +-- We finally get the group definition +π'Gr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Group ℓ +fst (π'Gr n A) = π' (suc n) A +1g (snd (π'Gr n A)) = 1π' (suc n) +GroupStr._·_ (snd (π'Gr n A)) = ·π' n +inv (snd (π'Gr n A)) = -π' n +is-set (isSemigroup (isMonoid (isGroup (snd (π'Gr n A))))) = squash₂ +assoc (isSemigroup (isMonoid (isGroup (snd (π'Gr n A))))) = π'-assoc n +identity (isMonoid (isGroup (snd (π'Gr n A)))) x = (π'-rUnit n x) , (π'-lUnit n x) +inverse (isGroup (snd (π'Gr n A))) x = (π'-rCancel n x) , (π'-lCancel n x) + +-- and finally, the Iso +π'Gr≅πGr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → GroupIso (π'Gr n A) (πGr n A) +fst (π'Gr≅πGr n A) = setTruncIso (IsoSphereMapΩ (suc n)) +snd (π'Gr≅πGr n A) = + makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) + λ p q i → ∣ IsoSphereMapΩ-pres∙Π n p q i ∣₂) + +{- +In file X we gave a filler of the following square + suspMapΩ +Ωⁿ A -------------------> Ωⁿ⁺¹ (Susp A) + | | + | | + | ≃ eq₁ | ≃ eq₂ + | | + v suspMap v + (Sⁿ →∙ A) -------------- > (Sⁿ⁺¹ →∙ Susp A) + +We would like this to preserve composition +(in order to deduce properties of suspMap from the corresponding +ones of suspMapΩ). For this, we need that both equivalences +on the sides are structure preserving +-} + +-- We define an alternative notion of composition (to ∙Π) on +-- Sⁿ →∙ Ω (Susp∙ (typ A)). It will be easier to work with +private + invComp : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → S₊∙ n →∙ Ω (Susp∙ (typ A)) + → S₊∙ n →∙ Ω (Susp∙ (typ A)) + → S₊∙ n →∙ Ω (Susp∙ (typ A)) + fst (invComp n f g) x = (fst f x) ∙ (fst g x) + snd (invComp n f g) = cong₂ _∙_ (snd f) (snd g) ∙ sym (rUnit refl) + + -- We prove that it agrees with ∙Π + ∙Π≡invComp : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (f g : S₊∙ (suc n) →∙ Ω (Susp∙ (typ A))) + → ∙Π f g ≡ invComp {A = A} (suc n) f g + ∙Π≡invComp zero f g = + →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt (λ { base → rUnit refl + ∙ sym (cong (_∙ fst g (snd (S₊∙ 1))) (snd f) + ∙ cong (refl ∙_) (snd g)) + ; (loop i) j → + hcomp (λ k → + λ { (i = i0) → (rUnit refl + ∙ sym (cong (_∙ fst g (snd (S₊∙ 1))) (snd f) + ∙ cong (refl ∙_) (snd g))) j + ; (i = i1) → (rUnit refl + ∙ sym (cong (_∙ fst g (snd (S₊∙ 1))) (snd f) + ∙ cong (refl ∙_) (snd g))) j + ; (j = i0) → ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) + ∙ (sym (snd g) ∙∙ cong (fst g) loop ∙∙ snd g)) i + ; (j = i1) → cong₂Funct _∙_ + (cong (fst f) loop) (cong (fst g) loop) (~ k) i}) + (hcomp (λ k → + λ { (i = i0) → (rUnit refl + ∙ sym (cong (_∙ snd g (~ k)) (λ j → snd f (j ∨ ~ k)) + ∙ cong (refl ∙_) (λ j → snd g (j ∨ ~ k)))) j + ; (i = i1) → (rUnit refl + ∙ sym (cong (_∙ snd g (~ k)) (λ j → snd f (j ∨ ~ k)) + ∙ cong (refl ∙_) (λ j → snd g (j ∨ ~ k)))) j + ; (j = i0) → ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) + ∙ (sym (snd g) ∙∙ cong (fst g) loop ∙∙ snd g)) i + ; (j = i1) → (cong (_∙ snd g (~ k)) + (doubleCompPath-filler (sym (snd f)) + (cong (fst f) loop) + (snd f) (~ k)) ∙ + cong ((snd f (~ k)) ∙_) + (doubleCompPath-filler (sym (snd g)) + (cong (fst g) loop) (snd g) (~ k))) i}) + (hcomp (λ k → + λ { (i = i0) → (rUnit (rUnit refl) + ∙ cong (rUnit refl ∙_) (cong sym (rUnit refl))) k j + ; (i = i1) → (rUnit (rUnit refl) + ∙ cong (rUnit refl ∙_) (cong sym (rUnit refl))) k j + ; (j = i0) → ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) + ∙ (sym (snd g) ∙∙ cong (fst g) loop ∙∙ snd g)) i + ; (j = i1) → (cong (_∙ refl) + ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) + ∙ cong (refl ∙_) + (sym (snd g) ∙∙ cong (fst g) loop ∙∙ snd g)) i}) + ((cong (λ x → rUnit x j) + (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) + ∙ cong (λ x → lUnit x j) + (sym (snd g) ∙∙ cong (fst g) loop ∙∙ snd g)) i)))})) + ∙Π≡invComp {A = A} (suc n) f g = + →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt λ { north → rUnit refl + ∙ sym (cong (fst f north ∙_) (snd g) + ∙ cong (_∙ refl) (snd f)) + ; south → rUnit refl + ∙ sym (cong₂ _∙_ + (cong (fst f) (sym (merid (ptSn _))) ∙ snd f) + (cong (fst g) (sym (merid (ptSn _))) ∙ snd g)) + ; (merid a i) j → p a i j}) + where + module _ (a : S₊ (suc n)) where + f-n = fst f north + g-n = fst g north + cong-f = (cong (fst f) (merid a ∙ sym (merid (ptSn _)))) + cong-g = (cong (fst g) (merid a ∙ sym (merid (ptSn _)))) + c-f = sym (snd f) ∙∙ cong-f ∙∙ snd f + c-g = sym (snd g) ∙∙ cong-g ∙∙ snd g + + p : I → I → fst (Ω (Susp∙ (typ A))) + p i j = + hcomp (λ k → + λ { (i = i0) → (rUnit (λ _ → snd (Susp∙ (typ A))) + ∙ sym ((cong (fst f north ∙_) (snd g) + ∙ cong (_∙ refl) (snd f)))) j + ; (i = i1) → (rUnit refl + ∙ sym (cong₂ _∙_ + (compPath-filler' + (cong (fst f) (sym (merid (ptSn _)))) (snd f) k) + (compPath-filler' + (cong (fst g) (sym (merid (ptSn _)))) (snd g) k))) j + ; (j = i0) → (c-f ∙ c-g) i + ; (j = i1) → fst f (compPath-filler + (merid a) + (sym (merid (ptSn _))) (~ k) i) + ∙ fst g (compPath-filler + (merid a) + (sym (merid (ptSn _))) (~ k) i)}) + (hcomp (λ k → + λ {(i = i0) → (rUnit (λ _ → snd (Susp∙ (typ A))) + ∙ sym ((cong (fst f north ∙_) (snd g) + ∙ cong (_∙ refl) (snd f)))) j + ; (i = i1) → (rUnit refl ∙ sym (cong₂ _∙_ (snd f) (snd g))) j + ; (j = i0) → (c-f ∙ c-g) i + ; (j = i1) → cong₂Funct _∙_ cong-f cong-g (~ k) i}) + (hcomp (λ k → + λ {(i = i0) → (rUnit refl + ∙ sym (compPath-filler' + ((cong (fst f north ∙_) (snd g))) + (cong (_∙ refl) (snd f)) k)) j + ; (i = i1) → (rUnit refl + ∙ sym (cong₂ _∙_ (λ i → snd f (i ∨ ~ k)) (snd g))) j + ; (j = i0) → (c-f ∙ c-g) i + ; (j = i1) → (cong (λ x → x ∙ snd g (~ k)) + (doubleCompPath-filler refl + cong-f (snd f) (~ k)) + ∙ cong ((snd f (~ k)) ∙_) + (doubleCompPath-filler + (sym (snd g)) cong-g refl (~ k))) i}) + (hcomp (λ k → + λ {(i = i0) → compPath-filler + (rUnit (λ _ → snd (Susp∙ (typ A)))) + (sym ((cong (_∙ refl) (snd f)))) k j + ; (i = i1) → compPath-filler + (rUnit refl) + (sym (cong (refl ∙_) (snd g))) k j + ; (j = i0) → (c-f ∙ c-g) i + ; (j = i1) → (cong (_∙ refl) + ((λ j → snd f (~ j ∧ ~ k)) ∙∙ cong-f ∙∙ snd f) + ∙ cong (refl ∙_) + (sym (snd g) ∙∙ cong-g ∙∙ (λ j → snd g (j ∧ ~ k)))) i}) + (((cong (λ x → rUnit x j) c-f) ∙ (cong (λ x → lUnit x j) c-g)) i)))) + + hom-botᵣ⁻ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (f g : S₊∙ (suc n) →∙ Susp∙ (typ A)) + → botᵣ⁻ {A = A} n (∙Π f g) + ≡ invComp {A = A} n (botᵣ⁻ {A = A} n f) (botᵣ⁻ {A = A} n g) + +-- We ge that +hom-botᵣ⁻ zero f g = + ΣPathP ((funExt (λ { false → sym (rUnit _) + ; true → (rUnit _)})) + , ((λ i j → rUnit refl (i ∧ ~ j)) + ▷ lUnit (sym (rUnit refl)))) +hom-botᵣ⁻ (suc n) f g = + →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt (λ x → (sym (rUnit (cong (fst (∙Π f g)) (merid x ∙ sym (merid (ptSn _)))))) + ∙∙ cong-∙ (fst (∙Π f g)) (merid x) (sym (merid (ptSn _))) + ∙∙ cong (cong (fst (∙Π f g)) (merid x) ∙_) (cong sym lem) + ∙ sym (rUnit (cong (fst (∙Π f g)) (merid x))))) + where + lem : cong (fst (∙Π f g)) (merid (ptSn (suc n))) ≡ refl + lem = (λ i → (sym (snd f) ∙∙ cong (fst f) (rCancel (merid (ptSn _)) i) ∙∙ snd f) + ∙ (sym (snd g) ∙∙ cong (fst g) (rCancel (merid (ptSn _)) i) ∙∙ snd g)) + ∙ (λ i → ∙∙lCancel (snd f) i ∙ ∙∙lCancel (snd g) i) + ∙ sym (rUnit refl) + +-- We get that botᵣ⁻ (and hence botᵣ) is homomorphism +hom-botᵣ⁻' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (f g : S₊∙ (suc (suc n)) →∙ Susp∙ (typ A)) + → botᵣ⁻ {A = A} (suc n) (∙Π f g) + ≡ ∙Π (botᵣ⁻ {A = A} (suc n) f) (botᵣ⁻ {A = A} (suc n) g) +hom-botᵣ⁻' {A = A} n f g = + hom-botᵣ⁻ {A = A} (suc n) f g + ∙ sym (∙Π≡invComp {A = A} _ (botᵣ⁻ {A = A} _ f) (botᵣ⁻ {A = A} _ g)) + + +hom-rMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (p q : typ ((Ω^ (suc (suc n))) (Susp∙ (typ A)))) + → rMap (suc n) {A = A} (p ∙ q) + ≡ ∙Π (rMap (suc n) {A = A} p) (rMap (suc n) {A = A} q) +hom-rMap {A = A} n p q = + cong (lMap (suc n) {A = Ω (Susp∙ (typ A))}) + (morphLemmas.isMorphInv _∙_ _∙_ + (Iso.inv (flipΩIso (suc n))) (flipΩIsopres· n) + (Iso.fun (flipΩIso (suc n))) + (Iso.leftInv (flipΩIso (suc n))) (Iso.rightInv (flipΩIso (suc n))) + p q) + ∙ isHom-lMap _ (fun (flipΩIso (suc n)) p) (fun (flipΩIso (suc n)) q) + +isHom-IsoΩSphereMap' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + (p q : typ (Ω ((Ω^ n) (Susp∙ (typ A))))) + → Iso.fun (IsoΩSphereMap' {A = A} n) (p ∙ q) + ≡ ∙Π {n = (suc n)} (Iso.fun (IsoΩSphereMap' {A = A} n) p) + (Iso.fun (IsoΩSphereMap' {A = A} n) q) +isHom-IsoΩSphereMap' {A = A} zero p q = + (cong (botᵣ {A = A} 0) + (cong ((lMap 0 {A = Ω (Susp∙ (typ A))})) + (transportRefl (p ∙ q)))) + ∙ ΣPathP (funExt (λ { base → refl + ; (loop i) j + → lem j i}) + , refl) + where + lem : p ∙ q ≡ (cong (fst (fun (IsoΩSphereMap' {A = A} zero) p)) loop ∙ refl) + ∙ (cong (fst (fun (IsoΩSphereMap' {A = A} zero) q)) loop ∙ refl) + lem = cong₂ _∙_ + (sym (transportRefl p) + ∙ rUnit (cong (fst (fun (IsoΩSphereMap' {A = A} zero) p)) loop)) + (sym (transportRefl q) + ∙ rUnit (cong (fst (fun (IsoΩSphereMap' {A = A} zero) q)) loop)) +isHom-IsoΩSphereMap' {A = A} (suc n) p q = + cong (botᵣ {A = A} (suc n) ∘ lMap (suc n) {A = Ω (Susp∙ (typ A))}) + (morphLemmas.isMorphInv + _∙_ _∙_ + (Iso.inv (flipΩIso (suc n))) + (flipΩIsopres· n) + (Iso.fun (flipΩIso (suc n))) + (Iso.leftInv (flipΩIso (suc n))) + (Iso.rightInv (flipΩIso (suc n))) p q) + ∙ (cong (botᵣ {A = A} (suc n)) + (isHom-lMap _ {A = Ω (Susp∙ (typ A))} + (fun (flipΩIso (suc n)) p) (fun (flipΩIso (suc n)) q)) + ∙ morphLemmas.isMorphInv ∙Π ∙Π (botᵣ⁻ {A = A} (suc n)) + (hom-botᵣ⁻' {A = A} n) + (botᵣ {A = A} (suc n)) + (Iso.leftInv (botᵣIso {A = A} (suc n))) + (Iso.rightInv (botᵣIso {A = A} (suc n))) + (lMap (suc n) (fun (flipΩIso (suc n)) p)) + (lMap (suc n) (fun (flipΩIso (suc n)) q))) + +suspMapΩ→hom : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p q : typ ((Ω^ (suc n)) A)) + → suspMapΩ∙ (suc n) .fst (p ∙ q) + ≡ suspMapΩ∙ (suc n) .fst p ∙ suspMapΩ∙ (suc n) .fst q +suspMapΩ→hom {A = A} n p q = + cong (sym (snd (suspMapΩ∙ {A = A} n)) ∙∙_∙∙ snd (suspMapΩ∙ {A = A} n)) + (cong-∙ (fst (suspMapΩ∙ {A = A} n)) p q) + ∙ help (snd (suspMapΩ∙ {A = A} n)) _ _ + where + help : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) (q s : x ≡ x) + → sym p ∙∙ (q ∙ s) ∙∙ p + ≡ (sym p ∙∙ q ∙∙ p) ∙ (sym p ∙∙ s ∙∙ p) + help {x = x} = + J (λ y p → (q s : x ≡ x) + → sym p ∙∙ (q ∙ s) ∙∙ p + ≡ (sym p ∙∙ q ∙∙ p ) ∙ (sym p ∙∙ s ∙∙ p)) + λ q s → sym (rUnit (q ∙ s)) + ∙ cong₂ _∙_ (rUnit q) (rUnit s) + +private + transpLemTyp : ∀ {ℓ} → {A1 A2 B1 B2 : Type ℓ} + (AB1 : A1 ≡ B1) (AB2 : A2 ≡ B2) + → (f : A1 → A2) (g : B1 → B2) + → PathP (λ i → AB1 i → AB2 i) f g + → Type ℓ + transpLemTyp {ℓ} {A1} {A2} {B1} {B2} AB1 AB2 f g p = + (+A1 : A1 → A1 → A1) + → (+A2 : A2 → A2 → A2) + → (+B1 : B1 → B1 → B1) + → (+B2 : B2 → B2 → B2) + → ((x y : A1) → transport AB1 (+A1 x y) + ≡ +B1 (transport AB1 x) (transport AB1 y)) + → ((x y : A2) → transport AB2 (+A2 x y) + ≡ +B2 (transport AB2 x) (transport AB2 y)) + → ((x y : A1) → f (+A1 x y) ≡ +A2 (f x) (f y)) + → ((x y : B1) → g (+B1 x y) ≡ +B2 (g x) (g y)) + + transpLem : ∀ {ℓ} → {A1 A2 B1 B2 : Type ℓ} (AB1 : A1 ≡ B1) (AB2 : A2 ≡ B2) + → (f : A1 → A2) (g : B1 → B2) + → (eq : PathP (λ i → AB1 i → AB2 i) f g) + → transpLemTyp AB1 AB2 f g eq + transpLem {ℓ} {A1} {A2} {B1} {B2} = + J (λ B1 AB1 → (AB2 : A2 ≡ B2) + → (f : A1 → A2) (g : B1 → B2) + → (eq : PathP (λ i → AB1 i → AB2 i) f g) + → transpLemTyp AB1 AB2 f g eq) + (J (λ B2 AB2 → (f : A1 → A2) (g : A1 → B2) + → (eq : PathP (λ i → A1 → AB2 i) f g) + → transpLemTyp refl AB2 f g eq) + λ f g → J (λ g p → transpLemTyp refl refl f g p) + λ +A1 +A2 +B1 +B2 trId1 trId2 fhom1 + → λ x y → cong f (cong₂ +B1 (sym (transportRefl x)) (sym (transportRefl y)) + ∙∙ sym (trId1 x y) + ∙∙ transportRefl (+A1 x y)) + ∙∙ fhom1 x y + ∙∙ ((sym (transportRefl _) ∙ (trId2 (f x) (f y))) + ∙ cong₂ +B2 (transportRefl (f x)) (transportRefl (f y)))) + +open import Cubical.Foundations.Univalence +suspMap→hom : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f g : S₊∙ (suc n) →∙ A) + → suspMap n (∙Π f g) ≡ ∙Π (suspMap n f) (suspMap n g) +suspMap→hom {A = A} n = + transpLem (Ω≡SphereMap {A = A} (suc n)) (Ω≡SphereMap' {A = A} (suc n)) + (suspMapΩ {A = A} (suc n)) (suspMap {A = A} n) + (suspMap→ {A = A} n) + _∙_ _∙_ ∙Π ∙Π + (λ x y → uaβ (isoToEquiv (IsoΩSphereMap {A = A} (suc n))) (x ∙ y) + ∙∙ isHom-lMap n x y + ∙∙ cong₂ ∙Π + (sym (uaβ (isoToEquiv (IsoΩSphereMap {A = A} (suc n))) x)) + (sym (uaβ (isoToEquiv (IsoΩSphereMap {A = A} (suc n))) y))) + (λ x y → uaβ (isoToEquiv (IsoΩSphereMap' {A = A} (suc n))) (x ∙ y) + ∙∙ isHom-IsoΩSphereMap' {A = A} (suc n) x y + ∙∙ cong₂ ∙Π + (sym (uaβ (isoToEquiv (IsoΩSphereMap' {A = A} (suc n))) x)) + (sym (uaβ (isoToEquiv (IsoΩSphereMap' {A = A} (suc n))) y))) + (suspMapΩ→hom {A = A} n) + +private + isConnectedPres : ∀ {ℓ} {A : Pointed ℓ} (con n : ℕ) + → isConnectedFun con (suspMapΩ∙ {A = A} (suc n) .fst) + → isConnectedFun con (suspMap {A = A} n) + isConnectedPres {A = A} con n hyp = + transport (λ i → isConnectedFun con (suspMap→ {A = A} n i)) hyp + +isConnectedSuspMap : (n m : ℕ) + → isConnectedFun ((m + suc m) ∸ n) (suspMap {A = S₊∙ (suc m)} n) +isConnectedSuspMap n m = + isConnectedPres _ _ (suspMapΩ-connected m (suc n) (sphereConnected (suc m))) + +suspMapπ' : (n m : ℕ) + → π' (suc n) (S₊∙ (suc m)) + → π' (suc (suc n)) (S₊∙ (suc (suc m))) +suspMapπ' n m = sMap (suspMap {A = S₊∙ (suc m)} n) + +suspMapHom : (n m : ℕ) + → GroupHom (π'Gr n (S₊∙ (suc m))) + (π'Gr (suc n) (S₊∙ (suc (suc m)))) +fst (suspMapHom n m) = suspMapπ' n m +snd (suspMapHom n m) = + makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) + λ f g → cong ∣_∣₂ (suspMap→hom n f g)) + +isSurjectiveSuspMap : (n : ℕ) → isSurjective (suspMapHom (2 + n) (suc n)) +isSurjectiveSuspMap n = + sElim (λ _ → isProp→isSet squash) + λ f → trRec (subst (λ x → isOfHLevel x (isInIm (suspMapHom (2 + n) (suc n)) ∣ f ∣₂)) + (sym (snd (lem n n))) + (isProp→isOfHLevelSuc {A = isInIm (suspMapHom (2 + n) (suc n)) ∣ f ∣₂} + (fst (lem n n)) squash)) + (λ p → ∣ ∣ fst p ∣₂ , (cong ∣_∣₂ (snd p)) ∣) + (fst (isConnectedSuspMap (2 + n) (suc n) f)) + where + lem : (m n : ℕ) → Σ[ x ∈ ℕ ] ((m + suc (suc n) ∸ suc n) ≡ suc x) + lem zero zero = 0 , refl + lem (suc m) zero = suc m , +-comm m 2 + lem zero (suc n) = lem zero n + lem (suc m) (suc n) = fst (lem (suc m) n) , (cong (_∸ (suc n)) (+-comm m (3 + n)) + ∙∙ cong (_∸ n) (+-comm (suc (suc n)) m) + ∙∙ snd (lem (suc m) n)) diff --git a/Cubical/Homotopy/Group/S3.agda b/Cubical/Homotopy/Group/S3.agda index 3462594ca..fa2b358a0 100644 --- a/Cubical/Homotopy/Group/S3.agda +++ b/Cubical/Homotopy/Group/S3.agda @@ -409,6 +409,28 @@ fst (botMap n {A = A} f) x = merid (fst f x) ∙ sym (merid (pt A)) snd (botMap n {A = A} f) = cong (_∙ merid (pt A) ⁻¹) (cong merid (snd f)) ∙ rCancel (merid (pt A)) +flipΩPath : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → ((Ω^ (suc n)) A) ≡ (Ω^ n) (Ω A) +flipΩPath {A = A} zero = refl +flipΩPath {A = A} (suc n) = cong Ω (flipΩPath {A = A} n) + +flipΩIso : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → Iso (fst ((Ω^ (suc n)) A)) (fst ((Ω^ n) (Ω A))) +flipΩIso {A = A} n = pathToIso (cong fst (flipΩPath n)) + +flipΩIsopres· : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (f g : fst ((Ω^ (suc n)) (Ω A))) + → Iso.inv (flipΩIso (suc n)) (f ∙ g) + ≡ (Iso.inv (flipΩIso (suc n)) f) + ∙ (Iso.inv (flipΩIso (suc n)) g) +flipΩIsopres· {A = A} n f g i = + transp (λ j → flipΩPath {A = A} n (~ i ∧ ~ j) .snd + ≡ flipΩPath n (~ i ∧ ~ j) .snd) i + (transp (λ j → flipΩPath {A = A} n (~ i ∨ ~ j) .snd + ≡ flipΩPath n (~ i ∨ ~ j) .snd) (~ i) f + ∙ transp (λ j → flipΩPath {A = A} n (~ i ∨ ~ j) .snd + ≡ flipΩPath n (~ i ∨ ~ j) .snd) (~ i) g) + rMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → typ ((Ω^ (suc n)) (Susp∙ (typ A))) → (S₊∙ n →∙ Ω (Susp∙ (typ A))) @@ -827,404 +849,3 @@ suspMap→ {A = A} n = (funExt⁻ (filler-top□ (suc n)) (invEq (_ , isEquiv-lMap n) f)) ∙∙ sym (filler▿ (suc n) (lMap (suc n) {A = A} (invEq (lMap (suc n) , isEquiv-lMap n) f))) ∙ cong (suspMap n) (secEq ((lMap (suc n) , isEquiv-lMap n)) f)) - -{- -We would like for the above dependent path to preserve composition -(in order to deduce properties of suspMap from the corresponding -ones of suspMapΩ). For this, we need that both Ω≡SphereMap and -Ω≡SphereMap' are structure preserving --} - --- We define an alternative notion of composition (to ∙Π) on --- Sⁿ →∙ Ω (Susp∙ (typ A)). It will be easier to work with -private - invComp : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → S₊∙ n →∙ Ω (Susp∙ (typ A)) - → S₊∙ n →∙ Ω (Susp∙ (typ A)) - → S₊∙ n →∙ Ω (Susp∙ (typ A)) - fst (invComp n f g) x = (fst f x) ∙ (fst g x) - snd (invComp n f g) = cong₂ _∙_ (snd f) (snd g) ∙ sym (rUnit refl) - - -- We prove that it agrees with ∙Π - ∙Π≡invComp : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → (f g : S₊∙ (suc n) →∙ Ω (Susp∙ (typ A))) - → ∙Π f g ≡ invComp {A = A} (suc n) f g - ∙Π≡invComp zero f g = - →∙Homogeneous≡ (isHomogeneousPath _ _) - (funExt (λ { base → rUnit refl - ∙ sym (cong (_∙ fst g (snd (S₊∙ 1))) (snd f) - ∙ cong (refl ∙_) (snd g)) - ; (loop i) j → - hcomp (λ k → - λ { (i = i0) → (rUnit refl - ∙ sym (cong (_∙ fst g (snd (S₊∙ 1))) (snd f) - ∙ cong (refl ∙_) (snd g))) j - ; (i = i1) → (rUnit refl - ∙ sym (cong (_∙ fst g (snd (S₊∙ 1))) (snd f) - ∙ cong (refl ∙_) (snd g))) j - ; (j = i0) → ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) - ∙ (sym (snd g) ∙∙ cong (fst g) loop ∙∙ snd g)) i - ; (j = i1) → cong₂Funct _∙_ - (cong (fst f) loop) (cong (fst g) loop) (~ k) i}) - (hcomp (λ k → - λ { (i = i0) → (rUnit refl - ∙ sym (cong (_∙ snd g (~ k)) (λ j → snd f (j ∨ ~ k)) - ∙ cong (refl ∙_) (λ j → snd g (j ∨ ~ k)))) j - ; (i = i1) → (rUnit refl - ∙ sym (cong (_∙ snd g (~ k)) (λ j → snd f (j ∨ ~ k)) - ∙ cong (refl ∙_) (λ j → snd g (j ∨ ~ k)))) j - ; (j = i0) → ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) - ∙ (sym (snd g) ∙∙ cong (fst g) loop ∙∙ snd g)) i - ; (j = i1) → (cong (_∙ snd g (~ k)) - (doubleCompPath-filler (sym (snd f)) - (cong (fst f) loop) - (snd f) (~ k)) ∙ - cong ((snd f (~ k)) ∙_) - (doubleCompPath-filler (sym (snd g)) - (cong (fst g) loop) (snd g) (~ k))) i}) - (hcomp (λ k → - λ { (i = i0) → (rUnit (rUnit refl) - ∙ cong (rUnit refl ∙_) (cong sym (rUnit refl))) k j - ; (i = i1) → (rUnit (rUnit refl) - ∙ cong (rUnit refl ∙_) (cong sym (rUnit refl))) k j - ; (j = i0) → ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) - ∙ (sym (snd g) ∙∙ cong (fst g) loop ∙∙ snd g)) i - ; (j = i1) → (cong (_∙ refl) - ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) - ∙ cong (refl ∙_) - (sym (snd g) ∙∙ cong (fst g) loop ∙∙ snd g)) i}) - ((cong (λ x → rUnit x j) - (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) - ∙ cong (λ x → lUnit x j) - (sym (snd g) ∙∙ cong (fst g) loop ∙∙ snd g)) i)))})) - ∙Π≡invComp {A = A} (suc n) f g = - →∙Homogeneous≡ (isHomogeneousPath _ _) - (funExt λ { north → rUnit refl - ∙ sym (cong (fst f north ∙_) (snd g) - ∙ cong (_∙ refl) (snd f)) - ; south → rUnit refl - ∙ sym (cong₂ _∙_ - (cong (fst f) (sym (merid (ptSn _))) ∙ snd f) - (cong (fst g) (sym (merid (ptSn _))) ∙ snd g)) - ; (merid a i) j → p a i j}) - where - module _ (a : S₊ (suc n)) where - f-n = fst f north - g-n = fst g north - cong-f = (cong (fst f) (merid a ∙ sym (merid (ptSn _)))) - cong-g = (cong (fst g) (merid a ∙ sym (merid (ptSn _)))) - c-f = sym (snd f) ∙∙ cong-f ∙∙ snd f - c-g = sym (snd g) ∙∙ cong-g ∙∙ snd g - - p : I → I → fst (Ω (Susp∙ (typ A))) - p i j = - hcomp (λ k → - λ { (i = i0) → (rUnit (λ _ → snd (Susp∙ (typ A))) - ∙ sym ((cong (fst f north ∙_) (snd g) - ∙ cong (_∙ refl) (snd f)))) j - ; (i = i1) → (rUnit refl - ∙ sym (cong₂ _∙_ - (compPath-filler' - (cong (fst f) (sym (merid (ptSn _)))) (snd f) k) - (compPath-filler' - (cong (fst g) (sym (merid (ptSn _)))) (snd g) k))) j - ; (j = i0) → (c-f ∙ c-g) i - ; (j = i1) → fst f (compPath-filler - (merid a) - (sym (merid (ptSn _))) (~ k) i) - ∙ fst g (compPath-filler - (merid a) - (sym (merid (ptSn _))) (~ k) i)}) - (hcomp (λ k → - λ {(i = i0) → (rUnit (λ _ → snd (Susp∙ (typ A))) - ∙ sym ((cong (fst f north ∙_) (snd g) - ∙ cong (_∙ refl) (snd f)))) j - ; (i = i1) → (rUnit refl ∙ sym (cong₂ _∙_ (snd f) (snd g))) j - ; (j = i0) → (c-f ∙ c-g) i - ; (j = i1) → cong₂Funct _∙_ cong-f cong-g (~ k) i}) - (hcomp (λ k → - λ {(i = i0) → (rUnit refl - ∙ sym (compPath-filler' - ((cong (fst f north ∙_) (snd g))) - (cong (_∙ refl) (snd f)) k)) j - ; (i = i1) → (rUnit refl - ∙ sym (cong₂ _∙_ (λ i → snd f (i ∨ ~ k)) (snd g))) j - ; (j = i0) → (c-f ∙ c-g) i - ; (j = i1) → (cong (λ x → x ∙ snd g (~ k)) - (doubleCompPath-filler refl - cong-f (snd f) (~ k)) - ∙ cong ((snd f (~ k)) ∙_) - (doubleCompPath-filler - (sym (snd g)) cong-g refl (~ k))) i}) - (hcomp (λ k → - λ {(i = i0) → compPath-filler - (rUnit (λ _ → snd (Susp∙ (typ A)))) - (sym ((cong (_∙ refl) (snd f)))) k j - ; (i = i1) → compPath-filler - (rUnit refl) - (sym (cong (refl ∙_) (snd g))) k j - ; (j = i0) → (c-f ∙ c-g) i - ; (j = i1) → (cong (_∙ refl) - ((λ j → snd f (~ j ∧ ~ k)) ∙∙ cong-f ∙∙ snd f) - ∙ cong (refl ∙_) - (sym (snd g) ∙∙ cong-g ∙∙ (λ j → snd g (j ∧ ~ k)))) i}) - (((cong (λ x → rUnit x j) c-f) ∙ (cong (λ x → lUnit x j) c-g)) i)))) - - hom-botᵣ⁻ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → (f g : S₊∙ (suc n) →∙ Susp∙ (typ A)) - → botᵣ⁻ {A = A} n (∙Π f g) - ≡ invComp {A = A} n (botᵣ⁻ {A = A} n f) (botᵣ⁻ {A = A} n g) - --- We ge that -hom-botᵣ⁻ zero f g = - ΣPathP ((funExt (λ { false → sym (rUnit _) - ; true → (rUnit _)})) - , ((λ i j → rUnit refl (i ∧ ~ j)) - ▷ lUnit (sym (rUnit refl)))) -hom-botᵣ⁻ (suc n) f g = - →∙Homogeneous≡ (isHomogeneousPath _ _) - (funExt (λ x → (sym (rUnit (cong (fst (∙Π f g)) (merid x ∙ sym (merid (ptSn _)))))) - ∙∙ cong-∙ (fst (∙Π f g)) (merid x) (sym (merid (ptSn _))) - ∙∙ cong (cong (fst (∙Π f g)) (merid x) ∙_) (cong sym lem) - ∙ sym (rUnit (cong (fst (∙Π f g)) (merid x))))) - where - lem : cong (fst (∙Π f g)) (merid (ptSn (suc n))) ≡ refl - lem = (λ i → (sym (snd f) ∙∙ cong (fst f) (rCancel (merid (ptSn _)) i) ∙∙ snd f) - ∙ (sym (snd g) ∙∙ cong (fst g) (rCancel (merid (ptSn _)) i) ∙∙ snd g)) - ∙ (λ i → ∙∙lCancel (snd f) i ∙ ∙∙lCancel (snd g) i) - ∙ sym (rUnit refl) - --- We get that botᵣ⁻ (and hence botᵣ) is homomorphism -hom-botᵣ⁻' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → (f g : S₊∙ (suc (suc n)) →∙ Susp∙ (typ A)) - → botᵣ⁻ {A = A} (suc n) (∙Π f g) - ≡ ∙Π (botᵣ⁻ {A = A} (suc n) f) (botᵣ⁻ {A = A} (suc n) g) -hom-botᵣ⁻' {A = A} n f g = - hom-botᵣ⁻ {A = A} (suc n) f g - ∙ sym (∙Π≡invComp {A = A} _ (botᵣ⁻ {A = A} _ f) (botᵣ⁻ {A = A} _ g)) - -isHom-lMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (p q : _) - → lMap (suc n) {A = A} (p ∙ q) - ≡ ∙Π (lMap (suc n) p) (lMap (suc n) q) -isHom-lMap zero p q = - ΣPathP ((funExt (λ { base → refl - ; (loop i) j → (rUnit p j ∙ rUnit q j) i})) - , refl) -isHom-lMap (suc n) {A = A} p q = - ΣPathP ((funExt (λ { north → refl - ; south → refl - ; (merid a i) j → h a j i})) - , refl) - where - doubleComp-lem : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) (q r : y ≡ y) - → (p ∙∙ q ∙∙ sym p) ∙ (p ∙∙ r ∙∙ sym p) - ≡ (p ∙∙ (q ∙ r) ∙∙ sym p) - doubleComp-lem p q r i j = - hcomp (λ k → λ { (i = i0) → (doubleCompPath-filler p q (sym p) k - ∙ doubleCompPath-filler p r (sym p) k) j - ; (i = i1) → doubleCompPath-filler p (q ∙ r) (sym p) k j - ; (j = i0) → p (~ k) - ; (j = i1) → p (~ k)}) - ((q ∙ r) j) - - help2 : (p : typ ((Ω^ (suc (suc n))) A)) - → cong (fst (lMap (suc (suc n)) p)) (merid (ptSn _)) ≡ refl - help2 p = - cong (sym (lMapId (suc n) (ptSn _)) ∙∙_∙∙ lMapId (suc n) (ptSn _)) - (rUnit _ ∙ (λ j → (λ i → lMap (suc n) {A = A} refl .snd (i ∧ j)) - ∙∙ (λ i → lMap (suc n) {A = A} (p i) .snd j) - ∙∙ λ i → lMap (suc n) {A = A} refl .snd (~ i ∧ j)) - ∙ ∙∙lCancel _) - ∙ ∙∙lCancel _ - - h : (a : S₊ (suc n)) - → sym (lMapId (suc n) a) - ∙∙ funExt⁻ (cong fst (cong (lMap (suc n)) (p ∙ q))) a - ∙∙ lMapId (suc n) a - ≡ cong (fst (∙Π (lMap (suc (suc n)) p) (lMap (suc (suc n)) q))) (merid a) - h a = (cong (sym (lMapId (suc n) a) ∙∙_∙∙ (lMapId (suc n) a)) - (cong-∙ (λ x → lMap (suc n) x .fst a) p q) - ∙ sym (doubleComp-lem (sym (lMapId (suc n) a)) _ _)) - ∙∙ cong₂ _∙_ (sym (cong (cong (fst (lMap (suc (suc n)) p)) (merid a) ∙_) - (cong sym (help2 p)) ∙ sym (rUnit _))) - (sym (cong (cong (fst (lMap (suc (suc n)) q)) (merid a) ∙_) - (cong sym (help2 q)) ∙ sym (rUnit _))) - ∙∙ λ i → (rUnit (cong-∙ (fst (lMap (suc (suc n)) p)) - (merid a) (sym (merid (ptSn _))) (~ i)) i) - ∙ (rUnit (cong-∙ (fst (lMap (suc (suc n)) q)) - (merid a) (sym (merid (ptSn _)))(~ i)) i) - -hom-rMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → (p q : typ ((Ω^ (suc (suc n))) (Susp∙ (typ A)))) - → rMap (suc n) {A = A} (p ∙ q) - ≡ ∙Π (rMap (suc n) {A = A} p) (rMap (suc n) {A = A} q) -hom-rMap {A = A} n p q = - cong (lMap (suc n) {A = Ω (Susp∙ (typ A))}) - (morphLemmas.isMorphInv _∙_ _∙_ - (Iso.inv (flipΩIso (suc n))) (flipΩIsopres· n) - (Iso.fun (flipΩIso (suc n))) - (Iso.leftInv (flipΩIso (suc n))) (Iso.rightInv (flipΩIso (suc n))) - p q) - ∙ isHom-lMap _ (fun (flipΩIso (suc n)) p) (fun (flipΩIso (suc n)) q) - -isHom-IsoΩSphereMap' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - (p q : typ (Ω ((Ω^ n) (Susp∙ (typ A))))) - → Iso.fun (IsoΩSphereMap' {A = A} n) (p ∙ q) - ≡ ∙Π {n = (suc n)} (Iso.fun (IsoΩSphereMap' {A = A} n) p) - (Iso.fun (IsoΩSphereMap' {A = A} n) q) -isHom-IsoΩSphereMap' {A = A} zero p q = - (cong (botᵣ {A = A} 0) - (cong ((lMap 0 {A = Ω (Susp∙ (typ A))})) - (transportRefl (p ∙ q)))) - ∙ ΣPathP (funExt (λ { base → refl - ; (loop i) j - → lem j i}) - , refl) - where - lem : p ∙ q ≡ (cong (fst (fun (IsoΩSphereMap' {A = A} zero) p)) loop ∙ refl) - ∙ (cong (fst (fun (IsoΩSphereMap' {A = A} zero) q)) loop ∙ refl) - lem = cong₂ _∙_ - (sym (transportRefl p) - ∙ rUnit (cong (fst (fun (IsoΩSphereMap' {A = A} zero) p)) loop)) - (sym (transportRefl q) - ∙ rUnit (cong (fst (fun (IsoΩSphereMap' {A = A} zero) q)) loop)) -isHom-IsoΩSphereMap' {A = A} (suc n) p q = - cong (botᵣ {A = A} (suc n) ∘ lMap (suc n) {A = Ω (Susp∙ (typ A))}) - (morphLemmas.isMorphInv - _∙_ _∙_ - (Iso.inv (flipΩIso (suc n))) - (flipΩIsopres· n) - (Iso.fun (flipΩIso (suc n))) - (Iso.leftInv (flipΩIso (suc n))) - (Iso.rightInv (flipΩIso (suc n))) p q) - ∙ (cong (botᵣ {A = A} (suc n)) - (isHom-lMap _ {A = Ω (Susp∙ (typ A))} - (fun (flipΩIso (suc n)) p) (fun (flipΩIso (suc n)) q)) - ∙ morphLemmas.isMorphInv ∙Π ∙Π (botᵣ⁻ {A = A} (suc n)) - (hom-botᵣ⁻' {A = A} n) - (botᵣ {A = A} (suc n)) - (Iso.leftInv (botᵣIso {A = A} (suc n))) - (Iso.rightInv (botᵣIso {A = A} (suc n))) - (lMap (suc n) (fun (flipΩIso (suc n)) p)) - (lMap (suc n) (fun (flipΩIso (suc n)) q))) - - -suspMapΩ→hom : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p q : typ ((Ω^ (suc n)) A)) - → suspMapΩ∙ (suc n) .fst (p ∙ q) - ≡ suspMapΩ∙ (suc n) .fst p ∙ suspMapΩ∙ (suc n) .fst q -suspMapΩ→hom {A = A} n p q = - cong (sym (snd (suspMapΩ∙ {A = A} n)) ∙∙_∙∙ snd (suspMapΩ∙ {A = A} n)) - (cong-∙ (fst (suspMapΩ∙ {A = A} n)) p q) - ∙ help (snd (suspMapΩ∙ {A = A} n)) _ _ - where - help : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) (q s : x ≡ x) - → sym p ∙∙ (q ∙ s) ∙∙ p - ≡ (sym p ∙∙ q ∙∙ p) ∙ (sym p ∙∙ s ∙∙ p) - help {x = x} = - J (λ y p → (q s : x ≡ x) - → sym p ∙∙ (q ∙ s) ∙∙ p - ≡ (sym p ∙∙ q ∙∙ p ) ∙ (sym p ∙∙ s ∙∙ p)) - λ q s → sym (rUnit (q ∙ s)) - ∙ cong₂ _∙_ (rUnit q) (rUnit s) - -private - transpLemTyp : ∀ {ℓ} → {A1 A2 B1 B2 : Type ℓ} - (AB1 : A1 ≡ B1) (AB2 : A2 ≡ B2) - → (f : A1 → A2) (g : B1 → B2) - → PathP (λ i → AB1 i → AB2 i) f g - → Type ℓ - transpLemTyp {ℓ} {A1} {A2} {B1} {B2} AB1 AB2 f g p = - (+A1 : A1 → A1 → A1) - → (+A2 : A2 → A2 → A2) - → (+B1 : B1 → B1 → B1) - → (+B2 : B2 → B2 → B2) - → ((x y : A1) → transport AB1 (+A1 x y) - ≡ +B1 (transport AB1 x) (transport AB1 y)) - → ((x y : A2) → transport AB2 (+A2 x y) - ≡ +B2 (transport AB2 x) (transport AB2 y)) - → ((x y : A1) → f (+A1 x y) ≡ +A2 (f x) (f y)) - → ((x y : B1) → g (+B1 x y) ≡ +B2 (g x) (g y)) - - transpLem : ∀ {ℓ} → {A1 A2 B1 B2 : Type ℓ} (AB1 : A1 ≡ B1) (AB2 : A2 ≡ B2) - → (f : A1 → A2) (g : B1 → B2) - → (eq : PathP (λ i → AB1 i → AB2 i) f g) - → transpLemTyp AB1 AB2 f g eq - transpLem {ℓ} {A1} {A2} {B1} {B2} = - J (λ B1 AB1 → (AB2 : A2 ≡ B2) - → (f : A1 → A2) (g : B1 → B2) - → (eq : PathP (λ i → AB1 i → AB2 i) f g) - → transpLemTyp AB1 AB2 f g eq) - (J (λ B2 AB2 → (f : A1 → A2) (g : A1 → B2) - → (eq : PathP (λ i → A1 → AB2 i) f g) - → transpLemTyp refl AB2 f g eq) - λ f g → J (λ g p → transpLemTyp refl refl f g p) - λ +A1 +A2 +B1 +B2 trId1 trId2 fhom1 - → λ x y → cong f (cong₂ +B1 (sym (transportRefl x)) (sym (transportRefl y)) - ∙∙ sym (trId1 x y) - ∙∙ transportRefl (+A1 x y)) - ∙∙ fhom1 x y - ∙∙ ((sym (transportRefl _) ∙ (trId2 (f x) (f y))) - ∙ cong₂ +B2 (transportRefl (f x)) (transportRefl (f y)))) - -open import Cubical.Foundations.Univalence -suspMap→hom : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f g : S₊∙ (suc n) →∙ A) - → suspMap n (∙Π f g) ≡ ∙Π (suspMap n f) (suspMap n g) -suspMap→hom {A = A} n = - transpLem (Ω≡SphereMap {A = A} (suc n)) (Ω≡SphereMap' {A = A} (suc n)) - (suspMapΩ {A = A} (suc n)) (suspMap {A = A} n) - (suspMap→ {A = A} n) - _∙_ _∙_ ∙Π ∙Π - (λ x y → uaβ (isoToEquiv (IsoΩSphereMap {A = A} (suc n))) (x ∙ y) - ∙∙ isHom-lMap n x y - ∙∙ cong₂ ∙Π - (sym (uaβ (isoToEquiv (IsoΩSphereMap {A = A} (suc n))) x)) - (sym (uaβ (isoToEquiv (IsoΩSphereMap {A = A} (suc n))) y))) - (λ x y → uaβ (isoToEquiv (IsoΩSphereMap' {A = A} (suc n))) (x ∙ y) - ∙∙ isHom-IsoΩSphereMap' {A = A} (suc n) x y - ∙∙ cong₂ ∙Π - (sym (uaβ (isoToEquiv (IsoΩSphereMap' {A = A} (suc n))) x)) - (sym (uaβ (isoToEquiv (IsoΩSphereMap' {A = A} (suc n))) y))) - (suspMapΩ→hom {A = A} n) - -private - isConnectedPres : ∀ {ℓ} {A : Pointed ℓ} (con n : ℕ) - → isConnectedFun con (suspMapΩ∙ {A = A} (suc n) .fst) - → isConnectedFun con (suspMap {A = A} n) - isConnectedPres {A = A} con n hyp = - transport (λ i → isConnectedFun con (suspMap→ {A = A} n i)) hyp - -isConnectedSuspMap : (n m : ℕ) - → isConnectedFun ((m + suc m) ∸ n) (suspMap {A = S₊∙ (suc m)} n) -isConnectedSuspMap n m = - isConnectedPres _ _ (suspMapΩ-connected m (suc n) (sphereConnected (suc m))) - -suspMapπ' : (n m : ℕ) - → π' (suc n) (S₊∙ (suc m)) - → π' (suc (suc n)) (S₊∙ (suc (suc m))) -suspMapπ' n m = sMap (suspMap {A = S₊∙ (suc m)} n) - -suspMapHom : (n m : ℕ) - → GroupHom (π'Gr n (S₊∙ (suc m))) - (π'Gr (suc n) (S₊∙ (suc (suc m)))) -fst (suspMapHom n m) = suspMapπ' n m -snd (suspMapHom n m) = - makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) - λ f g → cong ∣_∣₂ (suspMap→hom n f g)) - -isSurjectiveSuspMap : (n : ℕ) → isSurjective (suspMapHom (2 + n) (suc n)) -isSurjectiveSuspMap n = - sElim (λ _ → isProp→isSet squash) - λ f → trRec (subst (λ x → isOfHLevel x (isInIm (suspMapHom (2 + n) (suc n)) ∣ f ∣₂)) - (sym (snd (lem n n))) - (isProp→isOfHLevelSuc {A = isInIm (suspMapHom (2 + n) (suc n)) ∣ f ∣₂} - (fst (lem n n)) squash)) - (λ p → ∣ ∣ fst p ∣₂ , (cong ∣_∣₂ (snd p)) ∣) - (fst (isConnectedSuspMap (2 + n) (suc n) f)) - where - lem : (m n : ℕ) → Σ[ x ∈ ℕ ] ((m + suc (suc n) ∸ suc n) ≡ suc x) - lem zero zero = 0 , refl - lem (suc m) zero = suc m , +-comm m 2 - lem zero (suc n) = lem zero n - lem (suc m) (suc n) = fst (lem (suc m) n) , (cong (_∸ (suc n)) (+-comm m (3 + n)) - ∙∙ cong (_∸ n) (+-comm (suc (suc n)) m) - ∙∙ snd (lem (suc m) n)) From 8e7aab924790d48af44a0fe903c91d5a4e56ab7e Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 1 Nov 2021 14:32:07 +0100 Subject: [PATCH 76/89] miinor --- Cubical/Algebra/Group/GroupPath.agda | 7 +- Cubical/Algebra/Group/ZModule.agda | 88 ++++++++------------- Cubical/Homotopy/HopfInvariant/HopfMap.agda | 4 +- 3 files changed, 41 insertions(+), 58 deletions(-) diff --git a/Cubical/Algebra/Group/GroupPath.agda b/Cubical/Algebra/Group/GroupPath.agda index 2ba25af6f..88ca42ece 100644 --- a/Cubical/Algebra/Group/GroupPath.agda +++ b/Cubical/Algebra/Group/GroupPath.agda @@ -135,10 +135,11 @@ JGroupEquiv : {G : Group ℓ} (P : (H : Group ℓ) → GroupEquiv G H → Type JGroupEquiv {G = G} P p {H} e = transport (λ i → P (GroupPath G H .fst e i) (transp (λ j → GroupEquiv G (GroupPath G H .fst e (i ∨ ~ j))) i e)) - (subst (P G) (sym l) p) + (subst (P G) (sym lem) p) where - l : (transp (λ j → GroupEquiv G (GroupPath G H .fst e (~ j))) i0 e) ≡ idGroupEquiv - l = Σ≡Prop (λ _ → isPropIsGroupHom _ _) + lem : transport (λ j → GroupEquiv G (GroupPath G H .fst e (~ j))) e + ≡ idGroupEquiv + lem = Σ≡Prop (λ _ → isPropIsGroupHom _ _) (Σ≡Prop (λ _ → isPropIsEquiv _) (funExt λ x → (λ i → fst (fst (fst e .snd .equiv-proof (transportRefl (fst (fst e) (transportRefl x i)) i)))) diff --git a/Cubical/Algebra/Group/ZModule.agda b/Cubical/Algebra/Group/ZModule.agda index a0e6aa647..ca326ad1f 100644 --- a/Cubical/Algebra/Group/ZModule.agda +++ b/Cubical/Algebra/Group/ZModule.agda @@ -7,8 +7,10 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv open import Cubical.Data.Sigma -open import Cubical.Data.Int renaming (_·_ to _*_ ; _+_ to _+ℤ_ ; _-_ to _-ℤ_) -open import Cubical.Data.Nat hiding (_·_) renaming (_+_ to _+ℕ_) +open import Cubical.Data.Int + renaming (_·_ to _*_ ; _+_ to _+ℤ_ ; _-_ to _-ℤ_) +open import Cubical.Data.Nat + hiding (_·_) renaming (_+_ to _+ℕ_) open import Cubical.Data.Empty renaming (rec to ⊥-rec) open import Cubical.Data.Sum @@ -56,7 +58,7 @@ homPresℤ· (ϕ , ϕhom) x (negsuc zero) = presinv ϕhom _ homPresℤ· {H = H} (ϕ , ϕhom) x (negsuc (suc n)) = pres· ϕhom _ _ ∙ cong₂ (_·_ (snd H)) (presinv ϕhom x) - ((homPresℤ· (ϕ , ϕhom) x (negsuc n))) + (homPresℤ· (ϕ , ϕhom) x (negsuc n)) rUnitℤ· : (G : Group ℓ) (x : ℤ) → (x ℤ[ G ]· 1g (snd G)) ≡ 1g (snd G) rUnitℤ· G (pos zero) = refl @@ -97,11 +99,13 @@ private module _ (G : Group ℓ) (g : fst G) where private - lem : (y : ℤ) → (predℤ y ℤ[ G ]· g) ≡ (snd G · inv (snd G) g) (y ℤ[ G ]· g) + lem : (y : ℤ) → (predℤ y ℤ[ G ]· g) + ≡ (snd G · inv (snd G) g) (y ℤ[ G ]· g) lem (pos zero) = sym (rid (snd G) _) lem (pos (suc n)) = sym (lid (snd G) ((pos n ℤ[ G ]· g))) - ∙∙ cong (λ x → _·_ (snd G) x (pos n ℤ[ G ]· g)) (sym (invl (snd G) g)) + ∙∙ cong (λ x → _·_ (snd G) x (pos n ℤ[ G ]· g)) + (sym (invl (snd G) g)) ∙∙ sym (assoc (snd G) _ _ _) lem (negsuc n) = refl @@ -111,7 +115,8 @@ module _ (G : Group ℓ) (g : fst G) where ∙ sym (lid (snd G) _) distrℤ· (pos (suc n)) (pos n₁) = cong (_ℤ[ G ]· g) (sym (pos+ (suc n) n₁)) - ∙ cong (_·_ (snd G) g) (cong (_ℤ[ G ]· g) (pos+ n n₁) ∙ distrℤ· (pos n) (pos n₁)) + ∙ cong (_·_ (snd G) g) + (cong (_ℤ[ G ]· g) (pos+ n n₁) ∙ distrℤ· (pos n) (pos n₁)) ∙ assoc (snd G) _ _ _ distrℤ· (pos (suc n)) (negsuc zero) = distrℤ· (pos n) 0 @@ -133,7 +138,7 @@ module _ (G : Group ℓ) (g : fst G) where ((sym (rid (snd G) (pos n ℤ[ G ]· g)) ∙ cong (_·_ (snd G) (pos n ℤ[ G ]· g)) (sym (invr (snd G) g))) ∙∙ assoc (snd G) _ _ _ - ∙∙ cong (λ x → _·_ (snd G) x (inv (snd G) g)) (precommℤ G g (pos n) )) + ∙∙ cong (λ x → _·_ (snd G) x (inv (snd G) g)) (precommℤ G g (pos n))) ∙ sym (assoc (snd G) _ _ _)) distrℤ· (negsuc zero) y = cong (_ℤ[ G ]· g) (+Comm -1 y) ∙ lem y @@ -146,7 +151,8 @@ module _ (G : Group ℓ) (g : fst G) where GroupHomℤ→ℤpres- : (e : GroupHom ℤGroup ℤGroup) (a : ℤ) → fst e (- a) ≡ - fst e a -GroupHomℤ→ℤpres- e a = cong (fst e) (minus≡0- a) ∙∙ presinv (snd e) a ∙∙ sym (minus≡0- _) +GroupHomℤ→ℤpres- e a = + cong (fst e) (minus≡0- a) ∙∙ presinv (snd e) a ∙∙ sym (minus≡0- _) ℤ·≡· : (a b : ℤ) → a * b ≡ (a ℤ[ ℤGroup ]· b) ℤ·≡· (pos zero) b = refl @@ -154,8 +160,10 @@ GroupHomℤ→ℤpres- e a = cong (fst e) (minus≡0- a) ∙∙ presinv (snd e) ℤ·≡· (negsuc zero) b = minus≡0- b ℤ·≡· (negsuc (suc n)) b = cong₂ _+ℤ_ (minus≡0- b) (ℤ·≡· (negsuc n) b) -GroupHomℤ→ℤPres· : (e : GroupHom ℤGroup ℤGroup) (a b : ℤ) → fst e (a * b) ≡ a * fst e b -GroupHomℤ→ℤPres· e a b = cong (fst e) (ℤ·≡· a b) ∙∙ homPresℤ· e b a ∙∙ sym (ℤ·≡· a (fst e b)) +GroupHomℤ→ℤPres· : (e : GroupHom ℤGroup ℤGroup) (a b : ℤ) + → fst e (a * b) ≡ a * fst e b +GroupHomℤ→ℤPres· e a b = + cong (fst e) (ℤ·≡· a b) ∙∙ homPresℤ· e b a ∙∙ sym (ℤ·≡· a (fst e b)) -- Generators in terms of ℤ-multiplication -- Todo : generalise @@ -234,10 +242,10 @@ GroupEquivℤ-pres1 e (pos (suc (suc n))) p = ∙∙ (cong (fst (fst (invGroupEquiv e))) (p ∙ ·Comm 1 (pos (suc (suc n))))) ∙∙ GroupHomℤ→ℤPres· (_ , snd (invGroupEquiv e)) (pos (suc (suc n))) 1 GroupEquivℤ-pres1 e (negsuc zero) p = cong abs p -GroupEquivℤ-pres1 e (negsuc (suc n)) p = ⊥-rec (¬1=x·suc-suc _ _ l33) +GroupEquivℤ-pres1 e (negsuc (suc n)) p = ⊥-rec (¬1=x·suc-suc _ _ lem₂) where - h3 : fst (fst e) (negsuc zero) ≡ pos (suc (suc n)) - h3 = GroupHomℤ→ℤpres- (_ , snd e) (pos 1) ∙ cong -_ p + lem₁ : fst (fst e) (negsuc zero) ≡ pos (suc (suc n)) + lem₁ = GroupHomℤ→ℤpres- (_ , snd e) (pos 1) ∙ cong -_ p compGroup : GroupEquiv ℤGroup ℤGroup fst compGroup = isoToEquiv (iso -_ -_ -Involutive -Involutive) @@ -246,56 +254,30 @@ GroupEquivℤ-pres1 e (negsuc (suc n)) p = ⊥-rec (¬1=x·suc-suc _ _ l33) compHom : GroupEquiv ℤGroup ℤGroup compHom = compGroupEquiv compGroup e - l32 : fst (fst compHom) (pos 1) ≡ pos (suc (suc n)) - l32 = h3 - - l33 : pos 1 ≡ invEq (fst compHom) (pos 1) * pos (suc (suc n)) - l33 = sym (retEq (fst compHom) (pos 1)) - ∙∙ cong (invEq (fst compHom)) l32 + lem₂ : pos 1 ≡ invEq (fst compHom) (pos 1) * pos (suc (suc n)) + lem₂ = sym (retEq (fst compHom) (pos 1)) + ∙∙ cong (invEq (fst compHom)) lem₁ ∙∙ (cong (invEq (fst compHom)) (·Comm (pos 1) (pos (suc (suc n)))) - ∙ GroupHomℤ→ℤPres· (_ , (snd (invGroupEquiv compHom))) (pos (suc (suc n))) (pos 1) + ∙ GroupHomℤ→ℤPres· (_ , (snd (invGroupEquiv compHom))) + (pos (suc (suc n))) (pos 1) ∙ ·Comm (pos (suc (suc n))) (invEq (fst compHom) (pos 1))) --- really needed? -{- -sesIsoPresGen : - ∀ (G : Group ℓ-zero) - → (iso : GroupEquiv ℤGroup G) - → (H : Group ℓ-zero) - → (iso2 : GroupEquiv ℤGroup H) - → (e : fst G) - → invEq (fst iso) e ≡ 1 - → (hom : GroupEquiv G H) - → abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1 -sesIsoPresGen G = JGroupEquiv (λ G iso → - (H : Group ℓ-zero) - (iso2 : GroupEquiv ℤGroup H) - → (e : fst G) - → invEq (fst iso) e ≡ 1 - → (hom : GroupEquiv G H) - → abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1) - λ H → JGroupEquiv (λ H iso2 → (e : ℤ) → e ≡ pos 1 → - (hom : GroupEquiv ℤGroup H) → - abs (invEq (fst iso2) (fst (fst hom) e)) ≡ 1) - λ e p hom → cong (abs ∘ (fst (fst hom))) p ∙ autoPres1 hom - where - autoPres1 : (e : GroupEquiv ℤGroup ℤGroup) - → abs (fst (fst e) 1) ≡ 1 - autoPres1 e = GroupEquivℤ-pres1 e _ refl --} - -allisoPresGen : ∀ {ℓ} (G : Group ℓ) (ϕ : GroupEquiv G ℤGroup) (x : fst G) +groupEquivPresGen : ∀ {ℓ} (G : Group ℓ) (ϕ : GroupEquiv G ℤGroup) (x : fst G) → (fst (fst ϕ) x ≡ 1) ⊎ (fst (fst ϕ) x ≡ - 1) → (ψ : GroupEquiv G ℤGroup) → (fst (fst ψ) x ≡ 1) ⊎ (fst (fst ψ) x ≡ - 1) -allisoPresGen G (ϕeq , ϕhom) x (inl r) (ψeq , ψhom) = +groupEquivPresGen G (ϕeq , ϕhom) x (inl r) (ψeq , ψhom) = abs→⊎ _ _ (cong abs (cong (fst ψeq) (sym (retEq ϕeq x) ∙ cong (invEq ϕeq) r)) ∙ GroupEquivℤ-pres1 (compGroupEquiv (invGroupEquiv (ϕeq , ϕhom)) (ψeq , ψhom)) _ refl) -allisoPresGen G (ϕeq , ϕhom) x (inr r) (ψeq , ψhom) = +groupEquivPresGen G (ϕeq , ϕhom) x (inr r) (ψeq , ψhom) = abs→⊎ _ _ (cong abs (cong (fst ψeq) (sym (retEq ϕeq x) ∙ cong (invEq ϕeq) r)) - ∙ cong abs (IsGroupHom.presinv (snd (compGroupEquiv (invGroupEquiv (ϕeq , ϕhom)) (ψeq , ψhom))) 1) - ∙ absLem _ (GroupEquivℤ-pres1 (compGroupEquiv (invGroupEquiv (ϕeq , ϕhom)) (ψeq , ψhom)) _ refl)) + ∙ cong abs (IsGroupHom.presinv + (snd (compGroupEquiv (invGroupEquiv (ϕeq , ϕhom)) + (ψeq , ψhom))) 1) + ∙ absLem _ (GroupEquivℤ-pres1 + (compGroupEquiv (invGroupEquiv (ϕeq , ϕhom)) (ψeq , ψhom)) + _ refl)) where absLem : ∀ x → abs x ≡ 1 → abs (pos 0 -ℤ x) ≡ 1 absLem (pos zero) p = ⊥-rec (snotz (sym p)) diff --git a/Cubical/Homotopy/HopfInvariant/HopfMap.agda b/Cubical/Homotopy/HopfInvariant/HopfMap.agda index c14c7979e..2e3a6db42 100644 --- a/Cubical/Homotopy/HopfInvariant/HopfMap.agda +++ b/Cubical/Homotopy/HopfInvariant/HopfMap.agda @@ -542,14 +542,14 @@ CP2≡CP²' = → abs (fst (fst 3rdEq) (h₁ ⌣ h₁)) ≡ 1 help g h (ϕ , idg) (ψ , idh) ⌣eq ξ = ⊎→abs _ _ - (allisoPresGen _ + (groupEquivPresGen _ (compGroupEquiv main (compGroupEquiv (invGroupEquiv main) ψ)) h (abs→⊎ _ _ (cong abs (cong (fst (fst ψ)) (retEq (fst main) h)) ∙ idh)) (compGroupEquiv main ξ)) where lem₁ : ((fst (fst ψ) h) ≡ 1) ⊎ (fst (fst ψ) h ≡ -1) → abs (fst (fst ϕ) h) ≡ 1 - lem₁ p = ⊎→abs _ _ (allisoPresGen _ ψ h p ϕ) + lem₁ p = ⊎→abs _ _ (groupEquivPresGen _ ψ h p ϕ) lem₂ : ((fst (fst ϕ) h) ≡ 1) ⊎ (fst (fst ϕ) h ≡ -1) → ((fst (fst ϕ) g) ≡ 1) ⊎ (fst (fst ϕ) g ≡ -1) From 0c169263b23d531fa73dfc57b71d0c0a73565ffa Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Mon, 1 Nov 2021 14:40:35 +0100 Subject: [PATCH 77/89] minor --- Cubical/Algebra/Group/ZModule.agda | 16 +++++--- Cubical/Data/Sigma/Properties.agda | 7 ---- Cubical/Foundations/Pointed/Homogeneous.agda | 43 -------------------- 3 files changed, 11 insertions(+), 55 deletions(-) diff --git a/Cubical/Algebra/Group/ZModule.agda b/Cubical/Algebra/Group/ZModule.agda index ca326ad1f..d2360b843 100644 --- a/Cubical/Algebra/Group/ZModule.agda +++ b/Cubical/Algebra/Group/ZModule.agda @@ -209,9 +209,12 @@ private lem zero zero = 1 , refl lem zero (suc y) = (suc (suc y)) , +Comm (negsuc zero) (negsuc (suc y)) lem (suc x) zero = (suc (suc x)) , refl - lem (suc x) (suc y) = (lem (suc (suc x)) y .fst) , (predℤ+negsuc y (negsuc (suc x)) ∙ snd ((lem (suc (suc x))) y)) + lem (suc x) (suc y) = + (lem (suc (suc x)) y .fst) + , (predℤ+negsuc y (negsuc (suc x)) ∙ snd ((lem (suc (suc x))) y)) - intLem₂ : (n x : ℕ) → Σ[ a ∈ ℕ ] ((pos (suc x)) * pos (suc (suc n)) ≡ pos (suc (suc a))) + intLem₂ : (n x : ℕ) + → Σ[ a ∈ ℕ ] ((pos (suc x)) * pos (suc (suc n)) ≡ pos (suc (suc a))) intLem₂ n zero = n , refl intLem₂ n (suc x) = h3 _ _ (intLem₂ n x) where @@ -239,7 +242,8 @@ GroupEquivℤ-pres1 e (pos (suc (suc n))) p = where h3 : pos 1 ≡ _ h3 = sym (retEq (fst e) 1) - ∙∙ (cong (fst (fst (invGroupEquiv e))) (p ∙ ·Comm 1 (pos (suc (suc n))))) + ∙∙ cong (fst (fst (invGroupEquiv e))) + (p ∙ ·Comm 1 (pos (suc (suc n)))) ∙∙ GroupHomℤ→ℤPres· (_ , snd (invGroupEquiv e)) (pos (suc (suc n))) 1 GroupEquivℤ-pres1 e (negsuc zero) p = cong abs p GroupEquivℤ-pres1 e (negsuc (suc n)) p = ⊥-rec (¬1=x·suc-suc _ _ lem₂) @@ -267,8 +271,10 @@ groupEquivPresGen : ∀ {ℓ} (G : Group ℓ) (ϕ : GroupEquiv G ℤGroup) (x : → (ψ : GroupEquiv G ℤGroup) → (fst (fst ψ) x ≡ 1) ⊎ (fst (fst ψ) x ≡ - 1) groupEquivPresGen G (ϕeq , ϕhom) x (inl r) (ψeq , ψhom) = - abs→⊎ _ _ (cong abs (cong (fst ψeq) (sym (retEq ϕeq x) ∙ cong (invEq ϕeq) r)) - ∙ GroupEquivℤ-pres1 (compGroupEquiv (invGroupEquiv (ϕeq , ϕhom)) (ψeq , ψhom)) _ refl) + abs→⊎ _ _ (cong abs (cong (fst ψeq) (sym (retEq ϕeq x) + ∙ cong (invEq ϕeq) r)) + ∙ GroupEquivℤ-pres1 (compGroupEquiv + (invGroupEquiv (ϕeq , ϕhom)) (ψeq , ψhom)) _ refl) groupEquivPresGen G (ϕeq , ϕhom) x (inr r) (ψeq , ψhom) = abs→⊎ _ _ (cong abs (cong (fst ψeq) (sym (retEq ϕeq x) ∙ cong (invEq ϕeq) r)) diff --git a/Cubical/Data/Sigma/Properties.agda b/Cubical/Data/Sigma/Properties.agda index 607a86415..7ea7d146e 100644 --- a/Cubical/Data/Sigma/Properties.agda +++ b/Cubical/Data/Sigma/Properties.agda @@ -359,10 +359,3 @@ module _ {A : Type ℓ} {B : A → Type ℓ'} {C : ∀ a → B a → Type ℓ''} Iso.leftInv curryIso f = refl unquoteDecl curryEquiv = declStrictIsoToEquiv curryEquiv curryIso - -ΣΣ : Iso (Σ[ x ∈ A ] (Σ[ y ∈ B x ] (C x y))) - (Σ[ x ∈ Σ A B ] C (fst x) (snd x)) -fun ΣΣ (x , y , p) = (x , y) , p -inv ΣΣ ((x , y) , p) = x , y , p -rightInv ΣΣ _ = refl -leftInv ΣΣ _ = refl diff --git a/Cubical/Foundations/Pointed/Homogeneous.agda b/Cubical/Foundations/Pointed/Homogeneous.agda index 7bb614d26..5315e2685 100644 --- a/Cubical/Foundations/Pointed/Homogeneous.agda +++ b/Cubical/Foundations/Pointed/Homogeneous.agda @@ -46,49 +46,6 @@ isHomogeneous {ℓ} (A , x) = ∀ y → Path (Pointed ℓ) (A , x) (A , y) }) (sym (h (pt B∙)) ∙ h ((sym f₀ ∙∙ funExt⁻ p a₀ ∙∙ g₀) i)) -→∙Homogeneous≡Path : ∀ {ℓ ℓ'} {A∙ : Pointed ℓ} {B∙ : Pointed ℓ'} {f∙ g∙ : A∙ →∙ B∙} - (h : isHomogeneous B∙) → (p q : f∙ ≡ g∙) → cong fst p ≡ cong fst q → p ≡ q -→∙Homogeneous≡Path {A∙ = A∙@(A , a₀)} {B∙@(B , b)} {f∙@(f , f₀)} {g∙@(g , g₀)} h p q r = - transport (λ k - → PathP (λ i - → PathP (λ j → (A , a₀) →∙ newPath-refl p q r i j (~ k)) - (f , f₀) (g , g₀)) p q) - (badPath p q r) - where - newPath : (p q : f∙ ≡ g∙) (r : cong fst p ≡ cong fst q) - → Square (refl {x = b}) refl refl refl - newPath p q r i j = - hcomp (λ k → λ {(i = i0) → cong snd p j k - ; (i = i1) → cong snd q j k - ; (j = i0) → f₀ k - ; (j = i1) → g₀ k}) - (r i j a₀) - - newPath-refl : (p q : f∙ ≡ g∙) (r : cong fst p ≡ cong fst q) - → PathP (λ i → (PathP (λ j → B∙ ≡ (B , newPath p q r i j))) refl refl) refl refl - newPath-refl p q r i j k = - hcomp (λ w → λ { (i = i0) → lCancel (h b) w k - ; (i = i1) → lCancel (h b) w k - ; (j = i0) → lCancel (h b) w k - ; (j = i1) → lCancel (h b) w k - ; (k = i0) → lCancel (h b) w k - ; (k = i1) → B , newPath p q r i j}) - ((sym (h b) ∙ h (newPath p q r i j)) k) - - badPath : (p q : f∙ ≡ g∙) (r : cong fst p ≡ cong fst q) - → PathP (λ i → - PathP (λ j → A∙ →∙ (B , newPath p q r i j)) - (f , f₀) (g , g₀)) - p q - fst (badPath p q r i j) = r i j - snd (badPath p q s i j) k = - hcomp (λ r → λ { (i = i0) → snd (p j) (r ∧ k) - ; (i = i1) → snd (q j) (r ∧ k) - ; (j = i0) → f₀ (k ∧ r) - ; (j = i1) → g₀ (k ∧ r) - ; (k = i0) → s i j a₀}) - (s i j a₀) - isHomogeneousPi : ∀ {ℓ ℓ'} {A : Type ℓ} {B∙ : A → Pointed ℓ'} → (∀ a → isHomogeneous (B∙ a)) → isHomogeneous (Πᵘ∙ A B∙) isHomogeneousPi h f i .fst = ∀ a → typ (h a (f a) i) From b2b4cad5d74922c12630753d63b676d87f128386 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 2 Nov 2021 13:40:16 +0100 Subject: [PATCH 78/89] fixes --- Cubical/HITs/Hopf.agda | 41 - Cubical/HITs/Join/Properties.agda | 56 +- Cubical/HITs/Sn/Properties.agda | 11 +- Cubical/Homotopy/Connected.agda | 12 + Cubical/Homotopy/Freudenthal.agda | 19 +- Cubical/Homotopy/Group/Base.agda | 13 +- Cubical/Homotopy/Group/Properties.agda | 68 +- .../Homotopy/Group/SuspensionMapPathP.agda | 36 +- Cubical/Homotopy/HSpace.agda | 52 ++ Cubical/Homotopy/HomotopyGroup.agda | 724 ------------------ Cubical/Homotopy/Hopf.agda | 485 ++++++------ .../Homotopy/HopfInvariant/Homomorphism.agda | 2 - Cubical/Homotopy/HopfInvariant/HopfMap.agda | 46 +- 13 files changed, 394 insertions(+), 1171 deletions(-) create mode 100644 Cubical/Homotopy/HSpace.agda delete mode 100644 Cubical/Homotopy/HomotopyGroup.agda diff --git a/Cubical/HITs/Hopf.agda b/Cubical/HITs/Hopf.agda index 7c26352e6..0b72621fb 100644 --- a/Cubical/HITs/Hopf.agda +++ b/Cubical/HITs/Hopf.agda @@ -338,44 +338,3 @@ leftInv IsoS³TotalHopf x = (JoinS¹S¹→TotalHopf→JoinS¹S¹ (S³→joinS¹S¹ (inv IsoS³S3 x))) ∙∙ cong (fun IsoS³S3) (S³→joinS¹S¹→S³ (inv IsoS³S3 x)) ∙∙ Iso.rightInv IsoS³S3 x - --- kebab : TotalHopf → TotalHopf → TotalHopf --- kebab (north , b) (c , d) = c , {!b!} --- kebab (south , b) (c , d) = {!!} --- kebab (merid a i , b) (c , d) = {!!} - --- S¹→Space : (x : S¹) → Type --- S¹→Space base = TotalHopf --- S¹→Space (loop i) = {!!} - --- TotalHopf→TotalHopf : TotalHopf → TotalHopf --- TotalHopf→TotalHopf (north , p) = north , invLooper p --- TotalHopf→TotalHopf (south , p) = south , p --- TotalHopf→TotalHopf (merid a i , p) = {!transport (λ i → Glue S¹ (Border a i) → TotalHopf) (λ x → north , invLooper x)!} --- where --- kebab2 : PathP (λ i → Glue S¹ (Border a i) → TotalHopf) (λ x → north , invLooper x) λ p → south , p --- kebab2 = toPathP (funExt λ x → transportRefl _ ∙ ΣPathP (refl , {!p!}) ∙ {!!}) - - --- open import Cubical.Data.Sum - --- code : ∀ {ℓ} {A : Type ℓ} {B : Type ℓ} (x : A) → A ⊎ B → Type _ --- code {A = A} {B = B} x (inl y) = x ≡ y --- code {A = A} {B = B} y (inr x) = Path (A ⊎ B) (inl y) (inr x) - --- decode' : ∀ {ℓ} {A : Type ℓ} {B : Type ℓ} → (y : A ⊎ B) (x : A) → code x y → inl x ≡ y --- decode' (inl x₁) x p = cong inl p --- decode' (inr x₁) x p = p - --- encode' : ∀ {ℓ} {A : Type ℓ} {B : Type ℓ} (x : A) (y : A ⊎ B) → inl x ≡ y → code x y --- encode' x y = J (λ y _ → code x y) refl - --- encode-decode : ∀ {ℓ} {A : Type ℓ} {B : Type ℓ} (x : A) (y : A) → (s : code {A = A} {B = B} x (inl y)) → encode' {A = A} {B = B} x (inl y) (decode' (inl y) x s) ≡ s --- encode-decode {A = A} {B = B} x y = J (λ y s → encode' {A = A} {B = B} x (inl y) (λ i → inl (s i)) ≡ s) (transportRefl refl) - --- en - --- lal : S₊ 1 → TotalHopf → TotalHopf --- lal a (north , p) = {!merid a ?!} , {!!} --- lal a (south , p) = {!!} --- lal a (merid a₁ i , p) = {!!} diff --git a/Cubical/HITs/Join/Properties.agda b/Cubical/HITs/Join/Properties.agda index 24a6fe649..7ea9d7b36 100644 --- a/Cubical/HITs/Join/Properties.agda +++ b/Cubical/HITs/Join/Properties.agda @@ -29,17 +29,19 @@ private variable ℓ ℓ' : Level +open Iso + -- Characterisation of function type join A B → C IsoFunSpaceJoin : ∀ {ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} → Iso (join A B → C) (Σ[ f ∈ (A → C) ] Σ[ g ∈ (B → C) ] ((a : A) (b : B) → f a ≡ g b)) -Iso.fun IsoFunSpaceJoin f = (f ∘ inl) , ((f ∘ inr) , (λ a b → cong f (push a b))) -Iso.inv IsoFunSpaceJoin (f , g , p) (inl x) = f x -Iso.inv IsoFunSpaceJoin (f , g , p) (inr x) = g x -Iso.inv IsoFunSpaceJoin (f , g , p) (push a b i) = p a b i -Iso.rightInv IsoFunSpaceJoin (f , g , p) = refl -Iso.leftInv IsoFunSpaceJoin f = +fun IsoFunSpaceJoin f = (f ∘ inl) , ((f ∘ inr) , (λ a b → cong f (push a b))) +inv IsoFunSpaceJoin (f , g , p) (inl x) = f x +inv IsoFunSpaceJoin (f , g , p) (inr x) = g x +inv IsoFunSpaceJoin (f , g , p) (push a b i) = p a b i +rightInv IsoFunSpaceJoin (f , g , p) = refl +leftInv IsoFunSpaceJoin f = funExt λ { (inl x) → refl ; (inr x) → refl ; (push a b i) → refl} -- Alternative definition of the join using a pushout @@ -444,24 +446,28 @@ join-commFun² (inl x) = refl join-commFun² (inr x) = refl join-commFun² (push a b i) = refl -join-comm : ∀ {ℓ'} {A : Type ℓ} {B : Type ℓ'} → Iso (join A B) (join B A) -Iso.fun join-comm = join-commFun -Iso.inv join-comm = join-commFun -Iso.rightInv join-comm = join-commFun² -Iso.leftInv join-comm = join-commFun² +join-comm : ∀ {ℓ'} {A : Type ℓ} {B : Type ℓ'} + → Iso (join A B) (join B A) +fun join-comm = join-commFun +inv join-comm = join-commFun +rightInv join-comm = join-commFun² +leftInv join-comm = join-commFun² -- Applying Isos to joins (more efficient than transports) -Iso→joinIso : ∀ {ℓ' ℓ'' ℓ'''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} - → Iso A C → Iso B D → Iso (join A B) (join C D) -Iso.fun (Iso→joinIso is1 is2) (inl x) = inl (Iso.fun is1 x) -Iso.fun (Iso→joinIso is1 is2) (inr x) = inr (Iso.fun is2 x) -Iso.fun (Iso→joinIso is1 is2) (push a b i) = push (Iso.fun is1 a) (Iso.fun is2 b) i -Iso.inv (Iso→joinIso is1 is2) (inl x) = inl (Iso.inv is1 x) -Iso.inv (Iso→joinIso is1 is2) (inr x) = inr (Iso.inv is2 x) -Iso.inv (Iso→joinIso is1 is2) (push a b i) = push (Iso.inv is1 a) (Iso.inv is2 b) i -Iso.rightInv (Iso→joinIso is1 is2) (inl x) i = inl (Iso.rightInv is1 x i) -Iso.rightInv (Iso→joinIso is1 is2) (inr x) i = inr (Iso.rightInv is2 x i) -Iso.rightInv (Iso→joinIso is1 is2) (push a b j) i = push (Iso.rightInv is1 a i) (Iso.rightInv is2 b i) j -Iso.leftInv (Iso→joinIso is1 is2) (inl x) i = inl (Iso.leftInv is1 x i) -Iso.leftInv (Iso→joinIso is1 is2) (inr x) i = inr (Iso.leftInv is2 x i) -Iso.leftInv (Iso→joinIso is1 is2) (push a b i) j = push (Iso.leftInv is1 a j) (Iso.leftInv is2 b j) i +Iso→joinIso : ∀ {ℓ'' ℓ'''} + {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} + → Iso A C → Iso B D → Iso (join A B) (join C D) +fun (Iso→joinIso is1 is2) (inl x) = inl (fun is1 x) +fun (Iso→joinIso is1 is2) (inr x) = inr (fun is2 x) +fun (Iso→joinIso is1 is2) (push a b i) = push (fun is1 a) (fun is2 b) i +inv (Iso→joinIso is1 is2) (inl x) = inl (inv is1 x) +inv (Iso→joinIso is1 is2) (inr x) = inr (inv is2 x) +inv (Iso→joinIso is1 is2) (push a b i) = push (inv is1 a) (inv is2 b) i +rightInv (Iso→joinIso is1 is2) (inl x) i = inl (rightInv is1 x i) +rightInv (Iso→joinIso is1 is2) (inr x) i = inr (rightInv is2 x i) +rightInv (Iso→joinIso is1 is2) (push a b j) i = + push (rightInv is1 a i) (rightInv is2 b i) j +leftInv (Iso→joinIso is1 is2) (inl x) i = inl (leftInv is1 x i) +leftInv (Iso→joinIso is1 is2) (inr x) i = inr (leftInv is2 x i) +leftInv (Iso→joinIso is1 is2) (push a b i) j = + push (leftInv is1 a j) (leftInv is2 b j) i diff --git a/Cubical/HITs/Sn/Properties.agda b/Cubical/HITs/Sn/Properties.agda index 6af47f72f..7fbeef2d3 100644 --- a/Cubical/HITs/Sn/Properties.agda +++ b/Cubical/HITs/Sn/Properties.agda @@ -317,17 +317,22 @@ isConnectedPathSⁿ n x y = -- Equivalence Sⁿ*Sᵐ≃Sⁿ⁺ᵐ⁺¹ -IsoSphereJoin : (n m : ℕ) → Iso (join (S₊ n) (S₊ m)) (S₊ (suc (n + m))) +IsoSphereJoin : (n m : ℕ) + → Iso (join (S₊ n) (S₊ m)) (S₊ (suc (n + m))) IsoSphereJoin zero m = compIso join-comm (compIso (invIso Susp-iso-joinBool) (invIso (IsoSucSphereSusp m))) IsoSphereJoin (suc n) m = compIso (Iso→joinIso - (subst (λ x → Iso (S₊ (suc x)) (join (S₊ n) Bool)) (+-comm n 0) (invIso (IsoSphereJoin n 0))) + (subst (λ x → Iso (S₊ (suc x)) (join (S₊ n) Bool)) + (+-comm n 0) (invIso (IsoSphereJoin n 0))) idIso) (compIso (equivToIso joinAssocDirect) - (compIso (Iso→joinIso idIso (compIso join-comm (compIso (invIso Susp-iso-joinBool) (invIso (IsoSucSphereSusp m))))) + (compIso (Iso→joinIso idIso + (compIso join-comm + (compIso (invIso Susp-iso-joinBool) + (invIso (IsoSucSphereSusp m))))) (compIso (IsoSphereJoin n (suc m)) (pathToIso λ i → S₊ (suc (+-suc n m i)))))) diff --git a/Cubical/Homotopy/Connected.agda b/Cubical/Homotopy/Connected.agda index 492cfacbc..3e4b6dbdb 100644 --- a/Cubical/Homotopy/Connected.agda +++ b/Cubical/Homotopy/Connected.agda @@ -186,6 +186,18 @@ isConnectedCong n f cf {a₀} {a₁} q = (sym (fiberCong f q)) (isConnectedPath n (cf (f a₁)) (a₀ , q) (a₁ , refl)) +isConnectedCong' : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x : A} {y : B} + (n : ℕ) (f : A → B) + → isConnectedFun (suc n) f + → (p : f x ≡ y) + → isConnectedFun n + λ (q : x ≡ x) → sym p ∙∙ cong f q ∙∙ p +isConnectedCong' {x = x} n f conf p = + transport (λ i → (isConnectedFun n + λ (q : x ≡ x) + → doubleCompPath-filler (sym p) (cong f q) p i)) + (isConnectedCong n f conf) + isConnectedRetract : ∀ {ℓ ℓ'} (n : HLevel) {A : Type ℓ} {B : Type ℓ'} (f : A → B) (g : B → A) diff --git a/Cubical/Homotopy/Freudenthal.agda b/Cubical/Homotopy/Freudenthal.agda index a69038af8..d471c16f9 100644 --- a/Cubical/Homotopy/Freudenthal.agda +++ b/Cubical/Homotopy/Freudenthal.agda @@ -10,6 +10,7 @@ open import Cubical.Foundations.Everything open import Cubical.Data.Nat open import Cubical.Data.Nat.Order open import Cubical.Data.Sigma +open import Cubical.Data.Empty renaming (rec to ⊥-rec) open import Cubical.HITs.Nullification open import Cubical.HITs.Susp renaming (toSusp to σ) open import Cubical.HITs.Truncation as Trunc renaming (rec to trRec ; elim to trElim) @@ -17,14 +18,11 @@ open import Cubical.Homotopy.Connected open import Cubical.Homotopy.WedgeConnectivity open import Cubical.Homotopy.Loopspace open import Cubical.HITs.SmashProduct -open import Cubical.Relation.Nullary -open import Cubical.Data.Empty renaming (rec to ⊥-rec) open import Cubical.HITs.S1 hiding (encode) open import Cubical.HITs.Sn open import Cubical.HITs.S2 open import Cubical.HITs.S3 -open import Cubical.Foundations.Equiv.HalfAdjoint module _ {ℓ} (n : HLevel) {A : Pointed ℓ} (connA : isConnected (suc (suc n)) (typ A)) where @@ -143,25 +141,12 @@ suspMapΩ∙ : ∀ {ℓ} {A : Pointed ℓ}(n : ℕ) →∙ ((Ω^ (suc n)) (Susp∙ (typ A))) fst (suspMapΩ∙ {A = A} zero) a = merid a ∙ sym (merid (pt A)) snd (suspMapΩ∙ {A = A} zero) = rCancel (merid (pt A)) -fst (suspMapΩ∙ {A = A} (suc n)) p = sym (snd (suspMapΩ∙ n)) ∙∙ cong (fst (suspMapΩ∙ n)) p ∙∙ snd (suspMapΩ∙ n) -snd (suspMapΩ∙ {A = A} (suc n)) = ∙∙lCancel (snd (suspMapΩ∙ n)) +suspMapΩ∙ {A = A} (suc n) = Ω→ (suspMapΩ∙ {A = A} n) suspMapΩ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → typ ((Ω^ n) A) → typ ((Ω^ (suc n)) (Susp∙ (typ A))) suspMapΩ {A = A} n = suspMapΩ∙ {A = A} n .fst -isConnectedCong' : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x : A} {y : B} - (n : ℕ) (f : A → B) - → isConnectedFun (suc n) f - → (p : f x ≡ y) - → isConnectedFun n - λ (q : x ≡ x) → sym p ∙∙ cong f q ∙∙ p -isConnectedCong' {x = x} n f conf p = - transport (λ i → (isConnectedFun n - λ (q : x ≡ x) - → doubleCompPath-filler (sym p) (cong f q) p i)) - (isConnectedCong n f conf) - suspMapΩ-connected : ∀ {ℓ} (n : HLevel) (m : ℕ) {A : Pointed ℓ} (connA : isConnected (suc (suc n)) (typ A)) → isConnectedFun ((suc n + suc n) ∸ m) (suspMapΩ {A = A} m) diff --git a/Cubical/Homotopy/Group/Base.agda b/Cubical/Homotopy/Group/Base.agda index b63241ffc..d2f85a816 100644 --- a/Cubical/Homotopy/Group/Base.agda +++ b/Cubical/Homotopy/Group/Base.agda @@ -7,18 +7,15 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed open import Cubical.Foundations.HLevels open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Function -open import Cubical.Foundations.Path -open import Cubical.Foundations.Transport open import Cubical.HITs.SetTruncation - renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim - ; elim2 to sElim2 ; elim3 to sElim3 ; map to sMap) + renaming (rec to sRec ; rec2 to sRec2 + ; elim to sElim ; elim2 to sElim2 ; elim3 to sElim3 + ; map to sMap) open import Cubical.HITs.Sn -open import Cubical.Data.Bool open import Cubical.HITs.Susp open import Cubical.HITs.S1 + open import Cubical.Data.Sigma open import Cubical.Data.Nat @@ -26,7 +23,6 @@ open import Cubical.Algebra.Group open import Cubical.Algebra.Semigroup open import Cubical.Algebra.Monoid -open Iso open IsGroup open IsSemigroup open IsMonoid @@ -77,6 +73,7 @@ open GroupStr → ·π (suc n) x y ≡ ·π (suc n) y x π-comm n = sElim2 (λ _ _ → isSetPathImplicit) λ p q i → ∣ EH n p q i ∣₂ +-- πₙ₊₁ πGr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Group ℓ fst (πGr n A) = π (suc n) A 1g (snd (πGr n A)) = 1π (suc n) diff --git a/Cubical/Homotopy/Group/Properties.agda b/Cubical/Homotopy/Group/Properties.agda index 670131bc1..1de45df29 100644 --- a/Cubical/Homotopy/Group/Properties.agda +++ b/Cubical/Homotopy/Group/Properties.agda @@ -23,28 +23,31 @@ open import Cubical.Homotopy.Group.SuspensionMapPathP open import Cubical.Homotopy.Group.SuspensionMapPathP using (IsoΩSphereMap) public open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Freudenthal +open import Cubical.Homotopy.Connected open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous open import Cubical.Foundations.HLevels open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Function -open import Cubical.Foundations.Path -open import Cubical.Foundations.Transport -open import Cubical.Functions.Embedding +open import Cubical.Foundations.Univalence + open import Cubical.Functions.Morphism -open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; rec2 to pRec2 ; elim to pElim) open import Cubical.HITs.SetTruncation renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; elim3 to sElim3 ; map to sMap) -open import Cubical.HITs.PropositionalTruncation - renaming (rec to pRec ; rec2 to pRec2 ; elim to pElim) +open import Cubical.HITs.Truncation + renaming (rec to trRec) +open import Cubical.HITs.S1 open import Cubical.HITs.Sn -open import Cubical.Data.Bool open import Cubical.HITs.Susp -open import Cubical.HITs.S1 +open import Cubical.Data.Bool open import Cubical.Data.Sigma open import Cubical.Data.Nat @@ -58,18 +61,6 @@ open IsSemigroup open IsMonoid open GroupStr -open import Cubical.Homotopy.Freudenthal -open import Cubical.Homotopy.Connected -open import Cubical.HITs.Truncation - renaming (rec to trRec) - - -open import Cubical.Data.Sum -open import Cubical.Relation.Nullary -open import Cubical.Data.Empty - -open import Cubical.Foundations.Equiv.HalfAdjoint - -- We finally arrive at the main result IsoSphereMapΩ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) @@ -143,8 +134,8 @@ IsoSphereMapΩ-pres∙Π n = morphLemmas.isMorphInv _∙_ ∙Π (Ω→SphereMap (suc n)) (isHom-lMap n) (SphereMap→Ω (suc n)) - (Iso.rightInv (IsoΩSphereMap (suc n))) - (Iso.leftInv (IsoΩSphereMap (suc n))) + (rightInv (IsoΩSphereMap (suc n))) + (leftInv (IsoΩSphereMap (suc n))) -- It is useful to define the ``Group Structure'' on (S₊∙ n →∙ A) -- before doing it on π'. These will be the equivalents of the @@ -296,23 +287,23 @@ snd (∙Π-lCancel {A = A} {n = suc n} f i) = refl → (f g h : S₊∙ (suc n) →∙ A) → ∙Π f (∙Π g h) ≡ ∙Π (∙Π f g) h ∙Π-assoc {n = n} f g h = - sym (Iso.leftInv (IsoSphereMapΩ (suc n)) (∙Π f (∙Π g h))) + sym (leftInv (IsoSphereMapΩ (suc n)) (∙Π f (∙Π g h))) ∙∙ cong (Ω→SphereMap (suc n)) (IsoSphereMapΩ-pres∙Π n f (∙Π g h) ∙∙ cong (SphereMap→Ω (suc n) f ∙_) (IsoSphereMapΩ-pres∙Π n g h) ∙∙ ∙assoc (SphereMap→Ω (suc n) f) (SphereMap→Ω (suc n) g) (SphereMap→Ω (suc n) h) ∙∙ cong (_∙ SphereMap→Ω (suc n) h) (sym (IsoSphereMapΩ-pres∙Π n f g)) ∙∙ sym (IsoSphereMapΩ-pres∙Π n (∙Π f g) h)) - ∙∙ Iso.leftInv (IsoSphereMapΩ (suc n)) (∙Π (∙Π f g) h) + ∙∙ leftInv (IsoSphereMapΩ (suc n)) (∙Π (∙Π f g) h) ∙Π-comm : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → (f g : S₊∙ (suc (suc n)) →∙ A) → ∙Π f g ≡ ∙Π g f ∙Π-comm {A = A} {n = n} f g = - sym (Iso.leftInv (IsoSphereMapΩ (suc (suc n))) (∙Π f g)) + sym (leftInv (IsoSphereMapΩ (suc (suc n))) (∙Π f g)) ∙∙ cong (Ω→SphereMap (suc (suc n))) (IsoSphereMapΩ-pres∙Π (suc n) f g ∙∙ EH _ _ _ ∙∙ sym (IsoSphereMapΩ-pres∙Π (suc n) g f)) - ∙∙ Iso.leftInv (IsoSphereMapΩ (suc (suc n))) (∙Π g f) + ∙∙ leftInv (IsoSphereMapΩ (suc (suc n))) (∙Π g f) {- π'' as a group -} 1π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' n A @@ -559,17 +550,17 @@ hom-rMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) hom-rMap {A = A} n p q = cong (lMap (suc n) {A = Ω (Susp∙ (typ A))}) (morphLemmas.isMorphInv _∙_ _∙_ - (Iso.inv (flipΩIso (suc n))) (flipΩIsopres· n) - (Iso.fun (flipΩIso (suc n))) - (Iso.leftInv (flipΩIso (suc n))) (Iso.rightInv (flipΩIso (suc n))) + (inv (flipΩIso (suc n))) (flipΩIsopres· n) + (fun (flipΩIso (suc n))) + (leftInv (flipΩIso (suc n))) (rightInv (flipΩIso (suc n))) p q) ∙ isHom-lMap _ (fun (flipΩIso (suc n)) p) (fun (flipΩIso (suc n)) q) isHom-IsoΩSphereMap' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p q : typ (Ω ((Ω^ n) (Susp∙ (typ A))))) - → Iso.fun (IsoΩSphereMap' {A = A} n) (p ∙ q) - ≡ ∙Π {n = (suc n)} (Iso.fun (IsoΩSphereMap' {A = A} n) p) - (Iso.fun (IsoΩSphereMap' {A = A} n) q) + → fun (IsoΩSphereMap' {A = A} n) (p ∙ q) + ≡ ∙Π {n = (suc n)} (fun (IsoΩSphereMap' {A = A} n) p) + (fun (IsoΩSphereMap' {A = A} n) q) isHom-IsoΩSphereMap' {A = A} zero p q = (cong (botᵣ {A = A} 0) (cong ((lMap 0 {A = Ω (Susp∙ (typ A))})) @@ -590,19 +581,19 @@ isHom-IsoΩSphereMap' {A = A} (suc n) p q = cong (botᵣ {A = A} (suc n) ∘ lMap (suc n) {A = Ω (Susp∙ (typ A))}) (morphLemmas.isMorphInv _∙_ _∙_ - (Iso.inv (flipΩIso (suc n))) + (inv (flipΩIso (suc n))) (flipΩIsopres· n) - (Iso.fun (flipΩIso (suc n))) - (Iso.leftInv (flipΩIso (suc n))) - (Iso.rightInv (flipΩIso (suc n))) p q) + (fun (flipΩIso (suc n))) + (leftInv (flipΩIso (suc n))) + (rightInv (flipΩIso (suc n))) p q) ∙ (cong (botᵣ {A = A} (suc n)) (isHom-lMap _ {A = Ω (Susp∙ (typ A))} (fun (flipΩIso (suc n)) p) (fun (flipΩIso (suc n)) q)) ∙ morphLemmas.isMorphInv ∙Π ∙Π (botᵣ⁻ {A = A} (suc n)) (hom-botᵣ⁻' {A = A} n) (botᵣ {A = A} (suc n)) - (Iso.leftInv (botᵣIso {A = A} (suc n))) - (Iso.rightInv (botᵣIso {A = A} (suc n))) + (leftInv (botᵣIso {A = A} (suc n))) + (rightInv (botᵣIso {A = A} (suc n))) (lMap (suc n) (fun (flipΩIso (suc n)) p)) (lMap (suc n) (fun (flipΩIso (suc n)) q))) @@ -663,7 +654,6 @@ private ∙∙ ((sym (transportRefl _) ∙ (trId2 (f x) (f y))) ∙ cong₂ +B2 (transportRefl (f x)) (transportRefl (f y)))) -open import Cubical.Foundations.Univalence suspMap→hom : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f g : S₊∙ (suc n) →∙ A) → suspMap n (∙Π f g) ≡ ∙Π (suspMap n f) (suspMap n g) suspMap→hom {A = A} n = diff --git a/Cubical/Homotopy/Group/SuspensionMapPathP.agda b/Cubical/Homotopy/Group/SuspensionMapPathP.agda index 73018c60e..8487e5dc7 100644 --- a/Cubical/Homotopy/Group/SuspensionMapPathP.agda +++ b/Cubical/Homotopy/Group/SuspensionMapPathP.agda @@ -27,28 +27,27 @@ module Cubical.Homotopy.Group.SuspensionMapPathP where open import Cubical.Homotopy.Group.Base open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Freudenthal +open import Cubical.Homotopy.Connected open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous open import Cubical.Foundations.HLevels open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Function open import Cubical.Foundations.Path +open import Cubical.Foundations.Equiv open import Cubical.Foundations.Transport -open import Cubical.Functions.Embedding + open import Cubical.Functions.Morphism -open import Cubical.Foundations.Pointed.Homogeneous -open import Cubical.HITs.SetTruncation - renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim - ; elim2 to sElim2 ; elim3 to sElim3 ; map to sMap) -open import Cubical.HITs.PropositionalTruncation - renaming (rec to pRec ; rec2 to pRec2 ; elim to pElim) open import Cubical.HITs.Sn -open import Cubical.Data.Bool hiding (_≟_) open import Cubical.HITs.Susp open import Cubical.HITs.S1 + +open import Cubical.Data.Bool hiding (_≟_) open import Cubical.Data.Sigma open import Cubical.Data.Nat @@ -62,19 +61,6 @@ open IsSemigroup open IsMonoid open GroupStr -open import Cubical.Homotopy.Freudenthal -open import Cubical.Homotopy.Connected -open import Cubical.Foundations.Equiv -open import Cubical.HITs.Truncation - renaming (rec to trRec ; elim to trElim ; map to trMap) - - -open import Cubical.Data.Sum -open import Cubical.Relation.Nullary -open import Cubical.Data.Empty - -open import Cubical.Foundations.Equiv.HalfAdjoint - -- Solves some termination issues private +nInd : ∀ {ℓ} {P : ℕ → Type ℓ} @@ -396,9 +382,9 @@ flipΩIso {A = A} n = pathToIso (cong fst (flipΩPath n)) flipΩIsopres· : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) → (f g : fst ((Ω^ (suc n)) (Ω A))) - → Iso.inv (flipΩIso (suc n)) (f ∙ g) - ≡ (Iso.inv (flipΩIso (suc n)) f) - ∙ (Iso.inv (flipΩIso (suc n)) g) + → inv (flipΩIso (suc n)) (f ∙ g) + ≡ (inv (flipΩIso (suc n)) f) + ∙ (inv (flipΩIso (suc n)) g) flipΩIsopres· {A = A} n f g i = transp (λ j → flipΩPath {A = A} n (~ i ∧ ~ j) .snd ≡ flipΩPath n (~ i ∧ ~ j) .snd) i @@ -410,7 +396,7 @@ flipΩIsopres· {A = A} n f g i = rMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → typ ((Ω^ (suc n)) (Susp∙ (typ A))) → (S₊∙ n →∙ Ω (Susp∙ (typ A))) -rMap n = lMap n ∘ Iso.fun (flipΩIso n) +rMap n = lMap n ∘ fun (flipΩIso n) private rMap1 : ∀ {ℓ} {A : Pointed ℓ} diff --git a/Cubical/Homotopy/HSpace.agda b/Cubical/Homotopy/HSpace.agda new file mode 100644 index 000000000..bcd2db4a4 --- /dev/null +++ b/Cubical/Homotopy/HSpace.agda @@ -0,0 +1,52 @@ +{-# OPTIONS --safe #-} +module Cubical.Homotopy.HSpace where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.HLevels +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn + +record HSpace {ℓ : Level} (A : Pointed ℓ) : Type ℓ where + constructor HSp + field + μ : typ A → typ A → typ A + μₗ : (x : typ A) → μ (pt A) x ≡ x + μᵣ : (x : typ A) → μ x (pt A) ≡ x + μₗᵣ : μₗ (pt A) ≡ μᵣ (pt A) + +record AssocHSpace {ℓ : Level} {A : Pointed ℓ} (e : HSpace A) : Type ℓ where + constructor AssocHSp + field + μ-assoc : (x y z : _) + → HSpace.μ e (HSpace.μ e x y) z ≡ HSpace.μ e x (HSpace.μ e y z) + + μ-assoc-filler : (y z : typ A) + → PathP (λ i → HSpace.μ e (HSpace.μₗ e y i) z + ≡ HSpace.μₗ e (HSpace.μ e y z) i) + (μ-assoc (pt A) y z) + refl + + +-- Instances +open HSpace +open AssocHSpace + +-- S¹ +S1-HSpace : HSpace (S₊∙ 1) +μ S1-HSpace = _·_ +μₗ S1-HSpace x = refl +μᵣ S1-HSpace base = refl +μᵣ S1-HSpace (loop i) = refl +μₗᵣ S1-HSpace x = refl + +S1-AssocHSpace : AssocHSpace S1-HSpace +μ-assoc S1-AssocHSpace base _ _ = refl +μ-assoc S1-AssocHSpace (loop i) x y j = h x y j i + where + h : (x y : S₊ 1) → cong (_· y) (rotLoop x) ≡ rotLoop (x · y) + h = wedgeconFun _ _ (λ _ _ → isOfHLevelPath 2 (isGroupoidS¹ _ _) _ _) + (λ x → refl) + (λ { base → refl ; (loop i) → refl}) + refl +μ-assoc-filler S1-AssocHSpace _ _ = refl diff --git a/Cubical/Homotopy/HomotopyGroup.agda b/Cubical/Homotopy/HomotopyGroup.agda deleted file mode 100644 index f9f0fc541..000000000 --- a/Cubical/Homotopy/HomotopyGroup.agda +++ /dev/null @@ -1,724 +0,0 @@ -{-# OPTIONS --safe --experimental-lossy-unification #-} -{- -This file contains -1. The definition of πₙ as a truncated loop space -2. The definition of πₙ as a truncated function space (Sⁿ →∙ A) -3. A structure preserving equivalence Ωⁿ A ≃ (Sⁿ →∙ A) -4. A proof that the two constructions of homotopy groups are isomorphic --} -module Cubical.Homotopy.HomotopyGroup where - -open import Cubical.Homotopy.Loopspace - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Function -open import Cubical.Foundations.Path -open import Cubical.Foundations.Transport - -open import Cubical.HITs.SetTruncation - renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim - ; elim2 to sElim2 ; elim3 to sElim3 ; map to sMap) -open import Cubical.HITs.Sn -open import Cubical.Data.Bool -open import Cubical.HITs.Susp -open import Cubical.HITs.S1 -open import Cubical.Data.Sigma -open import Cubical.Data.Nat - -open import Cubical.Algebra.Group -open import Cubical.Algebra.Semigroup -open import Cubical.Algebra.Monoid - -open Iso -open IsGroup -open IsSemigroup -open IsMonoid -open GroupStr - - -{- homotopy Group -} -π : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Type ℓ -π n A = ∥ typ ((Ω^ n) A) ∥₂ - -π' : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Type ℓ -π' n A = ∥ S₊∙ n →∙ A ∥₂ - - -{- π as a group -} -1π : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π n A -1π zero {A = A} = ∣ pt A ∣₂ -1π (suc n) = ∣ refl ∣₂ - -·π : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π (suc n) A → π (suc n) A → π (suc n) A -·π n = sRec2 squash₂ λ p q → ∣ p ∙ q ∣₂ - --π : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π (suc n) A → π (suc n) A --π n = sMap sym - -π-rUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π (suc n) A) - → (·π n x (1π (suc n))) ≡ x -π-rUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ rUnit p (~ i) ∣₂ - -π-lUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π (suc n) A) - → (·π n (1π (suc n)) x) ≡ x -π-lUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ lUnit p (~ i) ∣₂ - -π-rCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π (suc n) A) - → (·π n x (-π n x)) ≡ 1π (suc n) -π-rCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ rCancel p i ∣₂ - -π-lCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π (suc n) A) - → (·π n (-π n x) x) ≡ 1π (suc n) -π-lCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ lCancel p i ∣₂ - -π-assoc : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y z : π (suc n) A) - → ·π n x (·π n y z) ≡ ·π n (·π n x y) z -π-assoc n = sElim3 (λ _ _ _ → isSetPathImplicit) λ p q r i → ∣ ∙assoc p q r i ∣₂ - -π-comm : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y : π (suc (suc n)) A) - → ·π (suc n) x y ≡ ·π (suc n) y x -π-comm n = sElim2 (λ _ _ → isSetPathImplicit) λ p q i → ∣ EH n p q i ∣₂ - -πGr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Group ℓ -fst (πGr n A) = π (suc n) A -1g (snd (πGr n A)) = 1π (suc n) -GroupStr._·_ (snd (πGr n A)) = ·π n -inv (snd (πGr n A)) = -π n -is-set (isSemigroup (isMonoid (isGroup (snd (πGr n A))))) = squash₂ -assoc (isSemigroup (isMonoid (isGroup (snd (πGr n A))))) = π-assoc n -identity (isMonoid (isGroup (snd (πGr n A)))) x = (π-rUnit n x) , (π-lUnit n x) -inverse (isGroup (snd (πGr n A))) x = (π-rCancel n x) , (π-lCancel n x) - --- We now define the group structure on π'. --- We first define the corresponding structure on the untruncated --- (S₊∙ n →∙ A). We first give multiplication and inversion and use this --- to give a structure preserving equivalence (S₊∙ n →∙ A) ≃ Ωⁿ A. - -∙Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (S₊∙ n →∙ A) - → (S₊∙ n →∙ A) - → (S₊∙ n →∙ A) -∙Π {A = A} {n = zero} p q = (λ _ → pt A) , refl -fst (∙Π {A = A} {n = suc zero} (f , p) (g , q)) base = pt A -fst (∙Π {A = A} {n = suc zero} (f , p) (g , q)) (loop j) = - ((sym p ∙∙ cong f loop ∙∙ p) ∙ (sym q ∙∙ cong g loop ∙∙ q)) j -snd (∙Π {A = A} {n = suc zero} (f , p) (g , q)) = refl -fst (∙Π {A = A} {n = suc (suc n)} (f , p) (g , q)) north = pt A -fst (∙Π {A = A} {n = suc (suc n)} (f , p) (g , q)) south = pt A -fst (∙Π {A = A} {n = suc (suc n)} (f , p) (g , q)) (merid a j) = - ((sym p ∙∙ cong f (merid a ∙ sym (merid (ptSn (suc n)))) ∙∙ p) - ∙ (sym q ∙∙ cong g (merid a ∙ sym (merid (ptSn (suc n)))) ∙∙ q)) j -snd (∙Π {A = A} {n = suc (suc n)} (f , p) (g , q)) = refl - --Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (S₊∙ n →∙ A) - → (S₊∙ n →∙ A) --Π {n = zero} f = f -fst (-Π {A = A} {n = suc zero} f) base = fst f base -fst (-Π {A = A} {n = suc zero} f) (loop j) = fst f (loop (~ j)) -snd (-Π {A = A} {n = suc zero} f) = snd f -fst (-Π {A = A} {n = suc (suc n)} f) north = fst f north -fst (-Π {A = A} {n = suc (suc n)} f) south = fst f north -fst (-Π {A = A} {n = suc (suc n)} f) (merid a j) = - fst f ((merid a ∙ sym (merid (ptSn _))) (~ j)) -snd (-Π {A = A} {n = suc (suc n)} f) = snd f - -module _ {ℓ : Level} {A : Pointed ℓ} where -flipΩPath : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → ((Ω^ (suc n)) A) ≡ (Ω^ n) (Ω A) -flipΩPath {A = A} zero = refl -flipΩPath {A = A} (suc n) = cong Ω (flipΩPath {A = A} n) - -flipΩIso : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → Iso (fst ((Ω^ (suc n)) A)) (fst ((Ω^ n) (Ω A))) -flipΩIso {A = A} n = pathToIso (cong fst (flipΩPath n)) - -flipΩIsopres· : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → (f g : fst ((Ω^ (suc n)) (Ω A))) - → Iso.inv (flipΩIso (suc n)) (f ∙ g) - ≡ (Iso.inv (flipΩIso (suc n)) f) - ∙ (Iso.inv (flipΩIso (suc n)) g) -flipΩIsopres· {A = A} n f g i = - transp (λ j → flipΩPath {A = A} n (~ i ∧ ~ j) .snd - ≡ flipΩPath n (~ i ∧ ~ j) .snd) i - (transp (λ j → flipΩPath {A = A} n (~ i ∨ ~ j) .snd - ≡ flipΩPath n (~ i ∨ ~ j) .snd) (~ i) f - ∙ transp (λ j → flipΩPath {A = A} n (~ i ∨ ~ j) .snd - ≡ flipΩPath n (~ i ∨ ~ j) .snd) (~ i) g) - -module _ {ℓ : Level} {A : Pointed ℓ} where - suspension→ : (n : ℕ) → (S₊∙ (suc n) →∙ Ω A) → (S₊∙ (suc (suc n)) →∙ A) - fst (suspension→ n f) north = pt A - fst (suspension→ n f) south = pt A - fst (suspension→ n f) (merid a i₁) = fst f a i₁ - snd (suspension→ n f) = refl - - suspension-pres∙ : - (n : ℕ) (f g : (S₊∙ (suc n) →∙ Ω A)) - → suspension→ n (∙Π f g) ≡ ∙Π (suspension→ n f) (suspension→ n g) - suspension-pres∙ n f g = ΣPathP ((funExt main) , refl) - where - J2 : ∀ {ℓ ℓ'} {A : Type ℓ} {x : A} - → (P : (p q : x ≡ x) → (refl ≡ p) → (refl ≡ q) → Type ℓ') - → P refl refl refl refl - → (p q : x ≡ x) (P' : _) (Q : _) → P p q P' Q - J2 P b p q = - J (λ p P' → (Q₁ : refl ≡ q) → P p q P' Q₁) - (J (λ q Q → P refl q refl Q) b) - - J4 : ∀ {ℓ ℓ'} {A : Type ℓ} {x : A} - → (P : (p q r s : x ≡ x) - → (refl ≡ p) → (p ≡ q) → (refl ≡ r) → (r ≡ s) → Type ℓ') - → P refl refl refl refl refl refl refl refl - → (p q r s : x ≡ x) (P' : _) (Q : _) (R : _) (S : _) → P p q r s P' Q R S - J4 P b p q r s = - J (λ p P' → (Q₁ : p ≡ q) (R : refl ≡ r) (S₁ : r ≡ s) - → P p q r s P' Q₁ R S₁) - (J (λ q Q₁ → (R : refl ≡ r) (S₁ : r ≡ s) - → P refl q r s refl Q₁ R S₁) - (J (λ r R → (S₁ : r ≡ s) - → P refl refl r s refl refl R S₁) - (J (λ s S₁ → P refl refl refl s refl refl refl S₁) - b))) - - looper-help : (n : ℕ) (f g : (S₊∙ (suc n) →∙ Ω A)) (a : _) - → fst (∙Π f g) a ≡ fst f a ∙ fst g a - looper-help zero f g base = - rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g)) - looper-help zero f g (loop i) k = - lazy-fill _ _ - (sym (snd f)) (sym (snd g)) - (cong (fst f) loop) (cong (fst g) loop) k i - where - lazy-fill : - ∀ {ℓ} {A : Type ℓ} {x : A} (p q : x ≡ x) - (prefl : refl ≡ p) (qrefl : refl ≡ q) (fl : p ≡ p) (fr : q ≡ q) - → PathP (λ i → (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i - ≡ (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i) - ((prefl ∙∙ fl ∙∙ sym prefl) ∙ (qrefl ∙∙ fr ∙∙ sym qrefl)) - (cong₂ _∙_ fl fr) - lazy-fill {x = x} = - J2 (λ p q prefl qrefl → (fl : p ≡ p) (fr : q ≡ q) - → PathP (λ i → (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i - ≡ (rUnit (refl {x = x}) ∙ cong₂ _∙_ prefl qrefl) i) - ((prefl ∙∙ fl ∙∙ sym prefl) ∙ (qrefl ∙∙ fr ∙∙ sym qrefl)) - (cong₂ _∙_ fl fr)) - λ fl fr → cong₂ _∙_ (sym (rUnit fl)) (sym (rUnit fr)) - ◁ flipSquare (sym (rUnit (rUnit refl)) - ◁ (flipSquare ((λ j → cong (λ x → rUnit x j) fl - ∙ cong (λ x → lUnit x j) fr) - ▷ sym (cong₂Funct _∙_ fl fr)) - ▷ (rUnit (rUnit refl)))) - looper-help (suc n) f g north = - rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g)) - looper-help (suc n) f g south = - (rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) - ∙ cong₂ _∙_ (cong (fst f) (merid (ptSn _))) - (cong (fst g) (merid (ptSn _))) - looper-help (suc n) f g (merid a j) i = help i j - where - help : - PathP (λ i → (rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) i - ≡ ((rUnit refl ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g))) - ∙ cong₂ _∙_ (cong (fst f) (merid (ptSn _))) - (cong (fst g) (merid (ptSn _)))) i) - (cong (fst (∙Π f g)) (merid a)) - (cong₂ _∙_ (cong (fst f) (merid a)) - (cong (fst g) (merid a))) - help = (cong₂ _∙_ (cong (sym (snd f) ∙∙_∙∙ snd f) - (cong-∙ (fst f) - (merid a) (sym (merid (ptSn _))))) - ((cong (sym (snd g) ∙∙_∙∙ snd g) - (cong-∙ (fst g) - (merid a) (sym (merid (ptSn _)))))) - ◁ lazy-help _ _ _ _ (sym (snd f)) (cong (fst f) (merid (ptSn _))) - (sym (snd g)) (cong (fst g) (merid (ptSn _))) - (cong (fst f) (merid a)) (cong (fst g) (merid a))) - where - lazy-help : - ∀ {ℓ} {A : Type ℓ} {x : A} - (p-north p-south q-north q-south : x ≡ x) - (prefl : refl ≡ p-north) (p-merid : p-north ≡ p-south) - (qrefl : refl ≡ q-north) (q-merid : q-north ≡ q-south) - (fl : p-north ≡ p-south) (fr : q-north ≡ q-south) - → PathP (λ i → (rUnit refl ∙ cong₂ _∙_ prefl qrefl) i - ≡ ((rUnit refl ∙ cong₂ _∙_ prefl qrefl) - ∙ cong₂ _∙_ p-merid q-merid) i) - ((prefl ∙∙ fl ∙ sym p-merid ∙∙ sym prefl) - ∙ (qrefl ∙∙ fr ∙ sym q-merid ∙∙ sym qrefl)) - (cong₂ _∙_ fl fr) - lazy-help {x = x} = - J4 (λ p-north p-south q-north q-south prefl p-merid qrefl q-merid - → (fl : p-north ≡ p-south) (fr : q-north ≡ q-south) → - PathP (λ i₁ → (rUnit refl ∙ cong₂ _∙_ prefl qrefl) i₁ ≡ - ((rUnit refl ∙ cong₂ _∙_ prefl qrefl) - ∙ cong₂ _∙_ p-merid q-merid) i₁) - ((prefl ∙∙ fl ∙ sym p-merid ∙∙ sym prefl) ∙ - (qrefl ∙∙ fr ∙ sym q-merid ∙∙ sym qrefl)) - (cong₂ _∙_ fl fr)) - λ fl fr → (cong₂ _∙_ (sym (rUnit (fl ∙ refl)) ∙ sym (rUnit fl)) - (sym (rUnit (fr ∙ refl)) ∙ sym (rUnit fr)) - ◁ flipSquare ((sym (rUnit _) - ◁ flipSquare - ((λ j → cong (λ x → rUnit x j) fl - ∙ cong (λ x → lUnit x j) fr) - ▷ sym (cong₂Funct _∙_ fl fr))) - ▷ (rUnit _ ∙ rUnit _))) - main : (x : fst (S₊∙ (suc (suc n)))) → - fst (suspension→ n (∙Π f g)) x ≡ - fst (∙Π (suspension→ n f) (suspension→ n g)) x - main north = refl - main south = refl - main (merid a i₁) j = help j i₁ - where - help : fst (∙Π f g) a - ≡ cong (fst (∙Π (suspension→ n f) (suspension→ n g))) (merid a) - help = looper-help n f g a - ∙ cong₂ _∙_ (sym (cong-∙ (fst (suspension→ n f)) - (merid a) - (sym (merid (ptSn _))) - ∙∙ cong (fst f a ∙_) (cong sym (snd f)) - ∙∙ sym (rUnit _))) - (sym (cong-∙ (fst (suspension→ n g)) - (merid a) - (sym (merid (ptSn _))) - ∙∙ cong (fst g a ∙_) (cong sym (snd g)) - ∙∙ sym (rUnit _))) - ∙ λ i → rUnit (cong (fst (suspension→ n f)) - (merid a ∙ sym (merid (ptSn _)))) i - ∙ rUnit (cong (fst (suspension→ n g)) - (merid a ∙ sym (merid (ptSn _)))) i - - suspension←filler : (n : ℕ) → (S₊∙ (suc (suc n)) →∙ A) - → I → I → I → fst A - suspension←filler n f i j k = - hfill (λ k → λ { (i = i0) → doubleCompPath-filler - (sym (snd f)) - (cong (fst f) (merid (ptSn _) - ∙ sym (merid (ptSn _)))) - (snd f) k j - ; (i = i1) → snd f k - ; (j = i0) → snd f k - ; (j = i1) → snd f k}) - (inS ((fst f) (rCancel' (merid (ptSn _)) i j))) - k - - suspension← : (n : ℕ) → (S₊∙ (suc (suc n)) →∙ A) → (S₊∙ (suc n) →∙ Ω A) - fst (suspension← n f) a = - sym (snd f) ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f - snd (suspension← n f) i j = suspension←filler n f i j i1 - - suspensionIso : (n : ℕ) → Iso (S₊∙ (suc n) →∙ Ω A) (S₊∙ (suc (suc n)) →∙ A) - Iso.fun (suspensionIso n) f = suspension→ n f - Iso.inv (suspensionIso n) f = suspension← n f - Iso.rightInv (suspensionIso n) (f , p) = - ΣPathP (funExt help , λ i j → p (~ i ∨ j)) - where - help : (x : _) → fst (suspension→ n (suspension← n (f , p))) x ≡ f x - help north = sym p - help south = sym p ∙ cong f (merid (ptSn _)) - help (merid a j) i = - hcomp (λ k → λ { (i = i1) → f (merid a j) - ; (j = i0) → p (~ i ∧ k) - ; (j = i1) → compPath-filler' (sym p) - (cong f (merid (ptSn _))) k i}) - (f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) - Iso.leftInv (suspensionIso n) f = - ΣPathP (funExt lem , PathP-filler) - where - - side₁ : (x : S₊ (suc n)) → I → I → I → fst A - side₁ x i j k = - hfill (λ k → λ { (i = i0) → suspension→ n f .fst - (compPath-filler' (merid x) - (sym (merid (ptSn _))) k j) - ; (i = i1) → snd A - ; (j = i0) → fst f x (i ∨ ~ k) - ; (j = i1) → snd A}) - (inS (snd f i (~ j))) - k - - side₂ : (x : S₊ (suc n)) → I → I → I → fst A - side₂ x i j k = - hfill (λ k → λ { (i = i0) → doubleCompPath-filler - refl - (cong (suspension→ n f .fst) - (merid x ∙ sym (merid (ptSn _)))) - refl k j - ; (i = i1) → fst f x (j ∨ ~ k) - ; (j = i0) → fst f x (i ∧ ~ k) - ; (j = i1) → snd A}) - (inS (side₁ x i j i1)) - k - - compPath-filler'' : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) - → I → I → I → A - compPath-filler'' {z = z} p q k j i = - hfill (λ k → λ { (i = i0) → p (~ j) - ; (i = i1) → q k - ; (j = i0) → q (i ∧ k) }) - (inS (p (i ∨ ~ j))) - k - - lem : (x : _) → fst (suspension← n (suspension→ n f)) x ≡ fst f x - lem x i j = side₂ x i j i1 - - sideCube : Cube (λ j k → snd f j (~ k)) - (λ j k → fst (suspension→ n f) - (rCancel' (merid (ptSn _)) j k)) - (λ r k → suspension→ n f .fst - (compPath-filler' (merid (ptSn _)) - (sym (merid (ptSn _))) r k)) - (λ _ _ → pt A) - (λ r j → snd f j (~ r)) λ _ _ → pt A - sideCube r j k = - hcomp (λ i → λ {(r = i0) → snd f j (~ k ∨ ~ i) - ; (r = i1) → fst (suspension→ n f) - (rCancel-filler' (merid (ptSn _)) (~ i) j k) - ; (j = i0) → suspension→ n f .fst - (compPath-filler'' (merid (ptSn _)) - (sym (merid (ptSn _))) i r k) - ; (j = i1) → snd f (~ r) (~ i ∧ k) - ; (k = i0) → snd f j (~ r) - ; (k = i1) → snd f (~ r ∧ j) (~ i)}) - (hcomp (λ i → λ {(r = i0) → pt A - ; (r = i1) → snd f (~ i) k - ; (j = i0) → snd f (~ i) (k ∨ ~ r) - ; (j = i1) → snd f (~ r ∨ ~ i) k - ; (k = i0) → snd f (j ∨ ~ i) (~ r) - ; (k = i1) → pt A}) - (pt A)) - - PathP-filler : PathP _ _ _ - PathP-filler i j k = - hcomp (λ r → λ {(i = i0) → suspension←filler n (suspension→ n f) j k r - ; (i = i1) → snd f j (k ∨ ~ r) - ; (j = i0) → side₂ (ptSn _) i k r - ; (j = i1) → snd A - ; (k = i0) → snd f j (i ∧ ~ r) - ; (k = i1) → snd A}) - (hcomp ((λ r → λ {(i = i0) → sideCube r j k - ; (i = i1) → pt A - ; (j = i0) → side₁ (ptSn _) i k r - ; (j = i1) → pt A - ; (k = i0) → snd f j (i ∨ ~ r) - ; (k = i1) → pt A})) - (snd f (i ∨ j) (~ k))) - - suspension←-pres∙ : (n : ℕ) → (f g : S₊∙ (suc (suc n)) →∙ A) - → suspension← n (∙Π f g) - ≡ ∙Π (suspension← n f) (suspension← n g) - suspension←-pres∙ n f g = - sym (Iso.leftInv (suspensionIso n) (suspension← n (∙Π f g))) - ∙∙ cong (suspension← n) - (Iso.rightInv (suspensionIso n) (∙Π f g) - ∙∙ cong₂ ∙Π (sym (Iso.rightInv (suspensionIso n) f)) - (sym (Iso.rightInv (suspensionIso n) g)) - ∙∙ sym (suspension-pres∙ _ (suspension← n f) (suspension← n g))) - ∙∙ (Iso.leftInv (suspensionIso n) (∙Π (suspension← n f) (suspension← n g))) - -SphereMap→Ω : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → (S₊∙ (suc n) →∙ A) → (fst ((Ω^ (suc n)) A)) -SphereMap→Ω zero (f , p) = sym p ∙∙ cong f loop ∙∙ p -SphereMap→Ω {A = A} (suc n) (f , p) = - Iso.inv (flipΩIso _) - (SphereMap→Ω {A = Ω A} n (Iso.inv (suspensionIso _) (f , p))) - -SphereMap→Ωpres∙ : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → (p q : S₊∙ (suc n) →∙ A) - → SphereMap→Ω n (∙Π p q) ≡ SphereMap→Ω n p ∙ SphereMap→Ω n q -SphereMap→Ωpres∙ zero f g = sym (rUnit _) -SphereMap→Ωpres∙ (suc n) f g = - cong (Iso.inv (flipΩIso (suc n))) - (cong (SphereMap→Ω n) (suspension←-pres∙ n f g) - ∙ SphereMap→Ωpres∙ n (suspension← n f) (suspension← n g)) - ∙ flipΩIsopres· n _ _ - -Ω→SphereMap : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → (fst ((Ω^ (suc n)) A)) → (S₊∙ (suc n) →∙ A) -fst (Ω→SphereMap {A = A} zero f) base = pt A -fst (Ω→SphereMap {A = A} zero f) (loop i₁) = f i₁ -snd (Ω→SphereMap {A = A} zero f) = refl -Ω→SphereMap {A = A} (suc n) f = - Iso.fun (suspensionIso _) (Ω→SphereMap n (Iso.fun (flipΩIso _) f)) - -Ω→SphereMap→Ω : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → (x : fst ((Ω^ (suc n)) A)) - → SphereMap→Ω n (Ω→SphereMap n x) ≡ x -Ω→SphereMap→Ω zero x = sym (rUnit _) -Ω→SphereMap→Ω (suc n) x = - cong (Iso.inv (flipΩIso (suc n)) ∘ SphereMap→Ω n) - (Iso.leftInv (suspensionIso _) - (Ω→SphereMap n (Iso.fun (flipΩIso (suc n)) x))) - ∙∙ cong (Iso.inv (flipΩIso (suc n))) - (Ω→SphereMap→Ω n (Iso.fun (flipΩIso (suc n)) x)) - ∙∙ Iso.leftInv (flipΩIso (suc n)) x - -SphereMap→Ω→SphereMap : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → (x : (S₊∙ (suc n) →∙ A)) - → Ω→SphereMap n (SphereMap→Ω n x) ≡ x -SphereMap→Ω→SphereMap zero (f , p) = - ΣPathP (funExt fstp , λ i j → p (~ i ∨ j)) - where - fstp : (x : _) → _ ≡ f x - fstp base = sym p - fstp (loop i) j = doubleCompPath-filler (sym p) (cong f loop) p (~ j) i -SphereMap→Ω→SphereMap (suc n) f = - cong (Iso.fun (suspensionIso n) ∘ Ω→SphereMap n) - (Iso.rightInv (flipΩIso (suc n)) - (SphereMap→Ω n (Iso.inv (suspensionIso n) f))) - ∙∙ cong (Iso.fun (suspensionIso n)) - (SphereMap→Ω→SphereMap n (Iso.inv (suspensionIso n) f)) - ∙∙ Iso.rightInv (suspensionIso _) _ - - --- We finally arrive at the main result -IsoSphereMapΩ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → Iso (S₊∙ n →∙ A) (fst ((Ω^ n) A)) -Iso.fun (IsoSphereMapΩ {A = A} zero) (f , p) = f false -fst (Iso.inv (IsoSphereMapΩ {A = A} zero) a) false = a -fst (Iso.inv (IsoSphereMapΩ {A = A} zero) a) true = pt A -snd (Iso.inv (IsoSphereMapΩ {A = A} zero) a) = refl -Iso.rightInv (IsoSphereMapΩ {A = A} zero) a = refl -Iso.leftInv (IsoSphereMapΩ {A = A} zero) (f , p) = - ΣPathP ((funExt (λ { false → refl ; true → sym p})) - , λ i j → p (~ i ∨ j)) -Iso.fun (IsoSphereMapΩ {A = A} (suc n)) = SphereMap→Ω n -Iso.inv (IsoSphereMapΩ {A = A} (suc n)) = Ω→SphereMap n -Iso.rightInv (IsoSphereMapΩ {A = A} (suc n)) = Ω→SphereMap→Ω n -Iso.leftInv (IsoSphereMapΩ {A = A} (suc n)) = SphereMap→Ω→SphereMap n - --- The iso is structure preserving -IsoSphereMapΩ-pres∙Π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f g : S₊∙ (suc n) →∙ A) - → SphereMap→Ω n (∙Π f g) ≡ SphereMap→Ω n f ∙ SphereMap→Ω n g -IsoSphereMapΩ-pres∙Π zero f g = sym (rUnit _) -IsoSphereMapΩ-pres∙Π (suc n) f g = - cong (Iso.inv (flipΩIso (suc n))) - (cong (SphereMap→Ω n) (suspension←-pres∙ n f g) - ∙ SphereMap→Ωpres∙ n (suspension← n f) (suspension← n g)) - ∙ (flipΩIsopres· n (SphereMap→Ω n (suspension← n f)) - (SphereMap→Ω n (suspension← n g))) - --- It is useful to define the ``Group Structure'' on (S₊∙ n →∙ A) --- before doing it on π'. These will be the equivalents of the --- usual groupoid laws on Ω A. -1Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → (S₊∙ n →∙ A) -fst (1Π {A = A}) _ = pt A -snd (1Π {A = A}) = refl - -∙Π-rUnit : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f : S₊∙ (suc n) →∙ A) - → ∙Π f 1Π ≡ f -fst (∙Π-rUnit {A = A} {n = zero} f i) base = snd f (~ i) -fst (∙Π-rUnit {A = A} {n = zero} f i) (loop j) = help i j - where - help : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) - (((sym (snd f)) ∙∙ (cong (fst f) loop) ∙∙ snd f) - ∙ (refl ∙ refl)) - (cong (fst f) loop) - help = (cong ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) ∙_) - (sym (rUnit refl)) ∙ sym (rUnit _)) - ◁ λ i j → doubleCompPath-filler (sym (snd f)) - (cong (fst f) loop) (snd f) (~ i) j -snd (∙Π-rUnit {A = A} {n = zero} f i) j = snd f (~ i ∨ j) -fst (∙Π-rUnit {A = A} {n = suc n} f i) north = snd f (~ i) -fst (∙Π-rUnit {A = A} {n = suc n} f i) south = - (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i -fst (∙Π-rUnit {A = A} {n = suc n} f i) (merid a j) = help i j - where - help : PathP (λ i → snd f (~ i) - ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) - (((sym (snd f)) - ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) - ∙∙ snd f) - ∙ (refl ∙ refl)) - (cong (fst f) (merid a)) - help = (cong (((sym (snd f)) - ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) - ∙∙ snd f) ∙_) - (sym (rUnit refl)) - ∙ sym (rUnit _)) - ◁ λ i j → hcomp (λ k → - λ { (j = i0) → snd f (~ i ∧ k) - ; (j = i1) → compPath-filler' (sym (snd f)) - (cong (fst f) (merid (ptSn (suc n)))) k i - ; (i = i0) → doubleCompPath-filler (sym (snd f)) - (cong (fst f) - (merid a ∙ sym (merid (ptSn (suc n))))) - (snd f) k j - ; (i = i1) → fst f (merid a j)}) - (fst f (compPath-filler (merid a) - (sym (merid (ptSn _))) (~ i) j)) -snd (∙Π-rUnit {A = A} {n = suc n} f i) j = snd f (~ i ∨ j) - -∙Π-lUnit : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f : S₊∙ (suc n) →∙ A) - → ∙Π 1Π f ≡ f -fst (∙Π-lUnit {n = zero} f i) base = snd f (~ i) -fst (∙Π-lUnit {n = zero} f i) (loop j) = s i j - where - s : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) - ((refl ∙ refl) ∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) - (cong (fst f) loop) - s = (cong (_∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) - (sym (rUnit refl)) ∙ sym (lUnit _)) - ◁ λ i j → doubleCompPath-filler (sym (snd f)) - (cong (fst f) loop) (snd f) (~ i) j -snd (∙Π-lUnit {n = zero} f i) j = snd f (~ i ∨ j) -fst (∙Π-lUnit {n = suc n} f i) north = snd f (~ i) -fst (∙Π-lUnit {n = suc n} f i) south = - (sym (snd f) ∙ cong (fst f) (merid (ptSn _))) i -fst (∙Π-lUnit {n = suc n} f i) (merid a j) = help i j - where - help : PathP (λ i → snd f (~ i) - ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) - ((refl ∙ refl) ∙ ((sym (snd f)) - ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) - ∙∙ snd f)) - (cong (fst f) (merid a)) - help = - (cong (_∙ ((sym (snd f)) - ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) - ∙∙ snd f)) - (sym (rUnit refl)) - ∙ sym (lUnit _)) - ◁ λ i j → hcomp (λ k → - λ { (j = i0) → snd f (~ i ∧ k) - ; (j = i1) → compPath-filler' (sym (snd f)) - (cong (fst f) (merid (ptSn (suc n)))) k i - ; (i = i0) → doubleCompPath-filler (sym (snd f)) - (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) - (snd f) k j - ; (i = i1) → fst f (merid a j)}) - (fst f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) -snd (∙Π-lUnit {n = suc n} f i) j = snd f (~ i ∨ j) - -∙Π-rCancel : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f : S₊∙ (suc n) →∙ A) - → ∙Π f (-Π f) ≡ 1Π -fst (∙Π-rCancel {A = A} {n = zero} f i) base = pt A -fst (∙Π-rCancel {A = A} {n = zero} f i) (loop j) = - rCancel (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) i j -snd (∙Π-rCancel {A = A} {n = zero} f i) = refl -fst (∙Π-rCancel {A = A} {n = suc n} f i) north = pt A -fst (∙Π-rCancel {A = A} {n = suc n} f i) south = pt A -fst (∙Π-rCancel {A = A} {n = suc n} f i) (merid a i₁) = sl i i₁ - where - pl = (sym (snd f) - ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) - ∙∙ snd f) - sl : pl - ∙ ((sym (snd f) - ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) - ∙∙ snd f)) ≡ refl - sl = cong (pl ∙_) (cong (sym (snd f) ∙∙_∙∙ (snd f)) - (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) - ∙∙ cong₂ _∙_ refl - (cong (cong (fst f)) (rCancel (merid (ptSn _)))) - ∙∙ sym (rUnit _))) - ∙ rCancel pl -snd (∙Π-rCancel {A = A} {n = suc n} f i) = refl - -∙Π-lCancel : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f : S₊∙ (suc n) →∙ A) - → ∙Π (-Π f) f ≡ 1Π -fst (∙Π-lCancel {A = A} {n = zero} f i) base = pt A -fst (∙Π-lCancel {A = A} {n = zero} f i) (loop j) = - rCancel (sym (snd f) ∙∙ cong (fst f) (sym loop) ∙∙ snd f) i j -fst (∙Π-lCancel {A = A} {n = suc n} f i) north = pt A -fst (∙Π-lCancel {A = A} {n = suc n} f i) south = pt A -fst (∙Π-lCancel {A = A} {n = suc n} f i) (merid a j) = sl i j - where - pl = (sym (snd f) - ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) - ∙∙ snd f) - - sl : (sym (snd f) - ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) - ∙∙ snd f) ∙ pl - ≡ refl - sl = cong (_∙ pl) (cong (sym (snd f) ∙∙_∙∙ (snd f)) - (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) - ∙∙ cong₂ _∙_ refl (cong (cong (fst f)) (rCancel (merid (ptSn _)))) - ∙∙ sym (rUnit _))) - ∙ lCancel pl -snd (∙Π-lCancel {A = A} {n = zero} f i) = refl -snd (∙Π-lCancel {A = A} {n = suc n} f i) = refl - -∙Π-assoc : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f g h : S₊∙ (suc n) →∙ A) - → ∙Π f (∙Π g h) ≡ ∙Π (∙Π f g) h -∙Π-assoc {n = n} f g h = - sym (Iso.leftInv (IsoSphereMapΩ (suc n)) (∙Π f (∙Π g h))) - ∙∙ cong (Ω→SphereMap n) (SphereMap→Ωpres∙ n f (∙Π g h) - ∙∙ cong (SphereMap→Ω n f ∙_) (SphereMap→Ωpres∙ n g h) - ∙∙ ∙assoc (SphereMap→Ω n f) (SphereMap→Ω n g) (SphereMap→Ω n h) - ∙∙ cong (_∙ SphereMap→Ω n h) (sym (SphereMap→Ωpres∙ n f g)) - ∙∙ sym (SphereMap→Ωpres∙ n (∙Π f g) h)) - ∙∙ Iso.leftInv (IsoSphereMapΩ (suc n)) (∙Π (∙Π f g) h) - -∙Π-comm : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f g : S₊∙ (suc (suc n)) →∙ A) - → ∙Π f g ≡ ∙Π g f -∙Π-comm {A = A} {n = n} f g = - sym (Iso.leftInv (IsoSphereMapΩ (suc (suc n))) (∙Π f g)) - ∙∙ cong (Ω→SphereMap (suc n)) (SphereMap→Ωpres∙ (suc n) f g - ∙∙ EH _ _ _ - ∙∙ sym (SphereMap→Ωpres∙ (suc n) g f)) - ∙∙ Iso.leftInv (IsoSphereMapΩ (suc (suc n))) (∙Π g f) - - -{- π'' as a group -} -1π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' n A -1π' n {A = A} = ∣ 1Π ∣₂ - -·π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' (suc n) A → π' (suc n) A → π' (suc n) A -·π' n = sRec2 squash₂ λ p q → ∣ ∙Π p q ∣₂ - --π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' (suc n) A → π' (suc n) A --π' n = sMap -Π - -π'-rUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) - → (·π' n x (1π' (suc n))) ≡ x -π'-rUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-rUnit p i ∣₂ - -π'-lUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) - → (·π' n (1π' (suc n)) x) ≡ x -π'-lUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-lUnit p i ∣₂ - -π'-rCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) - → (·π' n x (-π' n x)) ≡ 1π' (suc n) -π'-rCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-rCancel p i ∣₂ - -π'-lCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) - → (·π' n (-π' n x) x) ≡ 1π' (suc n) -π'-lCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-lCancel p i ∣₂ - -π'-assoc : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y z : π' (suc n) A) - → ·π' n x (·π' n y z) ≡ ·π' n (·π' n x y) z -π'-assoc n = sElim3 (λ _ _ _ → isSetPathImplicit) λ p q r i → ∣ ∙Π-assoc p q r i ∣₂ - -π'-comm : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y : π' (suc (suc n)) A) - → ·π' (suc n) x y ≡ ·π' (suc n) y x -π'-comm n = sElim2 (λ _ _ → isSetPathImplicit) λ p q i → ∣ ∙Π-comm p q i ∣₂ - -π'Gr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Group ℓ -fst (π'Gr n A) = π' (suc n) A -1g (snd (π'Gr n A)) = 1π' (suc n) -GroupStr._·_ (snd (π'Gr n A)) = ·π' n -inv (snd (π'Gr n A)) = -π' n -is-set (isSemigroup (isMonoid (isGroup (snd (π'Gr n A))))) = squash₂ -assoc (isSemigroup (isMonoid (isGroup (snd (π'Gr n A))))) = π'-assoc n -identity (isMonoid (isGroup (snd (π'Gr n A)))) x = (π'-rUnit n x) , (π'-lUnit n x) -inverse (isGroup (snd (π'Gr n A))) x = (π'-rCancel n x) , (π'-lCancel n x) - --- Finally, the Iso -π'Gr≅πGr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → GroupIso (π'Gr n A) (πGr n A) -fst (π'Gr≅πGr n A) = setTruncIso (IsoSphereMapΩ (suc n)) -snd (π'Gr≅πGr n A) = - makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) - λ p q i → ∣ IsoSphereMapΩ-pres∙Π n p q i ∣₂) diff --git a/Cubical/Homotopy/Hopf.agda b/Cubical/Homotopy/Hopf.agda index 1e30fd1ec..0353f8dff 100644 --- a/Cubical/Homotopy/Hopf.agda +++ b/Cubical/Homotopy/Hopf.agda @@ -1,150 +1,91 @@ {-# OPTIONS --safe #-} module Cubical.Homotopy.Hopf where -open import Cubical.ZCohomology.Base -open import Cubical.ZCohomology.Groups.Connected -open import Cubical.ZCohomology.GroupStructure -open import Cubical.ZCohomology.Properties -open import Cubical.ZCohomology.MayerVietorisUnreduced -open import Cubical.ZCohomology.Groups.Unit -open import Cubical.ZCohomology.Groups.Sn -open import Cubical.ZCohomology.RingStructure.CupProduct -open import Cubical.ZCohomology.RingStructure.RingLaws -open import Cubical.ZCohomology.RingStructure.GradedCommutativity - -open import Cubical.Functions.Embedding - -open import Cubical.Data.Fin +open import Cubical.Homotopy.HSpace +open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Transport open import Cubical.Foundations.Function -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Pointed.Homogeneous - open import Cubical.Foundations.Univalence -open import Cubical.Data.Empty renaming (rec to ⊥-rec) open import Cubical.Data.Sigma -open import Cubical.Data.Int hiding (_+'_) -open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) -open import Cubical.Data.Nat.Order open import Cubical.Data.Unit -open import Cubical.Data.Bool -open import Cubical.Algebra.Group - renaming (ℤ to ℤGroup ; Unit to UnitGroup) hiding (Bool) open import Cubical.HITs.Pushout.Flattening -open import Cubical.Homotopy.Connected open import Cubical.HITs.Pushout open import Cubical.HITs.Sn open import Cubical.HITs.Susp open import Cubical.HITs.S1 renaming (_·_ to _*_) -open import Cubical.HITs.Truncation - renaming (rec to trRec ; elim to trElim ; elim2 to trElim2) -open import Cubical.HITs.SetTruncation - renaming (rec to sRec ; elim to sElim ; elim2 to sElim2 ; map to sMap) open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim) +open import Cubical.HITs.Join -open import Cubical.Algebra.AbGroup - -open import Cubical.Homotopy.Loopspace +open Iso +open HSpace +open AssocHSpace -open import Cubical.HITs.Join +private + retEq≡secEq : ∀ {ℓ} {A B : Type ℓ} (e : A ≃ B) + → (x : _) → secEq e (e .fst x) ≡ cong (e .fst) (retEq e x) + retEq≡secEq {A = A} = + EquivJ (λ B e → (x : _) → secEq e (e .fst x) ≡ cong (e .fst) (retEq e x)) + λ _ → refl -retEq≡secEq : ∀ {ℓ} {A B : Type ℓ} (e : A ≃ B) - → (x : _) → secEq e (e .fst x) ≡ cong (e .fst) (retEq e x) -retEq≡secEq {A = A} = - EquivJ (λ B e → (x : _) → secEq e (e .fst x) ≡ cong (e .fst) (retEq e x) ) - λ _ → refl - -record HSpace {ℓ : Level} (A : Pointed ℓ) : Type ℓ where - constructor HSp - field - μ : typ A → typ A → typ A - μₗ : (x : typ A) → μ (pt A) x ≡ x - μᵣ : (x : typ A) → μ x (pt A) ≡ x - μₗᵣ : μₗ (pt A) ≡ μᵣ (pt A) - -record assocHSpace {ℓ : Level} {A : Pointed ℓ} (e : HSpace A) : Type ℓ where - constructor assocHSp - field - μ-assoc : (x y z : _) → HSpace.μ e (HSpace.μ e x y) z - ≡ HSpace.μ e x (HSpace.μ e y z) - - μ-assoc-filler : (y z : typ A) - → PathP (λ i → HSpace.μ e (HSpace.μₗ e y i) z ≡ HSpace.μₗ e (HSpace.μ e y z) i) - (μ-assoc (pt A) y z) refl - -S1-HSpace : HSpace (S₊∙ 1) -HSpace.μ S1-HSpace = _*_ -HSpace.μₗ S1-HSpace x = refl -HSpace.μᵣ S1-HSpace base = refl -HSpace.μᵣ S1-HSpace (loop i) = refl -HSpace.μₗᵣ S1-HSpace x = refl - -S1-AssocHSpace : assocHSpace S1-HSpace -assocHSpace.μ-assoc S1-AssocHSpace base _ _ = refl -assocHSpace.μ-assoc S1-AssocHSpace (loop i) x y j = h x y j i - where - h : (x y : S₊ 1) → cong (_* y) (rotLoop x) ≡ rotLoop (x * y) - h = wedgeconFun _ _ (λ _ _ → isOfHLevelPath 2 (isGroupoidS¹ _ _) _ _) - (λ x → refl) - (λ { base → refl ; (loop i) → refl}) - refl -assocHSpace.μ-assoc-filler S1-AssocHSpace _ _ = refl - -module hopfBase {ℓ : Level} {A : Pointed ℓ} {e : HSpace A} (e-ass : assocHSpace e) (conA : ((x y : typ A) → ∥ x ≡ y ∥)) where - isEquiv-μ : (x : typ A) → isEquiv (λ z → (HSpace.μ e z x)) +module Hopf {ℓ : Level} {A : Pointed ℓ} {e : HSpace A} + (e-ass : AssocHSpace e) (conA : ((x y : typ A) → ∥ x ≡ y ∥)) where + isEquiv-μ : (x : typ A) → isEquiv (λ z → (μ e z x)) isEquiv-μ x = pRec (isPropIsEquiv _) - (J (λ x _ → isEquiv (λ z → HSpace.μ e z x)) - (subst isEquiv (funExt (λ z → sym (HSpace.μᵣ e z))) (idIsEquiv (typ A)))) + (J (λ x _ → isEquiv (λ z → μ e z x)) + (subst isEquiv (funExt (λ z → sym (μᵣ e z))) (idIsEquiv (typ A)))) (conA (pt A) x) - μ-eq : (x : typ A) → typ A ≃ typ A - μ-eq x = (λ z → HSpace.μ e z x) , (isEquiv-μ x) - - isEquiv-μ' : (x : typ A) → isEquiv (HSpace.μ e x) + isEquiv-μ' : (x : typ A) → isEquiv (μ e x) isEquiv-μ' x = pRec (isPropIsEquiv _) - (J (λ x _ → isEquiv (HSpace.μ e x)) - (subst isEquiv (funExt (λ x → sym (HSpace.μₗ e x))) (idIsEquiv (typ A)))) + (J (λ x _ → isEquiv (μ e x)) + (subst isEquiv (funExt (λ x → sym (μₗ e x))) (idIsEquiv (typ A)))) (conA (pt A) x) + μ-eq : (x : typ A) → typ A ≃ typ A + μ-eq x = (λ z → μ e z x) , (isEquiv-μ x) + μ-eq' : (x : typ A) → typ A ≃ typ A - μ-eq' x = HSpace.μ e x , isEquiv-μ' x + μ-eq' x = μ e x , isEquiv-μ' x Hopf : Susp (typ A) → Type ℓ Hopf north = typ A Hopf south = typ A Hopf (merid a i₁) = ua (μ-eq a) i₁ - TotalSpaceHopf' : Type _ - TotalSpaceHopf' = Pushout {A = typ A × typ A} fst λ x → HSpace.μ e (fst x) (snd x) + TotalSpaceHopfPush : Type _ + TotalSpaceHopfPush = + Pushout {A = typ A × typ A} fst λ x → μ e (fst x) (snd x) - TotalSpaceHopf'→TotalSpace : TotalSpaceHopf' → Σ[ x ∈ Susp (typ A) ] Hopf x - TotalSpaceHopf'→TotalSpace (inl x) = north , x - TotalSpaceHopf'→TotalSpace (inr x) = south , x - TotalSpaceHopf'→TotalSpace (push (x , y) i₁) = merid y i₁ , ua-gluePt (μ-eq y) i₁ x + TotalSpaceHopfPush→TotalSpace : + TotalSpaceHopfPush → Σ[ x ∈ Susp (typ A) ] Hopf x + TotalSpaceHopfPush→TotalSpace (inl x) = north , x + TotalSpaceHopfPush→TotalSpace (inr x) = south , x + TotalSpaceHopfPush→TotalSpace (push (x , y) i₁) = + merid y i₁ , ua-gluePt (μ-eq y) i₁ x - joinIso₁ : Iso (TotalSpaceHopf') (join (typ A) (typ A)) + joinIso₁ : Iso (TotalSpaceHopfPush) (join (typ A) (typ A)) joinIso₁ = iso F G s r where - F : TotalSpaceHopf' → join (typ A) (typ A) + F : TotalSpaceHopfPush → join (typ A) (typ A) F (inl x) = inl x F (inr x) = inr x - F (push (a , x) i) = push a (HSpace.μ e a x) i + F (push (a , x) i) = push a (μ e a x) i - G : join (typ A) (typ A) → TotalSpaceHopf' + G : join (typ A) (typ A) → TotalSpaceHopfPush G (inl x) = inl x G (inr x) = inr x - G (push a b i) = (push (a , invEq (μ-eq' a) b) ∙ cong inr (secEq (μ-eq' a) b)) i + G (push a b i) = + (push (a , invEq (μ-eq' a) b) ∙ cong inr (secEq (μ-eq' a) b)) i s : section F G s (inl x) = refl @@ -152,7 +93,9 @@ module hopfBase {ℓ : Level} {A : Pointed ℓ} {e : HSpace A} (e-ass : assocHSp s (push a b i) j = hcomp (λ k → λ { (i = i0) → inl a ; (i = i1) → inr (secEq (μ-eq' a) b (j ∨ k)) - ; (j = i0) → F (compPath-filler (push (a , invEq (μ-eq' a) b)) (cong inr (secEq (μ-eq' a) b)) k i) + ; (j = i0) → F (compPath-filler + (push (a , invEq (μ-eq' a) b)) + (cong inr (secEq (μ-eq' a) b)) k i) ; (j = i1) → push a b i}) (hcomp (λ k → λ { (i = i0) → inl a ; (i = i1) → inr (secEq (μ-eq' a) b (~ k ∨ j)) @@ -165,235 +108,255 @@ module hopfBase {ℓ : Level} {A : Pointed ℓ} {e : HSpace A} (e-ass : assocHSp r (inr x) = refl r (push (x , y) i) j = hcomp (λ k → λ { (i = i0) → inl x - ; (i = i1) → inr (HSpace.μ e x y) - ; (j = i0) → (push (x , invEq (μ-eq' x) (HSpace.μ e x y)) - ∙ (λ i₁ → inr (retEq≡secEq (μ-eq' x) y (~ k) i₁))) i + ; (i = i1) → inr (μ e x y) + ; (j = i0) → (push (x , invEq (μ-eq' x) (μ e x y)) + ∙ (λ i₁ → inr (retEq≡secEq (μ-eq' x) y (~ k) i₁))) i ; (j = i1) → push (x , y) i}) (hcomp (λ k → λ { (i = i0) → inl x - ; (i = i1) → inr (HSpace.μ e x (retEq (μ-eq' x) y k)) + ; (i = i1) → inr (μ e x (retEq (μ-eq' x) y k)) ; (j = i1) → push (x , retEq (μ-eq' x) y k) i}) - ((push (x , invEq (μ-eq' x) (HSpace.μ e x y))) i)) + ((push (x , invEq (μ-eq' x) (μ e x y))) i)) - isEquivTotalSpaceHopf'→TotalSpace : isEquiv TotalSpaceHopf'→TotalSpace - isEquivTotalSpaceHopf'→TotalSpace = - isoToIsEquiv (iso _ inv inv* retr) + isEquivTotalSpaceHopfPush→TotalSpace : + isEquiv TotalSpaceHopfPush→TotalSpace + isEquivTotalSpaceHopfPush→TotalSpace = + isoToIsEquiv (iso _ inv' sect retr) where - inv : _ → _ - inv (north , y) = inl y - inv (south , y) = inr y - inv (merid a i , y) = + inv' : _ → _ + inv' (north , y) = inl y + inv' (south , y) = inr y + inv' (merid a i , y) = hcomp (λ k → λ { (i = i0) → push (y , a) (~ k) ; (i = i1) → inr y}) (inr (ua-unglue (μ-eq a) i y)) where - - pp : PathP (λ i → ua (μ-eq a) i → TotalSpaceHopf') inl inr - pp = ua→ {e = μ-eq a} {B = λ _ → TotalSpaceHopf'} λ b → push (b , a) - - inv* : (x : _) → TotalSpaceHopf'→TotalSpace (inv x) ≡ x - inv* (north , x) = refl - inv* (south , x) = refl - inv* (merid a i , y) j = - hcomp (λ k → λ { (i = i0) → merid a (~ k ∧ ~ j) , ua-gluePt (μ-eq a) (~ k ∧ ~ j) y + + pp : PathP (λ i → ua (μ-eq a) i → TotalSpaceHopfPush) + inl inr + pp = ua→ {e = μ-eq a} {B = λ _ → TotalSpaceHopfPush} λ b → push (b , a) + + sect : (x : _) → TotalSpaceHopfPush→TotalSpace (inv' x) ≡ x + sect (north , x) = refl + sect (south , x) = refl + sect (merid a i , y) j = + hcomp (λ k → λ { (i = i0) → merid a (~ k ∧ ~ j) + , ua-gluePt (μ-eq a) (~ k ∧ ~ j) y ; (i = i1) → south , y - ; (j = i0) → TotalSpaceHopf'→TotalSpace - (hfill (λ k → λ { (i = i0) → push (y , a) (~ k) - ; (i = i1) → inr y}) - (inS (inr (ua-unglue (μ-eq a) i y))) - k) + ; (j = i0) → + TotalSpaceHopfPush→TotalSpace + (hfill (λ k → λ { (i = i0) → push (y , a) (~ k) + ; (i = i1) → inr y}) + (inS (inr (ua-unglue (μ-eq a) i y))) + k) ; (j = i1) → merid a i , y}) - ((merid a (i ∨ ~ j)) , s (μ-eq a) i j y) + ((merid a (i ∨ ~ j)) , lem (μ-eq a) i j y) where - s : ∀ {ℓ} {A B : Type ℓ} (e : A ≃ B) → + lem : ∀ {ℓ} {A B : Type ℓ} (e : A ≃ B) → PathP (λ i → PathP (λ j → (y : ua e i) → ua e (i ∨ ~ j)) (λ y → ua-unglue e i y) λ y → y) (λ j y → ua-gluePt e (~ j) y) refl - s {A = A} {B = B} = EquivJ (λ B e → PathP (λ i → PathP (λ j → (y : ua e i) → ua e (i ∨ ~ j)) - (λ y → ua-unglue e i y) - λ y → y) - (λ j y → ua-gluePt e (~ j) y) - refl) - h - where - h : _ - h i j a = ua-gluePt (idEquiv B) (i ∨ ~ j) (ua-unglue (idEquiv B) i a) - - - retr : retract TotalSpaceHopf'→TotalSpace inv + lem {A = A} {B = B} = + EquivJ (λ B e → PathP (λ i → PathP (λ j → (y : ua e i) → ua e (i ∨ ~ j)) + (λ y → ua-unglue e i y) + λ y → y) + (λ j y → ua-gluePt e (~ j) y) + refl) + λ i j a → ua-gluePt (idEquiv B) (i ∨ ~ j) (ua-unglue (idEquiv B) i a) + + retr : retract TotalSpaceHopfPush→TotalSpace inv' retr (inl x) = refl retr (inr x) = refl retr (push (x , y) i) j = hcomp (λ k → λ { (i = i0) → push (x , y) (~ k) - ; (i = i1) → inr (HSpace.μ e x y) + ; (i = i1) → inr (μ e x y) ; (j = i1) → push (x , y) (i ∨ ~ k)}) - (inr (HSpace.μ e x y)) + (inr (μ e x y)) IsoTotalSpaceJoin : Iso (Σ[ x ∈ Susp (typ A) ] Hopf x) (join (typ A) (typ A)) IsoTotalSpaceJoin = - compIso (equivToIso (invEquiv (_ , isEquivTotalSpaceHopf'→TotalSpace))) + compIso (equivToIso (invEquiv (_ , isEquivTotalSpaceHopfPush→TotalSpace))) joinIso₁ - induced : TotalSpaceHopf' → Susp (typ A) - induced = fst ∘ TotalSpaceHopf'→TotalSpace + induced : TotalSpaceHopfPush → Susp (typ A) + induced = fst ∘ TotalSpaceHopfPush→TotalSpace - lem : (x y z : typ A) → (i j : I) → ua (μ-eq y) i - lem x y z i j = + ua-lem : (x y z : typ A) → (i j : I) → ua (μ-eq y) i + ua-lem x y z i j = fill (λ k → ua (μ-eq y) i) - (λ j → λ { (i = i0) → HSpace.μ e z x - ; (i = i1) → assocHSpace.μ-assoc e-ass z x y j}) - (inS (ua-gluePt (μ-eq y) i (HSpace.μ e z x))) + (λ j → λ { (i = i0) → μ e z x + ; (i = i1) → μ-assoc e-ass z x y j}) + (inS (ua-gluePt (μ-eq y) i (μ e z x))) j - v : (x : TotalSpaceHopf') → typ A ≃ Hopf (induced x) - v (inl x) = μ-eq x - v (inr x) = μ-eq x - v (push (x , y) i₁) = pp x y i₁ + TotalSpaceHopfPush→≃Hopf : (x : TotalSpaceHopfPush) → typ A ≃ Hopf (induced x) + TotalSpaceHopfPush→≃Hopf (inl x) = μ-eq x + TotalSpaceHopfPush→≃Hopf (inr x) = μ-eq x + TotalSpaceHopfPush→≃Hopf (push (x , y) i₁) = pp x y i₁ where - pp : (x y : _) → PathP (λ i → typ A ≃ ua (μ-eq y) i) (μ-eq x) (μ-eq (HSpace.μ e x y)) + pp : (x y : _) → PathP (λ i → typ A ≃ ua (μ-eq y) i) (μ-eq x) (μ-eq (μ e x y)) pp x y = ΣPathP (P , help) where P : PathP (λ z → typ A → ua (μ-eq y) z) (fst (μ-eq x)) - (fst (μ-eq (HSpace.μ e x y))) - P i z = lem x y z i i1 + (fst (μ-eq (μ e x y))) + P i z = ua-lem x y z i i1 abstract help : PathP (λ i₂ → isEquiv (P i₂)) (snd (μ-eq x)) - (snd (μ-eq (HSpace.μ e x y))) + (snd (μ-eq (μ e x y))) help = toPathP (isPropIsEquiv _ _ _) - v' : (a : typ A) (x : TotalSpaceHopf') → Σ[ x ∈ Susp (typ A) ] Hopf x - v' a x = (induced x) , fst (v x) a + Push→TotalSpaceHopf : (a : typ A) (x : TotalSpaceHopfPush) + → Σ[ x ∈ Susp (typ A) ] Hopf x + Push→TotalSpaceHopf a x = (induced x) , fst (TotalSpaceHopfPush→≃Hopf x) a - v'-equiv : (a : typ A) → isEquiv (v' a) - v'-equiv a = pRec (isPropIsEquiv _) - (J (λ a _ → isEquiv (v' a)) + Push→TotalSpaceHopf-equiv : (a : typ A) → isEquiv (Push→TotalSpaceHopf a) + Push→TotalSpaceHopf-equiv a = pRec (isPropIsEquiv _) + (J (λ a _ → isEquiv (Push→TotalSpaceHopf a)) (subst isEquiv (sym main) - isEquivTotalSpaceHopf'→TotalSpace)) + isEquivTotalSpaceHopfPush→TotalSpace)) (conA (pt A) a) where - help1 : (x : _) → fst ((v' (pt A)) x) ≡ fst (TotalSpaceHopf'→TotalSpace x) - help1 (inl x) = refl - help1 (inr x) = refl - help1 (push a i) = refl - - help2 : (x : _) - → PathP (λ i → Hopf (help1 x i)) - (snd ((v' (pt A)) x)) - (snd (TotalSpaceHopf'→TotalSpace x)) - help2 (inl x) = HSpace.μₗ e x - help2 (inr x) = HSpace.μₗ e x - help2 (push (x , y) i) j = - hcomp (λ k → λ {(i = i0) → HSpace.μₗ e x j - ; (i = i1) → assocHSpace.μ-assoc-filler e-ass x y j k - ; (j = i0) → lem x y (pt A) i k + lem₁ : (x : _) → fst ((Push→TotalSpaceHopf (pt A)) x) + ≡ fst (TotalSpaceHopfPush→TotalSpace x) + lem₁ (inl x) = refl + lem₁ (inr x) = refl + lem₁ (push a i) = refl + + lem₂ : (x : _) + → PathP (λ i → Hopf (lem₁ x i)) + (snd ((Push→TotalSpaceHopf (pt A)) x)) + (snd (TotalSpaceHopfPush→TotalSpace x)) + lem₂ (inl x) = μₗ e x + lem₂ (inr x) = μₗ e x + lem₂ (push (x , y) i) j = + hcomp (λ k → λ {(i = i0) → μₗ e x j + ; (i = i1) → μ-assoc-filler e-ass x y j k + ; (j = i0) → ua-lem x y (pt A) i k ; (j = i1) → ua-gluePt (μ-eq y) i x}) - (ua-gluePt (μ-eq y) i (HSpace.μₗ e x j)) + (ua-gluePt (μ-eq y) i (μₗ e x j)) - main : v' (pt A) ≡ TotalSpaceHopf'→TotalSpace - main i x = (help1 x i) , (help2 x i) + main : Push→TotalSpaceHopf (pt A) ≡ TotalSpaceHopfPush→TotalSpace + main i x = (lem₁ x i) , (lem₂ x i) - megaPush : Type _ - megaPush = Pushout {A = TotalSpaceHopf'} (λ _ → tt) induced + TotalSpaceHopfPush² : Type _ + TotalSpaceHopfPush² = Pushout {A = TotalSpaceHopfPush} (λ _ → tt) induced - P : megaPush → Type _ + P : TotalSpaceHopfPush² → Type _ P (inl x) = typ A P (inr x) = Hopf x - P (push a i) = ua (v a) i + P (push a i) = ua (TotalSpaceHopfPush→≃Hopf a) i - totalSpaceMegaPush : Type _ - totalSpaceMegaPush = Σ[ x ∈ megaPush ] P x + TotalSpacePush² : Type _ + TotalSpacePush² = Σ[ x ∈ TotalSpaceHopfPush² ] P x - totalSpaceMegaPush' : Type _ - totalSpaceMegaPush' = - Pushout {A = typ A × TotalSpaceHopf'} + TotalSpacePush²' : Type _ + TotalSpacePush²' = + Pushout {A = typ A × TotalSpaceHopfPush} {C = Σ[ x ∈ Susp (typ A) ] Hopf x} fst - λ x → v' (fst x) (snd x) + λ x → Push→TotalSpaceHopf (fst x) (snd x) - z : Iso totalSpaceMegaPush totalSpaceMegaPush' - z = compIso whe2 (compIso (equivToIso fl.flatten) whe) + IsoTotalSpacePush²TotalSpacePush²' : Iso TotalSpacePush² TotalSpacePush²' + IsoTotalSpacePush²TotalSpacePush²' = + compIso iso₂ (compIso (equivToIso fl.flatten) iso₁) where - module fl = FlatteningLemma ((λ _ → tt)) induced (λ x → typ A) Hopf v - - whe : Iso (Pushout fl.Σf fl.Σg) totalSpaceMegaPush' - Iso.fun whe (inl x) = inl (snd x) - Iso.fun whe (inr x) = inr x - Iso.fun whe (push a i) = push ((snd a) , (fst a)) i - Iso.inv whe (inl x) = inl (tt , x) - Iso.inv whe (inr x) = inr x - Iso.inv whe (push a i) = push (snd a , fst a) i - Iso.rightInv whe (inl x) = refl - Iso.rightInv whe (inr x) = refl - Iso.rightInv whe (push a i) = refl - Iso.leftInv whe (inl x) = refl - Iso.leftInv whe (inr x) = refl - Iso.leftInv whe (push a i) = refl - - whe2 : Iso totalSpaceMegaPush (Σ (Pushout (λ _ → tt) induced) fl.E) - Iso.fun whe2 (inl x , y) = inl x , y - Iso.fun whe2 (inr x , y) = inr x , y - Iso.fun whe2 (push a i , y) = push a i , y - Iso.inv whe2 (inl x , y) = inl x , y - Iso.inv whe2 (inr x , y) = inr x , y - Iso.inv whe2 (push a i , y) = push a i , y - Iso.rightInv whe2 (inl x , snd₁) = refl - Iso.rightInv whe2 (inr x , snd₁) = refl - Iso.rightInv whe2 (push a i , snd₁) = refl - Iso.leftInv whe2 (inl x , snd₁) = refl - Iso.leftInv whe2 (inr x , snd₁) = refl - Iso.leftInv whe2 (push a i , snd₁) = refl - - F : totalSpaceMegaPush' - → (Pushout {A = typ A × Σ (Susp (typ A)) Hopf} - fst - snd) + module fl = + FlatteningLemma (λ _ → tt) induced (λ x → typ A) + Hopf TotalSpaceHopfPush→≃Hopf + + iso₁ : Iso (Pushout fl.Σf fl.Σg) TotalSpacePush²' + fun iso₁ (inl x) = inl (snd x) + fun iso₁ (inr x) = inr x + fun iso₁ (push a i) = push ((snd a) , (fst a)) i + inv iso₁ (inl x) = inl (tt , x) + inv iso₁ (inr x) = inr x + inv iso₁ (push a i) = push (snd a , fst a) i + rightInv iso₁ (inl x) = refl + rightInv iso₁ (inr x) = refl + rightInv iso₁ (push a i) = refl + leftInv iso₁ (inl x) = refl + leftInv iso₁ (inr x) = refl + leftInv iso₁ (push a i) = refl + + iso₂ : Iso TotalSpacePush² (Σ (Pushout (λ _ → tt) induced) fl.E) + fun iso₂ (inl x , y) = inl x , y + fun iso₂ (inr x , y) = inr x , y + fun iso₂ (push a i , y) = push a i , y + inv iso₂ (inl x , y) = inl x , y + inv iso₂ (inr x , y) = inr x , y + inv iso₂ (push a i , y) = push a i , y + rightInv iso₂ (inl x , snd₁) = refl + rightInv iso₂ (inr x , snd₁) = refl + rightInv iso₂ (push a i , snd₁) = refl + leftInv iso₂ (inl x , snd₁) = refl + leftInv iso₂ (inr x , snd₁) = refl + leftInv iso₂ (push a i , snd₁) = refl + + F : TotalSpacePush²' + → (Pushout {A = typ A × Σ (Susp (typ A)) Hopf} fst snd) F (inl x) = inl x F (inr x) = inr x - F (push (x , y) i) = push (x , v' x y) i + F (push (x , y) i) = push (x , Push→TotalSpaceHopf x y) i - G : (Pushout {A = typ A × Σ (Susp (typ A)) Hopf} - fst - snd) - → totalSpaceMegaPush' + G : (Pushout {A = typ A × Σ (Susp (typ A)) Hopf} fst snd) + → TotalSpacePush²' G (inl x) = inl x G (inr x) = inr x G (push (x , y) i) = hcomp (λ k → λ { (i = i0) → inl x - ; (i = i1) → inr (secEq (_ , v'-equiv x) y k)}) - (push (x , invEq (_ , v'-equiv x) y) i) - - zz : Iso totalSpaceMegaPush' - (Pushout {A = typ A × Σ (Susp (typ A)) Hopf} - fst - snd) - Iso.fun zz = F - Iso.inv zz = G - Iso.rightInv zz (inl x) = refl - Iso.rightInv zz (inr x) = refl - Iso.rightInv zz (push (x , y) i) j = + ; (i = i1) + → inr (secEq (_ , Push→TotalSpaceHopf-equiv x) y k)}) + (push (x , invEq (_ , Push→TotalSpaceHopf-equiv x) y) i) + + zz : Iso TotalSpacePush²' + (Pushout {A = typ A × Σ (Susp (typ A)) Hopf} fst snd) + fun zz = F + inv zz = G + rightInv zz (inl x) = refl + rightInv zz (inr x) = refl + rightInv zz (push (x , y) i) j = hcomp (λ k → λ { (i = i0) → inl x - ; (i = i1) → inr (secEq (_ , v'-equiv x) y k) + ; (i = i1) + → inr (secEq (_ , Push→TotalSpaceHopf-equiv x) y k) ; (j = i0) → F ( - hfill (λ k → λ { (i = i0) → inl x - ; (i = i1) → inr (secEq (_ , v'-equiv x) y k)}) - (inS (push (x , invEq (_ , v'-equiv x) y) i)) k) - ; (j = i1) → push (x , (secEq (_ , v'-equiv x) y k)) i}) - (push (x , (secEq (_ , v'-equiv x) y i0)) i) - Iso.leftInv zz (inl x) = refl - Iso.leftInv zz (inr x) = refl - Iso.leftInv zz (push (x , y) i) j = + hfill (λ k → + λ { (i = i0) → inl x + ; (i = i1) + → inr (secEq (_ + , Push→TotalSpaceHopf-equiv x) y k)}) + (inS (push (x + , invEq (_ + , Push→TotalSpaceHopf-equiv x) y) i)) k) + ; (j = i1) + → push (x + , (secEq (_ + , Push→TotalSpaceHopf-equiv x) y k)) i}) + (push (x , (secEq (_ , Push→TotalSpaceHopf-equiv x) y i0)) i) + leftInv zz (inl x) = refl + leftInv zz (inr x) = refl + leftInv zz (push (x , y) i) j = hcomp (λ k → λ { (i = i0) → inl x - ; (i = i1) → inr (secEq (v' x , v'-equiv x) (v' x y) (j ∨ k)) + ; (i = i1) → inr (secEq (Push→TotalSpaceHopf x + , Push→TotalSpaceHopf-equiv x) + (Push→TotalSpaceHopf x y) (j ∨ k)) ; (j = i1) → push (x , y) i}) (hcomp (λ k → λ { (i = i0) → inl x - ; (i = i1) → inr (retEq≡secEq (v' x , v'-equiv x) y (~ k) j) + ; (i = i1) → inr (retEq≡secEq + (Push→TotalSpaceHopf x + , Push→TotalSpaceHopf-equiv x) + y (~ k) j) ; (j = i1) → push (x , y) i - ; (j = i0) → push (x , invEq (v' x , v'-equiv x) (v' x y)) i}) - (push (x , retEq (v' x , v'-equiv x) y j) i)) - - IsoJoin₁ : Iso totalSpaceMegaPush (join (typ A) (join (typ A) (typ A))) - IsoJoin₁ = - compIso z (compIso zz (compIso (equivToIso (joinPushout≃join _ _)) + ; (j = i0) → push (x , invEq + (Push→TotalSpaceHopf x + , Push→TotalSpaceHopf-equiv x) + (Push→TotalSpaceHopf x y)) i}) + (push (x , retEq (Push→TotalSpaceHopf x + , Push→TotalSpaceHopf-equiv x) y j) i)) + + joinIso₂ : Iso TotalSpacePush² (join (typ A) (join (typ A) (typ A))) + joinIso₂ = + compIso IsoTotalSpacePush²TotalSpacePush²' + (compIso zz (compIso (equivToIso (joinPushout≃join _ _)) (pathToIso (cong (join (typ A)) (isoToPath IsoTotalSpaceJoin))))) diff --git a/Cubical/Homotopy/HopfInvariant/Homomorphism.agda b/Cubical/Homotopy/HopfInvariant/Homomorphism.agda index 869063c8d..08f78287c 100644 --- a/Cubical/Homotopy/HopfInvariant/Homomorphism.agda +++ b/Cubical/Homotopy/HopfInvariant/Homomorphism.agda @@ -5,8 +5,6 @@ open import Cubical.Homotopy.HopfInvariant.Base open import Cubical.Homotopy.Group.Base open import Cubical.Homotopy.Group.Properties -open import Cubical.Relation.Nullary - open import Cubical.ZCohomology.Base open import Cubical.ZCohomology.GroupStructure open import Cubical.ZCohomology.Properties diff --git a/Cubical/Homotopy/HopfInvariant/HopfMap.agda b/Cubical/Homotopy/HopfInvariant/HopfMap.agda index 2e3a6db42..95e727d90 100644 --- a/Cubical/Homotopy/HopfInvariant/HopfMap.agda +++ b/Cubical/Homotopy/HopfInvariant/HopfMap.agda @@ -8,26 +8,20 @@ module Cubical.Homotopy.HopfInvariant.HopfMap where open import Cubical.Homotopy.HopfInvariant.Base open import Cubical.Homotopy.Hopf open import Cubical.Homotopy.Connected - -open import Cubical.Relation.Nullary +open import Cubical.Homotopy.HSpace open import Cubical.ZCohomology.Base open import Cubical.ZCohomology.GroupStructure open import Cubical.ZCohomology.Properties -open import Cubical.ZCohomology.MayerVietorisUnreduced -open import Cubical.ZCohomology.Groups.Unit -open import Cubical.ZCohomology.Groups.Wedge open import Cubical.ZCohomology.Groups.Sn open import Cubical.ZCohomology.RingStructure.CupProduct open import Cubical.ZCohomology.RingStructure.RingLaws -open import Cubical.ZCohomology.RingStructure.GradedCommutativity open import Cubical.ZCohomology.Gysin +open import Cubical.Foundations.Prelude open import Cubical.Foundations.Path open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Transport open import Cubical.Foundations.Function -open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed open import Cubical.Foundations.Pointed.Homogeneous open import Cubical.Foundations.Isomorphism @@ -51,11 +45,11 @@ open import Cubical.HITs.Join open import Cubical.HITs.S1 renaming (_·_ to _*_) open import Cubical.HITs.Sn open import Cubical.HITs.Susp -open import Cubical.HITs.Wedge open import Cubical.HITs.Truncation renaming (rec to trRec ; elim to trElim) open import Cubical.HITs.SetTruncation - renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap) + renaming (rec to sRec ; rec2 to sRec2 + ; elim to sElim ; elim2 to sElim2 ; map to sMap) open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec) @@ -65,13 +59,13 @@ snd HopfMap = refl -- We use the Hopf fibration in order to connect it to the Gysin Sequence module hopfS¹ = - hopfBase S1-AssocHSpace (sphereElim2 _ (λ _ _ → squash) ∣ refl ∣) + Hopf S1-AssocHSpace (sphereElim2 _ (λ _ _ → squash) ∣ refl ∣) S¹Hopf = hopfS¹.Hopf -E* = hopfS¹.totalSpaceMegaPush -IsoE*join = hopfS¹.IsoJoin₁ +E* = hopfS¹.TotalSpacePush²' +IsoE*join = hopfS¹.joinIso₂ IsoTotalHopf' = hopfS¹.joinIso₁ -CP² = hopfS¹.megaPush +CP² = hopfS¹.TotalSpaceHopfPush² fibr = hopfS¹.P TotalHopf' : Type _ @@ -88,18 +82,18 @@ conCP² x y = sRec2 squash₂ (λ p q → ∣ p ∙ sym q ∣₂) (conCP²' x) ( conCP²' (inl x) = ∣ refl ∣₂ conCP²' (inr x) = sphereElim 1 {A = λ x → ∥ inr x ≡ inl tt ∥₂} (λ _ → squash₂) ∣ sym (push (inl base)) ∣₂ x - conCP²' (push a i₁) = main a i₁ + conCP²' (push a i) = main a i where - indLem : ∀ {ℓ} {A : hopfS¹.TotalSpaceHopf' → Type ℓ} + indLem : ∀ {ℓ} {A : hopfS¹.TotalSpaceHopfPush → Type ℓ} → ((a : _) → isProp (A a)) → A (inl base) - → ((a : hopfS¹.TotalSpaceHopf') → A a) + → ((a : hopfS¹.TotalSpaceHopfPush) → A a) indLem {A = A} p b = PushoutToProp p (sphereElim 0 (λ _ → p _) b) (sphereElim 0 (λ _ → p _) (subst A (push (base , base)) b)) - main : (a : hopfS¹.TotalSpaceHopf') + main : (a : hopfS¹.TotalSpaceHopfPush) → PathP (λ i → ∥ Path CP² (push a i) (inl tt) ∥₂) (conCP²' (inl tt)) (conCP²' (inr (hopfS¹.induced a))) main = indLem (λ _ → isOfHLevelPathP' 1 squash₂ _ _) @@ -139,7 +133,7 @@ joinS¹S¹→Groupoid P grp b = → P (transp (λ j → isoToPath (invIso (IsoSphereJoin 1 1)) (i ∨ j)) i x)) (sphereElim _ (λ _ → grp _) b) -TotalHopf→Gpd : ∀ {ℓ} (P : hopfS¹.TotalSpaceHopf' → Type ℓ) +TotalHopf→Gpd : ∀ {ℓ} (P : hopfS¹.TotalSpaceHopfPush → Type ℓ) → ((x : _) → isGroupoid (P x)) → P (inl base) → (x : _) → P x @@ -153,8 +147,8 @@ TotalHopf→Gpd' : ∀ {ℓ} (P : Σ (S₊ 2) hopfS¹.Hopf → Type ℓ) → P (north , base) → (x : _) → P x TotalHopf→Gpd' P grp b = - transport (λ i → (x : ua (_ , hopfS¹.isEquivTotalSpaceHopf'→TotalSpace) i) - → P (transp (λ j → ua (_ , hopfS¹.isEquivTotalSpaceHopf'→TotalSpace) + transport (λ i → (x : ua (_ , hopfS¹.isEquivTotalSpaceHopfPush→TotalSpace) i) + → P (transp (λ j → ua (_ , hopfS¹.isEquivTotalSpaceHopfPush→TotalSpace) (i ∨ j)) i x)) (TotalHopf→Gpd _ (λ _ → grp _) b) @@ -168,7 +162,7 @@ CP²→Groupoid {P = P} grp b s2 pp (inl x) = b CP²→Groupoid {P = P} grp b s2 pp (inr x) = s2 x CP²→Groupoid {P = P} grp b s2 pp (push a i₁) = lem a i₁ where - lem : (a : hopfS¹.TotalSpaceHopf') → PathP (λ i → P (push a i)) b (s2 _) + lem : (a : hopfS¹.TotalSpaceHopfPush) → PathP (λ i → P (push a i)) b (s2 _) lem = TotalHopf→Gpd _ (λ _ → isOfHLevelPathP' 3 (grp _) _ _) pp -- The function inducing the iso H²(S²) ≅ H²(CP²) @@ -177,7 +171,7 @@ H²S²→H²CP²-raw f (inl x) = 0ₖ _ H²S²→H²CP²-raw f (inr x) = f x -ₖ f north H²S²→H²CP²-raw f (push a i) = TotalHopf→Gpd (λ x → 0ₖ 2 - ≡ f (hopfS¹.TotalSpaceHopf'→TotalSpace x .fst) -ₖ f north) + ≡ f (hopfS¹.TotalSpaceHopfPush→TotalSpace x .fst) -ₖ f north) (λ _ → isOfHLevelTrunc 4 _ _) (sym (rCancelₖ 2 (f north))) a i @@ -355,7 +349,7 @@ private invLooperLem₁ a x = secEq (hopfS¹.μ-eq a) x ∙∙ cong (_* x) (lemmie a) - ∙∙ assocHSpace.μ-assoc S1-AssocHSpace a (invLooper a) x + ∙∙ AssocHSpace.μ-assoc S1-AssocHSpace a (invLooper a) x ∙ comm·S¹ _ _ invLooperLem₂ : (a x : S¹) → invEq (hopfS¹.μ-eq a) x ≡ invLooper a * x @@ -477,7 +471,7 @@ isGenerator≃ℤ-e = -- Alternative definition of the hopfMap HopfMap' : S₊ 3 → S₊ 2 HopfMap' x = - hopfS¹.TotalSpaceHopf'→TotalSpace + hopfS¹.TotalSpaceHopfPush→TotalSpace (Iso.inv IsoTotalHopf' (Iso.inv (IsoSphereJoin 1 1) x)) .fst @@ -491,7 +485,7 @@ hopfMap≡HopfMap' = (Iso.inv (IsoSphereJoin 1 1) x))))) , flipSquare (sym (rUnit refl) ◁ λ _ _ → north)) where - lem : (x : _) → hopfS¹.TotalSpaceHopf'→TotalSpace x .fst + lem : (x : _) → hopfS¹.TotalSpaceHopfPush→TotalSpace x .fst ≡ JoinS¹S¹→TotalHopf (Iso.fun IsoTotalHopf' x) .fst lem (inl x) = refl lem (inr x) = refl From 90de52705bf330bd4c0dcd6850d360d29b17e728 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 2 Nov 2021 13:53:48 +0100 Subject: [PATCH 79/89] stuff --- Cubical/ZCohomology/Groups/CP2.agda | 913 ++++-------------- Cubical/ZCohomology/Gysin2.agda | 61 -- .../RingStructure/GradedCommutativity.agda | 3 +- 3 files changed, 200 insertions(+), 777 deletions(-) delete mode 100644 Cubical/ZCohomology/Gysin2.agda diff --git a/Cubical/ZCohomology/Groups/CP2.agda b/Cubical/ZCohomology/Groups/CP2.agda index 039c4f174..f7a7f0483 100644 --- a/Cubical/ZCohomology/Groups/CP2.agda +++ b/Cubical/ZCohomology/Groups/CP2.agda @@ -9,7 +9,6 @@ open import Cubical.ZCohomology.MayerVietorisUnreduced open import Cubical.ZCohomology.Groups.Unit open import Cubical.ZCohomology.Groups.Sn open import Cubical.ZCohomology.RingStructure.CupProduct -open import Cubical.ZCohomology.RingStructure.RingLaws open import Cubical.Foundations.HLevels open import Cubical.Foundations.Function @@ -27,7 +26,7 @@ open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) open import Cubical.Data.Nat.Order open import Cubical.Data.Unit open import Cubical.Algebra.Group - renaming (ℤ to ℤGroup ; Unit to UnitGroup) hiding (Bool) + renaming (ℤ to ℤGroup ; Unit to UnitGroup) open import Cubical.HITs.Pushout open import Cubical.HITs.S1 @@ -36,727 +35,213 @@ open import Cubical.HITs.Susp open import Cubical.HITs.Hopf open import Cubical.HITs.Join open import Cubical.HITs.SetTruncation - renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap) + renaming (rec to sRec ; elim to sElim ; elim2 to sElim2 ; map to sMap) open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim2 to pElim2 ; ∣_∣ to ∣_∣₁ ; map to pMap) open import Cubical.HITs.Truncation - renaming (elim to trElim ; rec2 to trRec2 ; rec to trRec) open import Cubical.Relation.Nullary open IsGroupHom open Iso - - -diag : ∀ {ℓ} {A : Type ℓ} {x y z w : A} (p : x ≡ y) (r : z ≡ w) (q : y ≡ z) (s : x ≡ w) (P : PathP (λ i → p i ≡ r (~ i)) s q) - → (λ i → P i i) ≡ p ∙ q -diag {x = x} {y = y} {z = z} {w = w} p r q s P i j = - hcomp (λ k → λ {(i = i0) → P j (j ∧ k) - ; (j = i0) → x - ; (j = i1) → q k}) - (p j) - {- - ———— Boundary —————————————————————————————————————————————— -i = i0 ⊢ P j j -i = i1 ⊢ (p ∙ q) j -j = i0 ⊢ x -j = i1 ⊢ z - -} - - -testagain : (g : S₊ 3 → S₊ 2) (p : Σ[ f ∈ (S₊ 2 → coHomK 4) ] ((x : S₊ 3) → f (g x) ≡ 0ₖ 4)) - → (s : (λ x → 0ₖ 4) ≡ fst p) - → p ≡ ((λ _ → 0ₖ 4) , λ x → funExt⁻ s (g x) ∙ (snd p x)) -testagain g (f , k) s = help f s k +CP² : Type +CP² = Pushout {A = TotalHopf} fst λ _ → tt + +characFunSpaceCP² : ∀ {ℓ} {A : Type ℓ} + → Iso (CP² → A) (Σ[ x ∈ A ] Σ[ f ∈ (S₊ (suc (suc zero)) → A) ] + ((y : TotalHopf) → f (fst y) ≡ x)) +fun characFunSpaceCP² f = (f (inr tt)) , ((f ∘ inl ) , (λ a → cong f (push a))) +inv characFunSpaceCP² (a , f , p) (inl x) = f x +inv characFunSpaceCP² (a , f , p) (inr x) = a +inv characFunSpaceCP² (a , f , p) (push b i) = p b i +rightInv characFunSpaceCP² _ = refl +leftInv characFunSpaceCP² _ = + funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a i) → refl} + +H⁰CP²≅ℤ : GroupIso (coHomGr 0 CP²) ℤGroup +H⁰CP²≅ℤ = + H⁰-connected (inr tt) + (PushoutToProp (λ _ → squash) + (sphereElim _ (λ _ → isOfHLevelSuc 1 squash) + ∣ sym (push (north , base)) ∣₁) + λ _ → ∣ refl ∣₁) + +module M = MV (S₊ 2) Unit TotalHopf fst (λ _ → tt) + +H²CP²≅ℤ : GroupIso (coHomGr 2 CP²) ℤGroup +H²CP²≅ℤ = compGroupIso (BijectionIso→GroupIso bij) + (compGroupIso (invGroupIso trivIso) (Hⁿ-Sⁿ≅ℤ 1)) + where + isContrH¹TotalHopf : isContr (coHom 1 TotalHopf) + isContrH¹TotalHopf = + isOfHLevelRetractFromIso 0 (setTruncIso (domIso (invIso (IsoS³TotalHopf)))) + (isOfHLevelRetractFromIso 0 ((fst (H¹-Sⁿ≅0 1))) isContrUnit) + + isContrH²TotalHopf : isContr (coHom 2 TotalHopf) + isContrH²TotalHopf = + isOfHLevelRetractFromIso 0 (setTruncIso (domIso (invIso (IsoS³TotalHopf)))) + ((isOfHLevelRetractFromIso 0 + (fst (Hⁿ-Sᵐ≅0 1 2 λ p → snotz (sym (cong predℕ p)))) isContrUnit)) + + trivIso : GroupIso (coHomGr 2 (Susp S¹)) (×coHomGr 2 (Susp S¹) Unit) + fun (fst trivIso) x = x , 0ₕ _ + inv (fst trivIso) = fst + rightInv (fst trivIso) (x , p) = + ΣPathP (refl , isContr→isProp (isContrHⁿ-Unit 1) _ _) + leftInv (fst trivIso) x = refl + snd trivIso = makeIsGroupHom λ _ _ → refl + + bij : BijectionIso (coHomGr 2 CP²) (×coHomGr 2 (Susp S¹) Unit) + BijectionIso.fun bij = M.i 2 + BijectionIso.inj bij x p = + pRec (squash₂ _ _) + (uncurry (λ z q + → sym q + ∙∙ cong (fst (M.d 1)) (isContr→isProp isContrH¹TotalHopf z (0ₕ _)) + ∙∙ (M.d 1) .snd .pres1)) + (M.Ker-i⊂Im-d 1 x p) + where + help : isInIm (M.d 1) x + help = M.Ker-i⊂Im-d 1 x p + BijectionIso.surj bij y = + M.Ker-Δ⊂Im-i 2 y (isContr→isProp isContrH²TotalHopf _ _) + +H⁴CP²≅ℤ : GroupIso (coHomGr 4 CP²) ℤGroup +H⁴CP²≅ℤ = compGroupIso (invGroupIso (BijectionIso→GroupIso bij)) + (compGroupIso help (Hⁿ-Sⁿ≅ℤ 2)) where - help : (f : (S₊ 2 → coHomK 4)) (s : (λ x → 0ₖ 4) ≡ f) (k : (x : S₊ 3) → f (g x) ≡ 0ₖ 4) - → Path (Σ[ f ∈ (S₊ 2 → coHomK 4) ] ((x : S₊ 3) → f (g x) ≡ 0ₖ 4)) (f , k) - (((λ _ → 0ₖ 4) , λ x → funExt⁻ s (g x) ∙ (k x))) - help f = - J (λ f s → (k : (x : S₊ 3) → f (g x) ≡ 0ₖ 4) - → Path (Σ[ f ∈ (S₊ 2 → coHomK 4) ] ((x : S₊ 3) → f (g x) ≡ 0ₖ 4)) (f , k) - (((λ _ → 0ₖ 4) , λ x → funExt⁻ s (g x) ∙ (k x)))) - λ k → ΣPathP (refl , (funExt (λ x → lUnit (k x)))) - -open import Cubical.Homotopy.Loopspace -open import Cubical.HITs.SmashProduct - -hopfy : ∀ {ℓ} {A : Pointed ℓ} - → (Ω^ 2) A .fst → (Ω^ 3) A .fst -hopfy α l i = - hcomp - (λ m → λ - { (i = i0) → refl - ; (i = i1) → λ j → side l m j - ; (l = i0) → α (i ∧ ~ m) - ; (l = i1) → α (~ i ∨ m) - }) - (λ j → top l i j) + help : GroupIso (coHomGr 3 TotalHopf) (coHomGr 3 (S₊ 3)) + help = isoType→isoCohom 3 (invIso IsoS³TotalHopf) + + bij : BijectionIso (coHomGr 3 TotalHopf) (coHomGr 4 CP²) + BijectionIso.fun bij = M.d 3 + BijectionIso.inj bij x p = + pRec (squash₂ _ _) + (uncurry (λ z q → + sym q + ∙∙ cong (M.Δ 3 .fst) + (isOfHLevelΣ 1 (isContr→isProp + (isOfHLevelRetractFromIso 0 + (fst (Hⁿ-Sᵐ≅0 2 1 λ p → snotz (cong predℕ p))) isContrUnit)) + (λ _ → isContr→isProp (isContrHⁿ-Unit 2)) + z (0ₕ _ , 0ₕ _)) + ∙∙ M.Δ 3 .snd .pres1)) + (M.Ker-d⊂Im-Δ _ x p) + BijectionIso.surj bij y = + M.Ker-i⊂Im-d 3 y (isOfHLevelΣ 1 + (isContr→isProp (isOfHLevelRetractFromIso 0 + (fst (Hⁿ-Sᵐ≅0 3 1 λ p → snotz (cong predℕ p))) isContrUnit)) + (λ _ → isContr→isProp (isContrHⁿ-Unit _)) _ _) + + +-- Characterisations of the trivial groups + +private + elim-TotalHopf : (B : TotalHopf → Type) + → ((x : _) → isOfHLevel 3 (B x)) → B (north , base) + → (x : _) → B x + elim-TotalHopf = + transport (λ i → (B : isoToPath IsoS³TotalHopf i → Type) + → ((x : _) → isOfHLevel 3 (B x)) + → B (transp (λ j → isoToPath IsoS³TotalHopf (i ∨ ~ j)) i (north , base)) → (x : _) → B x) + λ B hLev elim-TotalHopf → sphereElim _ (λ _ → hLev _) elim-TotalHopf + +H¹-CP²≅0 : GroupIso (coHomGr 1 CP²) UnitGroup +H¹-CP²≅0 = + contrGroupIsoUnit + (isOfHLevelRetractFromIso 0 (setTruncIso characFunSpaceCP²) + (isOfHLevelRetractFromIso 0 lem₂ lem₃)) where - a₀ = α i0 i0 - - top : (l i j : I) → _ - top l i j = - hcomp - (λ k → λ - { (i = i0) → a₀ - ; (j = i0) → a₀ - ; (j = i1) → α (~ (i ∧ l)) k - ; (l = i0) → α i j - ; (l = i1) → α (~ i) (j ∧ k) - }) - (α (i ∧ ~ l) j) - - side : (l m j : I) → _ - side l m j = - hcomp - (λ k → λ - { (m = i1) → a₀ - ; (j = i0) → a₀ - ; (j = i1) → α (~ (m ∨ l)) k - ; (l = i0) → α (~ m) (j ∧ k) - ; (l = i1) → α m j - }) - (α (m ∨ ~ l) j) - -EH-funct : ∀ {ℓ} (p q : typ ((Ω^ 2) (S₊∙ 2))) → EH 0 (p ∙ q) (sym (p ∙ q)) ≡ {!EH 0 p (sym p)!} -EH-funct = {!!} - -zz : (p q r : typ (Ω (S₊∙ 2))) (b : r ≡ r) (t : p ∙ sym p ≡ sym p ∙ p) (l : r ≡ p) (r : r ≡ q) (P : PathP (λ i → {!!} ≡ {!p i!}) {!!} {!!}) - → Path (Square {!!} {!!} {!!} {!!}) (λ i j → P (i ∨ j) i j) {!!} -zz p q = {!!} - -Iso1 : Iso ∥ typ ((Ω^ 3) (S₊∙ 2)) ∥₂ ∥ typ ((Ω^ 2) (S₊∙ 2)) ∥₂ -fun Iso1 = {!!} -inv Iso1 = sMap hopfy -rightInv Iso1 = sElim {!!} λ p → cong ∣_∣₂ {!!} -- (l p) + lem₁ : (f : (Susp S¹ → coHomK 1)) → ∥ (λ _ → 0ₖ _) ≡ f ∥ + lem₁ f = pMap (λ p → p) + (Iso.fun PathIdTrunc₀Iso (isOfHLevelRetractFromIso 1 + (fst (Hⁿ-Sᵐ≅0 0 1 (λ p → snotz (sym p)))) isPropUnit (0ₕ _) ∣ f ∣₂)) + + lem₂ : Iso ∥ (Σ[ x ∈ coHomK 1 ] ( Σ[ f ∈ (Susp S¹ → coHomK 1) ] ((y : TotalHopf) → f (fst y) ≡ x))) ∥₂ + ∥ (Σ[ f ∈ (Susp S¹ → coHomK 1) ] ((y : TotalHopf) → f (fst y) ≡ 0ₖ 1)) ∥₂ + fun lem₂ = sMap (uncurry λ x → uncurry λ f p → (λ y → (-ₖ x) +ₖ f y) , λ y → cong ((-ₖ x) +ₖ_) (p y) ∙ lCancelₖ _ x) + inv lem₂ = sMap λ p → 0ₖ _ , p + rightInv lem₂ = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ {(f , p) → cong ∣_∣₂ (ΣPathP ((funExt (λ x → lUnitₖ _ (f x))) + , (funExt (λ y → sym (rUnit (λ i → (-ₖ 0ₖ 1) +ₖ p y i))) + ◁ λ j y i → lUnitₖ _ (p y i) j)))} + leftInv lem₂ = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + (uncurry (coHomK-elim _ (λ _ → isPropΠ (λ _ → squash₂ _ _)) + (uncurry λ f p → cong ∣_∣₂ (ΣPathP (refl , (ΣPathP ((funExt (λ x → lUnitₖ _ (f x))) + , ((funExt (λ y → sym (rUnit (λ i → (-ₖ 0ₖ 1) +ₖ p y i))) + ◁ λ j y i → lUnitₖ _ (p y i) j))))))))) + + lem₃ : isContr _ + fst lem₃ = ∣ (λ _ → 0ₖ 1) , (λ _ → refl) ∣₂ + snd lem₃ = + sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + (uncurry λ f → pRec (isPropΠ (λ _ → squash₂ _ _)) + (J (λ f _ → (y : (y₁ : TotalHopf) → f (fst y₁) ≡ 0ₖ 1) → + ∣ (λ _ → 0ₖ 1) , (λ _ _ → 0ₖ 1) ∣₂ ≡ ∣ f , y ∣₂) + (λ y → cong ∣_∣₂ (ΣPathP ((funExt (λ z → sym (y (north , base)))) , toPathP (s y))))) + (lem₁ f)) + + where + s : (y : TotalHopf → 0ₖ 1 ≡ 0ₖ 1) + → transport (λ i → (_ : TotalHopf) → y (north , base) (~ i) ≡ ∣ base ∣) + (λ _ _ → 0ₖ 1) ≡ y + s y = funExt (elim-TotalHopf _ (λ _ → isOfHLevelPath 3 (isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) _ _) + λ k → transp (λ i → y (north , base) (~ i ∧ ~ k) ≡ ∣ base ∣) k + λ j → y (north , base) (~ k ∨ j)) + +Hⁿ-CP²≅0-higher : (n : ℕ) → ¬ (n ≡ 1) → GroupIso (coHomGr (3 +ℕ n) CP²) UnitGroup +Hⁿ-CP²≅0-higher n p = contrGroupIsoUnit ((0ₕ _) , (λ x → sym (main x))) where - l : (p : _) → {!!} - l p k i j = - hcomp {!λ r → {(i = i0) → ? ; (i = i1) → ? ; (j = i0) → ? ; (j = i1) → ? ; (k = i1) → ? ; (k = i0) → ?}!} - {!!} -leftInv Iso1 = - {!k = i0 ⊢ ((λ i₁ → rCancel p (~ i₁)) ∙∙ EH 0 p (p ⁻¹) ∙∙ lCancel p) - (~ i ∨ j) (~ i ∧ j) j -k = i1 ⊢ p i j -i = i0 ⊢ snd (Ω (S₊∙ 2)) j -i = i1 ⊢ snd (Ω (S₊∙ 2)) j -j = i0 ⊢ snd (S₊∙ 2) -j = i1 ⊢ snd (S₊∙ 2)!} - -surf' : typ ((Ω^ 2) (S₊∙ 2)) -surf' i j = hcomp (λ k → λ { (i = i0) → north - ; (i = i1) → north - ; (j = i0) → rCancel (merid base) k i - ; (j = i1) → rCancel (merid base) k i}) - ((merid (loop j) ∙ sym (merid base)) i) - -test : ℤ -test = sRec isSetℤ (λ p → Iso.inv (Iso-Kn-ΩKn+1 0) (cong (Iso.inv (Iso-Kn-ΩKn+1 1)) (cong (cong ∣_∣ₕ) p))) - (Iso.fun Iso1 ∣ hopfy surf' ∣₂) - --- fun1 : ∥ typ ((Ω^ 3) (S₊∙ 2)) ∥₂ → coHom 2 (Smash (S₊∙ 1) (S₊∙ 1)) --- fun1 = {!!} - - --- open import Cubical.Data.Bool --- S¹* : Type --- S¹* = Susp Bool - --- S³* : Type --- S³* = Susp (Susp (Susp Bool)) - --- open import Cubical.HITs.S3 --- open import Cubical.HITs.Join --- S¹*S¹→S³ : join S¹ S¹ → S₊ 2 --- S¹*S¹→S³ (inl x) = north --- S¹*S¹→S³ (inr x) = north --- S¹*S¹→S³ (push base b i) = north --- S¹*S¹→S³ (push (loop j) base i) = north --- S¹*S¹→S³ (push (loop i) (loop j) k) = (sym (rCancel surf') ∙∙ EH 0 surf' (sym surf') ∙∙ lCancel surf') k i j - --- c : Iso S³* TotalHopf --- fun c north = north , base --- fun c south = north , base --- fun c (merid a i) = h1 a i , S³≡S³ a i --- where --- h1 : (a : Susp (Susp Bool)) → Path (S₊ 2) north north --- h1 north = refl --- h1 south = refl --- h1 (merid north i) j = (surf' ∙ sym (surf')) i j --- h1 (merid south i) j = (sym surf' ∙ surf') i j --- h1 (merid (merid false k) i) j = (rCancel surf' ∙ sym (lCancel surf')) k i j --- h1 (merid (merid true k) i) j = EH 0 surf' (sym surf') k i j - --- S³≡S³ : (a : Susp (Susp Bool)) → PathP (λ i → HopfSuspS¹ (h1 a i)) base base --- S³≡S³ north = refl --- S³≡S³ south = refl --- S³≡S³ (merid a i) = {!!} -- help a i --- where --- sl : {!(j : I) → !} --- sl = {!!} --- l : (j : I) (x : Susp (Susp Bool)) (p : north ≡ x) → Path (PathP {!!} {!!} {!!}) (λ j → transp (λ i₂ → HopfSuspS¹ (h1 (p i₂) j)) i0 base) {!!} --- l = {!!} - --- s : (a : _) → (λ j → transp (λ i₂ → HopfSuspS¹ (h1 (merid a i₂) j)) i0 base) ≡ refl --- s x = {!!} --- help : (a : _) → SquareP (λ i j → HopfSuspS¹ (h1 (merid a i) j)) refl refl refl refl --- help a = toPathP (transportRefl _ ∙ s a) - --- h : (a : S₊ 2) → Path TotalHopf (north , base) (north , base) --- h a = cong₂ _,_ refl {!!} --- where --- help : (a : S₊ 2) → PathP (λ i → Glue S¹ (Border base i)) base base --- help = {!!} --- inv c = {!!} --- rightInv c = {!!} --- leftInv c = {!!} - - --- hej : S₊ 2 → Type₀ --- hej north = S¹ --- hej south = S¹ --- hej (merid a i) = {!!} - --- sauto : S¹ → I → S¹ --- sauto base i = base --- sauto (loop j) i = --- hcomp (λ k → λ {(i = i0) → loop (j ∨ k) --- ; (i = i1) → base --- ; (j = i0) → loop (i ∨ k) --- ; (j = i1) → base}) --- (loop (i ∨ j)) - --- open import Cubical.HITs.S1 renaming (_·_ to _*_) --- S₊2→Type : S₊ 2 → Type --- S₊2→Type north = typ (Ω (S₊∙ 2)) --- S₊2→Type south = typ (Ω (S₊∙ 2)) --- S₊2→Type (merid a i) = l i --- where --- h : (a : S¹) → isEquiv λ p → p ∙ merid (sauto a i) ∙ sym (merid base) --- h = sphereElim 0 (λ _ → isPropIsEquiv _) (subst isEquiv (λ j p → ((cong (p ∙_) (rCancel (merid base))) ∙ sym (rUnit p)) (~ j)) --- (idEquiv _ .snd)) - --- l : typ (Ω (S₊∙ 2)) ≡ typ (Ω (S₊∙ 2)) --- l = ua ((λ p → p ∙ merid (sauto a i) ∙ sym (merid base)) , h a) -- ua (a *_ , (rotIsEquiv a)) - --- t2 : hLevelTrunc 3 ((typ ((Ω^ 2) (S₊∙ 2)))) → S¹ --- t2 = {!!} - --- test : ℤ --- test = Iso.fun (fst (Hⁿ-Sⁿ≅ℤ 1)) ∣ {!!} ∣₂ - --- S2-add : hLevelTrunc 5 (S₊ 2) → hLevelTrunc 5 (S₊ 2) → hLevelTrunc 4 (S₊ 2) --- S2-add = trRec2 {!!} (wedgeconFun 1 1 (λ _ _ → isOfHLevelTrunc {!!}) {!!} {!!} {!!}) - --- CP² : Type --- CP² = Pushout {A = TotalHopf} fst λ _ → tt - --- characFunSpaceCP² : ∀ {ℓ} {A : Type ℓ} --- → Iso (CP² → A) (Σ[ x ∈ A ] Σ[ f ∈ (S₊ (suc (suc zero)) → A) ] --- ((y : TotalHopf) → f (fst y) ≡ x)) --- fun characFunSpaceCP² f = (f (inr tt)) , ((f ∘ inl ) , (λ a → cong f (push a))) --- inv characFunSpaceCP² (a , f , p) (inl x) = f x --- inv characFunSpaceCP² (a , f , p) (inr x) = a --- inv characFunSpaceCP² (a , f , p) (push b i) = p b i --- rightInv characFunSpaceCP² _ = refl --- leftInv characFunSpaceCP² _ = --- funExt λ { (inl x) → refl --- ; (inr x) → refl --- ; (push a i) → refl} - --- H⁰CP²≅ℤ : GroupIso (coHomGr 0 CP²) ℤGroup --- H⁰CP²≅ℤ = --- H⁰-connected (inr tt) --- (PushoutToProp (λ _ → squash) --- (sphereElim _ (λ _ → isOfHLevelSuc 1 squash) --- ∣ sym (push (north , base)) ∣₁) --- λ _ → ∣ refl ∣₁) - --- module M = MV (S₊ 2) Unit TotalHopf fst (λ _ → tt) - --- H²CP²≅ℤ : GroupIso (coHomGr 2 CP²) ℤGroup --- H²CP²≅ℤ = compGroupIso (BijectionIso→GroupIso bij) --- (compGroupIso (invGroupIso trivIso) (Hⁿ-Sⁿ≅ℤ 1)) --- where --- isContrH¹TotalHopf : isContr (coHom 1 TotalHopf) --- isContrH¹TotalHopf = --- isOfHLevelRetractFromIso 0 (setTruncIso (domIso (invIso (IsoS³TotalHopf)))) --- (isOfHLevelRetractFromIso 0 ((fst (H¹-Sⁿ≅0 1))) isContrUnit) - --- isContrH²TotalHopf : isContr (coHom 2 TotalHopf) --- isContrH²TotalHopf = --- isOfHLevelRetractFromIso 0 (setTruncIso (domIso (invIso (IsoS³TotalHopf)))) --- ((isOfHLevelRetractFromIso 0 --- (fst (Hⁿ-Sᵐ≅0 1 2 λ p → snotz (sym (cong predℕ p)))) isContrUnit)) - --- trivIso : GroupIso (coHomGr 2 (Susp S¹)) (×coHomGr 2 (Susp S¹) Unit) --- fun (fst trivIso) x = x , 0ₕ _ --- inv (fst trivIso) = fst --- rightInv (fst trivIso) (x , p) = --- ΣPathP (refl , isContr→isProp (isContrHⁿ-Unit 1) _ _) --- leftInv (fst trivIso) x = refl --- snd trivIso = makeIsGroupHom λ _ _ → refl - --- bij : BijectionIso (coHomGr 2 CP²) (×coHomGr 2 (Susp S¹) Unit) --- BijectionIso.fun bij = M.i 2 --- BijectionIso.inj bij x p = --- pRec (squash₂ _ _) --- (uncurry (λ z q --- → sym q --- ∙∙ cong (fst (M.d 1)) (isContr→isProp isContrH¹TotalHopf z (0ₕ _)) --- ∙∙ (M.d 1) .snd .pres1)) --- (M.Ker-i⊂Im-d 1 x p) --- where --- help : isInIm (M.d 1) x --- help = M.Ker-i⊂Im-d 1 x p --- BijectionIso.surj bij y = --- M.Ker-Δ⊂Im-i 2 y (isContr→isProp isContrH²TotalHopf _ _) - --- H⁴CP²≅ℤ : GroupIso (coHomGr 4 CP²) ℤGroup --- H⁴CP²≅ℤ = compGroupIso (invGroupIso (BijectionIso→GroupIso bij)) --- (compGroupIso help (Hⁿ-Sⁿ≅ℤ 2)) --- where --- help : GroupIso (coHomGr 3 TotalHopf) (coHomGr 3 (S₊ 3)) --- help = isoType→isoCohom 3 (invIso IsoS³TotalHopf) - --- bij : BijectionIso (coHomGr 3 TotalHopf) (coHomGr 4 CP²) --- BijectionIso.fun bij = M.d 3 --- BijectionIso.inj bij x p = --- pRec (squash₂ _ _) --- (uncurry (λ z q → --- sym q --- ∙∙ cong (M.Δ 3 .fst) --- (isOfHLevelΣ 1 (isContr→isProp --- (isOfHLevelRetractFromIso 0 --- (fst (Hⁿ-Sᵐ≅0 2 1 λ p → snotz (cong predℕ p))) isContrUnit)) --- (λ _ → isContr→isProp (isContrHⁿ-Unit 2)) --- z (0ₕ _ , 0ₕ _)) --- ∙∙ M.Δ 3 .snd .pres1)) --- (M.Ker-d⊂Im-Δ _ x p) --- BijectionIso.surj bij y = --- M.Ker-i⊂Im-d 3 y (isOfHLevelΣ 1 --- (isContr→isProp (isOfHLevelRetractFromIso 0 --- (fst (Hⁿ-Sᵐ≅0 3 1 λ p → snotz (cong predℕ p))) isContrUnit)) --- (λ _ → isContr→isProp (isContrHⁿ-Unit _)) _ _) - - --- -- Characterisations of the trivial groups - --- private --- elim-TotalHopf : (B : TotalHopf → Type) --- → ((x : _) → isOfHLevel 3 (B x)) → B (north , base) --- → (x : _) → B x --- elim-TotalHopf = --- transport (λ i → (B : isoToPath IsoS³TotalHopf i → Type) --- → ((x : _) → isOfHLevel 3 (B x)) --- → B (transp (λ j → isoToPath IsoS³TotalHopf (i ∨ ~ j)) i (north , base)) → (x : _) → B x) --- λ B hLev elim-TotalHopf → sphereElim _ (λ _ → hLev _) elim-TotalHopf - --- H¹-CP²≅0 : GroupIso (coHomGr 1 CP²) UnitGroup --- H¹-CP²≅0 = --- contrGroupIsoUnit --- (isOfHLevelRetractFromIso 0 (setTruncIso characFunSpaceCP²) --- (isOfHLevelRetractFromIso 0 lem₂ lem₃)) --- where --- lem₁ : (f : (Susp S¹ → coHomK 1)) → ∥ (λ _ → 0ₖ _) ≡ f ∥ --- lem₁ f = pMap (λ p → p) --- (Iso.fun PathIdTrunc₀Iso (isOfHLevelRetractFromIso 1 --- (fst (Hⁿ-Sᵐ≅0 0 1 (λ p → snotz (sym p)))) isPropUnit (0ₕ _) ∣ f ∣₂)) - --- lem₂ : Iso ∥ (Σ[ x ∈ coHomK 1 ] ( Σ[ f ∈ (Susp S¹ → coHomK 1) ] ((y : TotalHopf) → f (fst y) ≡ x))) ∥₂ --- ∥ (Σ[ f ∈ (Susp S¹ → coHomK 1) ] ((y : TotalHopf) → f (fst y) ≡ 0ₖ 1)) ∥₂ --- fun lem₂ = sMap (uncurry λ x → uncurry λ f p → (λ y → (-ₖ x) +ₖ f y) , λ y → cong ((-ₖ x) +ₖ_) (p y) ∙ lCancelₖ _ x) --- inv lem₂ = sMap λ p → 0ₖ _ , p --- rightInv lem₂ = --- sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) --- λ {(f , p) → cong ∣_∣₂ (ΣPathP ((funExt (λ x → lUnitₖ _ (f x))) --- , (funExt (λ y → sym (rUnit (λ i → (-ₖ 0ₖ 1) +ₖ p y i))) --- ◁ λ j y i → lUnitₖ _ (p y i) j)))} --- leftInv lem₂ = --- sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) --- (uncurry (coHomK-elim _ (λ _ → isPropΠ (λ _ → squash₂ _ _)) --- (uncurry λ f p → cong ∣_∣₂ (ΣPathP (refl , (ΣPathP ((funExt (λ x → lUnitₖ _ (f x))) --- , ((funExt (λ y → sym (rUnit (λ i → (-ₖ 0ₖ 1) +ₖ p y i))) --- ◁ λ j y i → lUnitₖ _ (p y i) j))))))))) - --- lem₃ : isContr _ --- fst lem₃ = ∣ (λ _ → 0ₖ 1) , (λ _ → refl) ∣₂ --- snd lem₃ = --- sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) --- (uncurry λ f → pRec (isPropΠ (λ _ → squash₂ _ _)) --- (J (λ f _ → (y : (y₁ : TotalHopf) → f (fst y₁) ≡ 0ₖ 1) → --- ∣ (λ _ → 0ₖ 1) , (λ _ _ → 0ₖ 1) ∣₂ ≡ ∣ f , y ∣₂) --- (λ y → cong ∣_∣₂ (ΣPathP ((funExt (λ z → sym (y (north , base)))) , toPathP (s y))))) --- (lem₁ f)) - --- where --- s : (y : TotalHopf → 0ₖ 1 ≡ 0ₖ 1) --- → transport (λ i → (_ : TotalHopf) → y (north , base) (~ i) ≡ ∣ base ∣) --- (λ _ _ → 0ₖ 1) ≡ y --- s y = funExt (elim-TotalHopf _ (λ _ → isOfHLevelPath 3 (isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) _ _) --- λ k → transp (λ i → y (north , base) (~ i ∧ ~ k) ≡ ∣ base ∣) k --- λ j → y (north , base) (~ k ∨ j)) - --- Hⁿ-CP²≅0-higher : (n : ℕ) → ¬ (n ≡ 1) → GroupIso (coHomGr (3 +ℕ n) CP²) UnitGroup --- Hⁿ-CP²≅0-higher n p = contrGroupIsoUnit ((0ₕ _) , (λ x → sym (main x))) --- where --- h : GroupHom (coHomGr (2 +ℕ n) TotalHopf) (coHomGr (3 +ℕ n) CP²) --- h = M.d (2 +ℕ n) - --- propᵣ : isProp (fst (×coHomGr (3 +ℕ n) (Susp S¹) Unit)) --- propᵣ = --- isPropΣ --- (isOfHLevelRetractFromIso 1 --- (fst (Hⁿ-Sᵐ≅0 (2 +ℕ n) 1 λ p → ⊥-rec (snotz (cong predℕ p)))) isPropUnit) --- λ _ → isContr→isProp (isContrHⁿ-Unit _) - --- propₗ : isProp (coHom (2 +ℕ n) TotalHopf) --- propₗ = subst (λ x → isProp (coHom (2 +ℕ n) x)) (isoToPath IsoS³TotalHopf) --- (isOfHLevelRetractFromIso 1 --- (fst (Hⁿ-Sᵐ≅0 (suc n) 2 λ q → p (cong predℕ q))) isPropUnit) - --- inIm : (x : coHom (3 +ℕ n) CP²) → isInIm h x --- inIm x = M.Ker-i⊂Im-d (2 +ℕ n) x (propᵣ _ _) - --- main : (x : coHom (3 +ℕ n) CP²) → x ≡ 0ₕ _ --- main x = --- pRec (squash₂ _ _) --- (uncurry (λ f p → sym p ∙∙ cong (h .fst) (propₗ f (0ₕ _)) ∙∙ pres1 (snd h))) --- (inIm x) - --- -- All trivial groups: --- Hⁿ-CP²≅0 : (n : ℕ) → ¬ suc n ≡ 2 → ¬ suc n ≡ 4 --- → GroupIso (coHomGr (suc n) CP²) UnitGroup --- Hⁿ-CP²≅0 zero p q = H¹-CP²≅0 --- Hⁿ-CP²≅0 (suc zero) p q = ⊥-rec (p refl) --- Hⁿ-CP²≅0 (suc (suc zero)) p q = Hⁿ-CP²≅0-higher 0 λ p → snotz (sym p) --- Hⁿ-CP²≅0 (suc (suc (suc zero))) p q = ⊥-rec (q refl) --- Hⁿ-CP²≅0 (suc (suc (suc (suc n)))) p q = --- Hⁿ-CP²≅0-higher (suc (suc n)) --- λ p → snotz (cong predℕ p) - --- -- Another Brunerie number --- ℤ→HⁿCP²→ℤ : ℤ → ℤ --- ℤ→HⁿCP²→ℤ x = --- Iso.fun (fst H⁴CP²≅ℤ) --- (Iso.inv (fst H²CP²≅ℤ) x ⌣ Iso.inv (fst H²CP²≅ℤ) x) - --- brunerie2 : ℤ --- brunerie2 = ℤ→HⁿCP²→ℤ 1 - --- {- --- |brunerie2|≡1 : abs (ℤ→HⁿCP²→ℤ 1) ≡ 1 --- |brunerie2|≡1 = refl --- -} - --- t : ∀ (a : TotalHopf) → ∣ fst a ∣ ≡ 0ₖ 2 --- t = {!!} - --- one : coHom 2 CP² --- one = ∣ (λ { (inl x) → ∣ x ∣ --- ; (inr x) → 0ₖ _ --- ; (push a i) → t a i}) ∣₂ - - --- CP²→CP⁴ : (x : coHom 2 CP²) → coHom 2 CP² → coHom 4 CP² --- CP²→CP⁴ x f = f ⌣ x - --- zz : Iso.fun (fst H⁴CP²≅ℤ) (CP²→CP⁴ one one) ≡ 1 --- zz = {!pres1!} ∙∙ {!!} ∙∙ {!!} --- open import Cubical.ZCohomology.Properties --- CP4→CP2 : (x : coHom 2 CP²) → coHom 4 CP² → coHom 2 CP² --- CP4→CP2 = sRec2 squash₂ --- λ e f → ∣ (λ { (inl x) → {!!} --- ; (inr x) → {!!} --- ; (push a i) → {!!}}) ∣₂ - - --- CP' : (f : S₊∙ 3 →∙ S₊∙ 2) → Type₀ --- CP' f = Pushout {A = S₊ 3} (fst f) λ _ → tt - --- open import Cubical.Homotopy.Loopspace - --- contrH2 : (f : S₊ 2 → coHomK 4) → ∥ (λ x → 0ₖ 4) ≡ f ∥₂ --- contrH2 = {!!} - --- CP4→CP2'' : (f : S₊∙ 3 →∙ S₊∙ 2) → ∥ (Σ[ g ∈ (S₊ 2 → coHomK 4) ] ((a : S₊ 3) → g (fst f a) ≡ 0ₖ 4)) ∥₂ --- → ∥ ((S₊∙ 3 →∙ (Ω (coHomK-ptd 4)))) ∥₂ --- CP4→CP2'' (f , p) = --- sRec squash₂ --- (uncurry λ g → --- sRec (isSetΠ (λ _ → squash₂)) (J (λ g _ → (y : (a : Susp (Susp S¹)) → g (f a) ≡ 0ₖ 4) --- → ∥ ((S₊∙ 3 →∙ (Ω (coHomK-ptd 4)))) ∥₂) --- λ p → ∣ (λ a → p a ∙ sym (p north)) , rCancel (p north) ∣₂) --- (contrH2 g)) - --- CP²' = CP' (fst ∘ Iso.fun (IsoS³TotalHopf) , refl) - --- gen0 : CP²' → coHomK 2 --- gen0 = {!!} - --- gen1 : CP²' → coHomK 4 --- gen1 (inl x) = 0ₖ 4 --- gen1 (inr x) = 0ₖ 4 --- gen1 (push a i) = Kn→ΩKn+1 3 ∣ a ∣ i - - --- hehehe : S₊ 3 → S₊ 2 --- hehehe north = north --- hehehe south = north --- hehehe (merid north i) = north --- hehehe (merid south i) = north --- hehehe (merid (merid base i₁) i) = north --- hehehe (merid (merid (loop k) j) i) = --- (sym (rCancel surf') ∙∙ EH 0 surf' (sym surf') ∙∙ lCancel surf') i j k - --- wtf : (x : S₊ 2) → Σ[ y ∈ S₊ 3 ] hehehe y ≡ x --- wtf north = north , refl --- wtf south = north , merid base --- wtf (merid base i) = north , λ j → (merid base (i ∧ j)) --- wtf (merid (loop j) i) = sq i j , {!!} --- where --- sq : Square {A = S₊ 3} (refl {x = north}) refl refl refl --- sq i j = hcomp (λ k → λ {(i = i0) → north --- ; (i = i1) → merid north (~ k) --- ; (j = i0) → merid north (~ k ∧ i) --- ; (j = i1) → merid {!!} (i ∧ ~ k)}) --- (merid (merid (loop j) j) i) - - --- -- sl : S₊ 2 → Type --- -- sl north = S¹ --- -- sl south = S¹ --- -- sl (merid base i) = S¹ --- -- sl (merid (loop j) i) = h i j --- -- where --- -- h : Square (refl {x = S¹}) (λ _ → S¹) (λ _ → S¹) λ _ → S¹ --- -- h i j = Glue {!HopfSuspS¹ (merid base i)!} --- -- λ { (i = i0) → S¹ , {!!} --- -- ; (i = i1) → {!!} --- -- ; (j = i0) → {!!} --- -- ; (j = i1) → {!!}} - --- -- CP7 : Type₀ --- -- CP7 = Pushout hehehe λ _ → tt - --- -- indMap : typ ((Ω^ 4) (coHomK-ptd 4)) → CP7 → coHomK 4 --- -- indMap P (inl x) = 0ₖ 4 --- -- indMap P (inr x) = 0ₖ 4 --- -- indMap P (push north i) = 0ₖ 4 --- -- indMap P (push south i) = 0ₖ 4 --- -- indMap P (push (merid north j) i) = 0ₖ 4 --- -- indMap P (push (merid south j) i) = 0ₖ 4 --- -- indMap P (push (merid (merid base i₁) j) i) = 0ₖ 4 --- -- indMap P (push (merid (merid (loop l) k) j) i) = P i j k l - --- -- ss : Iso ∥ (CP7 → coHomK 4) ∥₂ ∥ (Σ[ x ∈ coHomK 4 ] Σ[ f ∈ (S₊ 2 → coHomK 4) ] ((a : S₊ 3) → f (hehehe a) ≡ x)) ∥₂ --- -- fun ss = sMap λ f → (f (inr tt)) , ((f ∘ inl) , (λ a → cong f (push a))) --- -- inv ss = sMap λ {(pt , f , p) → λ { (inl x) → f x ; (inr x) → pt ; (push a i) → p a i}} --- -- rightInv ss = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) (λ x → refl) --- -- leftInv ss = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) --- -- λ f → cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i) → refl}) - --- -- ss2 : Iso (∥ (Σ[ x ∈ coHomK 4 ] Σ[ f ∈ (S₊ 2 → coHomK 4) ] ((a : S₊ 3) → f (hehehe a) ≡ x)) ∥₂) --- -- ∥ (Σ[ f ∈ (S₊ 2 → coHomK 4) ] ((a : S₊ 3) → f (hehehe a) ≡ 0ₖ 4)) ∥₂ --- -- fun ss2 = --- -- sRec squash₂ (uncurry (coHomK-elim 3 (λ _ → isOfHLevelΠ 4 λ _ → isOfHLevelSuc 3 (isOfHLevelSuc 2 squash₂)) --- -- ∣_∣₂)) --- -- inv ss2 = sMap λ p → (0ₖ 4 , p) --- -- rightInv ss2 = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ _ → refl --- -- leftInv ss2 = --- -- sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) --- -- (uncurry (coHomK-elim 3 --- -- (λ _ → isOfHLevelΠ 4 --- -- λ _ → isOfHLevelPath 4 --- -- (isOfHLevelSuc 3 (isOfHLevelSuc 2 squash₂)) _ _) λ _ → refl)) --- -- open import Cubical.Homotopy.Connected - --- -- truncLem : (f : (S₊ 2 → coHomK 4)) → ∥ (λ x → 0ₖ 4) ≡ f ∥₂ --- -- truncLem f = --- -- sRec squash₂ --- -- (λ hs → sRec squash₂ --- -- (λ h → sRec squash₂ (λ m-b → sRec squash₂ (λ m-l → ∣ funExt (λ { north → h --- -- ; south → hs --- -- ; (merid base i) → m-b i --- -- ; (merid (loop j) i) k → m-l j i k} ) ∣₂) --- -- (final (f north) (f south) h hs (cong f (merid base)) m-b (λ j i → f (merid (loop j) i)))) --- -- (h-m-b (f north) (f south) h hs (cong f (merid base)))) --- -- (h (f north))) (h (f south)) --- -- where --- -- h : (x : coHomK 4) → ∥ 0ₖ 4 ≡ x ∥₂ --- -- h = coHomK-elim _ (λ _ → isOfHLevelSuc 3 (isOfHLevelSuc 2 squash₂)) ∣ refl ∣₂ - --- -- stupid : isConnected (1 +ℕ 4) (coHomK 4) --- -- stupid = isConnectedKn 3 - --- -- h-m-b : (x y : coHomK 4) → (p : 0ₖ 4 ≡ x) (q : 0ₖ 4 ≡ y) (z : x ≡ y) → ∥ PathP (λ i → 0ₖ 4 ≡ z i) p q ∥₂ --- -- h-m-b x y p q z = trElim {B = λ _ → ∥ PathP (λ i → 0ₖ 4 ≡ z i) p q ∥₂} (λ _ → squash₂) ∣_∣₂ --- -- (isConnectedPathP 2 {A = λ i → 0ₖ 4 ≡ z i} (isConnectedPath 3 (isConnectedSubtr 4 1 stupid) _ _) p q .fst) - --- -- final : (x y : coHomK 4) → (p : 0ₖ 4 ≡ x) (q : 0ₖ 4 ≡ y) (z : x ≡ y) (w : PathP (λ i → 0ₖ 4 ≡ z i) p q) (r : z ≡ z) --- -- → ∥ PathP (λ j → PathP (λ i → 0ₖ 4 ≡ r j i) p q) w w ∥₂ --- -- final x y p q z w r = --- -- trRec {B = ∥ PathP (λ j → PathP (λ i → 0ₖ 4 ≡ r j i) p q) w w ∥₂} --- -- squash₂ --- -- ∣_∣₂ --- -- (isConnectedPathP 2 {A = (λ j → PathP (λ i → 0ₖ 4 ≡ r j i) p q) } --- -- (isConnectedPathP 3 (isConnectedPath 4 (isConnectedKn 3) _ _) _ _) w w .fst) - --- -- truncLemHLev : (f : (S₊ 2 → coHomK 4)) → isContr ∥ (λ x → 0ₖ 4) ≡ f ∥₂ --- -- truncLemHLev f = --- -- isOfHLevelRetractFromIso 0 --- -- (compIso setTruncTrunc2Iso (invIso (PathIdTruncIso 2))) --- -- (isOfHLevelPath 0 (∣ (λ _ → 0ₖ 4) ∣ , (trElim (λ _ → isOfHLevelPath 3 (isOfHLevelTrunc 3) _ _) λ f → sRec (isOfHLevelTrunc 3 _ _) (cong ∣_∣) (truncLem f))) _ _) - --- -- FF : (f : (S₊ 2 → coHomK 4)) → ∥ (λ _ → 0ₖ 4) ≡ f ∥₂ → ((a : S₊ 3) → f (hehehe a) ≡ 0ₖ 4) → ∥ (S₊ 3 → 0ₖ 4 ≡ 0ₖ 4) ∥₂ --- -- FF f = sRec (isSetΠ (λ _ → squash₂)) λ id p → ∣ (λ x → funExt⁻ id (hehehe x) ∙ (p x)) ∣₂ - --- -- sll : (S₊ 3 → 0ₖ 4 ≡ 0ₖ 4) → CP7 → coHomK 4 --- -- sll f (inl x) = 0ₖ 4 --- -- sll f (inr x) = 0ₖ 4 --- -- sll f (push a i) = f a i - --- -- ssl* : ∥ (S₊ 3 → 0ₖ 4 ≡ 0ₖ 4) ∥₂ → ∥ (CP7 → coHomK 4) ∥₂ --- -- ssl* = sMap sll - --- -- ss3 : Iso ∥ (Σ[ f ∈ (S₊ 2 → coHomK 4) ] ((a : S₊ 3) → f (hehehe a) ≡ 0ₖ 4)) ∥₂ ∥ (S₊ 3 → 0ₖ 4 ≡ 0ₖ 4) ∥₂ --- -- fun ss3 = sRec squash₂ (uncurry (λ f → FF f (truncLem f))) --- -- inv ss3 = sMap λ f → (λ _ → 0ₖ 4) , f --- -- rightInv ss3 = sElim {!!} λ f → {!λ i → FF (λ _ → 0ₖ 4) (h i) f!} ∙ {!!} --- -- where --- -- h : truncLem (λ _ → 0ₖ _) ≡ ∣ refl ∣₂ --- -- h = {!cong ∣_∣₂ ?!} --- -- leftInv ss3 = {!!} - --- -- s : ∀ (f : CP7 → coHomK 4) → Σ[ P ∈ (S₊∙ 3 →∙ Ω (coHomK-ptd 4)) ] --- -- ∣ f ∣₂ ≡ ∣ sll (fst P) ∣₂ --- -- s f = (transport (λ i → Σ[ P ∈ (S₊∙ 3 →∙ Ω (coHomK-ptd 4)) ] --- -- Iso.leftInv myIso ∣ f ∣₂ i ≡ ∣ sll (fst P) ∣₂) --- -- ((sRec l (λ f → (λ x → f x ∙ sym (f north)) , rCancel (f north)) (Iso.fun myIso ∣ f ∣₂)) , --- -- mainLem (Iso.fun myIso ∣ f ∣₂))) --- -- where --- -- myIso : Iso _ _ --- -- myIso = compIso ss (compIso ss2 ss3) - --- -- ll : (f : Susp (Susp S¹) → 0ₖ 4 ≡ 0ₖ 4) → ∥ f north ≡ refl ∥₂ --- -- ll f = trRec squash₂ ∣_∣₂ --- -- (isConnectedPath 2 (isConnectedSubtr 3 1 (isConnectedPathKn 3 (0ₖ 4) (0ₖ 4))) (f north) refl .fst) - --- -- l : isSet (S₊∙ 3 →∙ Ω (coHomK-ptd 4)) --- -- l = subst isSet (λ i → S₊∙ 3 →∙ Kn≃ΩKn+1∙ {n = 3} i) --- -- (isOfHLevel↑∙' 0 2) - --- -- kaha : (f : _) → inv myIso ∣ f ∣₂ ≡ ∣ sll f ∣₂ --- -- kaha f = cong ∣_∣₂ (funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i) → refl}) - --- -- mainLem : (f : _) → inv myIso f ≡ --- -- ∣ sll (fst (sRec l (λ f₁ → (λ x → f₁ x ∙ (λ i → f₁ north (~ i))) , rCancel (f₁ (snd (S₊∙ 3)))) --- -- f)) ∣₂ --- -- mainLem = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) --- -- λ f → sRec (isOfHLevelPath 2 squash₂ _ _) --- -- (λ p → kaha f ∙ cong ∣_∣₂ (cong sll ((λ i x → rUnit (f x) i) ∙ λ i x → f x ∙ sym (p (~ i))))) --- -- (ll f) - --- -- genCP7 : coHom 4 CP7 --- -- genCP7 = ∣ (λ { (inl x) → 0ₖ 4 --- -- ; (inr x) → 0ₖ 4 --- -- ; (push a i) → Kn→ΩKn+1 3 ∣ a ∣ i}) ∣₂ - --- -- heheid : (a : _) → ∣ hehehe a ∣ ≡ 0ₖ (suc (suc zero)) --- -- heheid = sphereElim _ (λ _ → isOfHLevelTrunc 4 _ _) refl - --- -- gen2CP7 : coHom 4 CP7 --- -- gen2CP7 = ∣ (λ { (inl x) → _⌣ₖ_ {n = 2} {m = 2} ∣ x ∣ ∣ x ∣ --- -- ; (inr x) → 0ₖ 4 --- -- ; (push a i) → heheid a i ⌣ₖ heheid a i}) ∣₂ - - --- -- testS : gen2CP7 ≡ genCP7 --- -- testS = --- -- cong ∣_∣₂ --- -- (funExt λ { (inl north) → refl --- -- ; (inl south) → refl --- -- ; (inl (merid a i)) j → h a j i --- -- ; (inr x) → {!!} --- -- ; (push a i) → {!!}}) --- -- where --- -- h : (a : S¹) → cong₂ (_⌣ₖ_ {n = 2} {m = 2}) (cong ∣_∣ₕ (merid a)) (cong ∣_∣ₕ (merid a)) ≡ λ i → 0ₖ 4 --- -- h = {!!} - --- -- gen2 : coHom 2 CP7 --- -- gen2 = ∣ (λ { (inl x) → ∣ x ∣ ; (inr x) → 0ₖ 2 ; (push a i) → heheid a i}) ∣₂ - --- -- open import Cubical.HITs.S1 renaming (_·_ to _*_) --- -- hopfSurj' : (x : S₊ 2) → HopfSuspS¹ x --- -- hopfSurj' north = base --- -- hopfSurj' south = base --- -- hopfSurj' (merid a i) = {!!} --- -- where --- -- k : (a : _) → PathP (λ i → Glue S¹ (Border a i)) base base --- -- k a = toPathP ((λ i → a * base) ∙ {!!}) - --- -- -- hopfSurj : (x : S₊ 2) → Σ[ s ∈ S₊ 3 ] fst (JoinS¹S¹→TotalHopf (S³→joinS¹S¹ (inv IsoS³S3 s))) ≡ x --- -- -- hopfSurj x = (Iso.inv (IsoS³TotalHopf) (x , l x)) , cong fst (Iso.rightInv (IsoS³TotalHopf) (x , l x)) - --- -- -- gen2 : (p : (a : S₊ 3) → ∣ fst (JoinS¹S¹→TotalHopf (S³→joinS¹S¹ (inv IsoS³S3 a))) ∣ ≡ 0ₖ 2) --- -- -- → CP²' → coHomK 4 --- -- -- gen2 p (inl x) = _⌣ₖ_ {n = 2} {m = 2} ∣ x ∣ ∣ x ∣ --- -- -- gen2 p (inr x) = 0ₖ _ --- -- -- gen2 p (push a i) = p a i ⌣ₖ p a i - --- -- -- gen≡ : (p : _) (a : _) → gen1 a ≡ gen2 p a --- -- -- gen≡ p (inl x) = (λ i → p (fst (hopfSurj x)) (~ i) ⌣ₖ p (fst (hopfSurj x)) (~ i)) --- -- -- ∙ λ i → _⌣ₖ_ {n = 2} {m = 2} ∣ hopfSurj x .snd i ∣ ∣ hopfSurj x .snd i ∣ --- -- -- gen≡ p (inr x) = refl --- -- -- gen≡ p (push a i) = {!!} - - --- -- -- -- CP4→CP2''' : (f : S₊∙ 3 →∙ S₊∙ 2) → ∥ (S₊∙ 3 →∙ (Ω (coHomK-ptd 4))) ∥₂ → coHom 2 (CP' f) --- -- -- -- CP4→CP2''' f = sRec squash₂ λ f → ∣ (λ { (inl x) → ΩKn+1→Kn 2 (cong (ΩKn+1→Kn 3) (h x f)) --- -- -- -- ; (inr x) → 0ₖ 2 --- -- -- -- ; (push a i) → {!!}}) ∣₂ --- -- -- -- where --- -- -- -- h : (x : S₊ 2) (f : ((S₊∙ 3 →∙ (Ω (coHomK-ptd 4))))) → typ ((Ω^ 2) (coHomK-ptd 4)) --- -- -- -- h x f i j = --- -- -- -- hcomp (λ k → λ { (i = i0) → 0ₖ 4 --- -- -- -- ; (i = i1) → 0ₖ 4 --- -- -- -- ; (j = i0) → snd f k i --- -- -- -- ; (j = i1) → snd f k i}) --- -- -- -- (fst f ((merid x ∙ sym (merid north)) j) i) - --- -- -- -- CP4→CP2' : (f : S₊∙ 3 →∙ S₊∙ 2) (x : coHom 2 (CP' f)) → coHom 4 (CP' f) → coHom 2 (CP' f) --- -- -- -- CP4→CP2' f = --- -- -- -- sRec2 squash₂ --- -- -- -- λ e g → ∣ (λ { (inl x) → ΩKn+1→Kn 2 ({!g!} ∙∙ cong (ΩKn+1→Kn 3) {!λ i j → g (push (merid x i)) j!} ∙∙ {!!}) --- -- -- -- ; (inr x) → {!!} --- -- -- -- ; (push a i) → {!!}}) ∣₂ --- -- -- -- where --- -- -- -- h : (g : CP' f → coHomK 4) → g (inr tt) ≡ 0ₖ _ → g ∘ inl ≡ (λ x → 0ₖ _) → S₊ 2 → typ ((Ω^ 2) (coHomK-ptd 4)) --- -- -- -- h g gpd idf x i j = --- -- -- -- hcomp (λ k → λ { (i = i0) → rCancelₖ 4 (g (push north j)) k --- -- -- -- ; (i = i1) → rCancelₖ 4 (g (push south j)) k --- -- -- -- ; (j = i0) → {!idf k ((fst f (merid x i))) -ₖ idf k ((fst f (merid north i)))!} -- idf k ((fst f (merid x i))) -ₖ gpd k --- -- -- -- ; (j = i1) → {!!}}) --- -- -- -- (g (push (merid x i) j) -ₖ g (push (merid north i) j)) - - --- -- contrfstIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : A → Type ℓ'} → (s : isContr A) → Iso (Σ A B) (B (s .fst)) --- -- fun (contrfstIso {B = B} s) (a , b) = subst B (sym (snd s a)) b --- -- inv (contrfstIso s) b = (fst s) , b --- -- rightInv (contrfstIso {B = B} s) b = (λ i → subst B (sym (h i)) b) ∙ transportRefl b --- -- where --- -- h : snd s (fst s) ≡ refl --- -- h = isProp→isSet (isContr→isProp s) _ _ _ _ --- -- leftInv (contrfstIso{B = B} s) (a , b) = --- -- ΣPathP ((snd s a) , λ i → transp (λ j → B (snd s a (i ∨ ~ j))) i b) - --- -- miniFib : hLevelTrunc 3 (S₊ 2 → coHomK 4) → Type --- -- miniFib x = trRec {B = TypeOfHLevel ℓ-zero 2} (isOfHLevelTypeOfHLevel 2) (λ f → (∥ ((x : S₊ 3) → f (hehehe x) ≡ 0ₖ 4) ∥₂) , squash₂) x .fst - --- -- l3 : isContr (hLevelTrunc 3 (S₊ 2 → coHomK 4)) --- -- l3 = {!!} - --- -- l7 : {!!} --- -- l7 = {!!} + h : GroupHom (coHomGr (2 +ℕ n) TotalHopf) (coHomGr (3 +ℕ n) CP²) + h = M.d (2 +ℕ n) + + propᵣ : isProp (fst (×coHomGr (3 +ℕ n) (Susp S¹) Unit)) + propᵣ = + isPropΣ + (isOfHLevelRetractFromIso 1 + (fst (Hⁿ-Sᵐ≅0 (2 +ℕ n) 1 λ p → ⊥-rec (snotz (cong predℕ p)))) isPropUnit) + λ _ → isContr→isProp (isContrHⁿ-Unit _) + + propₗ : isProp (coHom (2 +ℕ n) TotalHopf) + propₗ = subst (λ x → isProp (coHom (2 +ℕ n) x)) (isoToPath IsoS³TotalHopf) + (isOfHLevelRetractFromIso 1 + (fst (Hⁿ-Sᵐ≅0 (suc n) 2 λ q → p (cong predℕ q))) isPropUnit) + + inIm : (x : coHom (3 +ℕ n) CP²) → isInIm h x + inIm x = M.Ker-i⊂Im-d (2 +ℕ n) x (propᵣ _ _) + + main : (x : coHom (3 +ℕ n) CP²) → x ≡ 0ₕ _ + main x = + pRec (squash₂ _ _) + (uncurry (λ f p → sym p ∙∙ cong (h .fst) (propₗ f (0ₕ _)) ∙∙ pres1 (snd h))) + (inIm x) + +-- All trivial groups: +Hⁿ-CP²≅0 : (n : ℕ) → ¬ suc n ≡ 2 → ¬ suc n ≡ 4 + → GroupIso (coHomGr (suc n) CP²) UnitGroup +Hⁿ-CP²≅0 zero p q = H¹-CP²≅0 +Hⁿ-CP²≅0 (suc zero) p q = ⊥-rec (p refl) +Hⁿ-CP²≅0 (suc (suc zero)) p q = Hⁿ-CP²≅0-higher 0 λ p → snotz (sym p) +Hⁿ-CP²≅0 (suc (suc (suc zero))) p q = ⊥-rec (q refl) +Hⁿ-CP²≅0 (suc (suc (suc (suc n)))) p q = + Hⁿ-CP²≅0-higher (suc (suc n)) + λ p → snotz (cong predℕ p) + +-- Another Brunerie number +ℤ→HⁿCP²→ℤ : ℤ → ℤ +ℤ→HⁿCP²→ℤ x = + Iso.fun (fst H⁴CP²≅ℤ) + (Iso.inv (fst H²CP²≅ℤ) x ⌣ Iso.inv (fst H²CP²≅ℤ) x) + +brunerie2 : ℤ +brunerie2 = ℤ→HⁿCP²→ℤ 1 + +{- +|brunerie2|≡1 : abs (ℤ→HⁿCP²→ℤ 1) ≡ 1 +|brunerie2|≡1 = refl +-} diff --git a/Cubical/ZCohomology/Gysin2.agda b/Cubical/ZCohomology/Gysin2.agda deleted file mode 100644 index 48293e586..000000000 --- a/Cubical/ZCohomology/Gysin2.agda +++ /dev/null @@ -1,61 +0,0 @@ -{-# OPTIONS --safe --experimental-lossy-unification #-} -module Cubical.ZCohomology.Gysin2 where - -open import Cubical.ZCohomology.Base -open import Cubical.ZCohomology.Groups.Connected -open import Cubical.ZCohomology.GroupStructure -open import Cubical.ZCohomology.Groups.Wedge -open import Cubical.ZCohomology.Properties -open import Cubical.ZCohomology.MayerVietorisUnreduced -open import Cubical.ZCohomology.Groups.Unit -open import Cubical.ZCohomology.Groups.Sn -open import Cubical.ZCohomology.RingStructure.CupProduct -open import Cubical.ZCohomology.RingStructure.RingLaws -open import Cubical.ZCohomology.RingStructure.GradedCommutativity - -open import Cubical.Data.Sum -open import Cubical.Relation.Nullary - -open import Cubical.Functions.Embedding - -open import Cubical.Data.Fin - -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Path -open import Cubical.Foundations.Transport -open import Cubical.Foundations.Function hiding (case_of_) -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.GroupoidLaws -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Pointed.Homogeneous - -open import Cubical.Foundations.Univalence - -open import Cubical.Relation.Nullary - -open import Cubical.Data.Empty renaming (rec to ⊥-rec) -open import Cubical.Data.Sigma -open import Cubical.Data.Int hiding (_+'_) -open import Cubical.Data.Nat renaming (_+_ to _+ℕ_ ; _·_ to _·ℕ_) -open import Cubical.Data.Nat.Order -open import Cubical.Data.Unit -open import Cubical.Data.Bool -open import Cubical.Algebra.Group - renaming (ℤ to ℤGroup ; Unit to UnitGroup) hiding (Bool) - -open import Cubical.HITs.Pushout.Flattening -open import Cubical.Homotopy.Connected -open import Cubical.Homotopy.EilenbergSteenrod -open import Cubical.HITs.Pushout -open import Cubical.HITs.Sn -open import Cubical.HITs.Susp -open import Cubical.HITs.S1 renaming (_·_ to _*_) -open import Cubical.HITs.Truncation - renaming (rec to trRec ; elim to trElim ; elim2 to trElim2 ; map to trMap) -open import Cubical.HITs.SetTruncation - renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; elim3 to sElim3 ; map to sMap) -open import Cubical.HITs.PropositionalTruncation - renaming (rec to pRec ; elim to pElim) diff --git a/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda b/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda index 8b20487de..68ea383b7 100644 --- a/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda +++ b/Cubical/ZCohomology/RingStructure/GradedCommutativity.agda @@ -22,6 +22,7 @@ open import Cubical.Foundations.HLevels open import Cubical.Foundations.Function open import Cubical.Foundations.Transport open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Pointed open import Cubical.Foundations.Pointed.Homogeneous open import Cubical.Foundations.GroupoidLaws hiding (assoc) @@ -151,8 +152,6 @@ private ; south → cong ∣_∣ₕ (merid (ptSn _)) ; (merid a i) k → ∣ compPath-filler (merid a) (sym (merid (ptSn _))) (~ k) i ∣ₕ} - -open import Cubical.Foundations.Isomorphism -ₖ-gen² : {k : ℕ} (n m : ℕ) (p : isEvenT n ⊎ isOddT n) (q : isEvenT m ⊎ isOddT m) From f90b33452d08692b28dd46821d6c2bcecee7eb51 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Tue, 2 Nov 2021 14:53:55 +0100 Subject: [PATCH 80/89] stuff --- Cubical/ZCohomology/Gysin.agda | 19 ++++--------------- 1 file changed, 4 insertions(+), 15 deletions(-) diff --git a/Cubical/ZCohomology/Gysin.agda b/Cubical/ZCohomology/Gysin.agda index 21cfabc4d..05d4dd311 100644 --- a/Cubical/ZCohomology/Gysin.agda +++ b/Cubical/ZCohomology/Gysin.agda @@ -74,13 +74,6 @@ open import Cubical.Algebra.AbGroup open import Cubical.Homotopy.Loopspace -open import Cubical.HITs.Join - -open import Cubical.Homotopy.Hopf - -open import Cubical.HITs.SetQuotients renaming (_/_ to _/'_) --- open import Cubical.ZCohomology.Gysin - -- There seems to be some problems with the termination checker. -- Spelling out integer induction with 3 base cases like this -- solves the issue. @@ -266,19 +259,15 @@ module g-base where →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ y → -ₖ-gen² i n _ _ (fst x y)) - rCancel'' : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) - → sym p ∙∙ refl ∙∙ p ≡ refl - rCancel'' p = (λ j → (λ i → p (~ i ∨ j)) ∙∙ refl ∙∙ λ i → p (i ∨ j)) - ∙ sym (rUnit refl) - transpPres0ₖ : ∀ {k m : ℕ} (p : k ≡ m) → subst coHomK p (0ₖ k) ≡ 0ₖ m transpPres0ₖ {k = k} = J (λ m p → subst coHomK p (0ₖ k) ≡ 0ₖ m) (transportRefl _) - -- There will be some index swapping going on. We statet this explicitly, since we will + -- There will be some index swapping going on. We state this explicitly, since we will -- need to trace the maps later indexSwap : (n : ℕ) (i : ℕ) - → (S₊∙ n →∙ coHomK-ptd (n +' i)) ≃ (S₊∙ n →∙ coHomK-ptd (i +' n)) + → (S₊∙ n →∙ coHomK-ptd (n +' i)) + ≃ (S₊∙ n →∙ coHomK-ptd (i +' n)) indexSwap n i = isoToEquiv (iso (λ f → (λ x → subst coHomK (+'-comm n i) (fst f x)) , cong (subst coHomK (+'-comm n i)) (snd f) ∙ transpPres0ₖ (+'-comm n i)) @@ -311,7 +300,7 @@ module g-base where suspKn-Iso-fun∙ (suc n) m (f , p) = cong (ΩKn+1→Kn m) (cong (sym p ∙∙_∙∙ p) (cong (cong f) (rCancel (merid (ptSn _))))) - ∙∙ cong (ΩKn+1→Kn m) (rCancel'' p) + ∙∙ cong (ΩKn+1→Kn m) (∙∙lCancel p) ∙∙ ΩKn+1→Kn-refl m suspKn-Iso-inv : (n m : ℕ) → (S₊∙ n →∙ coHomK-ptd m) From 98fa19bf05f4e2c6c46b2f035c46ffae34d3bb06 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Thu, 4 Nov 2021 03:22:00 +0100 Subject: [PATCH 81/89] stuff --- Cubical/Algebra/Group/Instances/Unit.agda | 41 +++++ Cubical/Algebra/Group/ZModule.agda | 11 +- Cubical/Homotopy/Freudenthal.agda | 12 +- Cubical/Homotopy/Group/S3.agda | 183 ++++++++++++++++++++++ Cubical/Homotopy/Hopf.agda | 27 ++-- Cubical/Homotopy/HopfInvariant/Base.agda | 40 ----- 6 files changed, 248 insertions(+), 66 deletions(-) create mode 100644 Cubical/Homotopy/Group/S3.agda diff --git a/Cubical/Algebra/Group/Instances/Unit.agda b/Cubical/Algebra/Group/Instances/Unit.agda index caceb6eac..0605898f3 100644 --- a/Cubical/Algebra/Group/Instances/Unit.agda +++ b/Cubical/Algebra/Group/Instances/Unit.agda @@ -4,11 +4,13 @@ module Cubical.Algebra.Group.Instances.Unit where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Structure +open import Cubical.Foundations.Equiv open import Cubical.Data.Unit renaming (Unit to UnitType) open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Group.DirProd open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec) open GroupStr open IsGroupHom @@ -54,3 +56,42 @@ snd (contrGroupIsoUnit contr) = makeIsGroupHom λ _ _ → refl contrGroupEquivUnit : {G : Group ℓ} → isContr ⟨ G ⟩ → GroupEquiv G Unit contrGroupEquivUnit contr = GroupIso→GroupEquiv (contrGroupIsoUnit contr) + +SES→isEquiv : ∀ {ℓ ℓ'} {L R : Group ℓ-zero} + → {G : Group ℓ} {H : Group ℓ'} + → Unit ≡ L + → Unit ≡ R + → (lhom : GroupHom L G) (midhom : GroupHom G H) (rhom : GroupHom H R) + → ((x : _) → isInKer midhom x → isInIm lhom x) + → ((x : _) → isInKer rhom x → isInIm midhom x) + → isEquiv (fst midhom) +SES→isEquiv {R = R} {G = G} {H = H} = + J (λ L _ → Unit ≡ R → + (lhom : GroupHom L G) (midhom : GroupHom G H) + (rhom : GroupHom H R) → + ((x : fst G) → isInKer midhom x → isInIm lhom x) → + ((x : fst H) → isInKer rhom x → isInIm midhom x) → + isEquiv (fst midhom)) + ((J (λ R _ → (lhom : GroupHom Unit G) (midhom : GroupHom G H) + (rhom : GroupHom H R) → + ((x : fst G) → isInKer midhom x → isInIm lhom x) → + ((x : _) → isInKer rhom x → isInIm midhom x) → + isEquiv (fst midhom)) + main)) + where + main : (lhom : GroupHom Unit G) (midhom : GroupHom G H) + (rhom : GroupHom H Unit) → + ((x : fst G) → isInKer midhom x → isInIm lhom x) → + ((x : fst H) → isInKer rhom x → isInIm midhom x) → + isEquiv (fst midhom) + main lhom midhom rhom lexact rexact = + BijectionIsoToGroupEquiv {G = G} {H = H} + bijIso' .fst .snd + where + bijIso' : BijectionIso G H + BijectionIso.fun bijIso' = midhom + BijectionIso.inj bijIso' x inker = + pRec (GroupStr.is-set (snd G) _ _) + (λ s → sym (snd s) ∙ IsGroupHom.pres1 (snd lhom)) + (lexact _ inker) + BijectionIso.surj bijIso' x = rexact x refl diff --git a/Cubical/Algebra/Group/ZModule.agda b/Cubical/Algebra/Group/ZModule.agda index d2360b843..9958abafb 100644 --- a/Cubical/Algebra/Group/ZModule.agda +++ b/Cubical/Algebra/Group/ZModule.agda @@ -14,11 +14,6 @@ open import Cubical.Data.Nat open import Cubical.Data.Empty renaming (rec to ⊥-rec) open import Cubical.Data.Sum -open import Cubical.Displayed.Base -open import Cubical.Displayed.Auto -open import Cubical.Displayed.Record -open import Cubical.Displayed.Universe - open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Group.Properties open import Cubical.Algebra.Group.Morphisms @@ -216,11 +211,11 @@ private intLem₂ : (n x : ℕ) → Σ[ a ∈ ℕ ] ((pos (suc x)) * pos (suc (suc n)) ≡ pos (suc (suc a))) intLem₂ n zero = n , refl - intLem₂ n (suc x) = h3 _ _ (intLem₂ n x) + intLem₂ n (suc x) = lem _ _ (intLem₂ n x) where - h3 : (x : ℤ) (n : ℕ) → Σ[ a ∈ ℕ ] (x ≡ pos (suc (suc a))) + lem : (x : ℤ) (n : ℕ) → Σ[ a ∈ ℕ ] (x ≡ pos (suc (suc a))) → Σ[ a ∈ ℕ ] pos n +ℤ x ≡ pos (suc (suc a)) - h3 x n (a , p) = n +ℕ a + lem x n (a , p) = n +ℕ a , cong (pos n +ℤ_) p ∙ cong sucℤ (sucℤ+pos a (pos n)) ∙ sucℤ+pos a (pos (suc n)) ∙ (sym (pos+ (suc (suc n)) a)) diff --git a/Cubical/Homotopy/Freudenthal.agda b/Cubical/Homotopy/Freudenthal.agda index d471c16f9..8e75fdf2e 100644 --- a/Cubical/Homotopy/Freudenthal.agda +++ b/Cubical/Homotopy/Freudenthal.agda @@ -153,14 +153,14 @@ suspMapΩ-connected : ∀ {ℓ} (n : HLevel) (m : ℕ) {A : Pointed ℓ} suspMapΩ-connected n zero {A = A} connA = isConnectedσ n connA suspMapΩ-connected n (suc m) {A = A} connA with ((n + suc n) ≟ m) ... | (lt p) = subst (λ x → isConnectedFun x (suspMapΩ {A = A} (suc m))) - (sym (h _ m p)) + (sym (help _ m p)) λ b → tt* , (λ {tt* → refl}) where - h : (n m : ℕ) → n < m → (n ∸ m) ≡ 0 - h zero zero p = refl - h (suc n) zero p = ⊥-rec (¬-<-zero p) - h zero (suc m) p = refl - h (suc n) (suc m) p = h n m (pred-≤-pred p) + help : (n m : ℕ) → n < m → (n ∸ m) ≡ 0 + help zero zero p = refl + help (suc n) zero p = ⊥-rec (¬-<-zero p) + help zero (suc m) p = refl + help (suc n) (suc m) p = help n m (pred-≤-pred p) ... | (eq q) = subst (λ x → isConnectedFun x (suspMapΩ {A = A} (suc m))) (sym (help m) ∙ cong (_∸ m) (sym q)) λ b → tt* , (λ {tt* → refl}) diff --git a/Cubical/Homotopy/Group/S3.agda b/Cubical/Homotopy/Group/S3.agda new file mode 100644 index 000000000..65f015bf2 --- /dev/null +++ b/Cubical/Homotopy/Group/S3.agda @@ -0,0 +1,183 @@ +{-# OPTIONS --safe --experimental-lossy-unification #-} +module Cubical.Homotopy.Group.S3 where + +{- +This file contains a summary of what remains for π₄S³≅ℤ/2 to be proved. +See the module π₄S³ at the end of this file. +-} + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function + +open import Cubical.Data.Nat +open import Cubical.Data.Sum +open import Cubical.Data.Int + renaming (_·_ to _·ℤ_ ; _+_ to _+ℤ_) + +open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.Group.Properties +open import Cubical.Homotopy.HopfInvariant.Base +open import Cubical.Homotopy.HopfInvariant.Homomorphism +open import Cubical.Homotopy.HopfInvariant.HopfMap +open import Cubical.Algebra.Group.Instances.IntMod +open import Cubical.Foundations.Isomorphism + +open import Cubical.HITs.Sn +open import Cubical.HITs.SetTruncation + +open import Cubical.Algebra.Group + renaming (ℤ to ℤGroup ; Bool to BoolGroup ; Unit to UnitGroup) +open import Cubical.Algebra.Group.ZModule + + +-- Some type abbreviations (unproved results) +∃Whitehead : ∀ {ℓ} {X : Pointed ℓ} (n m : ℕ) → Type _ +∃Whitehead {X = X} n m = GroupHom (DirProd (π'Gr n X) (π'Gr m X)) + (π'Gr (n + m) X) + +π₃S²-gen : Type +π₃S²-gen = gen₁-by (π'Gr 2 (S₊∙ 2)) ∣ HopfMap ∣₂ + +π₄S³≅ℤ/something : (ϕ : ∃Whitehead {X = S₊∙ 2} 1 1) + → GroupEquiv ℤGroup (π'Gr 2 (S₊∙ 2)) + → Type +π₄S³≅ℤ/something ϕ eq = + GroupIso (π'Gr 3 (S₊∙ 3)) + (ℤ/ abs (invEq (fst eq) + (fst ϕ (∣ idfun∙ _ ∣₂ , ∣ idfun∙ _ ∣₂)))) + +miniLem₁ : Type +miniLem₁ = (g : ℤ) → gen₁-by ℤGroup g → (g ≡ 1) ⊎ (g ≡ -1) + +miniLem₂ : Type +miniLem₂ = (ϕ : GroupEquiv ℤGroup ℤGroup) (g : ℤ) + → (abs g ≡ abs (fst (fst ϕ) g)) + +-- some minor group lemmas +groupLem-help : miniLem₁ → (g : ℤ) → + gen₁-by ℤGroup g → + (ϕ : GroupHom ℤGroup ℤGroup) → + (fst ϕ g ≡ pos 1) ⊎ (fst ϕ g ≡ negsuc 0) + → isEquiv (fst ϕ) +groupLem-help grlem1 g gen ϕ = main (grlem1 g gen) + where + isEquiv- : isEquiv (-_) + isEquiv- = isoToIsEquiv (iso -_ -_ -Involutive -Involutive) + + lem : fst ϕ (pos 1) ≡ pos 1 → fst ϕ ≡ idfun _ + lem p = funExt lem2 + where + lem₁ : (x₁ : ℕ) → fst ϕ (pos x₁) ≡ idfun ℤ (pos x₁) + lem₁ zero = IsGroupHom.pres1 (snd ϕ) + lem₁ (suc zero) = p + lem₁ (suc (suc n)) = + IsGroupHom.pres· (snd ϕ) (pos (suc n)) 1 + ∙ cong₂ _+ℤ_ (lem₁ (suc n)) p + + lem2 : (x₁ : ℤ) → fst ϕ x₁ ≡ idfun ℤ x₁ + lem2 (pos n) = lem₁ n + lem2 (negsuc zero) = + IsGroupHom.presinv (snd ϕ) 1 ∙ cong (λ x → pos 0 - x) p + lem2 (negsuc (suc n)) = + (cong (fst ϕ) (sym (+Comm (pos 0) (negsuc (suc n)))) + ∙ IsGroupHom.presinv (snd ϕ) (pos (suc (suc n)))) + ∙∙ +Comm (pos 0) _ + ∙∙ cong (-_) (lem₁ (suc (suc n))) + + lem₂ : fst ϕ (negsuc 0) ≡ pos 1 → fst ϕ ≡ -_ + lem₂ p = funExt lem2 + where + s = IsGroupHom.presinv (snd ϕ) (negsuc 0) + ∙∙ +Comm (pos 0) _ + ∙∙ cong -_ p + + lem2 : (n : ℤ) → fst ϕ n ≡ - n + lem2 (pos zero) = IsGroupHom.pres1 (snd ϕ) + lem2 (pos (suc zero)) = s + lem2 (pos (suc (suc n))) = + IsGroupHom.pres· (snd ϕ) (pos (suc n)) 1 + ∙ cong₂ _+ℤ_ (lem2 (pos (suc n))) s + lem2 (negsuc zero) = p + lem2 (negsuc (suc n)) = + IsGroupHom.pres· (snd ϕ) (negsuc n) (negsuc 0) + ∙ cong₂ _+ℤ_ (lem2 (negsuc n)) p + + main : (g ≡ pos 1) ⊎ (g ≡ negsuc 0) + → (fst ϕ g ≡ pos 1) ⊎ (fst ϕ g ≡ negsuc 0) + → isEquiv (fst ϕ) + main (inl p) = + J (λ g p → (fst ϕ g ≡ pos 1) + ⊎ (fst ϕ g ≡ negsuc 0) → isEquiv (fst ϕ)) + (λ { (inl x) → subst isEquiv (sym (lem x)) (snd (idEquiv _)) + ; (inr x) → subst isEquiv + (sym (lem₂ (IsGroupHom.presinv (snd ϕ) (pos 1) + ∙ (cong (λ x → pos 0 - x) x)))) + isEquiv- }) + (sym p) + main (inr p) = + J (λ g p → (fst ϕ g ≡ pos 1) + ⊎ (fst ϕ g ≡ negsuc 0) → isEquiv (fst ϕ)) + (λ { (inl x) → subst isEquiv (sym (lem₂ x)) isEquiv- + ; (inr x) → subst isEquiv + (sym (lem ( + IsGroupHom.presinv (snd ϕ) (negsuc 0) + ∙ cong (λ x → pos 0 - x) x))) + (snd (idEquiv _))}) + (sym p) + +groupLem : {G : Group₀} + → miniLem₁ + → GroupEquiv ℤGroup G + → (g : fst G) + → gen₁-by G g + → (ϕ : GroupHom G ℤGroup) + → (fst ϕ g ≡ 1) ⊎ (fst ϕ g ≡ -1) + → isEquiv (fst ϕ) +groupLem {G = G} s = + JGroupEquiv + (λ G _ → (g : fst G) + → gen₁-by G g + → (ϕ : GroupHom G ℤGroup) + → (fst ϕ g ≡ 1) ⊎ (fst ϕ g ≡ -1) + → isEquiv (fst ϕ)) + (groupLem-help s) + +-- summary +module π₄S³ + (mini-lem₁ : miniLem₁) + (mini-lem₂ : miniLem₂) + (ℤ≅π₃S² : GroupEquiv ℤGroup (π'Gr 2 (S₊∙ 2))) + (gen-by-HopfMap : π₃S²-gen) + (whitehead : ∃Whitehead {X = S₊∙ 2} 1 1) + (π₄S³≅ℤ/whitehead : π₄S³≅ℤ/something whitehead ℤ≅π₃S²) + (hopfWhitehead : + abs (HopfInvariant-π' 0 + (fst whitehead (∣ idfun∙ _ ∣₂ , ∣ idfun∙ _ ∣₂))) + ≡ 2) + where + π₄S³ = π'Gr 3 (S₊∙ 3) + + hopfInvariantEquiv : GroupEquiv (π'Gr 2 (S₊∙ 2)) ℤGroup + fst (fst hopfInvariantEquiv) = HopfInvariant-π' 0 + snd (fst hopfInvariantEquiv) = + groupLem mini-lem₁ ℤ≅π₃S² ∣ HopfMap ∣₂ + gen-by-HopfMap + (GroupHom-HopfInvariant-π' 0) + (abs→⊎ _ _ HopfInvariant-HopfMap) + snd hopfInvariantEquiv = snd (GroupHom-HopfInvariant-π' 0) + + lem : ∀ {G : Group₀} (ϕ ψ : GroupEquiv ℤGroup G) (g : fst G) + → abs (invEq (fst ϕ) g) ≡ abs (invEq (fst ψ) g) + lem = + JGroupEquiv + (λ G ϕ → (ψ : GroupEquiv ℤGroup G) (g : fst G) + → abs (invEq (fst ϕ) g) ≡ abs (invEq (fst ψ) g)) + λ ψ → mini-lem₂ (invGroupEquiv ψ) + + main : GroupIso π₄S³ (ℤ/ 2) + main = subst (GroupIso π₄S³) + (cong (ℤ/_) (lem ℤ≅π₃S² (invGroupEquiv (hopfInvariantEquiv)) _ + ∙ hopfWhitehead)) + π₄S³≅ℤ/whitehead diff --git a/Cubical/Homotopy/Hopf.agda b/Cubical/Homotopy/Hopf.agda index 0353f8dff..4b2db2b93 100644 --- a/Cubical/Homotopy/Hopf.agda +++ b/Cubical/Homotopy/Hopf.agda @@ -41,7 +41,8 @@ module Hopf {ℓ : Level} {A : Pointed ℓ} {e : HSpace A} isEquiv-μ : (x : typ A) → isEquiv (λ z → (μ e z x)) isEquiv-μ x = pRec (isPropIsEquiv _) (J (λ x _ → isEquiv (λ z → μ e z x)) - (subst isEquiv (funExt (λ z → sym (μᵣ e z))) (idIsEquiv (typ A)))) + (subst isEquiv (funExt (λ z → sym (μᵣ e z))) + (idIsEquiv (typ A)))) (conA (pt A) x) isEquiv-μ' : (x : typ A) → isEquiv (μ e x) @@ -310,13 +311,13 @@ module Hopf {ℓ : Level} {A : Pointed ℓ} {e : HSpace A} → inr (secEq (_ , Push→TotalSpaceHopf-equiv x) y k)}) (push (x , invEq (_ , Push→TotalSpaceHopf-equiv x) y) i) - zz : Iso TotalSpacePush²' + IsoTotalSpacePush²'ΣPush : Iso TotalSpacePush²' (Pushout {A = typ A × Σ (Susp (typ A)) Hopf} fst snd) - fun zz = F - inv zz = G - rightInv zz (inl x) = refl - rightInv zz (inr x) = refl - rightInv zz (push (x , y) i) j = + fun IsoTotalSpacePush²'ΣPush = F + inv IsoTotalSpacePush²'ΣPush = G + rightInv IsoTotalSpacePush²'ΣPush (inl x) = refl + rightInv IsoTotalSpacePush²'ΣPush (inr x) = refl + rightInv IsoTotalSpacePush²'ΣPush (push (x , y) i) j = hcomp (λ k → λ { (i = i0) → inl x ; (i = i1) → inr (secEq (_ , Push→TotalSpaceHopf-equiv x) y k) @@ -334,9 +335,9 @@ module Hopf {ℓ : Level} {A : Pointed ℓ} {e : HSpace A} , (secEq (_ , Push→TotalSpaceHopf-equiv x) y k)) i}) (push (x , (secEq (_ , Push→TotalSpaceHopf-equiv x) y i0)) i) - leftInv zz (inl x) = refl - leftInv zz (inr x) = refl - leftInv zz (push (x , y) i) j = + leftInv IsoTotalSpacePush²'ΣPush (inl x) = refl + leftInv IsoTotalSpacePush²'ΣPush (inr x) = refl + leftInv IsoTotalSpacePush²'ΣPush (push (x , y) i) j = hcomp (λ k → λ { (i = i0) → inl x ; (i = i1) → inr (secEq (Push→TotalSpaceHopf x , Push→TotalSpaceHopf-equiv x) @@ -358,5 +359,7 @@ module Hopf {ℓ : Level} {A : Pointed ℓ} {e : HSpace A} joinIso₂ : Iso TotalSpacePush² (join (typ A) (join (typ A) (typ A))) joinIso₂ = compIso IsoTotalSpacePush²TotalSpacePush²' - (compIso zz (compIso (equivToIso (joinPushout≃join _ _)) - (pathToIso (cong (join (typ A)) (isoToPath IsoTotalSpaceJoin))))) + (compIso IsoTotalSpacePush²'ΣPush + (compIso (equivToIso (joinPushout≃join _ _)) + (pathToIso (cong (join (typ A)) + (isoToPath IsoTotalSpaceJoin))))) diff --git a/Cubical/Homotopy/HopfInvariant/Base.agda b/Cubical/Homotopy/HopfInvariant/Base.agda index 7236ba8c5..dc2cde51b 100644 --- a/Cubical/Homotopy/HopfInvariant/Base.agda +++ b/Cubical/Homotopy/HopfInvariant/Base.agda @@ -59,46 +59,6 @@ GroupIsoUnitGroup→isContr : {G : Group ℓ-zero} GroupIsoUnitGroup→isContr is = isOfHLevelRetractFromIso 0 (invIso (fst is)) isContrUnit --- move to group stuff -SES→isEquiv : ∀ {ℓ ℓ'} {L R : Group ℓ-zero} - → {G : Group ℓ} {H : Group ℓ'} - → UnitGroup ≡ L - → UnitGroup ≡ R - → (lhom : GroupHom L G) (midhom : GroupHom G H) (rhom : GroupHom H R) - → ((x : _) → isInKer midhom x → isInIm lhom x) - → ((x : _) → isInKer rhom x → isInIm midhom x) - → isEquiv (fst midhom) -SES→isEquiv {R = R} {G = G} {H = H} = - J (λ L _ → UnitGroup ≡ R → - (lhom : GroupHom L G) (midhom : GroupHom G H) - (rhom : GroupHom H R) → - ((x : fst G) → isInKer midhom x → isInIm lhom x) → - ((x : fst H) → isInKer rhom x → isInIm midhom x) → - isEquiv (fst midhom)) - ((J (λ R _ → (lhom : GroupHom UnitGroup G) (midhom : GroupHom G H) - (rhom : GroupHom H R) → - ((x : fst G) → isInKer midhom x → isInIm lhom x) → - ((x : _) → isInKer rhom x → isInIm midhom x) → - isEquiv (fst midhom)) - main)) - where - main : (lhom : GroupHom UnitGroup G) (midhom : GroupHom G H) - (rhom : GroupHom H UnitGroup) → - ((x : fst G) → isInKer midhom x → isInIm lhom x) → - ((x : fst H) → isInKer rhom x → isInIm midhom x) → - isEquiv (fst midhom) - main lhom midhom rhom lexact rexact = - BijectionIsoToGroupEquiv {G = G} {H = H} - bijIso' .fst .snd - where - bijIso' : BijectionIso G H - BijectionIso.fun bijIso' = midhom - BijectionIso.inj bijIso' x inker = - pRec (GroupStr.is-set (snd G) _ _) - (λ s → sym (snd s) ∙ IsGroupHom.pres1 (snd lhom)) - (lexact _ inker) - BijectionIso.surj bijIso' x = rexact x refl - -- The pushout of the hopf invariant HopfInvariantPush : (n : ℕ) → (f : S₊ (3 +ℕ n +ℕ n) → S₊ (2 +ℕ n)) From e372454aaafc4152ba9f04494ea4383b41cfc06b Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Thu, 4 Nov 2021 11:01:41 +0100 Subject: [PATCH 82/89] minor --- Cubical/Homotopy/Group/Properties.agda | 29 +++++++++++++------------- Cubical/Homotopy/HSpace.agda | 1 - 2 files changed, 15 insertions(+), 15 deletions(-) diff --git a/Cubical/Homotopy/Group/Properties.agda b/Cubical/Homotopy/Group/Properties.agda index 1de45df29..26058f02d 100644 --- a/Cubical/Homotopy/Group/Properties.agda +++ b/Cubical/Homotopy/Group/Properties.agda @@ -85,7 +85,7 @@ isHom-lMap zero p q = isHom-lMap (suc n) {A = A} p q = ΣPathP ((funExt (λ { north → refl ; south → refl - ; (merid a i) j → h a j i})) + ; (merid a i) j → main a j i})) , refl) where doubleComp-lem : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) (q r : y ≡ y) @@ -99,9 +99,9 @@ isHom-lMap (suc n) {A = A} p q = ; (j = i1) → p (~ k)}) ((q ∙ r) j) - help2 : (p : typ ((Ω^ (suc (suc n))) A)) - → cong (fst (lMap (suc (suc n)) p)) (merid (ptSn _)) ≡ refl - help2 p = + lem : (p : typ ((Ω^ (suc (suc n))) A)) + → cong (fst (lMap (suc (suc n)) p)) (merid (ptSn _)) ≡ refl + lem p = cong (sym (lMapId (suc n) (ptSn _)) ∙∙_∙∙ lMapId (suc n) (ptSn _)) (rUnit _ ∙ (λ j → (λ i → lMap (suc n) {A = A} refl .snd (i ∧ j)) ∙∙ (λ i → lMap (suc n) {A = A} (p i) .snd j) @@ -109,18 +109,18 @@ isHom-lMap (suc n) {A = A} p q = ∙ ∙∙lCancel _) ∙ ∙∙lCancel _ - h : (a : S₊ (suc n)) + main : (a : S₊ (suc n)) → sym (lMapId (suc n) a) ∙∙ funExt⁻ (cong fst (cong (lMap (suc n)) (p ∙ q))) a ∙∙ lMapId (suc n) a ≡ cong (fst (∙Π (lMap (suc (suc n)) p) (lMap (suc (suc n)) q))) (merid a) - h a = (cong (sym (lMapId (suc n) a) ∙∙_∙∙ (lMapId (suc n) a)) + main a = (cong (sym (lMapId (suc n) a) ∙∙_∙∙ (lMapId (suc n) a)) (cong-∙ (λ x → lMap (suc n) x .fst a) p q) ∙ sym (doubleComp-lem (sym (lMapId (suc n) a)) _ _)) ∙∙ cong₂ _∙_ (sym (cong (cong (fst (lMap (suc (suc n)) p)) (merid a) ∙_) - (cong sym (help2 p)) ∙ sym (rUnit _))) + (cong sym (lem p)) ∙ sym (rUnit _))) (sym (cong (cong (fst (lMap (suc (suc n)) q)) (merid a) ∙_) - (cong sym (help2 q)) ∙ sym (rUnit _))) + (cong sym (lem q)) ∙ sym (rUnit _))) ∙∙ λ i → (rUnit (cong-∙ (fst (lMap (suc (suc n)) p)) (merid a) (sym (merid (ptSn _))) (~ i)) i) ∙ (rUnit (cong-∙ (fst (lMap (suc (suc n)) q)) @@ -240,16 +240,17 @@ fst (∙Π-rCancel {A = A} {n = zero} f i) (loop j) = snd (∙Π-rCancel {A = A} {n = zero} f i) = refl fst (∙Π-rCancel {A = A} {n = suc n} f i) north = pt A fst (∙Π-rCancel {A = A} {n = suc n} f i) south = pt A -fst (∙Π-rCancel {A = A} {n = suc n} f i) (merid a i₁) = sl i i₁ +fst (∙Π-rCancel {A = A} {n = suc n} f i) (merid a i₁) = lem i i₁ where pl = (sym (snd f) ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f) - sl : pl + + lem : pl ∙ ((sym (snd f) ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f)) ≡ refl - sl = cong (pl ∙_) (cong (sym (snd f) ∙∙_∙∙ (snd f)) + lem = cong (pl ∙_) (cong (sym (snd f) ∙∙_∙∙ (snd f)) (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) ∙∙ cong₂ _∙_ refl (cong (cong (fst f)) (rCancel (merid (ptSn _)))) @@ -265,17 +266,17 @@ fst (∙Π-lCancel {A = A} {n = zero} f i) (loop j) = rCancel (sym (snd f) ∙∙ cong (fst f) (sym loop) ∙∙ snd f) i j fst (∙Π-lCancel {A = A} {n = suc n} f i) north = pt A fst (∙Π-lCancel {A = A} {n = suc n} f i) south = pt A -fst (∙Π-lCancel {A = A} {n = suc n} f i) (merid a j) = sl i j +fst (∙Π-lCancel {A = A} {n = suc n} f i) (merid a j) = lem i j where pl = (sym (snd f) ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f) - sl : (sym (snd f) + lem : (sym (snd f) ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) ∙∙ snd f) ∙ pl ≡ refl - sl = cong (_∙ pl) (cong (sym (snd f) ∙∙_∙∙ (snd f)) + lem = cong (_∙ pl) (cong (sym (snd f) ∙∙_∙∙ (snd f)) (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) ∙∙ cong₂ _∙_ refl (cong (cong (fst f)) (rCancel (merid (ptSn _)))) ∙∙ sym (rUnit _))) diff --git a/Cubical/Homotopy/HSpace.agda b/Cubical/Homotopy/HSpace.agda index bcd2db4a4..055f37e2a 100644 --- a/Cubical/Homotopy/HSpace.agda +++ b/Cubical/Homotopy/HSpace.agda @@ -27,7 +27,6 @@ record AssocHSpace {ℓ : Level} {A : Pointed ℓ} (e : HSpace A) : Type ℓ whe (μ-assoc (pt A) y z) refl - -- Instances open HSpace open AssocHSpace From 706e11fac5c10e2138d271fd6d3d7f2c562a91cb Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Sat, 6 Nov 2021 02:19:00 +0100 Subject: [PATCH 83/89] fixes --- Cubical/Algebra/Group/DirProd.agda | 1 + Cubical/Algebra/Group/GroupPath.agda | 4 +- Cubical/Algebra/Group/Instances/Unit.agda | 13 + .../Group/{ZModule.agda => ZAction.agda} | 37 +- Cubical/Data/Bool/Properties.agda | 12 + Cubical/Data/Int/Properties.agda | 3 + Cubical/Data/Nat/Order.agda | 18 + Cubical/Experiments/Brunerie.agda | 3 +- Cubical/Experiments/Problem.agda | 4 +- .../Experiments/ZCohomology/Benchmarks.agda | 3 +- .../ZCohomologyOld/KcompPrelims.agda | 3 +- Cubical/Foundations/Prelude.agda | 3 +- Cubical/HITs/Hopf.agda | 337 ----------------- Cubical/HITs/Sn/Properties.agda | 60 ++- Cubical/Homotopy/Freudenthal.agda | 22 +- Cubical/Homotopy/Group/Properties.agda | 14 +- Cubical/Homotopy/Group/S3.agda | 6 +- .../Homotopy/Group/SuspensionMapPathP.agda | 21 +- Cubical/Homotopy/Hopf.agda | 345 +++++++++++++++++- Cubical/Homotopy/HopfInvariant/Base.agda | 43 +-- .../Homotopy/HopfInvariant/Homomorphism.agda | 33 +- Cubical/Homotopy/HopfInvariant/HopfMap.agda | 74 +--- Cubical/Papers/Synthetic.agda | 10 +- Cubical/Talks/EPA2020.agda | 3 +- Cubical/ZCohomology/Groups/CP2.agda | 4 +- Cubical/ZCohomology/Gysin.agda | 53 +-- .../ZCohomology/RingStructure/RingLaws.agda | 12 +- 27 files changed, 559 insertions(+), 582 deletions(-) rename Cubical/Algebra/Group/{ZModule.agda => ZAction.agda} (90%) diff --git a/Cubical/Algebra/Group/DirProd.agda b/Cubical/Algebra/Group/DirProd.agda index 352d5446f..0f8a68393 100644 --- a/Cubical/Algebra/Group/DirProd.agda +++ b/Cubical/Algebra/Group/DirProd.agda @@ -8,6 +8,7 @@ open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Monoid open import Cubical.Algebra.Semigroup + open GroupStr open IsGroup hiding (rid ; lid ; invr ; invl) open IsMonoid hiding (rid ; lid) diff --git a/Cubical/Algebra/Group/GroupPath.agda b/Cubical/Algebra/Group/GroupPath.agda index 88ca42ece..c734b13e6 100644 --- a/Cubical/Algebra/Group/GroupPath.agda +++ b/Cubical/Algebra/Group/GroupPath.agda @@ -129,10 +129,10 @@ uaCompGroupEquiv f g = caracGroup≡ _ _ ( cong ⟨_⟩ (uaGroup f ∙ uaGroup g) ∎) -- J-rule for GroupEquivs -JGroupEquiv : {G : Group ℓ} (P : (H : Group ℓ) → GroupEquiv G H → Type ℓ') +GroupEquivJ : {G : Group ℓ} (P : (H : Group ℓ) → GroupEquiv G H → Type ℓ') → P G idGroupEquiv → ∀ {H} e → P H e -JGroupEquiv {G = G} P p {H} e = +GroupEquivJ {G = G} P p {H} e = transport (λ i → P (GroupPath G H .fst e i) (transp (λ j → GroupEquiv G (GroupPath G H .fst e (i ∨ ~ j))) i e)) (subst (P G) (sym lem) p) diff --git a/Cubical/Algebra/Group/Instances/Unit.agda b/Cubical/Algebra/Group/Instances/Unit.agda index 0605898f3..dd0e1b7ff 100644 --- a/Cubical/Algebra/Group/Instances/Unit.agda +++ b/Cubical/Algebra/Group/Instances/Unit.agda @@ -4,6 +4,7 @@ module Cubical.Algebra.Group.Instances.Unit where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Structure +open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv open import Cubical.Data.Unit renaming (Unit to UnitType) open import Cubical.Algebra.Group.Base @@ -11,6 +12,7 @@ open import Cubical.Algebra.Group.DirProd open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec) +open import Cubical.Algebra.Group.GroupPath open GroupStr open IsGroupHom @@ -95,3 +97,14 @@ SES→isEquiv {R = R} {G = G} {H = H} = (λ s → sym (snd s) ∙ IsGroupHom.pres1 (snd lhom)) (lexact _ inker) BijectionIso.surj bijIso' x = rexact x refl + +isContr→≡UnitGroup : {G : Group ℓ-zero} → isContr (fst G) → Unit ≡ G +isContr→≡UnitGroup c = + fst (GroupPath _ _) + (invGroupEquiv ((isContr→≃Unit c) + , (makeIsGroupHom (λ _ _ → refl)))) + +GroupIsoUnitGroup→isContr : {G : Group ℓ-zero} + → GroupIso Unit G → isContr (fst G) +GroupIsoUnitGroup→isContr is = + isOfHLevelRetractFromIso 0 (invIso (fst is)) isContrUnit diff --git a/Cubical/Algebra/Group/ZModule.agda b/Cubical/Algebra/Group/ZAction.agda similarity index 90% rename from Cubical/Algebra/Group/ZModule.agda rename to Cubical/Algebra/Group/ZAction.agda index 9958abafb..020248927 100644 --- a/Cubical/Algebra/Group/ZModule.agda +++ b/Cubical/Algebra/Group/ZAction.agda @@ -1,6 +1,6 @@ -- Left ℤ-multiplication on groups and some of its properties {-# OPTIONS --safe --experimental-lossy-unification #-} -module Cubical.Algebra.Group.ZModule where +module Cubical.Algebra.Group.ZAction where open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism @@ -19,6 +19,7 @@ open import Cubical.Algebra.Group.Properties open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.Instances.Int renaming (ℤ to ℤGroup) +open import Cubical.Algebra.Group.DirProd open import Cubical.Relation.Nullary @@ -30,10 +31,6 @@ open Iso open GroupStr open IsGroupHom -private - minus≡0- : (x : ℤ) → - x ≡ (0 -ℤ x) - minus≡0- x = sym (lid (snd ℤGroup) (- x)) - _ℤ[_]·_ : ℤ → (G : Group ℓ) → fst G → fst G pos zero ℤ[ G' ]· g = 1g (snd G') pos (suc n) ℤ[ G' ]· g = _·_ (snd G') g (pos n ℤ[ G' ]· g) @@ -286,3 +283,33 @@ groupEquivPresGen G (ϕeq , ϕhom) x (inr r) (ψeq , ψhom) = absLem (pos (suc (suc n))) p = ⊥-rec (snotz (cong predℕ p)) absLem (negsuc zero) p = refl absLem (negsuc (suc n)) p = ⊥-rec (snotz (cong predℕ p)) + +gen₂ℤ×ℤ : gen₂-by (DirProd ℤGroup ℤGroup) (1 , 0) (0 , 1) +fst (gen₂ℤ×ℤ (x , y)) = x , y +snd (gen₂ℤ×ℤ (x , y)) = + ΣPathP ((cong₂ _+ℤ_ ((·Comm 1 x) ∙ cong fst (sym (distrLem 1 0 x))) + ((·Comm 0 y) ∙ cong fst (sym (distrLem 0 1 y)))) + , +Comm y 0 + ∙ cong₂ _+ℤ_ (·Comm 0 x ∙ cong snd (sym (distrLem 1 0 x))) + (·Comm 1 y ∙ cong snd (sym (distrLem 0 1 y)))) + where + ℤ×ℤ = DirProd ℤGroup ℤGroup + _+''_ = GroupStr._·_ (snd ℤ×ℤ) + + -lem : (x : ℤ) → - x ≡ 0 -ℤ x + -lem (pos zero) = refl + -lem (pos (suc zero)) = refl + -lem (pos (suc (suc n))) = + cong predℤ (-lem (pos (suc n))) + -lem (negsuc zero) = refl + -lem (negsuc (suc n)) = cong sucℤ (-lem (negsuc n)) + + distrLem : (x y : ℤ) (z : ℤ) + → Path (ℤ × ℤ) (z ℤ[ ℤ×ℤ ]· (x , y)) (z * x , z * y) + distrLem x y (pos zero) = refl + distrLem x y (pos (suc n)) = + (cong ((x , y) +''_) (distrLem x y (pos n))) + distrLem x y (negsuc zero) = ΣPathP (sym (-lem x) , sym (-lem y)) + distrLem x y (negsuc (suc n)) = + cong₂ _+''_ (ΣPathP (sym (-lem x) , sym (-lem y))) + (distrLem x y (negsuc n)) diff --git a/Cubical/Data/Bool/Properties.agda b/Cubical/Data/Bool/Properties.agda index be4cd01c9..c899f997d 100644 --- a/Cubical/Data/Bool/Properties.agda +++ b/Cubical/Data/Bool/Properties.agda @@ -10,9 +10,11 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Transport open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Pointed open import Cubical.Data.Bool.Base open import Cubical.Data.Empty +open import Cubical.Data.Sigma open import Cubical.Relation.Nullary open import Cubical.Relation.Nullary.DecidableEq @@ -207,3 +209,13 @@ module BoolReflection where reflectEquiv : Bool ≃ (Bool ≡ Bool) reflectEquiv = isoToEquiv reflectIso + +IsoBool→∙ : ∀ {ℓ} {A : Pointed ℓ} → Iso ((Bool , true) →∙ A) (typ A) +Iso.fun IsoBool→∙ f = fst f false +fst (Iso.inv IsoBool→∙ a) false = a +fst (Iso.inv (IsoBool→∙ {A = A}) a) true = pt A +snd (Iso.inv IsoBool→∙ a) = refl +Iso.rightInv IsoBool→∙ a = refl +Iso.leftInv IsoBool→∙ (f , p) = + ΣPathP ((funExt (λ { false → refl ; true → sym p})) + , λ i j → p (~ i ∨ j)) diff --git a/Cubical/Data/Int/Properties.agda b/Cubical/Data/Int/Properties.agda index f1342736d..9c8f316c8 100644 --- a/Cubical/Data/Int/Properties.agda +++ b/Cubical/Data/Int/Properties.agda @@ -457,6 +457,9 @@ private ∙∙ cong (_+ ((negsuc n · b) · c)) (-DistL· b c) ∙∙ sym (·DistL+ (- b) (negsuc n · b) c) +minus≡0- : (x : ℤ) → - x ≡ (0 - x) +minus≡0- x = +Comm (- x) 0 + -- Absolute values abs→⊎ : (x : ℤ) (n : ℕ) → abs x ≡ n → (x ≡ pos n) ⊎ (x ≡ - pos n) abs→⊎ x n = J (λ n _ → (x ≡ pos n) ⊎ (x ≡ - pos n)) (help x) diff --git a/Cubical/Data/Nat/Order.agda b/Cubical/Data/Nat/Order.agda index 0404a27b3..2d1d2dfdb 100644 --- a/Cubical/Data/Nat/Order.agda +++ b/Cubical/Data/Nat/Order.agda @@ -327,3 +327,21 @@ module <-Reasoning where _<≡⟨_⟩_ : ∀ k → k < l → l ≡ m → k < m _ <≡⟨ p ⟩ q = _ ≤≡⟨ p ⟩ q + + +-- Some lemmas about ∸ +suc∸-fst : (n m : ℕ) → m < n → suc (n ∸ m) ≡ (suc n) ∸ m +suc∸-fst zero zero p = refl +suc∸-fst zero (suc m) p = ⊥.rec (¬-<-zero p) +suc∸-fst (suc n) zero p = refl +suc∸-fst (suc n) (suc m) p = (suc∸-fst n m (pred-≤-pred p)) + +n∸m≡0 : (n m : ℕ) → n < m → (n ∸ m) ≡ 0 +n∸m≡0 zero zero p = refl +n∸m≡0 (suc n) zero p = ⊥.rec (¬-<-zero p) +n∸m≡0 zero (suc m) p = refl +n∸m≡0 (suc n) (suc m) p = n∸m≡0 n m (pred-≤-pred p) + +n∸n≡0 : (n : ℕ) → n ∸ n ≡ 0 +n∸n≡0 zero = refl +n∸n≡0 (suc n) = n∸n≡0 n diff --git a/Cubical/Experiments/Brunerie.agda b/Cubical/Experiments/Brunerie.agda index f32633564..b60c7afe8 100644 --- a/Cubical/Experiments/Brunerie.agda +++ b/Cubical/Experiments/Brunerie.agda @@ -10,11 +10,12 @@ open import Cubical.HITs.S1 open import Cubical.HITs.S2 open import Cubical.HITs.S3 open import Cubical.HITs.Join -open import Cubical.HITs.Hopf open import Cubical.HITs.SetTruncation as SetTrunc open import Cubical.HITs.GroupoidTruncation as GroupoidTrunc open import Cubical.HITs.2GroupoidTruncation as 2GroupoidTrunc open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Hopf +open S¹Hopf -- This code is adapted from examples/brunerie3.ctt on the pi4s3_nobug branch of cubicaltt diff --git a/Cubical/Experiments/Problem.agda b/Cubical/Experiments/Problem.agda index 1ee86657f..76a8876e1 100644 --- a/Cubical/Experiments/Problem.agda +++ b/Cubical/Experiments/Problem.agda @@ -10,7 +10,9 @@ open import Cubical.HITs.S1 open import Cubical.HITs.S2 open import Cubical.HITs.S3 open import Cubical.HITs.Join -open import Cubical.HITs.Hopf +open import Cubical.Homotopy.Hopf + +open S¹Hopf ptType : Type _ ptType = Σ Type₀ \ A → A diff --git a/Cubical/Experiments/ZCohomology/Benchmarks.agda b/Cubical/Experiments/ZCohomology/Benchmarks.agda index 61a2662f3..aeb08a96a 100644 --- a/Cubical/Experiments/ZCohomology/Benchmarks.agda +++ b/Cubical/Experiments/ZCohomology/Benchmarks.agda @@ -46,7 +46,8 @@ open import Cubical.HITs.KleinBottle open import Cubical.HITs.RPn.Base open import Cubical.HITs.SetTruncation open import Cubical.HITs.Pushout -open import Cubical.HITs.Hopf +open import Cubical.Homotopy.Hopf +open S¹Hopf open import Cubical.HITs.Truncation open import Cubical.HITs.Susp open import Cubical.HITs.S1 diff --git a/Cubical/Experiments/ZCohomologyOld/KcompPrelims.agda b/Cubical/Experiments/ZCohomologyOld/KcompPrelims.agda index df2464b9f..9ff19c27a 100644 --- a/Cubical/Experiments/ZCohomologyOld/KcompPrelims.agda +++ b/Cubical/Experiments/ZCohomologyOld/KcompPrelims.agda @@ -3,7 +3,8 @@ module Cubical.Experiments.ZCohomologyOld.KcompPrelims where open import Cubical.ZCohomology.Base open import Cubical.Homotopy.Connected -open import Cubical.HITs.Hopf +open import Cubical.Homotopy.Hopf +open S¹Hopf -- open import Cubical.Homotopy.Freudenthal hiding (encode) open import Cubical.HITs.Sn open import Cubical.HITs.S1 diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index c4368711a..9e80d8105 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -515,8 +515,7 @@ compPathR→PathP∙∙ {p = p} {q = q} {r = r} {s = s} P j i = ; (j = i1) → doubleCompPath-filler p s (sym q) (~ k) i}) (P j i) -doubleCompPath≡compPath : - ∀ {ℓ} {A : Type ℓ} {x y z w : A} +doubleCompPath≡compPath : {x y z w : A} (p : x ≡ y) (q : y ≡ z) (r : z ≡ w) → p ∙∙ q ∙∙ r ≡ p ∙ q ∙ r doubleCompPath≡compPath p q r i j = diff --git a/Cubical/HITs/Hopf.agda b/Cubical/HITs/Hopf.agda index 0b72621fb..cad8df08f 100644 --- a/Cubical/HITs/Hopf.agda +++ b/Cubical/HITs/Hopf.agda @@ -1,340 +1,3 @@ {-# OPTIONS --safe #-} module Cubical.HITs.Hopf where -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Univalence - -open import Cubical.Data.Int hiding (_·_) -open import Cubical.Data.Sigma -open import Cubical.Foundations.Function - -open import Cubical.HITs.S1 -open import Cubical.HITs.S2 -open import Cubical.HITs.S3 -open import Cubical.HITs.Sn -open import Cubical.HITs.Susp -open import Cubical.HITs.Join -open import Cubical.HITs.Interval - renaming ( zero to I0 ; one to I1 ) - -Border : (x : S¹) → (j : I) → Partial (j ∨ ~ j) (Σ Type₀ (λ T → T ≃ S¹)) -Border x j (j = i0) = S¹ , (x ·_) , rotIsEquiv x -Border x j (j = i1) = S¹ , idEquiv S¹ - --- Hopf fibration using SuspS¹ -HopfSuspS¹ : SuspS¹ → Type₀ -HopfSuspS¹ north = S¹ -HopfSuspS¹ south = S¹ -HopfSuspS¹ (merid x j) = Glue S¹ (Border x j) - --- Hopf fibration using S² --- TODO : prove that it is equivalent to HopfSuspS¹ -HopfS² : S² → Type₀ -HopfS² base = S¹ -HopfS² (surf i j) = Glue S¹ (λ { (i = i0) → _ , idEquiv S¹ - ; (i = i1) → _ , idEquiv S¹ - ; (j = i0) → _ , idEquiv S¹ - ; (j = i1) → _ , _ , rotIsEquiv (loop i) } ) - --- Hopf fibration using more direct definition of the rot equivalence --- TODO : prove that it is equivalent to HopfSuspS¹ -HopfS²' : S² → Type₀ -HopfS²' base = S¹ -HopfS²' (surf i j) = Glue S¹ (λ { (i = i0) → _ , rotLoopEquiv i0 - ; (i = i1) → _ , rotLoopEquiv i0 - ; (j = i0) → _ , rotLoopEquiv i0 - ; (j = i1) → _ , rotLoopEquiv i } ) - --- Total space of the fibration -TotalHopf : Type₀ -TotalHopf = Σ SuspS¹ HopfSuspS¹ - --- Forward direction -filler-1 : I → (j : I) → (y : S¹) → Glue S¹ (Border y j) → join S¹ S¹ -filler-1 i j y x = hfill (λ t → λ { (j = i0) → inl (rotInv-1 x y t) - ; (j = i1) → inr x }) - (inS (push ((unglue (j ∨ ~ j) x) · invLooper y) (unglue (j ∨ ~ j) x) j)) i - -TotalHopf→JoinS¹S¹ : TotalHopf → join S¹ S¹ -TotalHopf→JoinS¹S¹ (north , x) = inl x -TotalHopf→JoinS¹S¹ (south , x) = inr x -TotalHopf→JoinS¹S¹ (merid y j , x) = filler-1 i1 j y x - --- Backward direction -JoinS¹S¹→TotalHopf : join S¹ S¹ → TotalHopf -JoinS¹S¹→TotalHopf (inl x) = (north , x) -JoinS¹S¹→TotalHopf (inr x) = (south , x) -JoinS¹S¹→TotalHopf (push y x j) = - (merid (invLooper y · x) j - , glue (λ { (j = i0) → y ; (j = i1) → x }) (rotInv-2 x y j)) - --- Now for the homotopies, we will need to fill squares indexed by x y : S¹ with value in S¹ --- Some will be extremeley tough, but happen to be easy when x = y = base --- therefore, we fill them for x = y = base and then use the connectedness of S¹ × S¹ and --- the discreteness of ΩS¹ to get general fillers. - --- To proceed with that strategy, we first need a lemma : --- the sections of the trivial fibration λ (_ : S¹) (_ : S¹) → Int are constant - --- this should be generalized to a constant fibration over a connected space with --- discrete fiber -fibℤ : S¹ → S¹ → Type₀ -fibℤ _ _ = ℤ - -S¹→HSet : (A : Type₀) (p : isSet A) (F : S¹ → A) (x : S¹) → F base ≡ F x -S¹→HSet A p F base = refl {x = F base} -S¹→HSet A p F (loop i) = f' i - where - f : PathP (λ i → F base ≡ F (loop i)) refl (cong F loop) - f i = λ j → F (loop (i ∧ j)) - L : cong F loop ≡ refl - L = p (F base) (F base) (f i1) refl - f' : PathP (λ i → F base ≡ F (loop i)) (refl {x = F base}) (refl {x = F base}) - f' = transport (λ i → PathP (λ j → F base ≡ F (loop j)) refl (L i)) f - -constant-loop : (F : S¹ → S¹ → ℤ) → (x y : S¹) → F base base ≡ F x y -constant-loop F x y = L0 ∙ L1 - where - p : isSet (S¹ → ℤ) - p = isSetΠ (λ _ → isSetℤ) - L : F base ≡ F x - L = S¹→HSet (S¹ → ℤ) p F x - L0 : F base base ≡ F x base - L0 i = L i base - L1 : F x base ≡ F x y - L1 = S¹→HSet ℤ isSetℤ (F x) y - -discretefib : (F : S¹ → S¹ → Type₀) → Type₀ -discretefib F = (a : (x y : S¹) → F x y) → - (b : (x y : S¹) → F x y) → - (a base base ≡ b base base) → - (x y : S¹) → a x y ≡ b x y - -discretefib-fibℤ : discretefib fibℤ -discretefib-fibℤ a b h x y i = - hcomp (λ t → λ { (i = i0) → constant-loop a x y t - ; (i = i1) → constant-loop b x y t }) - (h i) - --- first homotopy - -assocFiller-3-aux : I → I → I → I → S¹ -assocFiller-3-aux x y j i = - hfill (λ t → λ { (i = i0) → rotInv-1 (loop y) (loop (~ y) · loop x) t - ; (i = i1) → rotInv-3 (loop y) (loop x) t - ; (x = i0) (y = i0) → base - ; (x = i0) (y = i1) → base - ; (x = i1) (y = i0) → base - ; (x = i1) (y = i1) → base }) - (inS ((rotInv-2 (loop x) (loop y) i) · (invLooper (loop (~ y) · loop x)))) j - --- assocFiller-3-endpoint is used only in the type of the next function, to specify the --- second endpoint. --- However, I only need the first endpoint, but I cannot specify only one of them as is. --- TODO : use cubical extension types when available to remove assocFiller-3-endpoint -assocFiller-3-endpoint : (x : S¹) → (y : S¹) → y ≡ y -assocFiller-3-endpoint base base i = base -assocFiller-3-endpoint (loop x) base i = assocFiller-3-aux x i0 i1 i -assocFiller-3-endpoint base (loop y) i = assocFiller-3-aux i0 y i1 i -assocFiller-3-endpoint (loop x) (loop y) i = assocFiller-3-aux x y i1 i - -assocFiller-3 : (x : S¹) → (y : S¹) → - PathP (λ j → rotInv-1 y (invLooper y · x) j ≡ rotInv-3 y x j) - (λ i → ((rotInv-2 x y i) · (invLooper (invLooper y · x)))) - (assocFiller-3-endpoint x y) -assocFiller-3 base base j i = base -assocFiller-3 (loop x) base j i = assocFiller-3-aux x i0 j i -assocFiller-3 base (loop y) j i = assocFiller-3-aux i0 y j i -assocFiller-3 (loop x) (loop y) j i = assocFiller-3-aux x y j i - -assoc-3 : (_ y : S¹) → basedΩS¹ y -assoc-3 x y i = assocFiller-3 x y i1 i - -fibℤ≡fibAssoc-3 : fibℤ ≡ (λ _ y → basedΩS¹ y) -fibℤ≡fibAssoc-3 i = λ x y → basedΩS¹≡ℤ y (~ i) - -discretefib-fibAssoc-3 : discretefib (λ _ y → basedΩS¹ y) -discretefib-fibAssoc-3 = - transp (λ i → discretefib (fibℤ≡fibAssoc-3 i)) i0 discretefib-fibℤ - -assocConst-3 : (x y : S¹) → assoc-3 x y ≡ refl -assocConst-3 x y = discretefib-fibAssoc-3 assoc-3 (λ _ _ → refl) refl x y - -assocSquare-3 : I → I → S¹ → S¹ → S¹ -assocSquare-3 i j x y = hcomp (λ t → λ { (i = i0) → assocFiller-3 x y j i0 - ; (i = i1) → assocFiller-3 x y j i1 - ; (j = i0) → assocFiller-3 x y i0 i - ; (j = i1) → assocConst-3 x y t i }) - (assocFiller-3 x y j i) - -filler-3 : I → I → S¹ → S¹ → join S¹ S¹ -filler-3 i j y x = - hcomp (λ t → λ { (i = i0) → filler-1 t j (invLooper y · x) - (glue (λ { (j = i0) → y ; (j = i1) → x }) - (rotInv-2 x y j)) - ; (i = i1) → push (rotInv-3 y x t) x j - ; (j = i0) → inl (assocSquare-3 i t x y) - ; (j = i1) → inr x }) - (push ((rotInv-2 x y (i ∨ j)) · (invLooper (invLooper y · x))) (rotInv-2 x y (i ∨ j)) j) - -JoinS¹S¹→TotalHopf→JoinS¹S¹ : ∀ x → TotalHopf→JoinS¹S¹ (JoinS¹S¹→TotalHopf x) ≡ x -JoinS¹S¹→TotalHopf→JoinS¹S¹ (inl x) i = inl x -JoinS¹S¹→TotalHopf→JoinS¹S¹ (inr x) i = inr x -JoinS¹S¹→TotalHopf→JoinS¹S¹ (push y x j) i = filler-3 i j y x - --- Second homotopy - --- This HIT is the total space of the Hopf fibration but the ends of SuspS¹ have not been --- glued together yet — which makes it into a cylinder. --- This allows to write compositions that do not properly match at the endpoints. However, --- I suspect it is unnecessary. TODO : do without PseudoHopf - -PseudoHopf : Type₀ -PseudoHopf = (S¹ × Interval) × S¹ - -PseudoHopf-π1 : PseudoHopf → S¹ -PseudoHopf-π1 ((y , _) , _) = y - -PseudoHopf-π2 : PseudoHopf → S¹ -PseudoHopf-π2 (_ , x) = x - -assocFiller-4-aux : I → I → I → I → S¹ -assocFiller-4-aux x y j i = - hfill (λ t → λ { (i = i0) → ((invLooper (loop y · loop x · loop (~ y))) · (loop y · loop x)) - · (rotInv-1 (loop x) (loop y) t) - ; (i = i1) → (rotInv-4 (loop y) (loop y · loop x) (~ t)) · loop x - ; (x = i0) (y = i0) → base - ; (x = i0) (y = i1) → base - ; (x = i1) (y = i0) → base - ; (x = i1) (y = i1) → base }) - (inS (rotInv-2 (loop y · loop x) (loop y · loop x · loop (~ y)) i)) j - --- See assocFiller-3-endpoint --- TODO : use cubical extension types when available to remove assocFiller-4-endpoint -assocFiller-4-endpoint : (x y : S¹) → basedΩS¹ (((invLooper (y · x · invLooper y)) · (y · x)) · x) -assocFiller-4-endpoint base base i = base -assocFiller-4-endpoint (loop x) base i = assocFiller-4-aux x i0 i1 i -assocFiller-4-endpoint base (loop y) i = assocFiller-4-aux i0 y i1 i -assocFiller-4-endpoint (loop x) (loop y) i = assocFiller-4-aux x y i1 i - -assocFiller-4 : (x y : S¹) → - PathP (λ j → ((invLooper (y · x · invLooper y)) · (y · x)) · (rotInv-1 x y j) ≡ (rotInv-4 y (y · x) (~ j)) · x) - (λ i → (rotInv-2 (y · x) (y · x · invLooper y) i)) - (assocFiller-4-endpoint x y) -assocFiller-4 base base j i = base -assocFiller-4 (loop x) base j i = assocFiller-4-aux x i0 j i -assocFiller-4 base (loop y) j i = assocFiller-4-aux i0 y j i -assocFiller-4 (loop x) (loop y) j i = assocFiller-4-aux x y j i - -assoc-4 : (x y : S¹) → basedΩS¹ (((invLooper (y · x · invLooper y)) · (y · x)) · x) -assoc-4 x y i = assocFiller-4 x y i1 i - -fibℤ≡fibAssoc-4 : fibℤ ≡ (λ x y → basedΩS¹ (((invLooper (y · x · invLooper y)) · (y · x)) · x)) -fibℤ≡fibAssoc-4 i = λ x y → basedΩS¹≡ℤ (((invLooper (y · x · invLooper y)) · (y · x)) · x) (~ i) - -discretefib-fibAssoc-4 : discretefib (λ x y → basedΩS¹ (((invLooper (y · x · invLooper y)) · (y · x)) · x)) -discretefib-fibAssoc-4 = - transp (λ i → discretefib (fibℤ≡fibAssoc-4 i)) i0 discretefib-fibℤ - -assocConst-4 : (x y : S¹) → assoc-4 x y ≡ refl -assocConst-4 x y = discretefib-fibAssoc-4 assoc-4 (λ _ _ → refl) refl x y - -assocSquare-4 : I → I → S¹ → S¹ → S¹ -assocSquare-4 i j x y = - hcomp (λ t → λ { (i = i0) → assocFiller-4 x y j i0 - ; (i = i1) → assocFiller-4 x y j i1 - ; (j = i0) → assocFiller-4 x y i0 i - ; (j = i1) → assocConst-4 x y t i }) - (assocFiller-4 x y j i) - -filler-4-0 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → PseudoHopf -filler-4-0 i j y x = - let x' = unglue (j ∨ ~ j) x in - hfill (λ t → λ { (j = i0) → ((invLooper (y · x · invLooper y) · (y · x) , I0) - , invLooper (y · x · invLooper y) · (y · x) · (rotInv-1 x y t)) - ; (j = i1) → ((invLooper (x · invLooper y) · x , I1) , x) }) - (inS ((invLooper (x' · invLooper y) · x' , seg j) , rotInv-2 x' (x' · invLooper y) j)) i - -filler-4-1 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → PseudoHopf -filler-4-1 i j y x = - let x' = unglue (j ∨ ~ j) x in - hfill (λ t → λ { (j = i0) → ((invLooper (y · x · invLooper y) · (y · x) , I0) - , (rotInv-4 y (y · x) (~ t)) · x) - ; (j = i1) → ((invLooper (x · invLooper y) · x , I1) , x) }) - (inS ((invLooper (x' · invLooper y) · x' , seg j) , unglue (j ∨ ~ j) x)) i - -filler-4-2 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → TotalHopf -filler-4-2 i j y x = - let x' = unglue (j ∨ ~ j) x in - hcomp (λ t → λ { (i = i0) → JoinS¹S¹→TotalHopf (filler-1 t j y x) - ; (i = i1) → (merid (PseudoHopf-π1 (filler-4-0 t j y x)) j - , glue (λ { (j = i0) → rotInv-1 x y t ; (j = i1) → x }) - (PseudoHopf-π2 (filler-4-0 t j y x))) - ; (j = i0) → (north , rotInv-1 x y t) - ; (j = i1) → (south , x) }) - (merid (invLooper (x' · invLooper y) · x') j - , glue (λ { (j = i0) → y · x · invLooper y ; (j = i1) → x }) (rotInv-2 x' (x' · invLooper y) j)) - -filler-4-3 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → PseudoHopf -filler-4-3 i j y x = - let x' = unglue (j ∨ ~ j) x in - hcomp (λ t → λ { (i = i0) → filler-4-0 t j y x - ; (i = i1) → filler-4-1 t j y x - ; (j = i0) → ((invLooper (y · x · invLooper y) · (y · x) , I0) , assocSquare-4 i t x y) - ; (j = i1) → ((invLooper (x · invLooper y) · x , I1) , x) }) - ((invLooper (x' · invLooper y) · x' , seg j) , rotInv-2 x' (x' · invLooper y) (i ∨ j)) - -filler-4-4 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → PseudoHopf -filler-4-4 i j y x = - let x' = unglue (j ∨ ~ j) x in - hcomp (λ t → λ { (i = i0) → filler-4-1 t j y x - ; (i = i1) → ((y , seg j) , unglue (j ∨ ~ j) x) - ; (j = i0) → ((rotInv-4 y (y · x) i , I0) - , (rotInv-4 y (y · x) (i ∨ ~ t)) · x) - ; (j = i1) → ((rotInv-4 y x i , I1) , x) }) - ((rotInv-4 y x' i , seg j) , x') - -filler-4-5 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → TotalHopf -filler-4-5 i j y x = - hcomp (λ t → λ { (i = i0) → filler-4-2 (~ t) j y x - ; (i = i1) → (merid (PseudoHopf-π1 (filler-4-4 t j y x)) j - , glue (λ { (j = i0) → x ; (j = i1) → x }) - (PseudoHopf-π2 (filler-4-4 t j y x))) - ; (j = i0) → (north , x) - ; (j = i1) → (south , x) }) - (merid (PseudoHopf-π1 (filler-4-3 i j y x)) j - , glue (λ { (j = i0) → x ; (j = i1) → x }) (PseudoHopf-π2 (filler-4-3 i j y x))) - -TotalHopf→JoinS¹S¹→TotalHopf : ∀ x → JoinS¹S¹→TotalHopf (TotalHopf→JoinS¹S¹ x) ≡ x -TotalHopf→JoinS¹S¹→TotalHopf (north , x) i = (north , x) -TotalHopf→JoinS¹S¹→TotalHopf (south , x) i = (south , x) -TotalHopf→JoinS¹S¹→TotalHopf (merid y j , x) i = filler-4-5 i j y x - - -JoinS¹S¹≡TotalHopf : join S¹ S¹ ≡ TotalHopf -JoinS¹S¹≡TotalHopf = isoToPath (iso JoinS¹S¹→TotalHopf - TotalHopf→JoinS¹S¹ - TotalHopf→JoinS¹S¹→TotalHopf - JoinS¹S¹→TotalHopf→JoinS¹S¹) - -S³≡TotalHopf : S³ ≡ TotalHopf -S³≡TotalHopf = S³≡joinS¹S¹ ∙ JoinS¹S¹≡TotalHopf - -open Iso -IsoS³TotalHopf : Iso (S₊ 3) TotalHopf -fun IsoS³TotalHopf x = JoinS¹S¹→TotalHopf (S³→joinS¹S¹ (inv IsoS³S3 x)) -inv IsoS³TotalHopf x = fun IsoS³S3 (joinS¹S¹→S³ (TotalHopf→JoinS¹S¹ x)) -rightInv IsoS³TotalHopf x = - cong (JoinS¹S¹→TotalHopf ∘ S³→joinS¹S¹) - (leftInv IsoS³S3 (joinS¹S¹→S³ (TotalHopf→JoinS¹S¹ x))) - ∙∙ cong JoinS¹S¹→TotalHopf - (joinS¹S¹→S³→joinS¹S¹ (TotalHopf→JoinS¹S¹ x)) - ∙∙ TotalHopf→JoinS¹S¹→TotalHopf x -leftInv IsoS³TotalHopf x = - cong (fun IsoS³S3 ∘ joinS¹S¹→S³) - (JoinS¹S¹→TotalHopf→JoinS¹S¹ (S³→joinS¹S¹ (inv IsoS³S3 x))) - ∙∙ cong (fun IsoS³S3) (S³→joinS¹S¹→S³ (inv IsoS³S3 x)) - ∙∙ Iso.rightInv IsoS³S3 x diff --git a/Cubical/HITs/Sn/Properties.agda b/Cubical/HITs/Sn/Properties.agda index 7fbeef2d3..92eadbe35 100644 --- a/Cubical/HITs/Sn/Properties.agda +++ b/Cubical/HITs/Sn/Properties.agda @@ -11,7 +11,7 @@ open import Cubical.Foundations.Transport open import Cubical.Foundations.Prelude open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Univalence -open import Cubical.HITs.S1 +open import Cubical.HITs.S1 renaming (_·_ to _*_) open import Cubical.Data.Nat hiding (elim) open import Cubical.Data.Sigma open import Cubical.HITs.Sn.Base @@ -336,3 +336,61 @@ IsoSphereJoin (suc n) m = (compIso (IsoSphereJoin n (suc m)) (pathToIso λ i → S₊ (suc (+-suc n m i)))))) + +-- Some lemmas on the H +rUnitS¹ : (x : S¹) → x * base ≡ x +rUnitS¹ base = refl +rUnitS¹ (loop i₁) = refl + +commS¹ : (a x : S¹) → a * x ≡ x * a +commS¹ = wedgeconFun _ _ (λ _ _ → isGroupoidS¹ _ _) + (sym ∘ rUnitS¹) + rUnitS¹ + refl + +SuspS¹-hom : (a x : S¹) + → Path (Path (hLevelTrunc 4 (S₊ 2)) _ _) + (cong ∣_∣ₕ (merid (a * x) ∙ sym (merid base))) + (cong ∣_∣ₕ (merid a ∙ sym (merid base)) + ∙ (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) +SuspS¹-hom = wedgeconFun _ _ (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) + (λ x → lUnit _ + ∙ cong (_∙ cong ∣_∣ₕ (merid x ∙ sym (merid base))) + (cong (cong ∣_∣ₕ) (sym (rCancel (merid base))))) + (λ x → (λ i → cong ∣_∣ₕ (merid (rUnitS¹ x i) ∙ sym (merid base))) + ∙∙ rUnit _ + ∙∙ cong (cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙_) + (cong (cong ∣_∣ₕ) (sym (rCancel (merid base))))) + (sym (l (cong ∣_∣ₕ (merid base ∙ sym (merid base))) + (cong (cong ∣_∣ₕ) (sym (rCancel (merid base)))))) + where + l : ∀ {ℓ} {A : Type ℓ} {x : A} (p : x ≡ x) (P : refl ≡ p) + → lUnit p ∙ cong (_∙ p) P ≡ rUnit p ∙ cong (p ∙_) P + l p = J (λ p P → lUnit p ∙ cong (_∙ p) P ≡ rUnit p ∙ cong (p ∙_) P) refl + +rCancelS¹ : (x : S¹) → ptSn 1 ≡ x * (invLooper x) +rCancelS¹ base = refl +rCancelS¹ (loop i) j = + hcomp (λ r → λ {(i = i0) → base ; (i = i1) → base ; (j = i0) → base}) + base + +SuspS¹-inv : (x : S¹) → Path (Path (hLevelTrunc 4 (S₊ 2)) _ _) + (cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base))) + (cong ∣_∣ₕ (sym (merid x ∙ sym (merid base)))) +SuspS¹-inv x = (lUnit _ + ∙∙ cong (_∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base))) + (sym (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base))))) + ∙∙ sym (assoc _ _ _)) + ∙∙ cong (sym (cong ∣_∣ₕ (merid x ∙ sym (merid base))) ∙_) lem + ∙∙ (assoc _ _ _ + ∙∙ cong (_∙ (cong ∣_∣ₕ (sym (merid x ∙ sym (merid base))))) + (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) + ∙∙ sym (lUnit _)) + where + lem : cong ∣_∣ₕ (merid x ∙ sym (merid base)) + ∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base)) + ≡ cong ∣_∣ₕ (merid x ∙ sym (merid base)) + ∙ cong ∣_∣ₕ (sym (merid x ∙ sym (merid base))) + lem = sym (SuspS¹-hom x (invLooper x)) + ∙ ((λ i → cong ∣_∣ₕ (merid (rCancelS¹ x (~ i)) ∙ sym (merid base))) + ∙ cong (cong ∣_∣ₕ) (rCancel (merid base))) ∙ sym (rCancel _) diff --git a/Cubical/Homotopy/Freudenthal.agda b/Cubical/Homotopy/Freudenthal.agda index 8e75fdf2e..087c22689 100644 --- a/Cubical/Homotopy/Freudenthal.agda +++ b/Cubical/Homotopy/Freudenthal.agda @@ -153,29 +153,13 @@ suspMapΩ-connected : ∀ {ℓ} (n : HLevel) (m : ℕ) {A : Pointed ℓ} suspMapΩ-connected n zero {A = A} connA = isConnectedσ n connA suspMapΩ-connected n (suc m) {A = A} connA with ((n + suc n) ≟ m) ... | (lt p) = subst (λ x → isConnectedFun x (suspMapΩ {A = A} (suc m))) - (sym (help _ m p)) + (sym (n∸m≡0 _ m p)) λ b → tt* , (λ {tt* → refl}) - where - help : (n m : ℕ) → n < m → (n ∸ m) ≡ 0 - help zero zero p = refl - help (suc n) zero p = ⊥-rec (¬-<-zero p) - help zero (suc m) p = refl - help (suc n) (suc m) p = help n m (pred-≤-pred p) ... | (eq q) = subst (λ x → isConnectedFun x (suspMapΩ {A = A} (suc m))) - (sym (help m) ∙ cong (_∸ m) (sym q)) + (sym (n∸n≡0 m) ∙ cong (_∸ m) (sym q)) λ b → tt* , (λ {tt* → refl}) - where - help : (n : ℕ) → n ∸ n ≡ 0 - help zero = refl - help (suc n) = help n ... | (gt p) = isConnectedCong' (n + suc n ∸ m) (suspMapΩ {A = A} m) (subst (λ x → isConnectedFun x (suspMapΩ {A = A} m)) - (sym (help (n + suc n) m p)) + (sym (suc∸-fst (n + suc n) m p)) (suspMapΩ-connected n m connA)) (snd (suspMapΩ∙ m)) - where - help : (n m : ℕ) → m < n → suc (n ∸ m) ≡ (suc n) ∸ m - help zero zero p = refl - help zero (suc m) p = ⊥-rec (¬-<-zero p) - help (suc n) zero p = refl - help (suc n) (suc m) p = (help n m (pred-≤-pred p)) diff --git a/Cubical/Homotopy/Group/Properties.agda b/Cubical/Homotopy/Group/Properties.agda index 26058f02d..585493035 100644 --- a/Cubical/Homotopy/Group/Properties.agda +++ b/Cubical/Homotopy/Group/Properties.agda @@ -359,7 +359,9 @@ snd (π'Gr≅πGr n A) = λ p q i → ∣ IsoSphereMapΩ-pres∙Π n p q i ∣₂) {- -In file X we gave a filler of the following square +In the file SuspensionMapPathP we gave a filler +of the following square: + suspMapΩ Ωⁿ A -------------------> Ωⁿ⁺¹ (Susp A) | | @@ -510,12 +512,10 @@ private (sym (snd g) ∙∙ cong-g ∙∙ (λ j → snd g (j ∧ ~ k)))) i}) (((cong (λ x → rUnit x j) c-f) ∙ (cong (λ x → lUnit x j) c-g)) i)))) - hom-botᵣ⁻ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → (f g : S₊∙ (suc n) →∙ Susp∙ (typ A)) - → botᵣ⁻ {A = A} n (∙Π f g) - ≡ invComp {A = A} n (botᵣ⁻ {A = A} n f) (botᵣ⁻ {A = A} n g) - --- We ge that +hom-botᵣ⁻ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (f g : S₊∙ (suc n) →∙ Susp∙ (typ A)) + → botᵣ⁻ {A = A} n (∙Π f g) + ≡ invComp {A = A} n (botᵣ⁻ {A = A} n f) (botᵣ⁻ {A = A} n g) hom-botᵣ⁻ zero f g = ΣPathP ((funExt (λ { false → sym (rUnit _) ; true → (rUnit _)})) diff --git a/Cubical/Homotopy/Group/S3.agda b/Cubical/Homotopy/Group/S3.agda index 65f015bf2..e798b5b62 100644 --- a/Cubical/Homotopy/Group/S3.agda +++ b/Cubical/Homotopy/Group/S3.agda @@ -29,7 +29,7 @@ open import Cubical.HITs.SetTruncation open import Cubical.Algebra.Group renaming (ℤ to ℤGroup ; Bool to BoolGroup ; Unit to UnitGroup) -open import Cubical.Algebra.Group.ZModule +open import Cubical.Algebra.Group.ZAction -- Some type abbreviations (unproved results) @@ -136,7 +136,7 @@ groupLem : {G : Group₀} → (fst ϕ g ≡ 1) ⊎ (fst ϕ g ≡ -1) → isEquiv (fst ϕ) groupLem {G = G} s = - JGroupEquiv + GroupEquivJ (λ G _ → (g : fst G) → gen₁-by G g → (ϕ : GroupHom G ℤGroup) @@ -171,7 +171,7 @@ module π₄S³ lem : ∀ {G : Group₀} (ϕ ψ : GroupEquiv ℤGroup G) (g : fst G) → abs (invEq (fst ϕ) g) ≡ abs (invEq (fst ψ) g) lem = - JGroupEquiv + GroupEquivJ (λ G ϕ → (ψ : GroupEquiv ℤGroup G) (g : fst G) → abs (invEq (fst ϕ) g) ≡ abs (invEq (fst ψ) g)) λ ψ → mini-lem₂ (invGroupEquiv ψ) diff --git a/Cubical/Homotopy/Group/SuspensionMapPathP.agda b/Cubical/Homotopy/Group/SuspensionMapPathP.agda index 8487e5dc7..42ba0841f 100644 --- a/Cubical/Homotopy/Group/SuspensionMapPathP.agda +++ b/Cubical/Homotopy/Group/SuspensionMapPathP.agda @@ -123,19 +123,6 @@ lMapId2 (suc (suc n)) {A = A} = ; (merid a i) j → ∙∙lCancel (lMapId (suc n) {A = A} a) j i})) , refl) -{- -Ω (Ωⁿ A) -----------> Ω (Ωⁿ⁺¹ ΣA) - | | - | | - | | - v v - Ω (Sⁿ →∙ A) ----> Ω (Sⁿ⁺¹ →∙ ΣA) - | | - | | - | | - v v - (Sⁿ⁺¹ →∙ A) ----> Sⁿ⁺² →∙ Σ A --} -- We define the following maps which will be used to -- show that lMap is an equivalence @@ -227,7 +214,7 @@ snd (rightInv (SphereMapΩIso {A = A} (suc n)) p k i) j = ; (j = i1) → pt A ; (k = i0) → help Ωp (merid (ptSn (suc n))) r j i ; (k = i1) → snd (p i) j}) - (cool2 p-refl (sym (flipSquare (cong snd p))) k j i))) + (lem p-refl (sym (flipSquare (cong snd p))) k j i))) where Ωp = (fst (ΩSphereMap {A = A} (suc n) p)) wrap-refl : ∀ {ℓ} {A : Type ℓ} {x : A} {r s : x ≡ x} (p : r ≡ s) → _ @@ -242,10 +229,10 @@ snd (rightInv (SphereMapΩIso {A = A} (suc n)) p k i) j = ; (i = i1) → f (r k) }) (inS (f (q i))) k - cool2 : ∀ {ℓ} {A : Pointed ℓ} (p : Path (typ A) (pt A) (pt A)) (q : refl ≡ p) + lem : ∀ {ℓ} {A : Pointed ℓ} (p : Path (typ A) (pt A) (pt A)) (q : refl ≡ p) → PathP (λ i → (cong (p ∙_) (cong sym (sym q)) ∙ sym (rUnit p)) i ≡ refl) (rCancel p) (sym q) - cool2 {A = A} p = + lem {A = A} p = J (λ p q → PathP (λ i → (cong (p ∙_) (cong sym (sym q)) ∙ sym (rUnit p)) i ≡ refl) (rCancel p) (sym q)) @@ -455,7 +442,7 @@ snd (botᵣ (suc n) (f , p)) = refl {- The goal now is to fill the following diagram. - suspMap + suspMapΩ Ωⁿ A -------------------> Ω¹⁺ⁿ (Susp A) | | | | diff --git a/Cubical/Homotopy/Hopf.agda b/Cubical/Homotopy/Hopf.agda index 4b2db2b93..406315361 100644 --- a/Cubical/Homotopy/Hopf.agda +++ b/Cubical/Homotopy/Hopf.agda @@ -15,15 +15,20 @@ open import Cubical.Foundations.Univalence open import Cubical.Data.Sigma open import Cubical.Data.Unit +open import Cubical.Data.Int hiding (_·_) open import Cubical.HITs.Pushout.Flattening open import Cubical.HITs.Pushout open import Cubical.HITs.Sn open import Cubical.HITs.Susp -open import Cubical.HITs.S1 renaming (_·_ to _*_) +open import Cubical.HITs.S1 +open import Cubical.HITs.S2 +open import Cubical.HITs.S3 open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim) open import Cubical.HITs.Join +open import Cubical.HITs.Interval + renaming ( zero to I0 ; one to I1 ) open Iso open HSpace @@ -75,7 +80,7 @@ module Hopf {ℓ : Level} {A : Pointed ℓ} {e : HSpace A} merid y i₁ , ua-gluePt (μ-eq y) i₁ x joinIso₁ : Iso (TotalSpaceHopfPush) (join (typ A) (typ A)) - joinIso₁ = iso F G s r + joinIso₁ = theIso where F : TotalSpaceHopfPush → join (typ A) (typ A) F (inl x) = inl x @@ -118,10 +123,16 @@ module Hopf {ℓ : Level} {A : Pointed ℓ} {e : HSpace A} ; (j = i1) → push (x , retEq (μ-eq' x) y k) i}) ((push (x , invEq (μ-eq' x) (μ e x y))) i)) + theIso : Iso TotalSpaceHopfPush (join (typ A) (typ A)) + fun theIso = F + inv theIso = G + rightInv theIso = s + leftInv theIso = r + isEquivTotalSpaceHopfPush→TotalSpace : isEquiv TotalSpaceHopfPush→TotalSpace isEquivTotalSpaceHopfPush→TotalSpace = - isoToIsEquiv (iso _ inv' sect retr) + isoToIsEquiv theIso where inv' : _ → _ inv' (north , y) = inl y @@ -175,6 +186,12 @@ module Hopf {ℓ : Level} {A : Pointed ℓ} {e : HSpace A} ; (j = i1) → push (x , y) (i ∨ ~ k)}) (inr (μ e x y)) + theIso : Iso TotalSpaceHopfPush (Σ (Susp (typ A)) Hopf) + fun theIso = TotalSpaceHopfPush→TotalSpace + inv theIso = inv' + rightInv theIso = sect + leftInv theIso = retr + IsoTotalSpaceJoin : Iso (Σ[ x ∈ Susp (typ A) ] Hopf x) (join (typ A) (typ A)) IsoTotalSpaceJoin = compIso (equivToIso (invEquiv (_ , isEquivTotalSpaceHopfPush→TotalSpace))) @@ -363,3 +380,325 @@ module Hopf {ℓ : Level} {A : Pointed ℓ} {e : HSpace A} (compIso (equivToIso (joinPushout≃join _ _)) (pathToIso (cong (join (typ A)) (isoToPath IsoTotalSpaceJoin))))) + + +-- Direct construction of Hopf fibration for S¹ +module S¹Hopf where + Border : (x : S¹) → (j : I) → Partial (j ∨ ~ j) (Σ Type₀ (λ T → T ≃ S¹)) + Border x j (j = i0) = S¹ , (x ·_) , rotIsEquiv x + Border x j (j = i1) = S¹ , idEquiv S¹ + + -- Hopf fibration using SuspS¹ + HopfSuspS¹ : SuspS¹ → Type₀ + HopfSuspS¹ north = S¹ + HopfSuspS¹ south = S¹ + HopfSuspS¹ (merid x j) = Glue S¹ (Border x j) + + -- Hopf fibration using S² + -- TODO : prove that it is equivalent to HopfSuspS¹ + HopfS² : S² → Type₀ + HopfS² base = S¹ + HopfS² (surf i j) = Glue S¹ (λ { (i = i0) → _ , idEquiv S¹ + ; (i = i1) → _ , idEquiv S¹ + ; (j = i0) → _ , idEquiv S¹ + ; (j = i1) → _ , _ , rotIsEquiv (loop i) } ) + + -- Hopf fibration using more direct definition of the rot equivalence + -- TODO : prove that it is equivalent to HopfSuspS¹ + HopfS²' : S² → Type₀ + HopfS²' base = S¹ + HopfS²' (surf i j) = Glue S¹ (λ { (i = i0) → _ , rotLoopEquiv i0 + ; (i = i1) → _ , rotLoopEquiv i0 + ; (j = i0) → _ , rotLoopEquiv i0 + ; (j = i1) → _ , rotLoopEquiv i } ) + + -- Total space of the fibration + TotalHopf : Type₀ + TotalHopf = Σ SuspS¹ HopfSuspS¹ + + -- Forward direction + filler-1 : I → (j : I) → (y : S¹) → Glue S¹ (Border y j) → join S¹ S¹ + filler-1 i j y x = hfill (λ t → λ { (j = i0) → inl (rotInv-1 x y t) + ; (j = i1) → inr x }) + (inS (push ((unglue (j ∨ ~ j) x) · invLooper y) (unglue (j ∨ ~ j) x) j)) i + + TotalHopf→JoinS¹S¹ : TotalHopf → join S¹ S¹ + TotalHopf→JoinS¹S¹ (north , x) = inl x + TotalHopf→JoinS¹S¹ (south , x) = inr x + TotalHopf→JoinS¹S¹ (merid y j , x) = filler-1 i1 j y x + + -- Backward direction + JoinS¹S¹→TotalHopf : join S¹ S¹ → TotalHopf + JoinS¹S¹→TotalHopf (inl x) = (north , x) + JoinS¹S¹→TotalHopf (inr x) = (south , x) + JoinS¹S¹→TotalHopf (push y x j) = + (merid (invLooper y · x) j + , glue (λ { (j = i0) → y ; (j = i1) → x }) (rotInv-2 x y j)) + + -- Now for the homotopies, we will need to fill squares indexed by x y : S¹ with value in S¹ + -- Some will be extremeley tough, but happen to be easy when x = y = base + -- therefore, we fill them for x = y = base and then use the connectedness of S¹ × S¹ and + -- the discreteness of ΩS¹ to get general fillers. + + -- To proceed with that strategy, we first need a lemma : + -- the sections of the trivial fibration λ (_ : S¹) (_ : S¹) → Int are constant + + -- this should be generalized to a constant fibration over a connected space with + -- discrete fiber + fibℤ : S¹ → S¹ → Type₀ + fibℤ _ _ = ℤ + + S¹→HSet : (A : Type₀) (p : isSet A) (F : S¹ → A) (x : S¹) → F base ≡ F x + S¹→HSet A p F base = refl {x = F base} + S¹→HSet A p F (loop i) = f' i + where + f : PathP (λ i → F base ≡ F (loop i)) refl (cong F loop) + f i = λ j → F (loop (i ∧ j)) + L : cong F loop ≡ refl + L = p (F base) (F base) (f i1) refl + f' : PathP (λ i → F base ≡ F (loop i)) (refl {x = F base}) (refl {x = F base}) + f' = transport (λ i → PathP (λ j → F base ≡ F (loop j)) refl (L i)) f + + constant-loop : (F : S¹ → S¹ → ℤ) → (x y : S¹) → F base base ≡ F x y + constant-loop F x y = L0 ∙ L1 + where + p : isSet (S¹ → ℤ) + p = isSetΠ (λ _ → isSetℤ) + L : F base ≡ F x + L = S¹→HSet (S¹ → ℤ) p F x + L0 : F base base ≡ F x base + L0 i = L i base + L1 : F x base ≡ F x y + L1 = S¹→HSet ℤ isSetℤ (F x) y + + discretefib : (F : S¹ → S¹ → Type₀) → Type₀ + discretefib F = (a : (x y : S¹) → F x y) → + (b : (x y : S¹) → F x y) → + (a base base ≡ b base base) → + (x y : S¹) → a x y ≡ b x y + + discretefib-fibℤ : discretefib fibℤ + discretefib-fibℤ a b h x y i = + hcomp (λ t → λ { (i = i0) → constant-loop a x y t + ; (i = i1) → constant-loop b x y t }) + (h i) + + -- first homotopy + + assocFiller-3-aux : I → I → I → I → S¹ + assocFiller-3-aux x y j i = + hfill (λ t → λ { (i = i0) → rotInv-1 (loop y) (loop (~ y) · loop x) t + ; (i = i1) → rotInv-3 (loop y) (loop x) t + ; (x = i0) (y = i0) → base + ; (x = i0) (y = i1) → base + ; (x = i1) (y = i0) → base + ; (x = i1) (y = i1) → base }) + (inS ((rotInv-2 (loop x) (loop y) i) · (invLooper (loop (~ y) · loop x)))) j + + -- assocFiller-3-endpoint is used only in the type of the next function, to specify the + -- second endpoint. + -- However, I only need the first endpoint, but I cannot specify only one of them as is. + -- TODO : use cubical extension types when available to remove assocFiller-3-endpoint + assocFiller-3-endpoint : (x : S¹) → (y : S¹) → y ≡ y + assocFiller-3-endpoint base base i = base + assocFiller-3-endpoint (loop x) base i = assocFiller-3-aux x i0 i1 i + assocFiller-3-endpoint base (loop y) i = assocFiller-3-aux i0 y i1 i + assocFiller-3-endpoint (loop x) (loop y) i = assocFiller-3-aux x y i1 i + + assocFiller-3 : (x : S¹) → (y : S¹) → + PathP (λ j → rotInv-1 y (invLooper y · x) j ≡ rotInv-3 y x j) + (λ i → ((rotInv-2 x y i) · (invLooper (invLooper y · x)))) + (assocFiller-3-endpoint x y) + assocFiller-3 base base j i = base + assocFiller-3 (loop x) base j i = assocFiller-3-aux x i0 j i + assocFiller-3 base (loop y) j i = assocFiller-3-aux i0 y j i + assocFiller-3 (loop x) (loop y) j i = assocFiller-3-aux x y j i + + assoc-3 : (_ y : S¹) → basedΩS¹ y + assoc-3 x y i = assocFiller-3 x y i1 i + + fibℤ≡fibAssoc-3 : fibℤ ≡ (λ _ y → basedΩS¹ y) + fibℤ≡fibAssoc-3 i = λ x y → basedΩS¹≡ℤ y (~ i) + + discretefib-fibAssoc-3 : discretefib (λ _ y → basedΩS¹ y) + discretefib-fibAssoc-3 = + transp (λ i → discretefib (fibℤ≡fibAssoc-3 i)) i0 discretefib-fibℤ + + assocConst-3 : (x y : S¹) → assoc-3 x y ≡ refl + assocConst-3 x y = discretefib-fibAssoc-3 assoc-3 (λ _ _ → refl) refl x y + + assocSquare-3 : I → I → S¹ → S¹ → S¹ + assocSquare-3 i j x y = hcomp (λ t → λ { (i = i0) → assocFiller-3 x y j i0 + ; (i = i1) → assocFiller-3 x y j i1 + ; (j = i0) → assocFiller-3 x y i0 i + ; (j = i1) → assocConst-3 x y t i }) + (assocFiller-3 x y j i) + + filler-3 : I → I → S¹ → S¹ → join S¹ S¹ + filler-3 i j y x = + hcomp (λ t → λ { (i = i0) → filler-1 t j (invLooper y · x) + (glue (λ { (j = i0) → y ; (j = i1) → x }) + (rotInv-2 x y j)) + ; (i = i1) → push (rotInv-3 y x t) x j + ; (j = i0) → inl (assocSquare-3 i t x y) + ; (j = i1) → inr x }) + (push ((rotInv-2 x y (i ∨ j)) · (invLooper (invLooper y · x))) (rotInv-2 x y (i ∨ j)) j) + + JoinS¹S¹→TotalHopf→JoinS¹S¹ : ∀ x → TotalHopf→JoinS¹S¹ (JoinS¹S¹→TotalHopf x) ≡ x + JoinS¹S¹→TotalHopf→JoinS¹S¹ (inl x) i = inl x + JoinS¹S¹→TotalHopf→JoinS¹S¹ (inr x) i = inr x + JoinS¹S¹→TotalHopf→JoinS¹S¹ (push y x j) i = filler-3 i j y x + + -- Second homotopy + + -- This HIT is the total space of the Hopf fibration but the ends of SuspS¹ have not been + -- glued together yet — which makes it into a cylinder. + -- This allows to write compositions that do not properly match at the endpoints. However, + -- I suspect it is unnecessary. TODO : do without PseudoHopf + + PseudoHopf : Type₀ + PseudoHopf = (S¹ × Interval) × S¹ + + PseudoHopf-π1 : PseudoHopf → S¹ + PseudoHopf-π1 ((y , _) , _) = y + + PseudoHopf-π2 : PseudoHopf → S¹ + PseudoHopf-π2 (_ , x) = x + + assocFiller-4-aux : I → I → I → I → S¹ + assocFiller-4-aux x y j i = + hfill (λ t → λ { (i = i0) → ((invLooper (loop y · loop x · loop (~ y))) · (loop y · loop x)) + · (rotInv-1 (loop x) (loop y) t) + ; (i = i1) → (rotInv-4 (loop y) (loop y · loop x) (~ t)) · loop x + ; (x = i0) (y = i0) → base + ; (x = i0) (y = i1) → base + ; (x = i1) (y = i0) → base + ; (x = i1) (y = i1) → base }) + (inS (rotInv-2 (loop y · loop x) (loop y · loop x · loop (~ y)) i)) j + + -- See assocFiller-3-endpoint + -- TODO : use cubical extension types when available to remove assocFiller-4-endpoint + assocFiller-4-endpoint : (x y : S¹) → basedΩS¹ (((invLooper (y · x · invLooper y)) · (y · x)) · x) + assocFiller-4-endpoint base base i = base + assocFiller-4-endpoint (loop x) base i = assocFiller-4-aux x i0 i1 i + assocFiller-4-endpoint base (loop y) i = assocFiller-4-aux i0 y i1 i + assocFiller-4-endpoint (loop x) (loop y) i = assocFiller-4-aux x y i1 i + + assocFiller-4 : (x y : S¹) → + PathP (λ j → ((invLooper (y · x · invLooper y)) · (y · x)) · (rotInv-1 x y j) ≡ (rotInv-4 y (y · x) (~ j)) · x) + (λ i → (rotInv-2 (y · x) (y · x · invLooper y) i)) + (assocFiller-4-endpoint x y) + assocFiller-4 base base j i = base + assocFiller-4 (loop x) base j i = assocFiller-4-aux x i0 j i + assocFiller-4 base (loop y) j i = assocFiller-4-aux i0 y j i + assocFiller-4 (loop x) (loop y) j i = assocFiller-4-aux x y j i + + assoc-4 : (x y : S¹) → basedΩS¹ (((invLooper (y · x · invLooper y)) · (y · x)) · x) + assoc-4 x y i = assocFiller-4 x y i1 i + + fibℤ≡fibAssoc-4 : fibℤ ≡ (λ x y → basedΩS¹ (((invLooper (y · x · invLooper y)) · (y · x)) · x)) + fibℤ≡fibAssoc-4 i = λ x y → basedΩS¹≡ℤ (((invLooper (y · x · invLooper y)) · (y · x)) · x) (~ i) + + discretefib-fibAssoc-4 : discretefib (λ x y → basedΩS¹ (((invLooper (y · x · invLooper y)) · (y · x)) · x)) + discretefib-fibAssoc-4 = + transp (λ i → discretefib (fibℤ≡fibAssoc-4 i)) i0 discretefib-fibℤ + + assocConst-4 : (x y : S¹) → assoc-4 x y ≡ refl + assocConst-4 x y = discretefib-fibAssoc-4 assoc-4 (λ _ _ → refl) refl x y + + assocSquare-4 : I → I → S¹ → S¹ → S¹ + assocSquare-4 i j x y = + hcomp (λ t → λ { (i = i0) → assocFiller-4 x y j i0 + ; (i = i1) → assocFiller-4 x y j i1 + ; (j = i0) → assocFiller-4 x y i0 i + ; (j = i1) → assocConst-4 x y t i }) + (assocFiller-4 x y j i) + + filler-4-0 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → PseudoHopf + filler-4-0 i j y x = + let x' = unglue (j ∨ ~ j) x in + hfill (λ t → λ { (j = i0) → ((invLooper (y · x · invLooper y) · (y · x) , I0) + , invLooper (y · x · invLooper y) · (y · x) · (rotInv-1 x y t)) + ; (j = i1) → ((invLooper (x · invLooper y) · x , I1) , x) }) + (inS ((invLooper (x' · invLooper y) · x' , seg j) , rotInv-2 x' (x' · invLooper y) j)) i + + filler-4-1 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → PseudoHopf + filler-4-1 i j y x = + let x' = unglue (j ∨ ~ j) x in + hfill (λ t → λ { (j = i0) → ((invLooper (y · x · invLooper y) · (y · x) , I0) + , (rotInv-4 y (y · x) (~ t)) · x) + ; (j = i1) → ((invLooper (x · invLooper y) · x , I1) , x) }) + (inS ((invLooper (x' · invLooper y) · x' , seg j) , unglue (j ∨ ~ j) x)) i + + filler-4-2 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → TotalHopf + filler-4-2 i j y x = + let x' = unglue (j ∨ ~ j) x in + hcomp (λ t → λ { (i = i0) → JoinS¹S¹→TotalHopf (filler-1 t j y x) + ; (i = i1) → (merid (PseudoHopf-π1 (filler-4-0 t j y x)) j + , glue (λ { (j = i0) → rotInv-1 x y t ; (j = i1) → x }) + (PseudoHopf-π2 (filler-4-0 t j y x))) + ; (j = i0) → (north , rotInv-1 x y t) + ; (j = i1) → (south , x) }) + (merid (invLooper (x' · invLooper y) · x') j + , glue (λ { (j = i0) → y · x · invLooper y ; (j = i1) → x }) (rotInv-2 x' (x' · invLooper y) j)) + + filler-4-3 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → PseudoHopf + filler-4-3 i j y x = + let x' = unglue (j ∨ ~ j) x in + hcomp (λ t → λ { (i = i0) → filler-4-0 t j y x + ; (i = i1) → filler-4-1 t j y x + ; (j = i0) → ((invLooper (y · x · invLooper y) · (y · x) , I0) , assocSquare-4 i t x y) + ; (j = i1) → ((invLooper (x · invLooper y) · x , I1) , x) }) + ((invLooper (x' · invLooper y) · x' , seg j) , rotInv-2 x' (x' · invLooper y) (i ∨ j)) + + filler-4-4 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → PseudoHopf + filler-4-4 i j y x = + let x' = unglue (j ∨ ~ j) x in + hcomp (λ t → λ { (i = i0) → filler-4-1 t j y x + ; (i = i1) → ((y , seg j) , unglue (j ∨ ~ j) x) + ; (j = i0) → ((rotInv-4 y (y · x) i , I0) + , (rotInv-4 y (y · x) (i ∨ ~ t)) · x) + ; (j = i1) → ((rotInv-4 y x i , I1) , x) }) + ((rotInv-4 y x' i , seg j) , x') + + filler-4-5 : (_ j : I) → (y : S¹) → Glue S¹ (Border y j) → TotalHopf + filler-4-5 i j y x = + hcomp (λ t → λ { (i = i0) → filler-4-2 (~ t) j y x + ; (i = i1) → (merid (PseudoHopf-π1 (filler-4-4 t j y x)) j + , glue (λ { (j = i0) → x ; (j = i1) → x }) + (PseudoHopf-π2 (filler-4-4 t j y x))) + ; (j = i0) → (north , x) + ; (j = i1) → (south , x) }) + (merid (PseudoHopf-π1 (filler-4-3 i j y x)) j + , glue (λ { (j = i0) → x ; (j = i1) → x }) (PseudoHopf-π2 (filler-4-3 i j y x))) + + TotalHopf→JoinS¹S¹→TotalHopf : ∀ x → JoinS¹S¹→TotalHopf (TotalHopf→JoinS¹S¹ x) ≡ x + TotalHopf→JoinS¹S¹→TotalHopf (north , x) i = (north , x) + TotalHopf→JoinS¹S¹→TotalHopf (south , x) i = (south , x) + TotalHopf→JoinS¹S¹→TotalHopf (merid y j , x) i = filler-4-5 i j y x + + + JoinS¹S¹≡TotalHopf : join S¹ S¹ ≡ TotalHopf + JoinS¹S¹≡TotalHopf = isoToPath (iso JoinS¹S¹→TotalHopf + TotalHopf→JoinS¹S¹ + TotalHopf→JoinS¹S¹→TotalHopf + JoinS¹S¹→TotalHopf→JoinS¹S¹) + + S³≡TotalHopf : S³ ≡ TotalHopf + S³≡TotalHopf = S³≡joinS¹S¹ ∙ JoinS¹S¹≡TotalHopf + + open Iso + IsoS³TotalHopf : Iso (S₊ 3) TotalHopf + fun IsoS³TotalHopf x = JoinS¹S¹→TotalHopf (S³→joinS¹S¹ (inv IsoS³S3 x)) + inv IsoS³TotalHopf x = fun IsoS³S3 (joinS¹S¹→S³ (TotalHopf→JoinS¹S¹ x)) + rightInv IsoS³TotalHopf x = + cong (JoinS¹S¹→TotalHopf ∘ S³→joinS¹S¹) + (leftInv IsoS³S3 (joinS¹S¹→S³ (TotalHopf→JoinS¹S¹ x))) + ∙∙ cong JoinS¹S¹→TotalHopf + (joinS¹S¹→S³→joinS¹S¹ (TotalHopf→JoinS¹S¹ x)) + ∙∙ TotalHopf→JoinS¹S¹→TotalHopf x + leftInv IsoS³TotalHopf x = + cong (fun IsoS³S3 ∘ joinS¹S¹→S³) + (JoinS¹S¹→TotalHopf→JoinS¹S¹ (S³→joinS¹S¹ (inv IsoS³S3 x))) + ∙∙ cong (fun IsoS³S3) (S³→joinS¹S¹→S³ (inv IsoS³S3 x)) + ∙∙ Iso.rightInv IsoS³S3 x diff --git a/Cubical/Homotopy/HopfInvariant/Base.agda b/Cubical/Homotopy/HopfInvariant/Base.agda index dc2cde51b..a064d4b26 100644 --- a/Cubical/Homotopy/HopfInvariant/Base.agda +++ b/Cubical/Homotopy/HopfInvariant/Base.agda @@ -35,7 +35,7 @@ open import Cubical.Data.Unit open import Cubical.Algebra.Group renaming (ℤ to ℤGroup ; Unit to UnitGroup) -open import Cubical.Algebra.Group.ZModule +open import Cubical.Algebra.Group.ZAction open import Cubical.HITs.Pushout open import Cubical.HITs.Sn @@ -48,17 +48,6 @@ open import Cubical.HITs.SetTruncation open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec) -isContr→≡UnitGroup : {G : Group ℓ-zero} → isContr (fst G) → UnitGroup ≡ G -isContr→≡UnitGroup c = - fst (GroupPath _ _) - (invGroupEquiv ((isContr→≃Unit c) - , (makeIsGroupHom (λ _ _ → refl)))) - -GroupIsoUnitGroup→isContr : {G : Group ℓ-zero} - → GroupIso UnitGroup G → isContr (fst G) -GroupIsoUnitGroup→isContr is = - isOfHLevelRetractFromIso 0 (invIso (fst is)) isContrUnit - -- The pushout of the hopf invariant HopfInvariantPush : (n : ℕ) → (f : S₊ (3 +ℕ n +ℕ n) → S₊ (2 +ℕ n)) @@ -135,22 +124,8 @@ module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) whe Hopfβ-Iso : GroupIso (coHomGr (suc (3 +ℕ n +ℕ n)) (HopfInvariantPush _ (fst f))) ℤGroup - Iso.fun (fst Hopfβ-Iso) x = - Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) - (invEq (fst SphereHopfCohomIso) x) - Iso.inv (fst Hopfβ-Iso) f = - (fst (fst SphereHopfCohomIso)) - (Iso.inv (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) f) - Iso.rightInv (fst Hopfβ-Iso) f = - cong (Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) - (retEq (fst SphereHopfCohomIso) - (Iso.inv (fst (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) f)) - ∙ Iso.rightInv (fst (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) f - Iso.leftInv (fst Hopfβ-Iso) x = - cong (fst (fst SphereHopfCohomIso)) - (Iso.leftInv (fst (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) - (invEq (fst SphereHopfCohomIso) x)) - ∙ secEq (fst SphereHopfCohomIso) x + fst Hopfβ-Iso = compIso (invIso (equivToIso (fst SphereHopfCohomIso))) + (fst (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) snd Hopfβ-Iso = grHom where abstract @@ -160,10 +135,9 @@ module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) whe (λ x → Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n)))))) (invEq (fst SphereHopfCohomIso) x)) (ℤGroup .snd) - grHom = makeIsGroupHom λ x y → - cong (Iso.fun (fst ((Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))))) - (IsGroupHom.pres· (isGroupHomInv SphereHopfCohomIso) x y) - ∙ IsGroupHom.pres· (snd (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) _ _ + grHom = snd (compGroupIso + (GroupEquiv→GroupIso (invGroupEquiv (SphereHopfCohomIso))) + (Hⁿ-Sⁿ≅ℤ (suc (suc (n +ℕ n))))) Hⁿ-Sⁿ≅ℤ-nice-generator : (n : ℕ) → Iso.inv (fst (Hⁿ-Sⁿ≅ℤ (suc n))) 1 ≡ ∣ ∣_∣ ∣₂ Hⁿ-Sⁿ≅ℤ-nice-generator zero = Iso.leftInv (fst (Hⁿ-Sⁿ≅ℤ (suc zero))) _ @@ -226,8 +200,9 @@ module _ (n : ℕ) (f : S₊∙ (3 +ℕ n +ℕ n) →∙ S₊∙ (2 +ℕ n)) whe conCohom2+n = coHomK-elim _ (λ _ → isProp→isOfHLevelSuc (suc n) squash) ∣ refl ∣ - HIPSphereCohomIso : Iso (coHom (2 +ℕ n) (HopfInvariantPush n (fst f))) - (coHom (2 +ℕ n) ((S₊ (suc (suc n))))) + HIPSphereCohomIso : + Iso (coHom (2 +ℕ n) (HopfInvariantPush n (fst f))) + (coHom (2 +ℕ n) (S₊ (2 +ℕ n))) Iso.fun HIPSphereCohomIso = H→Sphere Iso.inv HIPSphereCohomIso = Sphere→H Iso.rightInv HIPSphereCohomIso = diff --git a/Cubical/Homotopy/HopfInvariant/Homomorphism.agda b/Cubical/Homotopy/HopfInvariant/Homomorphism.agda index 08f78287c..5c360dd59 100644 --- a/Cubical/Homotopy/HopfInvariant/Homomorphism.agda +++ b/Cubical/Homotopy/HopfInvariant/Homomorphism.agda @@ -33,7 +33,7 @@ open import Cubical.Data.Unit open import Cubical.Algebra.Group renaming (ℤ to ℤGroup ; Unit to UnitGroup) -open import Cubical.Algebra.Group.ZModule +open import Cubical.Algebra.Group.ZAction open import Cubical.HITs.Pushout open import Cubical.HITs.Sn @@ -713,36 +713,6 @@ jᵣ-βᵣ n f g = ; (inr x) → refl ; (push a i₁) → refl})) -gen₂ℤ×ℤ : gen₂-by (DirProd ℤGroup ℤGroup) (1 , 0) (0 , 1) -fst (gen₂ℤ×ℤ (x , y)) = x , y -snd (gen₂ℤ×ℤ (x , y)) = - ΣPathP ((cong₂ _+_ ((·Comm 1 x) ∙ cong fst (sym (distrLem 1 0 x))) - ((·Comm 0 y) ∙ cong fst (sym (distrLem 0 1 y)))) - , +Comm y 0 - ∙ cong₂ _+_ (·Comm 0 x ∙ cong snd (sym (distrLem 1 0 x))) - (·Comm 1 y ∙ cong snd (sym (distrLem 0 1 y)))) - where - ℤ×ℤ = DirProd ℤGroup ℤGroup - _+''_ = GroupStr._·_ (snd ℤ×ℤ) - - -lem : (x : ℤ) → - x ≡ 0 - x - -lem (pos zero) = refl - -lem (pos (suc zero)) = refl - -lem (pos (suc (suc n))) = - cong predℤ (-lem (pos (suc n))) - -lem (negsuc zero) = refl - -lem (negsuc (suc n)) = cong sucℤ (-lem (negsuc n)) - - distrLem : (x y : ℤ) (z : ℤ) - → Path (ℤ × ℤ) (z ℤ[ ℤ×ℤ ]· (x , y)) (z · x , z · y) - distrLem x y (pos zero) = refl - distrLem x y (pos (suc n)) = - (cong ((x , y) +''_) (distrLem x y (pos n))) - distrLem x y (negsuc zero) = ΣPathP (sym (-lem x) , sym (-lem y)) - distrLem x y (negsuc (suc n)) = - cong₂ _+''_ (ΣPathP (sym (-lem x) , sym (-lem y))) - (distrLem x y (negsuc n)) - genH²ⁿC* : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → gen₂-by (coHomGr ((2 +ℕ n) +' (2 +ℕ n)) (C* n f g)) (βₗ n f g) @@ -762,6 +732,7 @@ private X : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → ℤ X n f g = (genH²ⁿC* n f g) (α* n f g ⌣ α* n f g) .fst .fst + Y : (n : ℕ) (f g : S₊∙ (3 +ℕ (n +ℕ n)) →∙ S₊∙ (2 +ℕ n)) → ℤ Y n f g = (genH²ⁿC* n f g) (α* n f g ⌣ α* n f g) .fst .snd diff --git a/Cubical/Homotopy/HopfInvariant/HopfMap.agda b/Cubical/Homotopy/HopfInvariant/HopfMap.agda index 95e727d90..f72113b15 100644 --- a/Cubical/Homotopy/HopfInvariant/HopfMap.agda +++ b/Cubical/Homotopy/HopfInvariant/HopfMap.agda @@ -7,6 +7,7 @@ module Cubical.Homotopy.HopfInvariant.HopfMap where open import Cubical.Homotopy.HopfInvariant.Base open import Cubical.Homotopy.Hopf +open S¹Hopf open import Cubical.Homotopy.Connected open import Cubical.Homotopy.HSpace @@ -37,10 +38,9 @@ open import Cubical.Data.Unit open import Cubical.Algebra.Group renaming (ℤ to ℤGroup ; Unit to UnitGroup) -open import Cubical.Algebra.Group.ZModule +open import Cubical.Algebra.Group.ZAction open import Cubical.HITs.Pushout -open import Cubical.HITs.Hopf open import Cubical.HITs.Join open import Cubical.HITs.S1 renaming (_·_ to _*_) open import Cubical.HITs.Sn @@ -286,71 +286,13 @@ inrInjective f g = pRec (squash₂ _ _) -- A couple of basic lemma concerning the hSpace structure on S¹ private - rUnit* : (x : S¹) → x * base ≡ x - rUnit* base = refl - rUnit* (loop i₁) = refl - - merid*-lem : (a x : S¹) - → Path (Path (coHomK 2) _ _) - (cong ∣_∣ₕ (merid (a * x) ∙ sym (merid base))) - ((cong ∣_∣ₕ (merid a ∙ sym (merid base))) - ∙ (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) - merid*-lem = wedgeconFun _ _ (λ _ _ → isOfHLevelTrunc 4 _ _ _ _) - (λ x → lUnit _ - ∙ cong (_∙ cong ∣_∣ₕ (merid x ∙ sym (merid base))) - (cong (cong ∣_∣ₕ) (sym (rCancel (merid base))))) - (λ x → (λ i → cong ∣_∣ₕ (merid (rUnit* x i) ∙ sym (merid base))) - ∙∙ rUnit _ - ∙∙ cong (cong ∣_∣ₕ (merid x ∙ sym (merid base)) ∙_) - (cong (cong ∣_∣ₕ) (sym (rCancel (merid base))))) - (sym (l (cong ∣_∣ₕ (merid base ∙ sym (merid base))) - (cong (cong ∣_∣ₕ) (sym (rCancel (merid base)))))) - where - l : ∀ {ℓ} {A : Type ℓ} {x : A} (p : x ≡ x) (P : refl ≡ p) - → lUnit p ∙ cong (_∙ p) P ≡ rUnit p ∙ cong (p ∙_) P - l p = J (λ p P → lUnit p ∙ cong (_∙ p) P ≡ rUnit p ∙ cong (p ∙_) P) refl - - lemmie : (x : S¹) → ptSn 1 ≡ x * (invLooper x) - lemmie base = refl - lemmie (loop i) j = - hcomp (λ r → λ {(i = i0) → base ; (i = i1) → base ; (j = i0) → base}) - base - - - meridInvLooperLem : (x : S¹) → Path (Path (coHomK 2) _ _) - (cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base))) - (cong ∣_∣ₕ (sym (merid x ∙ sym (merid base)))) - meridInvLooperLem x = (lUnit _ - ∙∙ cong (_∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base))) - (sym (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base))))) - ∙∙ sym (assoc _ _ _)) - ∙∙ cong (sym (cong ∣_∣ₕ (merid x ∙ sym (merid base))) ∙_) lem - ∙∙ (assoc _ _ _ - ∙∙ cong (_∙ (cong ∣_∣ₕ (sym (merid x ∙ sym (merid base))))) - (lCancel (cong ∣_∣ₕ (merid x ∙ sym (merid base)))) - ∙∙ sym (lUnit _)) - where - lem : cong ∣_∣ₕ (merid x ∙ sym (merid base)) - ∙ cong ∣_∣ₕ (merid (invLooper x) ∙ sym (merid base)) - ≡ cong ∣_∣ₕ (merid x ∙ sym (merid base)) - ∙ cong ∣_∣ₕ (sym (merid x ∙ sym (merid base))) - lem = sym (merid*-lem x (invLooper x)) - ∙ ((λ i → cong ∣_∣ₕ (merid (lemmie x (~ i)) ∙ sym (merid base))) - ∙ cong (cong ∣_∣ₕ) (rCancel (merid base))) ∙ sym (rCancel _) - - comm·S¹ : (a x : S¹) → a * x ≡ x * a - comm·S¹ = wedgeconFun _ _ (λ _ _ → isGroupoidS¹ _ _) - (sym ∘ rUnit*) - rUnit* - refl - invLooperLem₁ : (a x : S¹) → (invEq (hopfS¹.μ-eq a) x) * a ≡ (invLooper a * x) * a invLooperLem₁ a x = secEq (hopfS¹.μ-eq a) x - ∙∙ cong (_* x) (lemmie a) + ∙∙ cong (_* x) (rCancelS¹ a) ∙∙ AssocHSpace.μ-assoc S1-AssocHSpace a (invLooper a) x - ∙ comm·S¹ _ _ + ∙ commS¹ _ _ invLooperLem₂ : (a x : S¹) → invEq (hopfS¹.μ-eq a) x ≡ invLooper a * x invLooperLem₂ a x = sym (retEq (hopfS¹.μ-eq a) (invEq (hopfS¹.μ-eq a) x)) @@ -401,14 +343,14 @@ Gysin-e≡genCP² = ∙∙ cong ((cong ∣_∣ₕ) (sym (merid a)) ∙_) (cong (cong ∣_∣ₕ) (cong sym (symDistr (merid base) (sym (merid (invLooper a * x))))) - ∙ cong sym (merid*-lem (invLooper a) x) + ∙ cong sym (SuspS¹-hom (invLooper a) x) ∙ symDistr ((cong ∣_∣ₕ) (merid (invLooper a) ∙ sym (merid base))) ((cong ∣_∣ₕ) (merid x ∙ sym (merid base))) ∙ isCommΩK 2 (sym (λ i₁ → ∣ (merid x ∙ (λ i₂ → merid base (~ i₂))) i₁ ∣)) (sym (λ i₁ → ∣ (merid (invLooper a) ∙ (λ i₂ → merid base (~ i₂))) i₁ ∣)) - ∙ cong₂ _∙_ (cong sym (meridInvLooperLem a) + ∙ cong₂ _∙_ (cong sym (SuspS¹-inv a) ∙ cong-∙ ∣_∣ₕ (merid a) (sym (merid base))) (cong (cong ∣_∣ₕ) (symDistr (merid x) (sym (merid base))) ∙ cong-∙ ∣_∣ₕ (merid base) (sym (merid x)))) @@ -594,7 +536,7 @@ CP2≡CP²' = -- The hopf invariant is ±1 for both definitions of the hopf map HopfInvariant-HopfMap' : - abs (HopfInvariant zero (HopfMap' , λ _ → HopfMap' (snd (S₊∙ 3)))) ≡ suc zero + abs (HopfInvariant zero (HopfMap' , λ _ → HopfMap' (snd (S₊∙ 3)))) ≡ 1 HopfInvariant-HopfMap' = cong abs (cong (Iso.fun (fst (Hopfβ-Iso zero (HopfMap' , refl)))) (transportRefl (⌣-α 0 (HopfMap' , refl)))) @@ -611,6 +553,6 @@ HopfInvariant-HopfMap' = (abs (fst (fst ϕ) GysinS².e) ≡ 1) l p = (GroupIso→GroupEquiv (fst p)) , (snd p) -HopfInvariant-HopfMap : abs (HopfInvariant zero HopfMap) ≡ suc zero +HopfInvariant-HopfMap : abs (HopfInvariant zero HopfMap) ≡ 1 HopfInvariant-HopfMap = cong abs (cong (HopfInvariant zero) hopfMap≡HopfMap') ∙ HopfInvariant-HopfMap' diff --git a/Cubical/Papers/Synthetic.agda b/Cubical/Papers/Synthetic.agda index 5fe73c927..d834d8270 100644 --- a/Cubical/Papers/Synthetic.agda +++ b/Cubical/Papers/Synthetic.agda @@ -39,7 +39,7 @@ import Cubical.HITs.Pushout.Properties as PushProp import Cubical.HITs.Join as Join import Cubical.HITs.Join.Properties as JoinProp -- 5. -import Cubical.HITs.Hopf as Hopf +import Cubical.Homotopy.Hopf as Hopf -------------------------------------------------------------------------------- -- 2. Cubical Agda @@ -207,8 +207,8 @@ open JoinProp using (join-assoc) public -- rot (denoted by _·_ here) in the paper is substituted by a rot and rotLoop in S1 open S1 using (_·_ ; rotLoop) public -open Hopf renaming ( HopfSuspS¹ to Hopf - ; JoinS¹S¹→TotalHopf to j2h - ; TotalHopf→JoinS¹S¹ to h2j) - using (HopfS²) public +open Hopf.S¹Hopf renaming ( HopfSuspS¹ to Hopf + ; JoinS¹S¹→TotalHopf to j2h + ; TotalHopf→JoinS¹S¹ to h2j) + using (HopfS²) public open S1 renaming (rotInv-1 to lem-rot-inv) public diff --git a/Cubical/Talks/EPA2020.agda b/Cubical/Talks/EPA2020.agda index d79f8f580..dc1f2e0b5 100644 --- a/Cubical/Talks/EPA2020.agda +++ b/Cubical/Talks/EPA2020.agda @@ -279,7 +279,8 @@ open import Cubical.HITs.Pushout -- We also defined the Hopf fibration and proved that its total space -- is S³ in about 300LOC: -open import Cubical.HITs.Hopf +open import Cubical.Homotopy.Hopf +open S¹Hopf -- There is also some integer cohomology: open import Cubical.ZCohomology.Everything diff --git a/Cubical/ZCohomology/Groups/CP2.agda b/Cubical/ZCohomology/Groups/CP2.agda index f7a7f0483..0d7066092 100644 --- a/Cubical/ZCohomology/Groups/CP2.agda +++ b/Cubical/ZCohomology/Groups/CP2.agda @@ -32,7 +32,6 @@ open import Cubical.HITs.Pushout open import Cubical.HITs.S1 open import Cubical.HITs.Sn open import Cubical.HITs.Susp -open import Cubical.HITs.Hopf open import Cubical.HITs.Join open import Cubical.HITs.SetTruncation renaming (rec to sRec ; elim to sElim ; elim2 to sElim2 ; map to sMap) @@ -42,6 +41,9 @@ open import Cubical.HITs.Truncation open import Cubical.Relation.Nullary +open import Cubical.Homotopy.Hopf +open S¹Hopf + open IsGroupHom open Iso diff --git a/Cubical/ZCohomology/Gysin.agda b/Cubical/ZCohomology/Gysin.agda index 05d4dd311..ee91d7f1a 100644 --- a/Cubical/ZCohomology/Gysin.agda +++ b/Cubical/ZCohomology/Gysin.agda @@ -12,8 +12,8 @@ open import Cubical.ZCohomology.Groups.Sn open import Cubical.ZCohomology.RingStructure.CupProduct open import Cubical.ZCohomology.RingStructure.RingLaws open import Cubical.ZCohomology.RingStructure.GradedCommutativity + open import Cubical.Relation.Nullary -open import Cubical.Homotopy.Group.Base open import Cubical.Functions.Embedding @@ -30,10 +30,6 @@ open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Equiv open import Cubical.Foundations.Pointed.Homogeneous -open import Cubical.Foundations.Univalence - -open import Cubical.Relation.Nullary - open import Cubical.Data.Sum open import Cubical.Data.Fin open import Cubical.Data.Empty renaming (rec to ⊥-rec) @@ -45,11 +41,10 @@ open import Cubical.Data.Unit open import Cubical.Data.Bool open import Cubical.Algebra.Group renaming (ℤ to ℤGroup ; Unit to UnitGroup) hiding (Bool) -open import Cubical.Algebra.Group.ZModule +open import Cubical.Algebra.Group.ZAction +open import Cubical.Algebra.AbGroup open import Cubical.HITs.Pushout.Flattening -open import Cubical.Homotopy.Connected -open import Cubical.Homotopy.EilenbergSteenrod open import Cubical.HITs.Pushout open import Cubical.HITs.Sn open import Cubical.HITs.Susp @@ -60,19 +55,12 @@ open import Cubical.HITs.SetTruncation renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; map to sMap ; elim3 to sElim3) open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec ; elim to pElim) - -open import Cubical.Algebra.AbGroup - -open import Cubical.Homotopy.Loopspace - open import Cubical.HITs.Join +open import Cubical.Homotopy.Connected open import Cubical.Homotopy.Hopf - - -open import Cubical.Algebra.AbGroup - open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Group.Base -- There seems to be some problems with the termination checker. -- Spelling out integer induction with 3 base cases like this @@ -176,17 +164,8 @@ fst (genFunSpace (suc n)) = ∣_∣ snd (genFunSpace (suc zero)) = refl snd (genFunSpace (suc (suc n))) = refl -π₀S→ℤ : ∀ {ℓ} {A : Pointed ℓ} → Iso ((Bool , true) →∙ A) (typ A) -Iso.fun π₀S→ℤ f = fst f false -fst (Iso.inv π₀S→ℤ a) false = a -fst (Iso.inv (π₀S→ℤ {A = A}) a) true = pt A -snd (Iso.inv π₀S→ℤ a) = refl -Iso.rightInv π₀S→ℤ a = refl -Iso.leftInv π₀S→ℤ (f , p) = - ΣPathP ((funExt (λ { false → refl ; true → sym p})) , λ i j → p (~ i ∨ j)) - πS≅ℤ : (n : ℕ) → GroupIso (πS n) ℤGroup -fst (πS≅ℤ zero) = π₀S→ℤ +fst (πS≅ℤ zero) = IsoBool→∙ snd (πS≅ℤ zero) = makeIsGroupHom λ _ _ → refl πS≅ℤ (suc n) = compGroupIso @@ -195,12 +174,14 @@ snd (πS≅ℤ zero) = makeIsGroupHom λ _ _ → refl (GroupEquiv→GroupIso (coHomGr≅coHomRedGr n (S₊∙ (suc n)))) (Hⁿ-Sⁿ≅ℤ n)) - Iso-πS-ℤ : (n : ℕ) → Iso (S₊∙ (suc n) →∙ coHomK-ptd (suc n)) ℤ Iso-πS-ℤ n = compIso (invIso (setTruncIdempotentIso (isOfHLevel↑∙' 0 n))) (compIso (equivToIso (fst (coHomGr≅coHomRedGr n (S₊∙ (suc n))))) (fst (Hⁿ-Sⁿ≅ℤ n))) +Iso-πS-ℤ' : (n : ℕ) → Iso (S₊∙ n →∙ coHomK-ptd n) ℤ +Iso-πS-ℤ' n = fst (πS≅ℤ n) + Iso-πS-ℤPres1 : (n : ℕ) → Iso.fun (fst (πS≅ℤ (suc n))) (∣_∣ , refl) ≡ pos 1 Iso-πS-ℤPres1 zero = refl Iso-πS-ℤPres1 (suc n) = @@ -218,14 +199,6 @@ Iso-πS-ℤPres1 (suc n) = ∣)))) i) ∙ Iso.leftInv (Iso-Kn-ΩKn+1 (suc (suc n))) ∣ x ∣) -Iso-πS-ℤPres1← : (n : ℕ) - → Iso.inv (fst (πS≅ℤ (suc n))) (pos 1) ≡ (∣_∣ , refl) -Iso-πS-ℤPres1← n = sym (cong (Iso.inv (fst (πS≅ℤ (suc n)))) (Iso-πS-ℤPres1 n)) - ∙ Iso.leftInv (fst (πS≅ℤ (suc n))) (∣_∣ , refl) - -IsoFunSpace : (n : ℕ) → Iso (S₊∙ n →∙ coHomK-ptd n) ℤ -IsoFunSpace n = fst (πS≅ℤ n) - -- The first step of the Gysin sequence is to formulate -- an equivalence g : Kᵢ ≃ (Sⁿ →∙ Kᵢ₊ₙ) module g-base where @@ -424,10 +397,10 @@ module g-base where +'-suc (suc n) (suc m) = refl decomposeG : (i n : ℕ) (x : coHomK i) - → Path _ (G (suc n) i x) - (subst (λ x → S₊∙ (suc n) →∙ coHomK-ptd x) - (sym (+'-suc n i)) - ((Iso.inv (suspKn-Iso n _)) (G n i x))) + → G (suc n) i x + ≡ subst (λ x → S₊∙ (suc n) →∙ coHomK-ptd x) + (sym (+'-suc n i)) + (Iso.inv (suspKn-Iso n _) (G n i x)) decomposeG zero zero x = →∙Homogeneous≡ (isHomogeneousKn _) (funExt λ z → (λ i → x ·₀ ∣ z ∣) ∙ h3 x z ∙ sym (transportRefl _)) diff --git a/Cubical/ZCohomology/RingStructure/RingLaws.agda b/Cubical/ZCohomology/RingStructure/RingLaws.agda index 48b9ca067..c0f215de7 100644 --- a/Cubical/ZCohomology/RingStructure/RingLaws.agda +++ b/Cubical/ZCohomology/RingStructure/RingLaws.agda @@ -588,13 +588,16 @@ assoc-⌣ₖ (suc n) (suc m) (suc k) x y z = -- Ring laws for ⌣ module _ {A : Type ℓ} (n m : ℕ) where ⌣-0ₕ : (f : coHom n A) → (f ⌣ 0ₕ m) ≡ 0ₕ _ - ⌣-0ₕ = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ f → cong ∣_∣₂ (funExt λ x → ⌣ₖ-0ₖ n m (f x)) + ⌣-0ₕ = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt λ x → ⌣ₖ-0ₖ n m (f x)) 0ₕ-⌣ : (f : coHom m A) → (0ₕ n ⌣ f) ≡ 0ₕ _ - 0ₕ-⌣ = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ f → cong ∣_∣₂ (funExt λ x → 0ₖ-⌣ₖ n m (f x)) + 0ₕ-⌣ = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) + λ f → cong ∣_∣₂ (funExt λ x → 0ₖ-⌣ₖ n m (f x)) - leftDistr-⌣ : (f : coHom n A) (g h : coHom m A) → f ⌣ (g +ₕ h) ≡ f ⌣ g +ₕ f ⌣ h + leftDistr-⌣ : (f : coHom n A) (g h : coHom m A) + → f ⌣ (g +ₕ h) ≡ f ⌣ g +ₕ f ⌣ h leftDistr-⌣ = sElim (λ _ → isSetΠ2 λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ f → sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) @@ -658,7 +661,8 @@ rUnit⌣ (suc n) = sElim (λ _ → isOfHLevelPath 2 squash₂ _ _) λ f → cong ∣_∣₂ (funExt λ x → rUnitₖ _ (f x)) --ₕDistᵣ : ∀ {ℓ} {A : Type ℓ} (n m : ℕ) (x : coHom n A) (y : coHom m A) → (-ₕ (x ⌣ y)) ≡ x ⌣ (-ₕ y) +-ₕDistᵣ : ∀ {ℓ} {A : Type ℓ} (n m : ℕ) + (x : coHom n A) (y : coHom m A) → (-ₕ (x ⌣ y)) ≡ x ⌣ (-ₕ y) -ₕDistᵣ n m = sElim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) λ f g → cong ∣_∣₂ (funExt λ x → -Distᵣ n m (f x) (g x)) From a945d876b08d5e52c55dfe1fbc364bb10cb231fd Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Sat, 6 Nov 2021 02:19:24 +0100 Subject: [PATCH 84/89] fixes --- Cubical/HITs/Hopf.agda | 3 --- 1 file changed, 3 deletions(-) delete mode 100644 Cubical/HITs/Hopf.agda diff --git a/Cubical/HITs/Hopf.agda b/Cubical/HITs/Hopf.agda deleted file mode 100644 index cad8df08f..000000000 --- a/Cubical/HITs/Hopf.agda +++ /dev/null @@ -1,3 +0,0 @@ -{-# OPTIONS --safe #-} -module Cubical.HITs.Hopf where - From 70352de3fdaefe61a4362e2a4348291f1acb4138 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Sat, 6 Nov 2021 02:23:50 +0100 Subject: [PATCH 85/89] minor --- Cubical/Algebra/Group/DirProd.agda | 1 - Cubical/Homotopy/Group/Properties.agda | 2 +- 2 files changed, 1 insertion(+), 2 deletions(-) diff --git a/Cubical/Algebra/Group/DirProd.agda b/Cubical/Algebra/Group/DirProd.agda index 0f8a68393..352d5446f 100644 --- a/Cubical/Algebra/Group/DirProd.agda +++ b/Cubical/Algebra/Group/DirProd.agda @@ -8,7 +8,6 @@ open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Monoid open import Cubical.Algebra.Semigroup - open GroupStr open IsGroup hiding (rid ; lid ; invr ; invl) open IsMonoid hiding (rid ; lid) diff --git a/Cubical/Homotopy/Group/Properties.agda b/Cubical/Homotopy/Group/Properties.agda index 585493035..713d7b135 100644 --- a/Cubical/Homotopy/Group/Properties.agda +++ b/Cubical/Homotopy/Group/Properties.agda @@ -1,6 +1,6 @@ {-# OPTIONS --safe --experimental-lossy-unification #-} {- -This file contians: +This file contains: 1. A proof that the equivalence Ωⁿ A ≃ (Sⁿ →∙ A) is structure preserving From 44743f612a16926a08a98ac540aff85617818769 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Sat, 6 Nov 2021 02:57:33 +0100 Subject: [PATCH 86/89] SES --- Cubical/Algebra/Group/Exact.agda | 58 +++++++++++++++++++ Cubical/Algebra/Group/Instances/Unit.agda | 40 ------------- Cubical/Homotopy/HopfInvariant/Base.agda | 1 + .../Homotopy/HopfInvariant/Homomorphism.agda | 1 + Cubical/Homotopy/HopfInvariant/HopfMap.agda | 1 + 5 files changed, 61 insertions(+), 40 deletions(-) create mode 100644 Cubical/Algebra/Group/Exact.agda diff --git a/Cubical/Algebra/Group/Exact.agda b/Cubical/Algebra/Group/Exact.agda new file mode 100644 index 000000000..ca2b8c2f4 --- /dev/null +++ b/Cubical/Algebra/Group/Exact.agda @@ -0,0 +1,58 @@ +{-# OPTIONS --safe #-} +module Cubical.Algebra.Group.Exact where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Algebra.Group.Base +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec) +open import Cubical.Algebra.Group.GroupPath + +open import Cubical.Algebra.Group.Instances.Unit + +private + variable + ℓ : Level + +-- TODO : Define exact sequences +-- (perhaps short, finite, ℕ-indexed and ℤ-indexed) + +SES→isEquiv : ∀ {ℓ ℓ'} {L R : Group ℓ-zero} + → {G : Group ℓ} {H : Group ℓ'} + → Unit ≡ L + → Unit ≡ R + → (lhom : GroupHom L G) (midhom : GroupHom G H) (rhom : GroupHom H R) + → ((x : _) → isInKer midhom x → isInIm lhom x) + → ((x : _) → isInKer rhom x → isInIm midhom x) + → isEquiv (fst midhom) +SES→isEquiv {R = R} {G = G} {H = H} = + J (λ L _ → Unit ≡ R → + (lhom : GroupHom L G) (midhom : GroupHom G H) + (rhom : GroupHom H R) → + ((x : fst G) → isInKer midhom x → isInIm lhom x) → + ((x : fst H) → isInKer rhom x → isInIm midhom x) → + isEquiv (fst midhom)) + ((J (λ R _ → (lhom : GroupHom Unit G) (midhom : GroupHom G H) + (rhom : GroupHom H R) → + ((x : fst G) → isInKer midhom x → isInIm lhom x) → + ((x : _) → isInKer rhom x → isInIm midhom x) → + isEquiv (fst midhom)) + main)) + where + main : (lhom : GroupHom Unit G) (midhom : GroupHom G H) + (rhom : GroupHom H Unit) → + ((x : fst G) → isInKer midhom x → isInIm lhom x) → + ((x : fst H) → isInKer rhom x → isInIm midhom x) → + isEquiv (fst midhom) + main lhom midhom rhom lexact rexact = + BijectionIsoToGroupEquiv {G = G} {H = H} + bijIso' .fst .snd + where + bijIso' : BijectionIso G H + BijectionIso.fun bijIso' = midhom + BijectionIso.inj bijIso' x inker = + pRec (GroupStr.is-set (snd G) _ _) + (λ s → sym (snd s) ∙ IsGroupHom.pres1 (snd lhom)) + (lexact _ inker) + BijectionIso.surj bijIso' x = rexact x refl diff --git a/Cubical/Algebra/Group/Instances/Unit.agda b/Cubical/Algebra/Group/Instances/Unit.agda index dd0e1b7ff..daf64d930 100644 --- a/Cubical/Algebra/Group/Instances/Unit.agda +++ b/Cubical/Algebra/Group/Instances/Unit.agda @@ -11,7 +11,6 @@ open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Group.DirProd open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.HITs.PropositionalTruncation renaming (rec to pRec) open import Cubical.Algebra.Group.GroupPath open GroupStr @@ -59,45 +58,6 @@ snd (contrGroupIsoUnit contr) = makeIsGroupHom λ _ _ → refl contrGroupEquivUnit : {G : Group ℓ} → isContr ⟨ G ⟩ → GroupEquiv G Unit contrGroupEquivUnit contr = GroupIso→GroupEquiv (contrGroupIsoUnit contr) -SES→isEquiv : ∀ {ℓ ℓ'} {L R : Group ℓ-zero} - → {G : Group ℓ} {H : Group ℓ'} - → Unit ≡ L - → Unit ≡ R - → (lhom : GroupHom L G) (midhom : GroupHom G H) (rhom : GroupHom H R) - → ((x : _) → isInKer midhom x → isInIm lhom x) - → ((x : _) → isInKer rhom x → isInIm midhom x) - → isEquiv (fst midhom) -SES→isEquiv {R = R} {G = G} {H = H} = - J (λ L _ → Unit ≡ R → - (lhom : GroupHom L G) (midhom : GroupHom G H) - (rhom : GroupHom H R) → - ((x : fst G) → isInKer midhom x → isInIm lhom x) → - ((x : fst H) → isInKer rhom x → isInIm midhom x) → - isEquiv (fst midhom)) - ((J (λ R _ → (lhom : GroupHom Unit G) (midhom : GroupHom G H) - (rhom : GroupHom H R) → - ((x : fst G) → isInKer midhom x → isInIm lhom x) → - ((x : _) → isInKer rhom x → isInIm midhom x) → - isEquiv (fst midhom)) - main)) - where - main : (lhom : GroupHom Unit G) (midhom : GroupHom G H) - (rhom : GroupHom H Unit) → - ((x : fst G) → isInKer midhom x → isInIm lhom x) → - ((x : fst H) → isInKer rhom x → isInIm midhom x) → - isEquiv (fst midhom) - main lhom midhom rhom lexact rexact = - BijectionIsoToGroupEquiv {G = G} {H = H} - bijIso' .fst .snd - where - bijIso' : BijectionIso G H - BijectionIso.fun bijIso' = midhom - BijectionIso.inj bijIso' x inker = - pRec (GroupStr.is-set (snd G) _ _) - (λ s → sym (snd s) ∙ IsGroupHom.pres1 (snd lhom)) - (lexact _ inker) - BijectionIso.surj bijIso' x = rexact x refl - isContr→≡UnitGroup : {G : Group ℓ-zero} → isContr (fst G) → Unit ≡ G isContr→≡UnitGroup c = fst (GroupPath _ _) diff --git a/Cubical/Homotopy/HopfInvariant/Base.agda b/Cubical/Homotopy/HopfInvariant/Base.agda index a064d4b26..b5a09a59d 100644 --- a/Cubical/Homotopy/HopfInvariant/Base.agda +++ b/Cubical/Homotopy/HopfInvariant/Base.agda @@ -36,6 +36,7 @@ open import Cubical.Data.Unit open import Cubical.Algebra.Group renaming (ℤ to ℤGroup ; Unit to UnitGroup) open import Cubical.Algebra.Group.ZAction +open import Cubical.Algebra.Group.Exact open import Cubical.HITs.Pushout open import Cubical.HITs.Sn diff --git a/Cubical/Homotopy/HopfInvariant/Homomorphism.agda b/Cubical/Homotopy/HopfInvariant/Homomorphism.agda index 5c360dd59..99a8c6fec 100644 --- a/Cubical/Homotopy/HopfInvariant/Homomorphism.agda +++ b/Cubical/Homotopy/HopfInvariant/Homomorphism.agda @@ -34,6 +34,7 @@ open import Cubical.Data.Unit open import Cubical.Algebra.Group renaming (ℤ to ℤGroup ; Unit to UnitGroup) open import Cubical.Algebra.Group.ZAction +open import Cubical.Algebra.Group.Exact open import Cubical.HITs.Pushout open import Cubical.HITs.Sn diff --git a/Cubical/Homotopy/HopfInvariant/HopfMap.agda b/Cubical/Homotopy/HopfInvariant/HopfMap.agda index f72113b15..587d1e5f9 100644 --- a/Cubical/Homotopy/HopfInvariant/HopfMap.agda +++ b/Cubical/Homotopy/HopfInvariant/HopfMap.agda @@ -39,6 +39,7 @@ open import Cubical.Data.Unit open import Cubical.Algebra.Group renaming (ℤ to ℤGroup ; Unit to UnitGroup) open import Cubical.Algebra.Group.ZAction +open import Cubical.Algebra.Group.Exact open import Cubical.HITs.Pushout open import Cubical.HITs.Join From 1774c587b5b10f7c7555f8ead70ea73d68d3a5b3 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Wed, 10 Nov 2021 15:35:55 +0100 Subject: [PATCH 87/89] stuff --- Cubical/Foundations/Pointed/Base.agda | 9 + Cubical/Homotopy/Group/Properties.agda | 222 +++---- .../Homotopy/Group/SuspensionMapPathP.agda | 609 +++++++----------- Cubical/Homotopy/Loopspace.agda | 14 + 4 files changed, 337 insertions(+), 517 deletions(-) diff --git a/Cubical/Foundations/Pointed/Base.agda b/Cubical/Foundations/Pointed/Base.agda index 5ed40cc0c..3b9355634 100644 --- a/Cubical/Foundations/Pointed/Base.agda +++ b/Cubical/Foundations/Pointed/Base.agda @@ -9,6 +9,7 @@ open import Cubical.Foundations.Structure using (typ) public open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Function Pointed : (ℓ : Level) → Type (ℓ-suc ℓ) Pointed ℓ = TypeWithStr ℓ (λ x → x) @@ -36,6 +37,14 @@ ua∙ : ∀ {ℓ} {A B : Pointed ℓ} (e : fst A ≃ fst B) fst (ua∙ e p i) = ua e i snd (ua∙ {A = A} e p i) = ua-gluePath e p i +{- J for pointed function types -} +→∙J : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Type ℓ'} + → (P : (b₀ : B) (f : A →∙ (B , b₀)) → Type ℓ'') + → ((f : fst A → B) → P (f (pt A)) (f , refl)) + → {b₀ : B} (f : A →∙ (B , b₀)) → P b₀ f +→∙J {A = A} P ind = + uncurry λ f → J (λ b₀ y → P b₀ (f , y)) (ind f) + {- HIT allowing for pattern matching on pointed types -} data Pointer {ℓ} (A : Pointed ℓ) : Type ℓ where pt₀ : Pointer A diff --git a/Cubical/Homotopy/Group/Properties.agda b/Cubical/Homotopy/Group/Properties.agda index 713d7b135..b189b7958 100644 --- a/Cubical/Homotopy/Group/Properties.agda +++ b/Cubical/Homotopy/Group/Properties.agda @@ -28,6 +28,7 @@ open import Cubical.Homotopy.Connected open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Equiv open import Cubical.Foundations.Pointed.Homogeneous open import Cubical.Foundations.HLevels open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) @@ -543,64 +544,34 @@ hom-botᵣ⁻' {A = A} n f g = hom-botᵣ⁻ {A = A} (suc n) f g ∙ sym (∙Π≡invComp {A = A} _ (botᵣ⁻ {A = A} _ f) (botᵣ⁻ {A = A} _ g)) - -hom-rMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → (p q : typ ((Ω^ (suc (suc n))) (Susp∙ (typ A)))) - → rMap (suc n) {A = A} (p ∙ q) - ≡ ∙Π (rMap (suc n) {A = A} p) (rMap (suc n) {A = A} q) -hom-rMap {A = A} n p q = - cong (lMap (suc n) {A = Ω (Susp∙ (typ A))}) - (morphLemmas.isMorphInv _∙_ _∙_ - (inv (flipΩIso (suc n))) (flipΩIsopres· n) - (fun (flipΩIso (suc n))) - (leftInv (flipΩIso (suc n))) (rightInv (flipΩIso (suc n))) - p q) - ∙ isHom-lMap _ (fun (flipΩIso (suc n)) p) (fun (flipΩIso (suc n)) q) - -isHom-IsoΩSphereMap' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - (p q : typ (Ω ((Ω^ n) (Susp∙ (typ A))))) - → fun (IsoΩSphereMap' {A = A} n) (p ∙ q) - ≡ ∙Π {n = (suc n)} (fun (IsoΩSphereMap' {A = A} n) p) - (fun (IsoΩSphereMap' {A = A} n) q) -isHom-IsoΩSphereMap' {A = A} zero p q = - (cong (botᵣ {A = A} 0) - (cong ((lMap 0 {A = Ω (Susp∙ (typ A))})) - (transportRefl (p ∙ q)))) - ∙ ΣPathP (funExt (λ { base → refl - ; (loop i) j - → lem j i}) - , refl) - where - lem : p ∙ q ≡ (cong (fst (fun (IsoΩSphereMap' {A = A} zero) p)) loop ∙ refl) - ∙ (cong (fst (fun (IsoΩSphereMap' {A = A} zero) q)) loop ∙ refl) - lem = cong₂ _∙_ - (sym (transportRefl p) - ∙ rUnit (cong (fst (fun (IsoΩSphereMap' {A = A} zero) p)) loop)) - (sym (transportRefl q) - ∙ rUnit (cong (fst (fun (IsoΩSphereMap' {A = A} zero) q)) loop)) -isHom-IsoΩSphereMap' {A = A} (suc n) p q = - cong (botᵣ {A = A} (suc n) ∘ lMap (suc n) {A = Ω (Susp∙ (typ A))}) - (morphLemmas.isMorphInv - _∙_ _∙_ - (inv (flipΩIso (suc n))) - (flipΩIsopres· n) - (fun (flipΩIso (suc n))) - (leftInv (flipΩIso (suc n))) - (rightInv (flipΩIso (suc n))) p q) - ∙ (cong (botᵣ {A = A} (suc n)) - (isHom-lMap _ {A = Ω (Susp∙ (typ A))} - (fun (flipΩIso (suc n)) p) (fun (flipΩIso (suc n)) q)) - ∙ morphLemmas.isMorphInv ∙Π ∙Π (botᵣ⁻ {A = A} (suc n)) - (hom-botᵣ⁻' {A = A} n) - (botᵣ {A = A} (suc n)) - (leftInv (botᵣIso {A = A} (suc n))) - (rightInv (botᵣIso {A = A} (suc n))) - (lMap (suc n) (fun (flipΩIso (suc n)) p)) - (lMap (suc n) (fun (flipΩIso (suc n)) q))) +hom-botᵣ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (f g : S₊∙ (suc n) →∙ Ω (Susp∙ (typ A))) + → botᵣ {A = A} (suc n) (∙Π f g) + ≡ ∙Π (botᵣ {A = A} (suc n) f) (botᵣ {A = A} (suc n) g) +hom-botᵣ {A = A} n f g = + morphLemmas.isMorphInv ∙Π ∙Π (botᵣ⁻ {A = A} (suc n)) + (hom-botᵣ⁻' {A = A} n) + (botᵣ {A = A} (suc n)) + (leftInv (botᵣIso {A = A} (suc n))) + (rightInv (botᵣIso {A = A} (suc n))) + f g + +isHom-IsoΩSphereMapᵣ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + (p q : typ ((Ω^ (2 + n)) (Susp∙ (typ A)))) + → Iso.fun (IsoΩSphereMapᵣ (suc n)) (p ∙ q) + ≡ ∙Π (Iso.fun (IsoΩSphereMapᵣ (suc n)) p) + (Iso.fun (IsoΩSphereMapᵣ (suc n)) q) +isHom-IsoΩSphereMapᵣ {A = A} n p q = + cong (botᵣ {A = A} (suc n)) + (cong (lMap (suc n) {A = Ω (Susp∙ (typ A))}) + (flipΩIsopres· n p q) + ∙ isHom-lMap n (fun (flipΩIso (suc n)) p) + (fun (flipΩIso (suc n)) q)) + ∙ hom-botᵣ n _ _ suspMapΩ→hom : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p q : typ ((Ω^ (suc n)) A)) - → suspMapΩ∙ (suc n) .fst (p ∙ q) - ≡ suspMapΩ∙ (suc n) .fst p ∙ suspMapΩ∙ (suc n) .fst q + → suspMapΩ (suc n) (p ∙ q) + ≡ suspMapΩ (suc n) p ∙ suspMapΩ (suc n) q suspMapΩ→hom {A = A} n p q = cong (sym (snd (suspMapΩ∙ {A = A} n)) ∙∙_∙∙ snd (suspMapΩ∙ {A = A} n)) (cong-∙ (fst (suspMapΩ∙ {A = A} n)) p q) @@ -617,97 +588,84 @@ suspMapΩ→hom {A = A} n p q = ∙ cong₂ _∙_ (rUnit q) (rUnit s) private - transpLemTyp : ∀ {ℓ} → {A1 A2 B1 B2 : Type ℓ} - (AB1 : A1 ≡ B1) (AB2 : A2 ≡ B2) - → (f : A1 → A2) (g : B1 → B2) - → PathP (λ i → AB1 i → AB2 i) f g - → Type ℓ - transpLemTyp {ℓ} {A1} {A2} {B1} {B2} AB1 AB2 f g p = - (+A1 : A1 → A1 → A1) - → (+A2 : A2 → A2 → A2) - → (+B1 : B1 → B1 → B1) - → (+B2 : B2 → B2 → B2) - → ((x y : A1) → transport AB1 (+A1 x y) - ≡ +B1 (transport AB1 x) (transport AB1 y)) - → ((x y : A2) → transport AB2 (+A2 x y) - ≡ +B2 (transport AB2 x) (transport AB2 y)) - → ((x y : A1) → f (+A1 x y) ≡ +A2 (f x) (f y)) - → ((x y : B1) → g (+B1 x y) ≡ +B2 (g x) (g y)) - - transpLem : ∀ {ℓ} → {A1 A2 B1 B2 : Type ℓ} (AB1 : A1 ≡ B1) (AB2 : A2 ≡ B2) - → (f : A1 → A2) (g : B1 → B2) - → (eq : PathP (λ i → AB1 i → AB2 i) f g) - → transpLemTyp AB1 AB2 f g eq - transpLem {ℓ} {A1} {A2} {B1} {B2} = - J (λ B1 AB1 → (AB2 : A2 ≡ B2) - → (f : A1 → A2) (g : B1 → B2) - → (eq : PathP (λ i → AB1 i → AB2 i) f g) - → transpLemTyp AB1 AB2 f g eq) - (J (λ B2 AB2 → (f : A1 → A2) (g : A1 → B2) - → (eq : PathP (λ i → A1 → AB2 i) f g) - → transpLemTyp refl AB2 f g eq) - λ f g → J (λ g p → transpLemTyp refl refl f g p) - λ +A1 +A2 +B1 +B2 trId1 trId2 fhom1 - → λ x y → cong f (cong₂ +B1 (sym (transportRefl x)) (sym (transportRefl y)) - ∙∙ sym (trId1 x y) - ∙∙ transportRefl (+A1 x y)) - ∙∙ fhom1 x y - ∙∙ ((sym (transportRefl _) ∙ (trId2 (f x) (f y))) - ∙ cong₂ +B2 (transportRefl (f x)) (transportRefl (f y)))) - -suspMap→hom : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f g : S₊∙ (suc n) →∙ A) - → suspMap n (∙Π f g) ≡ ∙Π (suspMap n f) (suspMap n g) -suspMap→hom {A = A} n = - transpLem (Ω≡SphereMap {A = A} (suc n)) (Ω≡SphereMap' {A = A} (suc n)) - (suspMapΩ {A = A} (suc n)) (suspMap {A = A} n) - (suspMap→ {A = A} n) - _∙_ _∙_ ∙Π ∙Π - (λ x y → uaβ (isoToEquiv (IsoΩSphereMap {A = A} (suc n))) (x ∙ y) - ∙∙ isHom-lMap n x y - ∙∙ cong₂ ∙Π - (sym (uaβ (isoToEquiv (IsoΩSphereMap {A = A} (suc n))) x)) - (sym (uaβ (isoToEquiv (IsoΩSphereMap {A = A} (suc n))) y))) - (λ x y → uaβ (isoToEquiv (IsoΩSphereMap' {A = A} (suc n))) (x ∙ y) - ∙∙ isHom-IsoΩSphereMap' {A = A} (suc n) x y - ∙∙ cong₂ ∙Π - (sym (uaβ (isoToEquiv (IsoΩSphereMap' {A = A} (suc n))) x)) - (sym (uaβ (isoToEquiv (IsoΩSphereMap' {A = A} (suc n))) y))) - (suspMapΩ→hom {A = A} n) + transportLem : ∀ {ℓ} {A B : Type ℓ} + (_+A_ : A → A → A) (_+B_ : B → B → B) + → (e : Iso A B) + → ((x y : A) → fun e (x +A y) ≡ fun e x +B fun e y) + → PathP (λ i → isoToPath e i → isoToPath e i → isoToPath e i) + _+A_ _+B_ + transportLem _+A_ _+B_ e hom = + toPathP (funExt (λ p → funExt λ q → + (λ i → transportRefl + (fun e (inv e (transportRefl p i) + +A inv e (transportRefl q i))) i) + ∙∙ hom (inv e p) (inv e q) + ∙∙ cong₂ _+B_ (rightInv e p) (rightInv e q))) + + pₗ : ∀ {ℓ} (A : Pointed ℓ) (n : ℕ) + → typ (Ω ((Ω^ n) A)) ≡ (S₊∙ (suc n) →∙ A) + pₗ A n = isoToPath (IsoΩSphereMap {A = A} (suc n)) + + pᵣ : ∀ {ℓ} (A : Pointed ℓ) (n : ℕ) + → typ ((Ω^ (2 + n)) (Susp∙ (typ A))) + ≡ (S₊∙ (suc (suc n)) →∙ Susp∙ (typ A)) + pᵣ A n = isoToPath (IsoΩSphereMapᵣ {A = A} (suc n)) + +∙Π→∙ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → PathP (λ i → pₗ A n i → pₗ A n i → pₗ A n i) _∙_ ∙Π +∙Π→∙ {A = A} n = + transportLem _∙_ ∙Π (IsoΩSphereMap {A = A} (suc n)) (isHom-lMap n) + +∙Π→∙ᵣ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → PathP (λ i → pᵣ A n i → pᵣ A n i → pᵣ A n i) _∙_ ∙Π +∙Π→∙ᵣ {A = A} n = + transportLem _∙_ ∙Π (IsoΩSphereMapᵣ {A = A} (suc n)) (isHom-IsoΩSphereMapᵣ n) + +isHom-suspMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f g : S₊∙ (suc n) →∙ A) + → suspMap n (∙Π f g) + ≡ ∙Π (suspMap n f) (suspMap n g) +isHom-suspMap {A = A} n = + transport (λ i → (f g : isoToPath (IsoΩSphereMap {A = A} (suc n)) i) + → Ωσ→suspMap n i (∙Π→∙ n i f g) + ≡ ∙Π→∙ᵣ n i (Ωσ→suspMap n i f) (Ωσ→suspMap n i g)) + (suspMapΩ→hom n) + +suspMapπ' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} + → π' (suc n) A + → π' (suc (suc n)) (Susp∙ (typ A)) +suspMapπ' n = sMap (suspMap n) + +suspMapπ'Hom : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → GroupHom (π'Gr n A) (π'Gr (suc n) (Susp∙ (typ A))) +fst (suspMapπ'Hom {A = A} n) = suspMapπ' n +snd (suspMapπ'Hom {A = A} n) = + makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) + λ f g → cong ∣_∣₂ (isHom-suspMap n f g)) + private isConnectedPres : ∀ {ℓ} {A : Pointed ℓ} (con n : ℕ) → isConnectedFun con (suspMapΩ∙ {A = A} (suc n) .fst) → isConnectedFun con (suspMap {A = A} n) isConnectedPres {A = A} con n hyp = - transport (λ i → isConnectedFun con (suspMap→ {A = A} n i)) hyp + transport (λ i → isConnectedFun con (Ωσ→suspMap {A = A} n i)) hyp isConnectedSuspMap : (n m : ℕ) → isConnectedFun ((m + suc m) ∸ n) (suspMap {A = S₊∙ (suc m)} n) isConnectedSuspMap n m = isConnectedPres _ _ (suspMapΩ-connected m (suc n) (sphereConnected (suc m))) -suspMapπ' : (n m : ℕ) - → π' (suc n) (S₊∙ (suc m)) - → π' (suc (suc n)) (S₊∙ (suc (suc m))) -suspMapπ' n m = sMap (suspMap {A = S₊∙ (suc m)} n) - -suspMapHom : (n m : ℕ) - → GroupHom (π'Gr n (S₊∙ (suc m))) - (π'Gr (suc n) (S₊∙ (suc (suc m)))) -fst (suspMapHom n m) = suspMapπ' n m -snd (suspMapHom n m) = - makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) - λ f g → cong ∣_∣₂ (suspMap→hom n f g)) - -isSurjectiveSuspMap : (n : ℕ) → isSurjective (suspMapHom (2 + n) (suc n)) +isSurjectiveSuspMap : (n : ℕ) → isSurjective (suspMapπ'Hom {A = S₊∙ (2 + n)} (2 + n)) isSurjectiveSuspMap n = sElim (λ _ → isProp→isSet squash) - λ f → trRec (subst (λ x → isOfHLevel x (isInIm (suspMapHom (2 + n) (suc n)) ∣ f ∣₂)) + λ f → + trRec + ((subst (λ x → isOfHLevel x (isInIm (suspMapπ'Hom (2 + n)) ∣ f ∣₂)) (sym (snd (lem n n))) - (isProp→isOfHLevelSuc {A = isInIm (suspMapHom (2 + n) (suc n)) ∣ f ∣₂} - (fst (lem n n)) squash)) - (λ p → ∣ ∣ fst p ∣₂ , (cong ∣_∣₂ (snd p)) ∣) - (fst (isConnectedSuspMap (2 + n) (suc n) f)) + (isProp→isOfHLevelSuc {A = isInIm (suspMapπ'Hom (2 + n)) ∣ f ∣₂} + (fst (lem n n)) squash))) + (λ p → ∣ ∣ fst p ∣₂ , (cong ∣_∣₂ (snd p)) ∣) + (fst (isConnectedSuspMap (2 + n) (suc n) f)) where lem : (m n : ℕ) → Σ[ x ∈ ℕ ] ((m + suc (suc n) ∸ suc n) ≡ suc x) lem zero zero = 0 , refl diff --git a/Cubical/Homotopy/Group/SuspensionMapPathP.agda b/Cubical/Homotopy/Group/SuspensionMapPathP.agda index 42ba0841f..86c95ba7a 100644 --- a/Cubical/Homotopy/Group/SuspensionMapPathP.agda +++ b/Cubical/Homotopy/Group/SuspensionMapPathP.agda @@ -61,6 +61,48 @@ open IsSemigroup open IsMonoid open GroupStr +flipΩPath : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → ((Ω^ (suc n)) A) ≡ (Ω^ n) (Ω A) +flipΩPath {A = A} zero = refl +flipΩPath {A = A} (suc n) = cong Ω (flipΩPath {A = A} n) + +flipΩIso : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → Iso (fst ((Ω^ (suc n)) A)) (fst ((Ω^ n) (Ω A))) +flipΩIso {A = A} n = pathToIso (cong fst (flipΩPath n)) + +flipΩIso⁻pres· : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (f g : fst ((Ω^ (suc n)) (Ω A))) + → inv (flipΩIso (suc n)) (f ∙ g) + ≡ (inv (flipΩIso (suc n)) f) + ∙ (inv (flipΩIso (suc n)) g) +flipΩIso⁻pres· {A = A} n f g i = + transp (λ j → flipΩPath {A = A} n (~ i ∧ ~ j) .snd + ≡ flipΩPath n (~ i ∧ ~ j) .snd) i + (transp (λ j → flipΩPath {A = A} n (~ i ∨ ~ j) .snd + ≡ flipΩPath n (~ i ∨ ~ j) .snd) (~ i) f + ∙ transp (λ j → flipΩPath {A = A} n (~ i ∨ ~ j) .snd + ≡ flipΩPath n (~ i ∨ ~ j) .snd) (~ i) g) + +flipΩIsopres· : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (f g : fst (Ω ((Ω^ (suc n)) A))) + → fun (flipΩIso (suc n)) (f ∙ g) + ≡ (fun (flipΩIso (suc n)) f) + ∙ (fun (flipΩIso (suc n)) g) +flipΩIsopres· n = + morphLemmas.isMorphInv _∙_ _∙_ + (inv (flipΩIso (suc n))) + (flipΩIso⁻pres· n) + (fun (flipΩIso (suc n))) + (Iso.leftInv (flipΩIso (suc n))) + (Iso.rightInv (flipΩIso (suc n))) + +flipΩrefl : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → fun (flipΩIso {A = A} (suc n)) refl ≡ refl +flipΩrefl {A = A} n j = + transp (λ i₁ → fst (Ω (flipΩPath {A = A} n ((i₁ ∨ j))))) + j (snd (Ω (flipΩPath n j))) + + -- Solves some termination issues private +nInd : ∀ {ℓ} {P : ℕ → Type ℓ} @@ -73,6 +115,22 @@ private +nInd {P = P} 0c 1c indc (suc (suc n)) = indc n (+nInd {P = P} 0c 1c indc (suc n)) +-- move to loop +Ω^→ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) + → (A →∙ B) → ((Ω^ n) A →∙ (Ω^ n) B) +Ω^→ zero f = f +Ω^→ (suc n) f = Ω→ (Ω^→ n f) + +-- move to pointed +post∘∙ : ∀ {ℓX ℓ ℓ'} (X : Pointed ℓX) {A : Pointed ℓ} {B : Pointed ℓ'} + → (A →∙ B) → ((X →∙ A ∙) →∙ (X →∙ B ∙)) +post∘∙ X f .fst g = f ∘∙ g +post∘∙ X f .snd = + ΣPathP + ( (funExt λ _ → f .snd) + , (sym (lUnit (f .snd)) ◁ λ i j → f .snd (i ∨ j)) + ) + suspMap : ∀ {ℓ} {A : Pointed ℓ}(n : ℕ) → S₊∙ (suc n) →∙ A → S₊∙ (suc (suc n)) →∙ Susp∙ (typ A) @@ -124,12 +182,17 @@ lMapId2 (suc (suc n)) {A = A} = → ∙∙lCancel (lMapId (suc n) {A = A} a) j i})) , refl) +lMap∙ : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} + → ((Ω^ n) A) →∙ (S₊∙ n →∙ A ∙) +lMap∙ n .fst = lMap n +lMap∙ n .snd = lMapId2 n + -- We define the following maps which will be used to -- show that lMap is an equivalence lMapSplit₁ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → typ ((Ω^ (suc n)) A) → typ (Ω (S₊∙ n →∙ A ∙)) -lMapSplit₁ n = Ω→ (lMap n , lMapId2 n) .fst +lMapSplit₁ n = Ω→ (lMap∙ n) .fst ΩSphereMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → typ (Ω (S₊∙ n →∙ A ∙)) @@ -351,86 +414,118 @@ isEquiv-lMap (suc n) = (isEquiv-lMap n))) (invEquiv (isoToEquiv (SphereMapΩIso (suc n)))))) -botMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} - → (S₊∙ n →∙ A) - → S₊∙ n →∙ Ω (Susp∙ (typ A)) -fst (botMap n {A = A} f) x = merid (fst f x) ∙ sym (merid (pt A)) -snd (botMap n {A = A} f) = cong (_∙ merid (pt A) ⁻¹) (cong merid (snd f)) - ∙ rCancel (merid (pt A)) - -flipΩPath : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → ((Ω^ (suc n)) A) ≡ (Ω^ n) (Ω A) -flipΩPath {A = A} zero = refl -flipΩPath {A = A} (suc n) = cong Ω (flipΩPath {A = A} n) - -flipΩIso : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → Iso (fst ((Ω^ (suc n)) A)) (fst ((Ω^ n) (Ω A))) -flipΩIso {A = A} n = pathToIso (cong fst (flipΩPath n)) - -flipΩIsopres· : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → (f g : fst ((Ω^ (suc n)) (Ω A))) - → inv (flipΩIso (suc n)) (f ∙ g) - ≡ (inv (flipΩIso (suc n)) f) - ∙ (inv (flipΩIso (suc n)) g) -flipΩIsopres· {A = A} n f g i = - transp (λ j → flipΩPath {A = A} n (~ i ∧ ~ j) .snd - ≡ flipΩPath n (~ i ∧ ~ j) .snd) i - (transp (λ j → flipΩPath {A = A} n (~ i ∨ ~ j) .snd - ≡ flipΩPath n (~ i ∨ ~ j) .snd) (~ i) f - ∙ transp (λ j → flipΩPath {A = A} n (~ i ∨ ~ j) .snd - ≡ flipΩPath n (~ i ∨ ~ j) .snd) (~ i) g) - -rMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} - → typ ((Ω^ (suc n)) (Susp∙ (typ A))) - → (S₊∙ n →∙ Ω (Susp∙ (typ A))) -rMap n = lMap n ∘ fun (flipΩIso n) -private - rMap1 : ∀ {ℓ} {A : Pointed ℓ} - → typ ((Ω^ (suc 1)) (Susp∙ (typ A))) - → (S₊∙ 1 →∙ Ω (Susp∙ (typ A))) - rMap1 = lMap 1 - - rMap≡rMap1 : ∀ {ℓ} {A : Pointed ℓ} → rMap 1 {A = A} ≡ rMap1 {A = A} - rMap≡rMap1 = funExt λ x → cong (lMap 1) (transportRefl x) - -flipΩrefl : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → fun (flipΩIso {A = A} (suc n)) refl ≡ refl -flipΩrefl {A = A} n j = - transp (λ i₁ → fst (Ω (flipΩPath {A = A} n ((i₁ ∨ j))))) - j (snd (Ω (flipΩPath n j))) +-- The homogeneity assumption is not necessary but simplifying +module _ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') (homogB : isHomogeneous B) (f : A →∙ B) + where -cong-lMap-lem : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p : _) - → cong (lMap (suc n) {A = Ω A}) (fun (flipΩIso (suc (suc n))) p) - ≡ (cong (lMap (suc n)) (sym (flipΩrefl n)) - ∙∙ cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) p - ∙∙ cong (lMap (suc n)) (flipΩrefl n)) -cong-lMap-lem {A = A} n p = - (λ i → cong (lMap (suc n) {A = Ω A}) - (fun (flipΩIso (suc (suc n))) (rUnit p i))) - ∙∙ (λ i j → lMap (suc n) {A = Ω A} - (transp (λ k → fst (flipΩPath {A = A} (suc (suc n)) (k ∨ i))) i - (((λ j → transp (λ k → fst (Ω (flipΩPath {A = A} n ((k ∨ ~ j) ∧ i)))) - (~ i ∨ ~ j) - (snd (Ω (flipΩPath n (~ j ∧ i))))) - ∙∙ (λ j → transp (λ k → fst (flipΩPath {A = A} (suc n) - (k ∧ i))) (~ i) (p j)) - ∙∙ λ j → transp (λ k → fst (Ω (flipΩPath {A = A} n ((k ∨ j) ∧ i)))) - (~ i ∨ j) (snd (Ω (flipΩPath n (j ∧ i)))))) j)) - ∙∙ cong-∙∙ (lMap (suc n)) - (sym (flipΩrefl n)) - (cong (fun (flipΩIso (suc n))) p) - (flipΩrefl n) - -botₗ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → (S₊∙ n →∙ A) → S₊∙ (suc n) →∙ (Susp∙ (typ A)) -fst (botₗ zero f) base = north -fst (botₗ {A = A} zero f) (loop i) = (merid (fst f false) ∙ sym (merid (pt A))) i -snd (botₗ zero f) = refl -botₗ (suc n) f = suspMap n f + isNaturalΩSphereMap : (n : ℕ) + → ∀ g → f ∘∙ ΩSphereMap n g ≡ ΩSphereMap n (Ω→ (post∘∙ (S₊∙ n) f) .fst g) + isNaturalΩSphereMap 0 g = + →∙Homogeneous≡ homogB (funExt lem) + where + lem : ∀ x → f .fst (ΩSphereMap 0 g .fst x) ≡ ΩSphereMap 0 (Ω→ (post∘∙ (S₊∙ 0) f) .fst g) .fst x + lem base = f .snd + lem (loop i) j = + hfill + (λ j → λ + { (i = i0) → post∘∙ _ f .snd j + ; (i = i1) → post∘∙ _ f .snd j + }) + (inS (f ∘∙ g i)) + j .fst false + isNaturalΩSphereMap (n@(suc _)) g = + →∙Homogeneous≡ homogB (funExt lem) + where + lem : ∀ x → f .fst (ΩSphereMap n g .fst x) ≡ ΩSphereMap n (Ω→ (post∘∙ (S₊∙ n) f) .fst g) .fst x + lem north = f .snd + lem south = f .snd + lem (merid a i) j = + hfill + (λ j → λ + { (i = i0) → post∘∙ _ f .snd j + ; (i = i1) → post∘∙ _ f .snd j + }) + (inS (f ∘∙ g i)) + j .fst a + + mutual + isNatural-lMap : ∀ n p → f ∘∙ lMap n p ≡ lMap n (Ω^→ n f .fst p) + isNatural-lMap 0 p = + →∙Homogeneous≡ homogB (funExt λ {true → f .snd; false → refl}) + isNatural-lMap (n@(suc n')) p = + cong (f ∘∙_) (lMap-split n' p) + ∙ isNaturalΩSphereMap n' (lMapSplit₁ n' p) + ∙ cong (ΩSphereMap n') inner + ∙ sym (lMap-split n' (Ω^→ n f .fst p)) + where + inner : Ω→ (post∘∙ (S₊∙ n') f) .fst (Ω→ (lMap∙ n') .fst p) ≡ Ω→ (lMap∙ n') .fst (Ω^→ (suc n') f .fst p) + inner = + sym (Ω→∘ (post∘∙ (S₊∙ n') f) (lMap∙ n') p) + ∙ cong (λ g∙ → Ω→ g∙ .fst p) (isNatural-lMap∙ n') + ∙ Ω→∘ (lMap∙ n') (Ω^→ n' f) p + + isNatural-lMap∙ : ∀ n → post∘∙ (S₊∙ n) f ∘∙ (lMap∙ n) ≡ (lMap∙ n {A = B} ∘∙ Ω^→ n f) + isNatural-lMap∙ n = →∙Homogeneous≡ (isHomogeneous→∙ homogB) (funExt (isNatural-lMap n)) + + +σ : ∀ {ℓ} (A : Pointed ℓ) (a : typ A) → Path (Susp (typ A)) north north +σ A a = merid a ∙ sym (merid (pt A)) + +σ∙ : ∀ {ℓ} (A : Pointed ℓ) → A →∙ Ω (Susp∙ (typ A)) +fst (σ∙ A) = σ A +snd (σ∙ A) = rCancel _ + + +top□ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Ω^→ (suc n) (σ∙ A) + ≡ (((Iso.fun (flipΩIso (suc n))) , flipΩrefl n) + ∘∙ suspMapΩ∙ (suc n)) +top□ {A = A} zero = + →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt λ p → sym (transportRefl _)) +top□ {A = A} (suc n) = + cong Ω→ (top□ {A = A} n) + ∙ →∙Homogeneous≡ + (isHomogeneousPath _ _) + (funExt λ x + → Ω→∘ (fun (flipΩIso (suc n)) , flipΩrefl n) (suspMapΩ∙ (suc n)) x) + +{- suspMapΩ + Ωⁿ A --------------------> Ω¹⁺ⁿ (Susp A) + | | + | = | ≃ flipΩ + | Ωⁿ→ σ v +Ωⁿ A -------------------> Ωⁿ (Ω (Susp A)) + | | + | | + | ≃ lMap | ≃ lMap + | | + v post∘∙ . σ v + (Sⁿ →∙ A) -------------- > (Sⁿ →∙ Ω (Susp A)) + | | + | = | ≃ botᵣ + | | + v suspMap v + (Sⁿ →∙ A) -------------- > (Sⁿ⁺¹→∙ Susp A) + -} + + +mid□ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (p : typ ((Ω^ (suc n)) A)) + → fst (post∘∙ (S₊∙ (suc n)) (σ∙ A)) (lMap (suc n) p) + ≡ lMap (suc n) (fst (Ω^→ (suc n) (σ∙ A)) p) +mid□ {A = A} n p = + funExt⁻ + (cong fst + (isNatural-lMap∙ + A (Ω (Susp∙ (typ A))) + (isHomogeneousPath _ _) + (σ∙ A) (suc n))) p botᵣ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → (S₊∙ n →∙ Ω (Susp∙ (typ A))) → S₊∙ (suc n) →∙ Susp∙ (typ A) + → (S₊∙ n →∙ Ω (Susp∙ (typ A))) + → S₊∙ (suc n) →∙ Susp∙ (typ A) fst (botᵣ zero (f , p)) base = north fst (botᵣ zero (f , p)) (loop i) = f false i snd (botᵣ zero (f , p)) = refl @@ -439,52 +534,6 @@ fst (botᵣ (suc n) (f , p)) south = north fst (botᵣ (suc n) (f , p)) (merid a i) = f a i snd (botᵣ (suc n) (f , p)) = refl -{- -The goal now is to fill the following diagram. - - suspMapΩ - Ωⁿ A -------------------> Ω¹⁺ⁿ (Susp A) - | | - | | - lMap | ≃ ≃ | rMap - v v -(S₊∙ n →∙ A) ----------> (S₊∙ n →∙ Ω (Susp A)) - \ botMap / - \ / - \ / - botₗ \ / botᵣ - (suspMap) \ / - \ / - \ / - v - (Sⁿ⁺¹ →∙ Susp A) - --} - --- The bottom part is trival▿ -filler▿ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (x : (S₊∙ n →∙ A)) - → botₗ n x ≡ botᵣ {A = A} n (botMap n x) -filler▿ zero (f , p) = - ΣPathP ((funExt (λ { base → refl ; (loop i) → refl})) , refl) -filler▿ (suc n) (f , p) = - ΣPathP ((funExt (λ { north → refl - ; south → refl - ; (merid a i) → refl})) - , refl) - -IsoΩSphereMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → Iso (typ ((Ω^ n) A)) ((S₊∙ n →∙ A)) -fun (IsoΩSphereMap zero) = lMap zero -inv (IsoΩSphereMap zero) f = fst f false -rightInv (IsoΩSphereMap zero) f = - ΣPathP ((funExt (λ { false → refl - ; true → sym (snd f)})) - , λ i j → snd f (~ i ∨ j)) -leftInv (IsoΩSphereMap zero) p = refl -IsoΩSphereMap (suc n) = equivToIso (_ , isEquiv-lMap n) - -Ω≡SphereMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → (typ ((Ω^ n) A)) ≡ ((S₊∙ n →∙ A)) -Ω≡SphereMap n = isoToPath (IsoΩSphereMap n) botᵣ⁻' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → S₊∙ (suc n) →∙ Susp∙ (typ A) → (S₊ n → typ (Ω (Susp∙ (typ A)))) @@ -506,14 +555,6 @@ snd (botᵣ⁻ {A = A} (suc n) f) = (cong (cong (fst f)) (rCancel (merid (ptSn _)))) ∙ ∙∙lCancel (snd f) -isEquiv-rMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → isEquiv (rMap n {A = A}) -isEquiv-rMap zero = - compEquiv (isoToEquiv (flipΩIso zero)) - (isoToEquiv (IsoΩSphereMap zero)) .snd -isEquiv-rMap (suc n) = - compEquiv (isoToEquiv (flipΩIso (suc n))) - (isoToEquiv (IsoΩSphereMap (suc n))) .snd - -- botᵣ is an Iso botᵣIso : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → Iso (S₊∙ n →∙ Ω (Susp∙ (typ A))) @@ -552,258 +593,56 @@ botᵣIso {A = A} n = (iso (botᵣ {A = A} n) (botᵣ⁻ {A = A} n) (h n) (retr ∙∙ (λ i → f x ∙ sym (p i) ) ∙∙ sym (rUnit (f x))) -IsoΩSphereMap' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → Iso (typ ((Ω^ (suc n)) (Susp∙ (typ A)))) - (S₊∙ (suc n) →∙ Susp∙ (typ A)) -IsoΩSphereMap' {A = A} n = - compIso (equivToIso (_ , isEquiv-rMap {A = A} n)) - (botᵣIso {A = A} n) - -Ω≡SphereMap' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → (typ ((Ω^ (suc n)) (Susp∙ (typ A)))) - ≡ (S₊∙ (suc n) →∙ Susp∙ (typ A)) -Ω≡SphereMap' {A = A} n = isoToPath (IsoΩSphereMap' {A = A} n) - -filler-top□ : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} - → rMap n {A = A} ∘ (suspMapΩ {A = A} n) - ≡ botMap n ∘ lMap n {A = A} -filler-top□ {ℓ} = - +nInd (λ {A} → funExt λ p → →∙Homogeneous≡ (isHomogeneousPath _ _) - (funExt λ { false → transportRefl _ - ; true → sym (rCancel _)})) - (λ {A} → funExt λ p → →∙Homogeneous≡ (isHomogeneousPath _ _) - (funExt λ x - → (λ i → ((rMap≡rMap1 {A = A} i) ∘ suspMapΩ 1) p .fst x) - ∙ sym (lem₁ p x))) - cool* - where - lem₁ : ∀ {ℓ} {A : Pointed ℓ} (p : _) (x : S¹) - → (botMap 1 ∘ lMap 1) p .fst x - ≡ (rMap1 {A = A} ∘ (suspMapΩ {A = A} 1)) p .fst x - lem₁ {A = A} p base = rCancel (merid (pt A)) - lem₁ {A = A} p (loop i) j = - doubleCompPath-filler - (sym (rCancel (merid (pt A)))) - (cong (λ x → merid x ∙ sym (merid (pt A))) p) - (rCancel (merid (pt A))) j i - - cool* : (n : ℕ) → - ({A : Pointed ℓ} → - rMap (suc n) {A = A} ∘ suspMapΩ (suc n) ≡ - botMap (suc n) ∘ lMap (suc n)) → - {A : Pointed ℓ} → - rMap (suc (suc n)) {A = A} ∘ suspMapΩ (suc (suc n)) ≡ - botMap (suc (suc n)) ∘ lMap (suc (suc n)) - cool* n ind {A} = - funExt λ p → →∙Homogeneous≡ (isHomogeneousPath (Susp (typ A)) refl) - (funExt - λ { north → sym (rCancel (merid (pt A))) - ; south → sym (rCancel (merid (pt A))) - ; (merid a i) j - → hcomp (λ k - → λ {( i = i0) → rCancel (merid (pt A)) (~ j) - ; (i = i1) → rCancel (merid (pt A)) (~ j) - ; (j = i0) → main p a (~ k) i - ; (j = i1) → (botMap (suc (suc n)) ∘ lMap (suc (suc n))) p .fst - (merid a i)}) - (doubleCompPath-filler - (sym (rCancel (merid (pt A)))) - (cong ((botMap (suc (suc n)) ∘ lMap (suc (suc n))) p .fst) - (merid a)) - (rCancel (merid (pt A))) (~ j) i)}) - where - module _ (p : typ (Ω _)) (a : S₊ (suc n)) where - abstract - indHyp : Path ((a₁ : fst (Ω ((Ω^ n) A))) → - Σ (fst (S₊∙ (suc n)) → fst (Ω (Susp∙ (typ A)))) - (λ f → f (snd (S₊∙ (suc n))) ≡ snd (Ω (Susp∙ (typ A))))) - ((rMap (suc n) {A = A} ∘ (suspMapΩ {A = A} (suc n)))) - (botMap (suc n) ∘ lMap (suc n) {A = A}) - indHyp = - funExt λ a → →∙Homogeneous≡ (isHomogeneousPath _ _) - (funExt (λ x → funExt⁻ (cong fst (funExt⁻ (ind {A = A}) a)) x)) - - pathLem₁ : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) (r : z ≡ z) - → p ∙ (q ∙ r ∙ sym q) ∙ sym p ≡ (p ∙ q) ∙ r ∙ sym (p ∙ q) - pathLem₁ p q r = - ∙assoc p (q ∙ r ∙ sym q) (sym p) - ∙ cong (_∙ sym p) (∙assoc p q (r ∙ sym q)) - ∙ sym (∙assoc (p ∙ q) (r ∙ sym q) (sym p)) - ∙ cong ((p ∙ q) ∙_) (sym (∙assoc r (sym q) (sym p)) - ∙ cong (r ∙_) (sym (symDistr p q))) - - pathLem₂ : ∀ {ℓ} {A : Type ℓ} {x y : A} (coh : x ≡ y) (p : x ≡ x) - → p ≡ coh ∙ (sym coh ∙∙ p ∙∙ coh) ∙ sym coh - pathLem₂ {x = x} = - J (λ y coh → (p : x ≡ x) → p ≡ coh ∙ (sym coh ∙∙ p ∙∙ coh) ∙ sym coh) - λ p → lUnit _ ∙ cong (refl ∙_) (rUnit _ ∙ cong (_∙ refl) (rUnit p)) - - pathLem₃ : ∀ {ℓ} {A : Type ℓ} {x₀ x y z w : A} - (p : x₀ ≡ x) (q : x ≡ y) (r : y ≡ z) (s : z ≡ w) (mid : w ≡ w) - (coh : x₀ ≡ w) - → isComm∙ (A , x₀) - → (p ∙∙ q ∙∙ r - ∙∙ s ∙∙ mid ∙∙ sym s - ∙∙ sym r ∙∙ sym q ∙∙ sym p) - ≡ (coh ∙∙ mid ∙∙ sym coh) - pathLem₃ p q r s mid coh comm = - doubleCompPath≡compPath p _ (sym p) - ∙∙ cong (λ x → p ∙ x ∙ (sym p)) - (doubleCompPath≡compPath q _ (sym q) - ∙ cong (λ x → q ∙ x ∙ (sym q)) - (doubleCompPath≡compPath r _ (sym r) - ∙ cong (λ x → r ∙ x ∙ (sym r)) - (doubleCompPath≡compPath s _ (sym s) - ∙ cong (λ x → s ∙ x ∙ sym s) (pathLem₂ (sym coh) mid) - ∙ pathLem₁ s (sym coh) (coh ∙∙ mid ∙∙ sym coh)) - ∙ pathLem₁ r (s ∙ sym coh) (coh ∙∙ mid ∙∙ sym coh)) - ∙ pathLem₁ q (r ∙ s ∙ sym coh) (coh ∙∙ mid ∙∙ sym coh)) - ∙∙ pathLem₁ p (q ∙ r ∙ s ∙ sym coh) (coh ∙∙ mid ∙∙ sym coh) - ∙∙ cong ((p ∙ q ∙ r ∙ s ∙ sym coh) ∙_) - (comm (coh ∙∙ mid ∙∙ sym coh) (sym (p ∙ q ∙ r ∙ s ∙ sym coh))) - ∙∙ ∙assoc _ _ _ - ∙∙ cong (_∙ (coh ∙∙ mid ∙∙ sym coh)) (rCancel _) - ∙∙ sym (lUnit _) - - pathLem₄ : ∀ {ℓ} {A : Type ℓ} {x y z : A} - (p : z ≡ x) (q : y ≡ z) (r : y ≡ y) - → (sym (q ∙ p) ∙∙ r ∙∙ (q ∙ p)) - ≡ (sym p ∙∙ sym q ∙∙ r ∙∙ q ∙∙ p) - pathLem₄ p q r = - cong (λ x → x ∙∙ r ∙∙ (q ∙ p)) (symDistr q p) - ∙∙ doubleCompPath≡compPath (sym p ∙ sym q) r (q ∙ p) - ∙∙ (sym (∙assoc (sym p) (sym q) (r ∙ q ∙ p)) - ∙∙ cong (sym p ∙_) (∙assoc (sym q) r (q ∙ p) - ∙ ∙assoc (sym q ∙ r) q p) - ∙∙ sym (doubleCompPath≡compPath (sym p) ((sym q ∙ r) ∙ q) p) - ∙ cong (sym p ∙∙_∙∙ p) - (sym (∙assoc (sym q) r q) - ∙ sym (doubleCompPath≡compPath (sym q) r q))) - - cong-lMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (x : _) - (p : refl ≡ x) (q : _ ≡ _) (a : S₊ (suc n)) - → (cong (lMap (suc n) {A = Ω A}) - (fun (flipΩIso (suc (suc n))) (p ∙∙ q ∙∙ sym p))) - ≡ (cong (lMap (suc n)) (sym (flipΩrefl n)) - ∙∙ cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) p - ∙∙ cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) q - ∙∙ sym (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) p) - ∙∙ cong (lMap (suc n)) (flipΩrefl n)) - cong-lMap n x = - J (λ x p → (q : x ≡ x) - → S₊ (suc n) - → cong (lMap (suc n)) - (fun (flipΩIso (suc (suc n))) (p ∙∙ q ∙∙ sym p)) - ≡ - (cong (lMap (suc n)) (sym (flipΩrefl n)) ∙∙ - cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) p ∙∙ - cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) q ∙∙ - sym (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) p) - ∙∙ cong (lMap (suc n)) (flipΩrefl n))) - λ q a → (λ j i → lMap (suc n) - (fun (flipΩIso (suc (suc n))) (rUnit q (~ j)) i)) - ∙∙ cong-lMap-lem n q - ∙∙ cong (cong (lMap (suc n)) (sym (flipΩrefl n)) - ∙∙_∙∙ - cong (lMap (suc n)) (flipΩrefl n)) - (rUnit (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) q)) - - main : cong ((rMap (suc (suc n)) {A = A} - ∘ suspMapΩ (suc (suc n))) p .fst) (merid a) - ≡ (sym (rCancel (merid (pt A))) - ∙∙ (λ i → (botMap (suc (suc n)) ∘ lMap (suc (suc n)) {A = A}) - p .fst (merid a i)) - ∙∙ rCancel _) - main = cong (sym (lMapId (suc n) a) ∙∙_∙∙ (lMapId (suc n) a)) - (cong (cong (λ x → fst x a)) - (cong-lMap _ _ - (sym (∙∙lCancel (snd (suspMapΩ∙ n)))) - (cong (suspMapΩ (suc n)) p) a)) - ∙∙ cong (sym (lMapId (suc n) a) ∙∙_∙∙ (lMapId (suc n) a)) - ((cong-∙∙ (λ x → fst x a) - (cong (lMap (suc n)) (sym (flipΩrefl n))) - (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) - (sym (∙∙lCancel (snd (suspMapΩ∙ n)))) - ∙∙ cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) - (cong ((suspMapΩ∙ (suc n)) .fst) p) - ∙∙ cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) - (∙∙lCancel (snd (suspMapΩ∙ n)))) - (cong (lMap (suc n)) (flipΩrefl n)))) - ∙∙ cong (sym (lMapId (suc n) a) ∙∙_∙∙ (lMapId (suc n) a)) - (cong (cong (λ x → fst x a) - (cong (lMap (suc n)) (sym (flipΩrefl n))) - ∙∙_∙∙ - cong (λ x → fst x a) - (cong (lMap (suc n)) (flipΩrefl n))) - (cong-∙∙ (λ x → fst x a) - (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) - (sym (∙∙lCancel (snd (suspMapΩ∙ n))))) - (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) - (cong ((suspMapΩ∙ (suc n)) .fst) p)) - (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) - (∙∙lCancel (snd (suspMapΩ∙ n)))))) - ∙∙ cong (sym (lMapId (suc n) a) ∙∙_∙∙ (lMapId (suc n) a)) - (cong (cong (λ x → fst x a) - (cong (lMap (suc n)) (sym (flipΩrefl n))) - ∙∙_∙∙ - cong (λ x → fst x a) - (cong (lMap (suc n)) (flipΩrefl n))) - (cong (cong (λ x → fst x a) - (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) - (sym (∙∙lCancel (snd (suspMapΩ∙ n))))) - ∙∙_∙∙ - cong (λ x → fst x a) - (cong (lMap (suc n) ∘ fun (flipΩIso (suc n))) - (∙∙lCancel (snd (suspMapΩ∙ n))))) - ((λ i j → rMap (suc n) {A = A} - (suspMapΩ∙ (suc n) .fst (p j)) .fst a) - ∙ rUnit _ - ∙ λ i → (λ j → indHyp (i ∧ j) - (snd (Ω ((Ω^ n) A))) .fst a) - ∙∙ (λ j → indHyp i (p j) .fst a) - ∙∙ λ j → indHyp (i ∧ ~ j) - (snd (Ω ((Ω^ n) A))) .fst a))) - ∙∙ pathLem₃ - (sym (lMapId (suc n) a)) - (λ i₁ → fst (lMap (suc n) (flipΩrefl n (~ i₁))) a) - (λ i₁ → fst (lMap (suc n) (fun (flipΩIso (suc n)) - (∙∙lCancel (snd (suspMapΩ∙ {A = A} n)) (~ i₁)))) a) - (λ j₁ → indHyp j₁ (snd (Ω ((Ω^ n) A))) .fst a) - (λ j₁ → (botMap (suc n) ∘ lMap (suc n)) (p j₁) .fst a) - (sym (cong (_∙ sym (merid (pt A))) - (cong merid (lMapId (suc n) a)) ∙ rCancel _)) - (EH 0) - ∙∙ pathLem₄ (rCancel (merid (pt A))) - (cong (_∙ sym (merid (pt A))) - (cong merid (lMapId (suc n) a))) - (λ j₁ → (botMap (suc n) ∘ lMap (suc n)) (p j₁) .fst a) - ∙∙ cong (sym (rCancel (merid (pt A))) ∙∙_∙∙ rCancel (merid (pt A))) - (sym (cong-∙∙ (λ x → merid x ∙ sym (merid (pt A))) - (sym (lMapId (suc n) a)) - (λ i → lMap (suc n) (p i) .fst a) - (lMapId (suc n) a))) - - --- main results -suspMap→TranspType : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → (typ (Ω ((Ω^ n) A)) → typ (Ω (Ω ((Ω^ n) (Susp∙ (typ A)))))) - ≡ ((S₊∙ (suc n) →∙ A) → (S₊∙ (suc (suc n)) →∙ Susp∙ (typ A))) -suspMap→TranspType {A = A} n i = - Ω≡SphereMap {A = A} (suc n) i → Ω≡SphereMap' {A = A} (suc n) i - -suspMap→ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → PathP (λ i → suspMap→TranspType {A = A} n i) - (suspMapΩ∙ (suc n) .fst) - (suspMap n) -suspMap→ {A = A} n = - toPathP (funExt λ f → - (λ j → transportRefl (botᵣ {A = A} (suc n) - (rMap (suc n) {A = A} - (suspMapΩ∙ (suc n) .fst - ((invEq (_ , isEquiv-lMap n) - (transportRefl f j)))))) j) + +bot□ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f : (S₊∙ (suc n) →∙ A)) + → suspMap n f ≡ botᵣ {A = A} (suc n) (post∘∙ (S₊∙ (suc n)) (σ∙ A) .fst f) +bot□ {A = A} n f = + ΣPathP ((funExt (λ { north → refl + ; south → refl + ; (merid a i) → refl})) + , refl) + + +IsoΩSphereMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (typ ((Ω^ n) A)) ((S₊∙ n →∙ A)) +fun (IsoΩSphereMap zero) = lMap zero +inv (IsoΩSphereMap zero) f = fst f false +rightInv (IsoΩSphereMap zero) f = + ΣPathP ((funExt (λ { false → refl + ; true → sym (snd f)})) + , λ i j → snd f (~ i ∨ j)) +leftInv (IsoΩSphereMap zero) p = refl +IsoΩSphereMap (suc n) = equivToIso (_ , isEquiv-lMap n) + +IsoΩSphereMapᵣ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (typ ((Ω^ (suc n)) (Susp∙ (typ A)))) ((S₊∙ (suc n) →∙ Susp∙ (typ A))) +IsoΩSphereMapᵣ {A = A} n = + compIso (flipΩIso n) + (compIso (IsoΩSphereMap n) (botᵣIso {A = A} n)) + +suspMapPathP : ∀ {ℓ} (A : Pointed ℓ) (n : ℕ) + → (typ ((Ω^ (suc n)) A) → (typ ((Ω^ (suc (suc n))) (Susp∙ (typ A))))) + ≡ ((S₊∙ (suc n) →∙ A) → S₊∙ (suc (suc n)) →∙ (Susp∙ (typ A))) +suspMapPathP A n i = + isoToPath (IsoΩSphereMap {A = A} (suc n)) i + → isoToPath (IsoΩSphereMapᵣ {A = A} (suc n)) i + +Ωσ→suspMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → PathP (λ i → suspMapPathP A n i) + (suspMapΩ (suc n)) + (suspMap n) +Ωσ→suspMap {A = A} n = + toPathP (funExt (λ p → + (λ i → transportRefl + (Iso.fun (IsoΩSphereMapᵣ {A = A} (suc n)) + (suspMapΩ {A = A} (suc n) + (Iso.inv (IsoΩSphereMap {A = A} (suc n)) + (transportRefl p i)))) i) ∙∙ cong (botᵣ {A = A} (suc n)) - (funExt⁻ (filler-top□ (suc n)) (invEq (_ , isEquiv-lMap n) f)) - ∙∙ sym (filler▿ (suc n) (lMap (suc n) {A = A} - (invEq (lMap (suc n) , isEquiv-lMap n) f))) - ∙ cong (suspMap n) (secEq ((lMap (suc n) , isEquiv-lMap n)) f)) + (cong (lMap (suc n) {A = Ω (Susp∙ (typ A)) }) + ((sym (funExt⁻ (cong fst (top□ {A = A} n)) + (invEq (lMap (suc n) , isEquiv-lMap n) p)))) + ∙ (sym (mid□ n (invEq (lMap (suc n) , isEquiv-lMap n) p)) + ∙ cong (σ∙ (fst A , snd A) ∘∙_) + (secEq (lMap (suc n) , isEquiv-lMap n) p))) + ∙∙ sym (bot□ n p))) diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda index f1c2a3f28..93681bd77 100644 --- a/Cubical/Homotopy/Loopspace.agda +++ b/Cubical/Homotopy/Loopspace.agda @@ -35,6 +35,20 @@ open Iso fst (Ω→ {A = A} {B = B} (f , p)) q = sym p ∙∙ cong f q ∙∙ p snd (Ω→ {A = A} {B = B} (f , p)) = ∙∙lCancel p +{- loop space map functoriality (missing pointedness proof) -} +Ω→∘ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + (g : B →∙ C) (f : A →∙ B) + → ∀ p → Ω→ (g ∘∙ f) .fst p ≡ (Ω→ g ∘∙ Ω→ f) .fst p +Ω→∘ g f p k i = + hcomp + (λ j → λ + { (k = i0) → doubleCompPath-filler (sym (snd (g ∘∙ f))) (cong (fst (g ∘∙ f)) p) (snd (g ∘∙ f)) j i + ; (k = i1) → doubleCompPath-filler (sym (snd g)) (cong (fst g) (fst (Ω→ f) p)) (snd g) j i + ; (i = i0) → compPath-filler' (cong (g .fst) (f .snd)) (g .snd) (~ k) j + ; (i = i1) → compPath-filler' (cong (g .fst) (f .snd)) (g .snd) (~ k) j + }) + (g .fst (doubleCompPath-filler (sym (f .snd)) (cong (f .fst) p) (f .snd) k i)) + isEquivΩ→ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → (f : (A →∙ B)) → isEquiv (fst f) → isEquiv (Ω→ f .fst) From e994768937ac6147700923f85785e57f6fe9d16e Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Thu, 11 Nov 2021 00:05:48 +0100 Subject: [PATCH 88/89] move --- Cubical/Algebra/Group/Exact.agda | 4 - Cubical/Foundations/Pointed/Properties.agda | 9 + Cubical/HITs/Sn/Properties.agda | 2 +- Cubical/HITs/Susp/Properties.agda | 50 +- Cubical/Homotopy/Freudenthal.agda | 30 - Cubical/Homotopy/Group/Base.agda | 521 +++++++++++++- Cubical/Homotopy/Group/S3.agda | 1 - .../{Properties.agda => SuspensionMap.agda} | 632 ++++++++--------- .../Homotopy/Group/SuspensionMapPathP.agda | 648 ------------------ Cubical/Homotopy/HopfInvariant/Base.agda | 2 +- .../Homotopy/HopfInvariant/Homomorphism.agda | 2 +- Cubical/Homotopy/Loopspace.agda | 84 ++- 12 files changed, 931 insertions(+), 1054 deletions(-) rename Cubical/Homotopy/Group/{Properties.agda => SuspensionMap.agda} (51%) delete mode 100644 Cubical/Homotopy/Group/SuspensionMapPathP.agda diff --git a/Cubical/Algebra/Group/Exact.agda b/Cubical/Algebra/Group/Exact.agda index ca2b8c2f4..c2fa70aff 100644 --- a/Cubical/Algebra/Group/Exact.agda +++ b/Cubical/Algebra/Group/Exact.agda @@ -11,10 +11,6 @@ open import Cubical.Algebra.Group.GroupPath open import Cubical.Algebra.Group.Instances.Unit -private - variable - ℓ : Level - -- TODO : Define exact sequences -- (perhaps short, finite, ℕ-indexed and ℤ-indexed) diff --git a/Cubical/Foundations/Pointed/Properties.agda b/Cubical/Foundations/Pointed/Properties.agda index d56e5b595..fa8693c72 100644 --- a/Cubical/Foundations/Pointed/Properties.agda +++ b/Cubical/Foundations/Pointed/Properties.agda @@ -49,6 +49,15 @@ _∘∙_ : {A : Pointed ℓA} {B : Pointed ℓB} {C : Pointed ℓC} ((g , g∙) ∘∙ (f , f∙)) .fst x = g (f x) ((g , g∙) ∘∙ (f , f∙)) .snd = (cong g f∙) ∙ g∙ +-- post composition +post∘∙ : ∀ {ℓX ℓ ℓ'} (X : Pointed ℓX) {A : Pointed ℓ} {B : Pointed ℓ'} + → (A →∙ B) → ((X →∙ A ∙) →∙ (X →∙ B ∙)) +post∘∙ X f .fst g = f ∘∙ g +post∘∙ X f .snd = + ΣPathP + ( (funExt λ _ → f .snd) + , (sym (lUnit (f .snd)) ◁ λ i j → f .snd (i ∨ j))) + -- pointed identity id∙ : (A : Pointed ℓA) → (A →∙ A) id∙ A .fst x = x diff --git a/Cubical/HITs/Sn/Properties.agda b/Cubical/HITs/Sn/Properties.agda index 92eadbe35..663f3e902 100644 --- a/Cubical/HITs/Sn/Properties.agda +++ b/Cubical/HITs/Sn/Properties.agda @@ -17,7 +17,7 @@ open import Cubical.Data.Sigma open import Cubical.HITs.Sn.Base open import Cubical.HITs.Susp open import Cubical.HITs.Truncation -open import Cubical.Homotopy.Loopspace +-- open import Cubical.Homotopy.Loopspace open import Cubical.Homotopy.Connected open import Cubical.HITs.Join open import Cubical.Data.Bool diff --git a/Cubical/HITs/Susp/Properties.agda b/Cubical/HITs/Susp/Properties.agda index 6f7d7eed5..bc017d4a1 100644 --- a/Cubical/HITs/Susp/Properties.agda +++ b/Cubical/HITs/Susp/Properties.agda @@ -7,13 +7,15 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv open import Cubical.Foundations.Path open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous open import Cubical.Foundations.GroupoidLaws -open import Cubical.Homotopy.Loopspace - open import Cubical.Data.Bool +open import Cubical.Data.Sigma + open import Cubical.HITs.Join open import Cubical.HITs.Susp.Base +open import Cubical.Homotopy.Loopspace private variable @@ -127,7 +129,47 @@ Iso.rightInv funSpaceSuspIso f = funExt λ {north → refl Iso.leftInv funSpaceSuspIso _ = refl toSusp : (A : Pointed ℓ) → typ A → typ (Ω (Susp∙ (typ A))) -toSusp A a = merid a ∙ merid (pt A) ⁻¹ +toSusp A x = merid x ∙ merid (pt A) ⁻¹ toSuspPointed : (A : Pointed ℓ) → A →∙ Ω (Susp∙ (typ A)) -toSuspPointed A = (λ a → merid a ∙ merid (pt A) ⁻¹) , rCancel (merid (pt A)) +fst (toSuspPointed A) = toSusp A +snd (toSuspPointed A) = rCancel (merid (pt A)) + +module _ {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} where + fromSusp→toΩ : Susp∙ (typ A) →∙ B → (A →∙ Ω B) + fst (fromSusp→toΩ f) x = sym (snd f) ∙∙ cong (fst f) (toSusp A x) ∙∙ snd f + snd (fromSusp→toΩ f) = + cong (sym (snd f) ∙∙_∙∙ (snd f)) + (cong (cong (fst f)) + (rCancel (merid (pt A)))) + ∙ ∙∙lCancel (snd f) + + toΩ→fromSusp : A →∙ Ω B → Susp∙ (typ A) →∙ B + fst (toΩ→fromSusp f) north = pt B + fst (toΩ→fromSusp f) south = pt B + fst (toΩ→fromSusp f) (merid a i) = fst f a i + snd (toΩ→fromSusp f) = refl + + ΩSuspAdjointIso : Iso (A →∙ Ω B) (Susp∙ (typ A) →∙ B) + fun ΩSuspAdjointIso = toΩ→fromSusp + inv ΩSuspAdjointIso = fromSusp→toΩ + rightInv ΩSuspAdjointIso f = + ΣPathP (funExt (λ { north → sym (snd f) + ; south → sym (snd f) ∙ cong (fst f) (merid (pt A)) + ; (merid a i) j → + hcomp (λ k → λ { (i = i0) → snd f (~ j ∧ k) + ; (i = i1) → compPath-filler' + (sym (snd f)) + (cong (fst f) (merid (pt A))) k j + ; (j = i1) → fst f (merid a i)}) + (fst f (compPath-filler (merid a) (sym (merid (pt A))) (~ j) i))}) + , λ i j → snd f (~ i ∨ j)) + leftInv ΩSuspAdjointIso f = + →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt λ x → sym (rUnit _) + ∙ cong-∙ (fst (toΩ→fromSusp f)) (merid x) (sym (merid (pt A))) + ∙ cong (fst f x ∙_) (cong sym (snd f)) + ∙ sym (rUnit _)) + + IsoΩFunSuspFun : Iso (typ (Ω (A →∙ B ∙))) (Susp∙ (typ A) →∙ B) + IsoΩFunSuspFun = compIso (ΩfunExtIso A B) ΩSuspAdjointIso diff --git a/Cubical/Homotopy/Freudenthal.agda b/Cubical/Homotopy/Freudenthal.agda index 087c22689..ba05ca725 100644 --- a/Cubical/Homotopy/Freudenthal.agda +++ b/Cubical/Homotopy/Freudenthal.agda @@ -133,33 +133,3 @@ FreudenthalIso : ∀ {ℓ} (n : HLevel) (A : Pointed ℓ) → Iso (hLevelTrunc ((suc n) + (suc n)) (typ A)) (hLevelTrunc ((suc n) + (suc n)) (typ (Ω (Susp (typ A) , north)))) FreudenthalIso n A iscon = connectedTruncIso _ (σ A) (isConnectedσ _ iscon) - - -{- connectedness of congⁿ σ (called suspMapΩ here) -} -suspMapΩ∙ : ∀ {ℓ} {A : Pointed ℓ}(n : ℕ) - → ((Ω^ n) A) - →∙ ((Ω^ (suc n)) (Susp∙ (typ A))) -fst (suspMapΩ∙ {A = A} zero) a = merid a ∙ sym (merid (pt A)) -snd (suspMapΩ∙ {A = A} zero) = rCancel (merid (pt A)) -suspMapΩ∙ {A = A} (suc n) = Ω→ (suspMapΩ∙ {A = A} n) - -suspMapΩ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → typ ((Ω^ n) A) → typ ((Ω^ (suc n)) (Susp∙ (typ A))) -suspMapΩ {A = A} n = suspMapΩ∙ {A = A} n .fst - -suspMapΩ-connected : ∀ {ℓ} (n : HLevel) (m : ℕ) {A : Pointed ℓ} - (connA : isConnected (suc (suc n)) (typ A)) - → isConnectedFun ((suc n + suc n) ∸ m) (suspMapΩ {A = A} m) -suspMapΩ-connected n zero {A = A} connA = isConnectedσ n connA -suspMapΩ-connected n (suc m) {A = A} connA with ((n + suc n) ≟ m) -... | (lt p) = subst (λ x → isConnectedFun x (suspMapΩ {A = A} (suc m))) - (sym (n∸m≡0 _ m p)) - λ b → tt* , (λ {tt* → refl}) -... | (eq q) = subst (λ x → isConnectedFun x (suspMapΩ {A = A} (suc m))) - (sym (n∸n≡0 m) ∙ cong (_∸ m) (sym q)) - λ b → tt* , (λ {tt* → refl}) -... | (gt p) = isConnectedCong' (n + suc n ∸ m) (suspMapΩ {A = A} m) - (subst (λ x → isConnectedFun x (suspMapΩ {A = A} m)) - (sym (suc∸-fst (n + suc n) m p)) - (suspMapΩ-connected n m connA)) - (snd (suspMapΩ∙ m)) diff --git a/Cubical/Homotopy/Group/Base.agda b/Cubical/Homotopy/Group/Base.agda index d2f85a816..3d3c79a68 100644 --- a/Cubical/Homotopy/Group/Base.agda +++ b/Cubical/Homotopy/Group/Base.agda @@ -5,30 +5,37 @@ open import Cubical.Homotopy.Loopspace open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous open import Cubical.Foundations.HLevels open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) +open import Cubical.Foundations.Path +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Functions.Morphism open import Cubical.HITs.SetTruncation renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim ; elim2 to sElim2 ; elim3 to sElim3 ; map to sMap) open import Cubical.HITs.Sn -open import Cubical.HITs.Susp +open import Cubical.HITs.Susp renaming (toSusp to σ) open import Cubical.HITs.S1 open import Cubical.Data.Sigma open import Cubical.Data.Nat +open import Cubical.Data.Bool open import Cubical.Algebra.Group open import Cubical.Algebra.Semigroup open import Cubical.Algebra.Monoid +open Iso open IsGroup open IsSemigroup open IsMonoid open GroupStr - {- Homotopy group -} π : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Type ℓ π n A = ∥ typ ((Ω^ n) A) ∥₂ @@ -116,3 +123,513 @@ fst (-Π {A = A} {n = suc (suc n)} f) south = fst f north fst (-Π {A = A} {n = suc (suc n)} f) (merid a j) = fst f ((merid a ∙ sym (merid (ptSn _))) (~ j)) snd (-Π {A = A} {n = suc (suc n)} f) = snd f + + +-- to prove that this gives a group structure on π', we first +-- prove that Ωⁿ A ≃ (Sⁿ →∙ A). +-- We use the following map +mutual + Ω→SphereMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} + → typ ((Ω^ n) A) → (S₊∙ n →∙ A) + fst (Ω→SphereMap zero a) false = a + fst (Ω→SphereMap zero {A = A} a) true = pt A + snd (Ω→SphereMap zero a) = refl + fst (Ω→SphereMap (suc zero) {A = A} p) base = pt A + fst (Ω→SphereMap (suc zero) p) (loop i) = p i + snd (Ω→SphereMap (suc zero) p) = refl + fst (Ω→SphereMap (suc (suc n)) {A = A} p) north = pt A + fst (Ω→SphereMap (suc (suc n)) {A = A} p) south = pt A + fst (Ω→SphereMap (suc (suc n)) p) (merid a i) = + (sym (Ω→SphereMapId (suc n) a) + ∙∙ (λ i → Ω→SphereMap (suc n) (p i) .fst a) + ∙∙ Ω→SphereMapId (suc n) a) i + snd (Ω→SphereMap (suc (suc n)) p) = refl + + Ω→SphereMapId : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (a : _) + → Ω→SphereMap n {A = A} (pt ((Ω^ n) A)) .fst a ≡ pt A + Ω→SphereMapId zero false = refl + Ω→SphereMapId zero true = refl + Ω→SphereMapId (suc zero) base = refl + Ω→SphereMapId (suc zero) (loop i) = refl + Ω→SphereMapId (suc (suc n)) north = refl + Ω→SphereMapId (suc (suc n)) south = refl + Ω→SphereMapId (suc (suc n)) {A = A} (merid a i) j = + ∙∙lCancel (Ω→SphereMapId (suc n) {A = A} a) j i + +Ω→SphereMapId2 : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} + → Ω→SphereMap n {A = A} (pt ((Ω^ n) A)) ≡ ((λ _ → pt A) , refl) +fst (Ω→SphereMapId2 n {A = A} i) a = funExt (Ω→SphereMapId n {A = A}) i a +snd (Ω→SphereMapId2 zero {A = A} i) = refl +snd (Ω→SphereMapId2 (suc zero) {A = A} i) = refl +snd (Ω→SphereMapId2 (suc (suc n)) {A = A} i) = refl + +-- Pointed version +Ω→SphereMap∙ : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} + → ((Ω^ n) A) →∙ (S₊∙ n →∙ A ∙) +Ω→SphereMap∙ n .fst = Ω→SphereMap n +Ω→SphereMap∙ n .snd = Ω→SphereMapId2 n + +-- We define the following maps which will be used to +-- show that Ω→SphereMap is an equivalence +Ω→SphereMapSplit₁ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → typ ((Ω^ (suc n)) A) + → typ (Ω (S₊∙ n →∙ A ∙)) +Ω→SphereMapSplit₁ n = Ω→ (Ω→SphereMap∙ n) .fst + +ΩSphereMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → typ (Ω (S₊∙ n →∙ A ∙)) + → (S₊∙ (suc n) →∙ A) +fst (ΩSphereMap {A = A} zero p) base = p i0 .fst false +fst (ΩSphereMap {A = A} zero p) (loop i) = p i .fst false +snd (ΩSphereMap {A = A} zero p) = refl +ΩSphereMap {A = A} (suc n) = fun IsoΩFunSuspFun + +-- Functoriality +-- The homogeneity assumption is not necessary but simplifying +isNaturalΩSphereMap : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') + (homogB : isHomogeneous B) (f : A →∙ B) (n : ℕ) + → ∀ g → f ∘∙ ΩSphereMap n g ≡ ΩSphereMap n (Ω→ (post∘∙ (S₊∙ n) f) .fst g) +isNaturalΩSphereMap A B homogB f 0 g = + →∙Homogeneous≡ homogB (funExt lem) + where + lem : ∀ x → f .fst (ΩSphereMap 0 g .fst x) + ≡ ΩSphereMap 0 (Ω→ (post∘∙ (S₊∙ 0) f) .fst g) .fst x + lem base = f .snd + lem (loop i) j = + hfill + (λ j → λ + { (i = i0) → post∘∙ _ f .snd j + ; (i = i1) → post∘∙ _ f .snd j + }) + (inS (f ∘∙ g i)) + j .fst false +isNaturalΩSphereMap A B homogB f (n@(suc _)) g = + →∙Homogeneous≡ homogB (funExt lem) + where + lem : ∀ x → f .fst (ΩSphereMap n g .fst x) + ≡ ΩSphereMap n (Ω→ (post∘∙ (S₊∙ n) f) .fst g) .fst x + lem north = f .snd + lem south = f .snd + lem (merid a i) j = + hfill + (λ j → λ + { (i = i0) → post∘∙ _ f .snd j + ; (i = i1) → post∘∙ _ f .snd j + }) + (inS (f ∘∙ g i)) + j .fst a + +SphereMapΩ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (S₊∙ (suc n) →∙ A) + → typ (Ω (S₊∙ n →∙ A ∙)) +SphereMapΩ {A = A} zero (f , p) = + ΣPathP ((funExt λ { false → sym p ∙∙ cong f loop ∙∙ p + ; true → refl}) + , refl) +SphereMapΩ {A = A} (suc n) = inv IsoΩFunSuspFun + +SphereMapΩIso : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (S₊∙ (suc n) →∙ A) + (typ (Ω (S₊∙ n →∙ A ∙))) +fun (SphereMapΩIso n) = SphereMapΩ n +inv (SphereMapΩIso n) = ΩSphereMap n +fst (rightInv (SphereMapΩIso zero) f i j) false = rUnit (λ j → fst (f j) false) (~ i) j +fst (rightInv (SphereMapΩIso {A = A} zero) f i j) true = snd (f j) (~ i) +snd (rightInv (SphereMapΩIso {A = A} zero) f i j) k = snd (f j) (~ i ∨ k) +rightInv (SphereMapΩIso (suc n)) = leftInv IsoΩFunSuspFun +leftInv (SphereMapΩIso zero) f = + ΣPathP ((funExt (λ { base → sym (snd f) + ; (loop i) j → doubleCompPath-filler + (sym (snd f)) + (cong (fst f) loop) + (snd f) (~ j) i})) + , λ i j → snd f (~ i ∨ j)) +leftInv (SphereMapΩIso (suc n)) = rightInv IsoΩFunSuspFun + +{- +In order to show that Ω→SphereMap is an equivalence, we show that it factors + + Ω→SphereMapSplit₁ ΩSphereMap +Ωⁿ⁺¹(Sⁿ →∙ A) ----------------> Ω (Sⁿ →∙ A) -----------> (Sⁿ⁺¹ →∙ A) +-} + +Ω→SphereMap-split : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p : typ ((Ω^ (suc n)) A)) + → Ω→SphereMap (suc n) p ≡ ΩSphereMap n (Ω→SphereMapSplit₁ n p) +Ω→SphereMap-split {A = A} zero p = + ΣPathP ((funExt (λ { base → refl + ; (loop i) j → lem (~ j) i})) + , refl) + where + lem : funExt⁻ (cong fst (Ω→SphereMapSplit₁ zero p)) false ≡ p + lem = (λ i → funExt⁻ (cong-∙∙ fst (sym (Ω→SphereMapId2 zero)) + (cong (Ω→SphereMap zero) p) + (Ω→SphereMapId2 zero) i) false) + ∙ sym (rUnit _) +Ω→SphereMap-split {A = A} (suc n) p = + ΣPathP ((funExt (λ { north → refl + ; south → refl + ; (merid a i) j → lem₂ a j i})) + , refl) + where + lem : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (a : S₊ (suc n)) + → Ω→SphereMapId (suc n) {A = A} a + ≡ (λ i → fst (Ω→SphereMapId2 (suc n) {A = A} i) a) + lem zero base = refl + lem zero (loop i) = refl + lem (suc n) north = refl + lem (suc n) south = refl + lem (suc n) (merid a i) = refl + + lem₂ : (a : S₊ (suc n)) + → ((λ i₁ → Ω→SphereMapId (suc n) {A = A} a (~ i₁)) + ∙∙ (λ i₁ → Ω→SphereMap (suc n) (p i₁) .fst a) + ∙∙ Ω→SphereMapId (suc n) a) + ≡ (λ i → Ω→SphereMapSplit₁ (suc n) p i .fst a) + lem₂ a = cong (λ x → sym x + ∙∙ funExt⁻ (cong fst (λ i → Ω→SphereMap (suc n) (p i))) a + ∙∙ x) + (lem n a) + ∙∙ sym (cong-∙∙ (λ x → x a) + (cong fst (λ i → Ω→SphereMapId2 (suc n) (~ i))) + (cong fst (λ i → Ω→SphereMap (suc n) (p i))) + (cong fst (Ω→SphereMapId2 (suc n)))) + ∙∙ (λ i → funExt⁻ (cong-∙∙ fst (sym (Ω→SphereMapId2 (suc n))) + (cong (Ω→SphereMap (suc n)) p) + (Ω→SphereMapId2 (suc n)) (~ i)) a) + +isEquiv-Ω→SphereMap₀ : ∀ {ℓ} {A : Pointed ℓ} + → isEquiv (Ω→SphereMap 0 {A = A}) +isEquiv-Ω→SphereMap₀ {A = A} = + isoToIsEquiv + (iso _ (λ f → fst f false) + (λ f → ΣPathP ((funExt (λ { false → refl ; true → sym (snd f)})) + , λ i j → snd f (~ i ∨ j))) + λ p → refl) + +isEquiv-Ω→SphereMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} + → isEquiv (Ω→SphereMap n {A = A}) +isEquiv-Ω→SphereMap zero {A = A} = + (isoToIsEquiv + (iso _ (λ f → fst f false) + (λ f → ΣPathP ((funExt (λ { false → refl + ; true → sym (snd f)})) + , λ i j → snd f (~ i ∨ j))) + λ _ → refl)) +isEquiv-Ω→SphereMap (suc zero) {A = A} = + isoToIsEquiv (iso _ invFun sec λ p → sym (rUnit p)) + where + invFun : S₊∙ 1 →∙ A → typ (Ω A) + invFun (f , p) = sym p ∙∙ cong f loop ∙∙ p + + sec : section (Ω→SphereMap 1) invFun + sec (f , p) = + ΣPathP ((funExt (λ { base → sym p + ; (loop i) j → doubleCompPath-filler + (sym p) (cong f loop) p (~ j) i})) + , λ i j → p (~ i ∨ j)) + +isEquiv-Ω→SphereMap (suc (suc n)) = + subst isEquiv (sym (funExt (Ω→SphereMap-split (suc n)))) + (snd (compEquiv + ((Ω→SphereMapSplit₁ (suc n)) , + (isEquivΩ→ (Ω→SphereMap (suc n) , Ω→SphereMapId2 (suc n)) + (isEquiv-Ω→SphereMap (suc n)))) + (invEquiv (isoToEquiv (SphereMapΩIso (suc n)))))) + +IsoΩSphereMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (typ ((Ω^ n) A)) (S₊∙ n →∙ A) +IsoΩSphereMap n = equivToIso (_ , isEquiv-Ω→SphereMap n) + +IsoSphereMapΩ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (S₊∙ n →∙ A) (fst ((Ω^ n) A)) +IsoSphereMapΩ {A = A} n = + invIso (IsoΩSphereMap n) + +SphereMap→Ω : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → S₊∙ n →∙ A → fst ((Ω^ n) A) +SphereMap→Ω n = fun (IsoSphereMapΩ n) + +isHom-Ω→SphereMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (p q : _) + → Ω→SphereMap (suc n) {A = A} (p ∙ q) + ≡ ∙Π (Ω→SphereMap (suc n) {A = A} p) + (Ω→SphereMap (suc n) {A = A} q) +isHom-Ω→SphereMap zero {A = A} p q = + ΣPathP ((funExt (λ { base → refl + ; (loop i) j → (rUnit p j ∙ rUnit q j) i})) + , refl) +isHom-Ω→SphereMap (suc n) {A = A} p q = + ΣPathP ((funExt (λ { north → refl + ; south → refl + ; (merid a i) j → main a j i})) + , refl) + where + doubleComp-lem : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) (q r : y ≡ y) + → (p ∙∙ q ∙∙ sym p) ∙ (p ∙∙ r ∙∙ sym p) + ≡ (p ∙∙ (q ∙ r) ∙∙ sym p) + doubleComp-lem p q r i j = + hcomp (λ k → λ { (i = i0) → (doubleCompPath-filler p q (sym p) k + ∙ doubleCompPath-filler p r (sym p) k) j + ; (i = i1) → doubleCompPath-filler p (q ∙ r) (sym p) k j + ; (j = i0) → p (~ k) + ; (j = i1) → p (~ k)}) + ((q ∙ r) j) + + lem : (p : typ ((Ω^ (suc (suc n))) A)) + → cong (fst (Ω→SphereMap (suc (suc n)) p)) (merid (ptSn _)) ≡ refl + lem p = + cong (sym (Ω→SphereMapId (suc n) (ptSn _)) ∙∙_∙∙ Ω→SphereMapId (suc n) (ptSn _)) + (rUnit _ ∙ (λ j → (λ i → Ω→SphereMap (suc n) {A = A} refl .snd (i ∧ j)) + ∙∙ (λ i → Ω→SphereMap (suc n) {A = A} (p i) .snd j) + ∙∙ λ i → Ω→SphereMap (suc n) {A = A} refl .snd (~ i ∧ j)) + ∙ ∙∙lCancel _) + ∙ ∙∙lCancel _ + + main : (a : S₊ (suc n)) + → sym (Ω→SphereMapId (suc n) a) + ∙∙ funExt⁻ (cong fst (cong (Ω→SphereMap (suc n)) (p ∙ q))) a + ∙∙ Ω→SphereMapId (suc n) a + ≡ cong (fst (∙Π (Ω→SphereMap (suc (suc n)) p) (Ω→SphereMap (suc (suc n)) q))) (merid a) + main a = (cong (sym (Ω→SphereMapId (suc n) a) ∙∙_∙∙ (Ω→SphereMapId (suc n) a)) + (cong-∙ (λ x → Ω→SphereMap (suc n) x .fst a) p q) + ∙ sym (doubleComp-lem (sym (Ω→SphereMapId (suc n) a)) _ _)) + ∙∙ cong₂ _∙_ (sym (cong (cong (fst (Ω→SphereMap (suc (suc n)) p)) (merid a) ∙_) + (cong sym (lem p)) ∙ sym (rUnit _))) + (sym (cong (cong (fst (Ω→SphereMap (suc (suc n)) q)) (merid a) ∙_) + (cong sym (lem q)) ∙ sym (rUnit _))) + ∙∙ λ i → (rUnit (cong-∙ (fst (Ω→SphereMap (suc (suc n)) p)) + (merid a) (sym (merid (ptSn _))) (~ i)) i) + ∙ (rUnit (cong-∙ (fst (Ω→SphereMap (suc (suc n)) q)) + (merid a) (sym (merid (ptSn _)))(~ i)) i) + + +-- The iso is structure preserving +IsoSphereMapΩ-pres∙Π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f g : S₊∙ (suc n) →∙ A) + → SphereMap→Ω (suc n) (∙Π f g) + ≡ SphereMap→Ω (suc n) f ∙ SphereMap→Ω (suc n) g +IsoSphereMapΩ-pres∙Π n = + morphLemmas.isMorphInv _∙_ ∙Π (Ω→SphereMap (suc n)) + (isHom-Ω→SphereMap n) + (SphereMap→Ω (suc n)) + (leftInv (IsoSphereMapΩ (suc n))) + (rightInv (IsoSphereMapΩ (suc n))) + +-- It is useful to define the ``Group Structure'' on (S₊∙ n →∙ A) +-- before doing it on π'. These will be the equivalents of the +-- usual groupoid laws on Ω A. +1Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → (S₊∙ n →∙ A) +fst (1Π {A = A}) _ = pt A +snd (1Π {A = A}) = refl + +∙Π-rUnit : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ∙Π f 1Π ≡ f +fst (∙Π-rUnit {A = A} {n = zero} f i) base = snd f (~ i) +fst (∙Π-rUnit {A = A} {n = zero} f i) (loop j) = help i j + where + help : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) + (((sym (snd f)) ∙∙ (cong (fst f) loop) ∙∙ snd f) + ∙ (refl ∙ refl)) + (cong (fst f) loop) + help = (cong ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) ∙_) + (sym (rUnit refl)) ∙ sym (rUnit _)) + ◁ λ i j → doubleCompPath-filler (sym (snd f)) + (cong (fst f) loop) (snd f) (~ i) j +snd (∙Π-rUnit {A = A} {n = zero} f i) j = snd f (~ i ∨ j) +fst (∙Π-rUnit {A = A} {n = suc n} f i) north = snd f (~ i) +fst (∙Π-rUnit {A = A} {n = suc n} f i) south = + (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i +fst (∙Π-rUnit {A = A} {n = suc n} f i) (merid a j) = help i j + where + help : PathP (λ i → snd f (~ i) + ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) + (((sym (snd f)) + ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + ∙∙ snd f) + ∙ (refl ∙ refl)) + (cong (fst f) (merid a)) + help = (cong (((sym (snd f)) + ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + ∙∙ snd f) ∙_) + (sym (rUnit refl)) + ∙ sym (rUnit _)) + ◁ λ i j → hcomp (λ k → + λ { (j = i0) → snd f (~ i ∧ k) + ; (j = i1) → compPath-filler' (sym (snd f)) + (cong (fst f) (merid (ptSn (suc n)))) k i + ; (i = i0) → doubleCompPath-filler (sym (snd f)) + (cong (fst f) + (merid a ∙ sym (merid (ptSn (suc n))))) + (snd f) k j + ; (i = i1) → fst f (merid a j)}) + (fst f (compPath-filler (merid a) + (sym (merid (ptSn _))) (~ i) j)) +snd (∙Π-rUnit {A = A} {n = suc n} f i) j = snd f (~ i ∨ j) + +∙Π-lUnit : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ∙Π 1Π f ≡ f +fst (∙Π-lUnit {n = zero} f i) base = snd f (~ i) +fst (∙Π-lUnit {n = zero} f i) (loop j) = s i j + where + s : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) + ((refl ∙ refl) ∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) + (cong (fst f) loop) + s = (cong (_∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) + (sym (rUnit refl)) ∙ sym (lUnit _)) + ◁ λ i j → doubleCompPath-filler (sym (snd f)) + (cong (fst f) loop) (snd f) (~ i) j +snd (∙Π-lUnit {n = zero} f i) j = snd f (~ i ∨ j) +fst (∙Π-lUnit {n = suc n} f i) north = snd f (~ i) +fst (∙Π-lUnit {n = suc n} f i) south = + (sym (snd f) ∙ cong (fst f) (merid (ptSn _))) i +fst (∙Π-lUnit {n = suc n} f i) (merid a j) = help i j + where + help : PathP (λ i → snd f (~ i) + ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) + ((refl ∙ refl) ∙ ((sym (snd f)) + ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + ∙∙ snd f)) + (cong (fst f) (merid a)) + help = + (cong (_∙ ((sym (snd f)) + ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + ∙∙ snd f)) + (sym (rUnit refl)) + ∙ sym (lUnit _)) + ◁ λ i j → hcomp (λ k → + λ { (j = i0) → snd f (~ i ∧ k) + ; (j = i1) → compPath-filler' (sym (snd f)) + (cong (fst f) (merid (ptSn (suc n)))) k i + ; (i = i0) → doubleCompPath-filler (sym (snd f)) + (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) + (snd f) k j + ; (i = i1) → fst f (merid a j)}) + (fst f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) +snd (∙Π-lUnit {n = suc n} f i) j = snd f (~ i ∨ j) + +∙Π-rCancel : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ∙Π f (-Π f) ≡ 1Π +fst (∙Π-rCancel {A = A} {n = zero} f i) base = pt A +fst (∙Π-rCancel {A = A} {n = zero} f i) (loop j) = + rCancel (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) i j +snd (∙Π-rCancel {A = A} {n = zero} f i) = refl +fst (∙Π-rCancel {A = A} {n = suc n} f i) north = pt A +fst (∙Π-rCancel {A = A} {n = suc n} f i) south = pt A +fst (∙Π-rCancel {A = A} {n = suc n} f i) (merid a i₁) = lem i i₁ + where + pl = (sym (snd f) + ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) + ∙∙ snd f) + + lem : pl + ∙ ((sym (snd f) + ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) + ∙∙ snd f)) ≡ refl + lem = cong (pl ∙_) (cong (sym (snd f) ∙∙_∙∙ (snd f)) + (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) + ∙∙ cong₂ _∙_ refl + (cong (cong (fst f)) (rCancel (merid (ptSn _)))) + ∙∙ sym (rUnit _))) + ∙ rCancel pl +snd (∙Π-rCancel {A = A} {n = suc n} f i) = refl + +∙Π-lCancel : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f : S₊∙ (suc n) →∙ A) + → ∙Π (-Π f) f ≡ 1Π +fst (∙Π-lCancel {A = A} {n = zero} f i) base = pt A +fst (∙Π-lCancel {A = A} {n = zero} f i) (loop j) = + rCancel (sym (snd f) ∙∙ cong (fst f) (sym loop) ∙∙ snd f) i j +fst (∙Π-lCancel {A = A} {n = suc n} f i) north = pt A +fst (∙Π-lCancel {A = A} {n = suc n} f i) south = pt A +fst (∙Π-lCancel {A = A} {n = suc n} f i) (merid a j) = lem i j + where + pl = (sym (snd f) + ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) + ∙∙ snd f) + + lem : (sym (snd f) + ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) + ∙∙ snd f) ∙ pl + ≡ refl + lem = cong (_∙ pl) (cong (sym (snd f) ∙∙_∙∙ (snd f)) + (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) + ∙∙ cong₂ _∙_ refl (cong (cong (fst f)) (rCancel (merid (ptSn _)))) + ∙∙ sym (rUnit _))) + ∙ lCancel pl +snd (∙Π-lCancel {A = A} {n = zero} f i) = refl +snd (∙Π-lCancel {A = A} {n = suc n} f i) = refl + +∙Π-assoc : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f g h : S₊∙ (suc n) →∙ A) + → ∙Π f (∙Π g h) ≡ ∙Π (∙Π f g) h +∙Π-assoc {n = n} f g h = + sym (leftInv (IsoSphereMapΩ (suc n)) (∙Π f (∙Π g h))) + ∙∙ cong (Ω→SphereMap (suc n)) (IsoSphereMapΩ-pres∙Π n f (∙Π g h) + ∙∙ cong (SphereMap→Ω (suc n) f ∙_) (IsoSphereMapΩ-pres∙Π n g h) + ∙∙ ∙assoc (SphereMap→Ω (suc n) f) (SphereMap→Ω (suc n) g) (SphereMap→Ω (suc n) h) + ∙∙ cong (_∙ SphereMap→Ω (suc n) h) (sym (IsoSphereMapΩ-pres∙Π n f g)) + ∙∙ sym (IsoSphereMapΩ-pres∙Π n (∙Π f g) h)) + ∙∙ leftInv (IsoSphereMapΩ (suc n)) (∙Π (∙Π f g) h) + +∙Π-comm : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} + → (f g : S₊∙ (suc (suc n)) →∙ A) + → ∙Π f g ≡ ∙Π g f +∙Π-comm {A = A} {n = n} f g = + sym (leftInv (IsoSphereMapΩ (suc (suc n))) (∙Π f g)) + ∙∙ cong (Ω→SphereMap (suc (suc n))) (IsoSphereMapΩ-pres∙Π (suc n) f g + ∙∙ EH _ _ _ + ∙∙ sym (IsoSphereMapΩ-pres∙Π (suc n) g f)) + ∙∙ leftInv (IsoSphereMapΩ (suc (suc n))) (∙Π g f) + +{- π'' as a group -} +1π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' n A +1π' n {A = A} = ∣ 1Π ∣₂ + +·π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' (suc n) A → π' (suc n) A → π' (suc n) A +·π' n = sRec2 squash₂ λ p q → ∣ ∙Π p q ∣₂ + +-π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' (suc n) A → π' (suc n) A +-π' n = sMap -Π + +π'-rUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) + → (·π' n x (1π' (suc n))) ≡ x +π'-rUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-rUnit p i ∣₂ + +π'-lUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) + → (·π' n (1π' (suc n)) x) ≡ x +π'-lUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-lUnit p i ∣₂ + +π'-rCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) + → (·π' n x (-π' n x)) ≡ 1π' (suc n) +π'-rCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-rCancel p i ∣₂ + +π'-lCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) + → (·π' n (-π' n x) x) ≡ 1π' (suc n) +π'-lCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-lCancel p i ∣₂ + +π'-assoc : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y z : π' (suc n) A) + → ·π' n x (·π' n y z) ≡ ·π' n (·π' n x y) z +π'-assoc n = sElim3 (λ _ _ _ → isSetPathImplicit) λ p q r i → ∣ ∙Π-assoc p q r i ∣₂ + +π'-comm : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y : π' (suc (suc n)) A) + → ·π' (suc n) x y ≡ ·π' (suc n) y x +π'-comm n = sElim2 (λ _ _ → isSetPathImplicit) λ p q i → ∣ ∙Π-comm p q i ∣₂ + +-- We finally get the group definition +π'Gr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Group ℓ +fst (π'Gr n A) = π' (suc n) A +1g (snd (π'Gr n A)) = 1π' (suc n) +GroupStr._·_ (snd (π'Gr n A)) = ·π' n +inv (snd (π'Gr n A)) = -π' n +is-set (isSemigroup (isMonoid (isGroup (snd (π'Gr n A))))) = squash₂ +assoc (isSemigroup (isMonoid (isGroup (snd (π'Gr n A))))) = π'-assoc n +identity (isMonoid (isGroup (snd (π'Gr n A)))) x = (π'-rUnit n x) , (π'-lUnit n x) +inverse (isGroup (snd (π'Gr n A))) x = (π'-rCancel n x) , (π'-lCancel n x) + +-- and finally, the Iso +π'Gr≅πGr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → GroupIso (π'Gr n A) (πGr n A) +fst (π'Gr≅πGr n A) = setTruncIso (IsoSphereMapΩ (suc n)) +snd (π'Gr≅πGr n A) = + makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) + λ p q i → ∣ IsoSphereMapΩ-pres∙Π n p q i ∣₂) diff --git a/Cubical/Homotopy/Group/S3.agda b/Cubical/Homotopy/Group/S3.agda index e798b5b62..97246ba64 100644 --- a/Cubical/Homotopy/Group/S3.agda +++ b/Cubical/Homotopy/Group/S3.agda @@ -17,7 +17,6 @@ open import Cubical.Data.Int renaming (_·_ to _·ℤ_ ; _+_ to _+ℤ_) open import Cubical.Homotopy.Group.Base -open import Cubical.Homotopy.Group.Properties open import Cubical.Homotopy.HopfInvariant.Base open import Cubical.Homotopy.HopfInvariant.Homomorphism open import Cubical.Homotopy.HopfInvariant.HopfMap diff --git a/Cubical/Homotopy/Group/Properties.agda b/Cubical/Homotopy/Group/SuspensionMap.agda similarity index 51% rename from Cubical/Homotopy/Group/Properties.agda rename to Cubical/Homotopy/Group/SuspensionMap.agda index b189b7958..e1f75a1f3 100644 --- a/Cubical/Homotopy/Group/Properties.agda +++ b/Cubical/Homotopy/Group/SuspensionMap.agda @@ -1,385 +1,306 @@ {-# OPTIONS --safe --experimental-lossy-unification #-} -{- -This file contains: -1. A proof that the equivalence Ωⁿ A ≃ (Sⁿ →∙ A) -is structure preserving - -2. Using the above, the complete group structure on (Sⁿ →∙ A), -the alternative definition of homotopy groups - -4. A proof that the dependent path in Homotopy.Group.SuspensionMapPathP -is structure preserving. - -5. A group isomorphism of the different definitions of homotopy groups - -6. Connectivity of the suspension map - -7. Surjectivity of the suspension map π₂₊ₙ(S¹⁺ⁿ) → π₃₊ₙ(S²⁺ⁿ) --} -module Cubical.Homotopy.Group.Properties where +module Cubical.Homotopy.Group.SuspensionMap where open import Cubical.Homotopy.Group.Base -open import Cubical.Homotopy.Group.SuspensionMapPathP -open import Cubical.Homotopy.Group.SuspensionMapPathP - using (IsoΩSphereMap) public open import Cubical.Homotopy.Loopspace open import Cubical.Homotopy.Freudenthal open import Cubical.Homotopy.Connected open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Equiv open import Cubical.Foundations.Pointed.Homogeneous open import Cubical.Foundations.HLevels open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Function -open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Path +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Transport open import Cubical.Functions.Morphism -open import Cubical.HITs.PropositionalTruncation - renaming (rec to pRec ; rec2 to pRec2 ; elim to pElim) -open import Cubical.HITs.SetTruncation - renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim - ; elim2 to sElim2 ; elim3 to sElim3 ; map to sMap) -open import Cubical.HITs.Truncation - renaming (rec to trRec) -open import Cubical.HITs.S1 open import Cubical.HITs.Sn -open import Cubical.HITs.Susp -open import Cubical.Data.Bool +open import Cubical.HITs.Susp renaming (toSusp to σ ; toSuspPointed to σ∙) +open import Cubical.HITs.S1 + +open import Cubical.Data.Bool hiding (_≟_) open import Cubical.Data.Sigma open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order open import Cubical.Algebra.Group open import Cubical.Algebra.Semigroup open import Cubical.Algebra.Monoid + +open import Cubical.HITs.PropositionalTruncation + renaming (rec to pRec ; rec2 to pRec2 ; elim to pElim) +open import Cubical.HITs.SetTruncation + renaming (rec to sRec ; rec2 to sRec2 ; elim to sElim + ; elim2 to sElim2 ; elim3 to sElim3 ; map to sMap) +open import Cubical.HITs.Truncation + renaming (rec to trRec) + open Iso open IsGroup open IsSemigroup open IsMonoid open GroupStr --- We finally arrive at the main result +{- +This file concerns the suspension maps +suspMapΩ : πₙA → πₙ₊₁ΣA and +suspMap : π'ₙA → π'ₙ₊₁ΣA +For instance, we want to transport freudenthal +for suspMapΩ to suspMap by establishing a dependent +path between the two functions. + +This gives, in particular, surjectivity of +πₙ₊₁(Sⁿ) → πₙ₊₂(Sⁿ⁺¹) for n ≥ 2. +-} -IsoSphereMapΩ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → Iso (S₊∙ n →∙ A) (fst ((Ω^ n) A)) -IsoSphereMapΩ n = invIso (IsoΩSphereMap n) +-- Definition of the suspension functions +suspMap : ∀ {ℓ} {A : Pointed ℓ}(n : ℕ) + → S₊∙ (suc n) →∙ A + → S₊∙ (suc (suc n)) →∙ Susp∙ (typ A) +fst (suspMap n f) north = north +fst (suspMap n f) south = north +fst (suspMap {A = A} n f) (merid a i) = + (merid (fst f a) ∙ sym (merid (pt A))) i +snd (suspMap n f) = refl -SphereMap→Ω : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → (S₊∙ n →∙ A) → (fst ((Ω^ n) A)) -SphereMap→Ω n = fun (IsoSphereMapΩ n) +suspMapΩ∙ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → ((Ω^ n) A) + →∙ ((Ω^ (suc n)) (Susp∙ (typ A))) +fst (suspMapΩ∙ {A = A} zero) a = merid a ∙ sym (merid (pt A)) +snd (suspMapΩ∙ {A = A} zero) = rCancel (merid (pt A)) +suspMapΩ∙ {A = A} (suc n) = Ω→ (suspMapΩ∙ {A = A} n) -Ω→SphereMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → (fst ((Ω^ n) A)) → (S₊∙ n →∙ A) -Ω→SphereMap n = inv (IsoSphereMapΩ n) +suspMapΩ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → typ ((Ω^ n) A) → typ ((Ω^ (suc n)) (Susp∙ (typ A))) +suspMapΩ {A = A} n = suspMapΩ∙ {A = A} n .fst -isHom-lMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (p q : _) - → lMap (suc n) {A = A} (p ∙ q) - ≡ ∙Π (lMap (suc n) p) (lMap (suc n) q) -isHom-lMap zero p q = - ΣPathP ((funExt (λ { base → refl - ; (loop i) j → (rUnit p j ∙ rUnit q j) i})) - , refl) -isHom-lMap (suc n) {A = A} p q = - ΣPathP ((funExt (λ { north → refl - ; south → refl - ; (merid a i) j → main a j i})) - , refl) - where - doubleComp-lem : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) (q r : y ≡ y) - → (p ∙∙ q ∙∙ sym p) ∙ (p ∙∙ r ∙∙ sym p) - ≡ (p ∙∙ (q ∙ r) ∙∙ sym p) - doubleComp-lem p q r i j = - hcomp (λ k → λ { (i = i0) → (doubleCompPath-filler p q (sym p) k - ∙ doubleCompPath-filler p r (sym p) k) j - ; (i = i1) → doubleCompPath-filler p (q ∙ r) (sym p) k j - ; (j = i0) → p (~ k) - ; (j = i1) → p (~ k)}) - ((q ∙ r) j) - - lem : (p : typ ((Ω^ (suc (suc n))) A)) - → cong (fst (lMap (suc (suc n)) p)) (merid (ptSn _)) ≡ refl - lem p = - cong (sym (lMapId (suc n) (ptSn _)) ∙∙_∙∙ lMapId (suc n) (ptSn _)) - (rUnit _ ∙ (λ j → (λ i → lMap (suc n) {A = A} refl .snd (i ∧ j)) - ∙∙ (λ i → lMap (suc n) {A = A} (p i) .snd j) - ∙∙ λ i → lMap (suc n) {A = A} refl .snd (~ i ∧ j)) - ∙ ∙∙lCancel _) - ∙ ∙∙lCancel _ - - main : (a : S₊ (suc n)) - → sym (lMapId (suc n) a) - ∙∙ funExt⁻ (cong fst (cong (lMap (suc n)) (p ∙ q))) a - ∙∙ lMapId (suc n) a - ≡ cong (fst (∙Π (lMap (suc (suc n)) p) (lMap (suc (suc n)) q))) (merid a) - main a = (cong (sym (lMapId (suc n) a) ∙∙_∙∙ (lMapId (suc n) a)) - (cong-∙ (λ x → lMap (suc n) x .fst a) p q) - ∙ sym (doubleComp-lem (sym (lMapId (suc n) a)) _ _)) - ∙∙ cong₂ _∙_ (sym (cong (cong (fst (lMap (suc (suc n)) p)) (merid a) ∙_) - (cong sym (lem p)) ∙ sym (rUnit _))) - (sym (cong (cong (fst (lMap (suc (suc n)) q)) (merid a) ∙_) - (cong sym (lem q)) ∙ sym (rUnit _))) - ∙∙ λ i → (rUnit (cong-∙ (fst (lMap (suc (suc n)) p)) - (merid a) (sym (merid (ptSn _))) (~ i)) i) - ∙ (rUnit (cong-∙ (fst (lMap (suc (suc n)) q)) - (merid a) (sym (merid (ptSn _)))(~ i)) i) - --- The iso is structure preserving -IsoSphereMapΩ-pres∙Π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f g : S₊∙ (suc n) →∙ A) - → SphereMap→Ω (suc n) (∙Π f g) - ≡ SphereMap→Ω (suc n) f ∙ SphereMap→Ω (suc n) g -IsoSphereMapΩ-pres∙Π n = - morphLemmas.isMorphInv _∙_ ∙Π (Ω→SphereMap (suc n)) - (isHom-lMap n) - (SphereMap→Ω (suc n)) - (rightInv (IsoΩSphereMap (suc n))) - (leftInv (IsoΩSphereMap (suc n))) - --- It is useful to define the ``Group Structure'' on (S₊∙ n →∙ A) --- before doing it on π'. These will be the equivalents of the --- usual groupoid laws on Ω A. -1Π : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} → (S₊∙ n →∙ A) -fst (1Π {A = A}) _ = pt A -snd (1Π {A = A}) = refl - -∙Π-rUnit : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f : S₊∙ (suc n) →∙ A) - → ∙Π f 1Π ≡ f -fst (∙Π-rUnit {A = A} {n = zero} f i) base = snd f (~ i) -fst (∙Π-rUnit {A = A} {n = zero} f i) (loop j) = help i j - where - help : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) - (((sym (snd f)) ∙∙ (cong (fst f) loop) ∙∙ snd f) - ∙ (refl ∙ refl)) - (cong (fst f) loop) - help = (cong ((sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) ∙_) - (sym (rUnit refl)) ∙ sym (rUnit _)) - ◁ λ i j → doubleCompPath-filler (sym (snd f)) - (cong (fst f) loop) (snd f) (~ i) j -snd (∙Π-rUnit {A = A} {n = zero} f i) j = snd f (~ i ∨ j) -fst (∙Π-rUnit {A = A} {n = suc n} f i) north = snd f (~ i) -fst (∙Π-rUnit {A = A} {n = suc n} f i) south = - (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i -fst (∙Π-rUnit {A = A} {n = suc n} f i) (merid a j) = help i j - where - help : PathP (λ i → snd f (~ i) - ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) - (((sym (snd f)) - ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) - ∙∙ snd f) - ∙ (refl ∙ refl)) - (cong (fst f) (merid a)) - help = (cong (((sym (snd f)) - ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) - ∙∙ snd f) ∙_) - (sym (rUnit refl)) - ∙ sym (rUnit _)) - ◁ λ i j → hcomp (λ k → - λ { (j = i0) → snd f (~ i ∧ k) - ; (j = i1) → compPath-filler' (sym (snd f)) - (cong (fst f) (merid (ptSn (suc n)))) k i - ; (i = i0) → doubleCompPath-filler (sym (snd f)) - (cong (fst f) - (merid a ∙ sym (merid (ptSn (suc n))))) - (snd f) k j - ; (i = i1) → fst f (merid a j)}) - (fst f (compPath-filler (merid a) - (sym (merid (ptSn _))) (~ i) j)) -snd (∙Π-rUnit {A = A} {n = suc n} f i) j = snd f (~ i ∨ j) - -∙Π-lUnit : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f : S₊∙ (suc n) →∙ A) - → ∙Π 1Π f ≡ f -fst (∙Π-lUnit {n = zero} f i) base = snd f (~ i) -fst (∙Π-lUnit {n = zero} f i) (loop j) = s i j - where - s : PathP (λ i → snd f (~ i) ≡ snd f (~ i)) - ((refl ∙ refl) ∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) - (cong (fst f) loop) - s = (cong (_∙ (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f)) - (sym (rUnit refl)) ∙ sym (lUnit _)) - ◁ λ i j → doubleCompPath-filler (sym (snd f)) - (cong (fst f) loop) (snd f) (~ i) j -snd (∙Π-lUnit {n = zero} f i) j = snd f (~ i ∨ j) -fst (∙Π-lUnit {n = suc n} f i) north = snd f (~ i) -fst (∙Π-lUnit {n = suc n} f i) south = - (sym (snd f) ∙ cong (fst f) (merid (ptSn _))) i -fst (∙Π-lUnit {n = suc n} f i) (merid a j) = help i j - where - help : PathP (λ i → snd f (~ i) - ≡ (sym (snd f) ∙ cong (fst f) (merid (ptSn (suc n)))) i) - ((refl ∙ refl) ∙ ((sym (snd f)) - ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) - ∙∙ snd f)) - (cong (fst f) (merid a)) - help = - (cong (_∙ ((sym (snd f)) - ∙∙ (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) - ∙∙ snd f)) - (sym (rUnit refl)) - ∙ sym (lUnit _)) - ◁ λ i j → hcomp (λ k → - λ { (j = i0) → snd f (~ i ∧ k) - ; (j = i1) → compPath-filler' (sym (snd f)) - (cong (fst f) (merid (ptSn (suc n)))) k i - ; (i = i0) → doubleCompPath-filler (sym (snd f)) - (cong (fst f) (merid a ∙ sym (merid (ptSn (suc n))))) - (snd f) k j - ; (i = i1) → fst f (merid a j)}) - (fst f (compPath-filler (merid a) (sym (merid (ptSn _))) (~ i) j)) -snd (∙Π-lUnit {n = suc n} f i) j = snd f (~ i ∨ j) - -∙Π-rCancel : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f : S₊∙ (suc n) →∙ A) - → ∙Π f (-Π f) ≡ 1Π -fst (∙Π-rCancel {A = A} {n = zero} f i) base = pt A -fst (∙Π-rCancel {A = A} {n = zero} f i) (loop j) = - rCancel (sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f) i j -snd (∙Π-rCancel {A = A} {n = zero} f i) = refl -fst (∙Π-rCancel {A = A} {n = suc n} f i) north = pt A -fst (∙Π-rCancel {A = A} {n = suc n} f i) south = pt A -fst (∙Π-rCancel {A = A} {n = suc n} f i) (merid a i₁) = lem i i₁ - where - pl = (sym (snd f) - ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) - ∙∙ snd f) - - lem : pl - ∙ ((sym (snd f) - ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) - ∙∙ snd f)) ≡ refl - lem = cong (pl ∙_) (cong (sym (snd f) ∙∙_∙∙ (snd f)) - (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) - ∙∙ cong₂ _∙_ refl - (cong (cong (fst f)) (rCancel (merid (ptSn _)))) - ∙∙ sym (rUnit _))) - ∙ rCancel pl -snd (∙Π-rCancel {A = A} {n = suc n} f i) = refl - -∙Π-lCancel : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f : S₊∙ (suc n) →∙ A) - → ∙Π (-Π f) f ≡ 1Π -fst (∙Π-lCancel {A = A} {n = zero} f i) base = pt A -fst (∙Π-lCancel {A = A} {n = zero} f i) (loop j) = - rCancel (sym (snd f) ∙∙ cong (fst f) (sym loop) ∙∙ snd f) i j -fst (∙Π-lCancel {A = A} {n = suc n} f i) north = pt A -fst (∙Π-lCancel {A = A} {n = suc n} f i) south = pt A -fst (∙Π-lCancel {A = A} {n = suc n} f i) (merid a j) = lem i j - where - pl = (sym (snd f) - ∙∙ cong (fst f) (merid a ∙ sym (merid (ptSn _))) - ∙∙ snd f) - - lem : (sym (snd f) - ∙∙ cong (fst (-Π f)) (merid a ∙ sym (merid (ptSn _))) - ∙∙ snd f) ∙ pl - ≡ refl - lem = cong (_∙ pl) (cong (sym (snd f) ∙∙_∙∙ (snd f)) - (cong-∙ (fst (-Π f)) (merid a) (sym (merid (ptSn _))) - ∙∙ cong₂ _∙_ refl (cong (cong (fst f)) (rCancel (merid (ptSn _)))) - ∙∙ sym (rUnit _))) - ∙ lCancel pl -snd (∙Π-lCancel {A = A} {n = zero} f i) = refl -snd (∙Π-lCancel {A = A} {n = suc n} f i) = refl - -∙Π-assoc : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f g h : S₊∙ (suc n) →∙ A) - → ∙Π f (∙Π g h) ≡ ∙Π (∙Π f g) h -∙Π-assoc {n = n} f g h = - sym (leftInv (IsoSphereMapΩ (suc n)) (∙Π f (∙Π g h))) - ∙∙ cong (Ω→SphereMap (suc n)) (IsoSphereMapΩ-pres∙Π n f (∙Π g h) - ∙∙ cong (SphereMap→Ω (suc n) f ∙_) (IsoSphereMapΩ-pres∙Π n g h) - ∙∙ ∙assoc (SphereMap→Ω (suc n) f) (SphereMap→Ω (suc n) g) (SphereMap→Ω (suc n) h) - ∙∙ cong (_∙ SphereMap→Ω (suc n) h) (sym (IsoSphereMapΩ-pres∙Π n f g)) - ∙∙ sym (IsoSphereMapΩ-pres∙Π n (∙Π f g) h)) - ∙∙ leftInv (IsoSphereMapΩ (suc n)) (∙Π (∙Π f g) h) - -∙Π-comm : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} - → (f g : S₊∙ (suc (suc n)) →∙ A) - → ∙Π f g ≡ ∙Π g f -∙Π-comm {A = A} {n = n} f g = - sym (leftInv (IsoSphereMapΩ (suc (suc n))) (∙Π f g)) - ∙∙ cong (Ω→SphereMap (suc (suc n))) (IsoSphereMapΩ-pres∙Π (suc n) f g - ∙∙ EH _ _ _ - ∙∙ sym (IsoSphereMapΩ-pres∙Π (suc n) g f)) - ∙∙ leftInv (IsoSphereMapΩ (suc (suc n))) (∙Π g f) - -{- π'' as a group -} -1π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' n A -1π' n {A = A} = ∣ 1Π ∣₂ - -·π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' (suc n) A → π' (suc n) A → π' (suc n) A -·π' n = sRec2 squash₂ λ p q → ∣ ∙Π p q ∣₂ - --π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' (suc n) A → π' (suc n) A --π' n = sMap -Π - -π'-rUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) - → (·π' n x (1π' (suc n))) ≡ x -π'-rUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-rUnit p i ∣₂ - -π'-lUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) - → (·π' n (1π' (suc n)) x) ≡ x -π'-lUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-lUnit p i ∣₂ - -π'-rCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) - → (·π' n x (-π' n x)) ≡ 1π' (suc n) -π'-rCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-rCancel p i ∣₂ - -π'-lCancel : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) - → (·π' n (-π' n x) x) ≡ 1π' (suc n) -π'-lCancel n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-lCancel p i ∣₂ - -π'-assoc : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y z : π' (suc n) A) - → ·π' n x (·π' n y z) ≡ ·π' n (·π' n x y) z -π'-assoc n = sElim3 (λ _ _ _ → isSetPathImplicit) λ p q r i → ∣ ∙Π-assoc p q r i ∣₂ - -π'-comm : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x y : π' (suc (suc n)) A) - → ·π' (suc n) x y ≡ ·π' (suc n) y x -π'-comm n = sElim2 (λ _ _ → isSetPathImplicit) λ p q i → ∣ ∙Π-comm p q i ∣₂ - --- We finally get the group definition -π'Gr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → Group ℓ -fst (π'Gr n A) = π' (suc n) A -1g (snd (π'Gr n A)) = 1π' (suc n) -GroupStr._·_ (snd (π'Gr n A)) = ·π' n -inv (snd (π'Gr n A)) = -π' n -is-set (isSemigroup (isMonoid (isGroup (snd (π'Gr n A))))) = squash₂ -assoc (isSemigroup (isMonoid (isGroup (snd (π'Gr n A))))) = π'-assoc n -identity (isMonoid (isGroup (snd (π'Gr n A)))) x = (π'-rUnit n x) , (π'-lUnit n x) -inverse (isGroup (snd (π'Gr n A))) x = (π'-rCancel n x) , (π'-lCancel n x) - --- and finally, the Iso -π'Gr≅πGr : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) → GroupIso (π'Gr n A) (πGr n A) -fst (π'Gr≅πGr n A) = setTruncIso (IsoSphereMapΩ (suc n)) -snd (π'Gr≅πGr n A) = - makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) - λ p q i → ∣ IsoSphereMapΩ-pres∙Π n p q i ∣₂) -{- -In the file SuspensionMapPathP we gave a filler -of the following square: +suspMapπ' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} + → π' (suc n) A + → π' (suc (suc n)) (Susp∙ (typ A)) +suspMapπ' n = sMap (suspMap n) - suspMapΩ -Ωⁿ A -------------------> Ωⁿ⁺¹ (Susp A) +suspMapπ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → π n A → π (suc n) (Susp∙ (typ A)) +suspMapπ n = sMap (suspMapΩ n) + +{- suspMapΩ + Ωⁿ A --------------------> Ω¹⁺ⁿ (Susp A) | | + | = | ≃ flipΩ + | Ωⁿ→ σ v +Ωⁿ A -------------------> Ωⁿ (Ω (Susp A)) | | - | ≃ eq₁ | ≃ eq₂ | | - v suspMap v - (Sⁿ →∙ A) -------------- > (Sⁿ⁺¹ →∙ Susp A) - -We would like this to preserve composition -(in order to deduce properties of suspMap from the corresponding -ones of suspMapΩ). For this, we need that both equivalences -on the sides are structure preserving --} - --- We define an alternative notion of composition (to ∙Π) on --- Sⁿ →∙ Ω (Susp∙ (typ A)). It will be easier to work with + | ≃ Ω→SphereMap | ≃ Ω→SphereMap + | | + v post∘∙ . σ v + (Sⁿ →∙ A) -------------- > (Sⁿ →∙ Ω (Susp A)) + | | + | = | ≃ botᵣ + | | + v suspMap v + (Sⁿ →∙ A) -------------- > (Sⁿ⁺¹→∙ Susp A) + -} + +botᵣ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (S₊∙ n →∙ Ω (Susp∙ (typ A))) + → S₊∙ (suc n) →∙ Susp∙ (typ A) +fst (botᵣ zero (f , p)) base = north +fst (botᵣ zero (f , p)) (loop i) = f false i +snd (botᵣ zero (f , p)) = refl +fst (botᵣ (suc n) (f , p)) north = north +fst (botᵣ (suc n) (f , p)) south = north +fst (botᵣ (suc n) (f , p)) (merid a i) = f a i +snd (botᵣ (suc n) (f , p)) = refl + +----- Top square filler ----- +top□ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Ω^→ (suc n) (σ∙ A) + ≡ (((Iso.fun (flipΩIso (suc n))) , flipΩrefl n) + ∘∙ suspMapΩ∙ (suc n)) +top□ {A = A} zero = + →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt λ p → sym (transportRefl _)) +top□ {A = A} (suc n) = + cong Ω→ (top□ {A = A} n) + ∙ →∙Homogeneous≡ + (isHomogeneousPath _ _) + (funExt λ x + → Ω→∘ (fun (flipΩIso (suc n)) , flipΩrefl n) (suspMapΩ∙ (suc n)) x) + +----- Middle square filler ----- +module _ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') + (homogB : isHomogeneous B) (f : A →∙ B) where + nat = isNaturalΩSphereMap A B homogB f + mutual + isNatural-Ω→SphereMap : ∀ n p → f ∘∙ Ω→SphereMap n p + ≡ Ω→SphereMap n (Ω^→ n f .fst p) + isNatural-Ω→SphereMap 0 p = + →∙Homogeneous≡ homogB (funExt λ {true → f .snd; false → refl}) + isNatural-Ω→SphereMap (n@(suc n')) p = + cong (f ∘∙_) (Ω→SphereMap-split n' p) + ∙ nat n' (Ω→SphereMapSplit₁ n' p) + ∙ cong (ΩSphereMap n') inner + ∙ sym (Ω→SphereMap-split n' (Ω^→ n f .fst p)) + where + inner : Ω→ (post∘∙ (S₊∙ n') f) .fst (Ω→ (Ω→SphereMap∙ n') .fst p) + ≡ Ω→ (Ω→SphereMap∙ n') .fst (Ω^→ (suc n') f .fst p) + inner = + sym (Ω→∘ (post∘∙ (S₊∙ n') f) (Ω→SphereMap∙ n') p) + ∙ cong (λ g∙ → Ω→ g∙ .fst p) (isNatural-Ω→SphereMap∙ n') + ∙ Ω→∘ (Ω→SphereMap∙ n') (Ω^→ n' f) p + + isNatural-Ω→SphereMap∙ : ∀ n + → post∘∙ (S₊∙ n) f ∘∙ (Ω→SphereMap∙ n) + ≡ (Ω→SphereMap∙ n {A = B} ∘∙ Ω^→ n f) + isNatural-Ω→SphereMap∙ n = + →∙Homogeneous≡ (isHomogeneous→∙ homogB) + (funExt (isNatural-Ω→SphereMap n)) + +mid□ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → (p : typ ((Ω^ (suc n)) A)) + → fst (post∘∙ (S₊∙ (suc n)) (σ∙ A)) (Ω→SphereMap (suc n) p) + ≡ Ω→SphereMap (suc n) (fst (Ω^→ (suc n) (σ∙ A)) p) +mid□ {A = A} n p = + funExt⁻ + (cong fst + (isNatural-Ω→SphereMap∙ + A (Ω (Susp∙ (typ A))) + (isHomogeneousPath _ _) + (σ∙ A) (suc n))) p + +----- Bottom square filler ----- +bot□ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f : (S₊∙ (suc n) →∙ A)) + → suspMap n f + ≡ botᵣ {A = A} (suc n) (post∘∙ (S₊∙ (suc n)) (σ∙ A) .fst f) +bot□ {A = A} n f = + ΣPathP ((funExt (λ { north → refl + ; south → refl + ; (merid a i) → refl})) + , refl) + +-- We prove that botᵣ is an equivalence +botᵣ⁻' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → S₊∙ (suc n) →∙ Susp∙ (typ A) → (S₊ n → typ (Ω (Susp∙ (typ A)))) +botᵣ⁻' zero f false = + sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f +botᵣ⁻' zero f true = refl +botᵣ⁻' (suc n) f x = + sym (snd f) + ∙∙ cong (fst f) (merid x ∙ sym (merid (ptSn (suc n)))) + ∙∙ snd f + +botᵣ⁻ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → S₊∙ (suc n) →∙ Susp∙ (typ A) + → (S₊∙ n →∙ Ω (Susp∙ (typ A))) +fst (botᵣ⁻ {A = A} n f) = botᵣ⁻' {A = A} n f +snd (botᵣ⁻ {A = A} zero f) = refl +snd (botᵣ⁻ {A = A} (suc n) f) = + cong (sym (snd f) ∙∙_∙∙ snd f) + (cong (cong (fst f)) (rCancel (merid (ptSn _)))) + ∙ ∙∙lCancel (snd f) + +botᵣIso : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (S₊∙ n →∙ Ω (Susp∙ (typ A))) + (S₊∙ (suc n) →∙ Susp∙ (typ A)) +botᵣIso {A = A} n = (iso (botᵣ {A = A} n) (botᵣ⁻ {A = A} n) (h n) (retr n)) + where + h : (n : ℕ) → section (botᵣ {A = A} n) (botᵣ⁻ {A = A} n) + h zero (f , p) = + ΣPathP (funExt (λ { base → sym p + ; (loop i) j → doubleCompPath-filler + (sym p) (cong f loop) p (~ j) i}) + , λ i j → p (~ i ∨ j)) + h (suc n) (f , p) = + ΣPathP (funExt (λ { north → sym p + ; south → sym p ∙ cong f (merid (ptSn _)) + ; (merid a i) j + → hcomp (λ k + → λ { (i = i0) → p (~ j ∧ k) + ; (i = i1) → compPath-filler' + (sym p) (cong f (merid (ptSn _))) k j + ; (j = i1) → f (merid a i)}) + (f (compPath-filler + (merid a) (sym (merid (ptSn _))) (~ j) i))}) + , λ i j → p (~ i ∨ j)) + + retr : (n : ℕ) → retract (botᵣ {A = A} n) (botᵣ⁻ {A = A} n) + retr zero (f , p) = + ΣPathP ((funExt (λ { false → sym (rUnit _) ; true → sym p})) + , λ i j → p (~ i ∨ j)) + retr (suc n) (f , p) = + →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt λ x → (λ i + → rUnit (cong-∙ (fst ((botᵣ {A = A}(suc n) (f , p)))) + (merid x) + (sym (merid (ptSn (suc n)))) i) (~ i)) + ∙∙ (λ i → f x ∙ sym (p i) ) + ∙∙ sym (rUnit (f x))) + +-- Right hand composite iso +IsoΩSphereMapᵣ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (typ ((Ω^ (suc n)) (Susp∙ (typ A)))) ((S₊∙ (suc n) →∙ Susp∙ (typ A))) +IsoΩSphereMapᵣ {A = A} n = + compIso (flipΩIso n) + (compIso (IsoΩSphereMap n) (botᵣIso {A = A} n)) + +-- The dependent path between the two suspension functions +suspMapPathP : ∀ {ℓ} (A : Pointed ℓ) (n : ℕ) + → (typ ((Ω^ (suc n)) A) → (typ ((Ω^ (suc (suc n))) (Susp∙ (typ A))))) + ≡ ((S₊∙ (suc n) →∙ A) → S₊∙ (suc (suc n)) →∙ (Susp∙ (typ A))) +suspMapPathP A n i = + isoToPath (IsoΩSphereMap {A = A} (suc n)) i + → isoToPath (IsoΩSphereMapᵣ {A = A} (suc n)) i + +Ωσ→suspMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → PathP (λ i → suspMapPathP A n i) + (suspMapΩ (suc n)) + (suspMap n) +Ωσ→suspMap {A = A} n = + toPathP (funExt (λ p → + (λ i → transportRefl + (Iso.fun (IsoΩSphereMapᵣ {A = A} (suc n)) + (suspMapΩ {A = A} (suc n) + (Iso.inv (IsoΩSphereMap {A = A} (suc n)) + (transportRefl p i)))) i) + ∙∙ cong (botᵣ {A = A} (suc n)) + (cong (Ω→SphereMap (suc n) {A = Ω (Susp∙ (typ A)) }) + ((sym (funExt⁻ (cong fst (top□ {A = A} n)) + (invEq (Ω→SphereMap (suc n) + , isEquiv-Ω→SphereMap (suc n)) p)))) + ∙ (sym (mid□ n (invEq (Ω→SphereMap (suc n) + , isEquiv-Ω→SphereMap (suc n)) p)) + ∙ cong (σ∙ (fst A , snd A) ∘∙_) + (secEq (Ω→SphereMap (suc n) + , isEquiv-Ω→SphereMap (suc n)) p))) + ∙∙ sym (bot□ n p))) + +-- Connectedness of suspFunΩ (Freudenthal) +suspMapΩ-connected : ∀ {ℓ} (n : HLevel) (m : ℕ) {A : Pointed ℓ} + (connA : isConnected (suc (suc n)) (typ A)) + → isConnectedFun ((suc n + suc n) ∸ m) (suspMapΩ {A = A} m) +suspMapΩ-connected n zero {A = A} connA = isConnectedσ n connA +suspMapΩ-connected n (suc m) {A = A} connA with ((n + suc n) ≟ m) +... | (lt p) = subst (λ x → isConnectedFun x (suspMapΩ {A = A} (suc m))) + (sym (n∸m≡0 _ m p)) + λ b → tt* , (λ {tt* → refl}) +... | (eq q) = subst (λ x → isConnectedFun x (suspMapΩ {A = A} (suc m))) + (sym (n∸n≡0 m) ∙ cong (_∸ m) (sym q)) + λ b → tt* , (λ {tt* → refl}) +... | (gt p) = isConnectedCong' (n + suc n ∸ m) (suspMapΩ {A = A} m) + (subst (λ x → isConnectedFun x (suspMapΩ {A = A} m)) + (sym (suc∸-fst (n + suc n) m p)) + (suspMapΩ-connected n m connA)) + (snd (suspMapΩ∙ m)) + +-- We prove that the right iso is structure preserving private invComp : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → S₊∙ n →∙ Ω (Susp∙ (typ A)) @@ -563,9 +484,9 @@ isHom-IsoΩSphereMapᵣ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (Iso.fun (IsoΩSphereMapᵣ (suc n)) q) isHom-IsoΩSphereMapᵣ {A = A} n p q = cong (botᵣ {A = A} (suc n)) - (cong (lMap (suc n) {A = Ω (Susp∙ (typ A))}) + (cong (Ω→SphereMap (suc n) {A = Ω (Susp∙ (typ A))}) (flipΩIsopres· n p q) - ∙ isHom-lMap n (fun (flipΩIso (suc n)) p) + ∙ isHom-Ω→SphereMap n (fun (flipΩIso (suc n)) p) (fun (flipΩIso (suc n)) q)) ∙ hom-botᵣ n _ _ @@ -614,7 +535,7 @@ private ∙Π→∙ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → PathP (λ i → pₗ A n i → pₗ A n i → pₗ A n i) _∙_ ∙Π ∙Π→∙ {A = A} n = - transportLem _∙_ ∙Π (IsoΩSphereMap {A = A} (suc n)) (isHom-lMap n) + transportLem _∙_ ∙Π (IsoΩSphereMap {A = A} (suc n)) (isHom-Ω→SphereMap n) ∙Π→∙ᵣ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → PathP (λ i → pᵣ A n i → pᵣ A n i → pᵣ A n i) _∙_ ∙Π @@ -630,10 +551,13 @@ isHom-suspMap {A = A} n = ≡ ∙Π→∙ᵣ n i (Ωσ→suspMap n i f) (Ωσ→suspMap n i g)) (suspMapΩ→hom n) -suspMapπ' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} - → π' (suc n) A - → π' (suc (suc n)) (Susp∙ (typ A)) -suspMapπ' n = sMap (suspMap n) +suspMapπHom : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → GroupHom (πGr n A) (πGr (suc n) (Susp∙ (typ A))) +fst (suspMapπHom {A = A} n) = suspMapπ (suc n) +snd (suspMapπHom {A = A} n) = + makeIsGroupHom + (sElim2 (λ _ _ → isSetPathImplicit) + λ p q → cong ∣_∣₂ (suspMapΩ→hom n p q)) suspMapπ'Hom : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → GroupHom (π'Gr n A) (π'Gr (suc n) (Susp∙ (typ A))) @@ -642,6 +566,13 @@ snd (suspMapπ'Hom {A = A} n) = makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) λ f g → cong ∣_∣₂ (isHom-suspMap n f g)) +πGr≅π'Grᵣ : ∀ {ℓ} (n : ℕ) (A : Pointed ℓ) + → GroupIso (πGr (suc n) (Susp∙ (typ A))) + (π'Gr (suc n) (Susp∙ (typ A))) +fst (πGr≅π'Grᵣ n A) = setTruncIso (IsoΩSphereMapᵣ (suc n)) +snd (πGr≅π'Grᵣ n A) = + makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) + λ f g → cong ∣_∣₂ (isHom-IsoΩSphereMapᵣ n f g)) private isConnectedPres : ∀ {ℓ} {A : Pointed ℓ} (con n : ℕ) @@ -655,7 +586,8 @@ isConnectedSuspMap : (n m : ℕ) isConnectedSuspMap n m = isConnectedPres _ _ (suspMapΩ-connected m (suc n) (sphereConnected (suc m))) -isSurjectiveSuspMap : (n : ℕ) → isSurjective (suspMapπ'Hom {A = S₊∙ (2 + n)} (2 + n)) +isSurjectiveSuspMap : (n : ℕ) + → isSurjective (suspMapπ'Hom {A = S₊∙ (2 + n)} (2 + n)) isSurjectiveSuspMap n = sElim (λ _ → isProp→isSet squash) λ f → diff --git a/Cubical/Homotopy/Group/SuspensionMapPathP.agda b/Cubical/Homotopy/Group/SuspensionMapPathP.agda deleted file mode 100644 index 86c95ba7a..000000000 --- a/Cubical/Homotopy/Group/SuspensionMapPathP.agda +++ /dev/null @@ -1,648 +0,0 @@ -{-# OPTIONS --safe #-} -{- -The goal of this file is to prove that the function -suspMapΩ : Ωⁿ A → Ωⁿ⁺¹ (Susp A), induced by -the Freudenthal function σ : A → ΩΣA, gets taken to -the canonical suspension map -suspMap : (Sⁿ →∙ A) → (Sⁿ⁺¹ →∙ Susp A) -given some suitable structure preserving equivalences -Ωⁿ A ≃ (Sⁿ →∙ A). - -The idea is to fill the following diagram - - suspMapΩ -Ωⁿ A -------------------> Ωⁿ⁺¹ (Susp A) - | | - | | - | ≃ eq₁ | ≃ eq₂ - | | - v suspMap v - (Sⁿ →∙ A) -------------- > (Sⁿ⁺¹ →∙ Susp A) - -(we choose eq₁ and eq₂ (intensionally) different for techinical reasons) - -Many results in this file are technical. See the end for the main results. --} -module Cubical.Homotopy.Group.SuspensionMapPathP where - -open import Cubical.Homotopy.Group.Base -open import Cubical.Homotopy.Loopspace -open import Cubical.Homotopy.Freudenthal -open import Cubical.Homotopy.Connected - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.Pointed.Homogeneous -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.GroupoidLaws renaming (assoc to ∙assoc) -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Function -open import Cubical.Foundations.Path -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Transport - -open import Cubical.Functions.Morphism - -open import Cubical.HITs.Sn -open import Cubical.HITs.Susp -open import Cubical.HITs.S1 - -open import Cubical.Data.Bool hiding (_≟_) -open import Cubical.Data.Sigma -open import Cubical.Data.Nat - -open import Cubical.Algebra.Group -open import Cubical.Algebra.Semigroup -open import Cubical.Algebra.Monoid - -open Iso -open IsGroup -open IsSemigroup -open IsMonoid -open GroupStr - -flipΩPath : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → ((Ω^ (suc n)) A) ≡ (Ω^ n) (Ω A) -flipΩPath {A = A} zero = refl -flipΩPath {A = A} (suc n) = cong Ω (flipΩPath {A = A} n) - -flipΩIso : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → Iso (fst ((Ω^ (suc n)) A)) (fst ((Ω^ n) (Ω A))) -flipΩIso {A = A} n = pathToIso (cong fst (flipΩPath n)) - -flipΩIso⁻pres· : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → (f g : fst ((Ω^ (suc n)) (Ω A))) - → inv (flipΩIso (suc n)) (f ∙ g) - ≡ (inv (flipΩIso (suc n)) f) - ∙ (inv (flipΩIso (suc n)) g) -flipΩIso⁻pres· {A = A} n f g i = - transp (λ j → flipΩPath {A = A} n (~ i ∧ ~ j) .snd - ≡ flipΩPath n (~ i ∧ ~ j) .snd) i - (transp (λ j → flipΩPath {A = A} n (~ i ∨ ~ j) .snd - ≡ flipΩPath n (~ i ∨ ~ j) .snd) (~ i) f - ∙ transp (λ j → flipΩPath {A = A} n (~ i ∨ ~ j) .snd - ≡ flipΩPath n (~ i ∨ ~ j) .snd) (~ i) g) - -flipΩIsopres· : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) - → (f g : fst (Ω ((Ω^ (suc n)) A))) - → fun (flipΩIso (suc n)) (f ∙ g) - ≡ (fun (flipΩIso (suc n)) f) - ∙ (fun (flipΩIso (suc n)) g) -flipΩIsopres· n = - morphLemmas.isMorphInv _∙_ _∙_ - (inv (flipΩIso (suc n))) - (flipΩIso⁻pres· n) - (fun (flipΩIso (suc n))) - (Iso.leftInv (flipΩIso (suc n))) - (Iso.rightInv (flipΩIso (suc n))) - -flipΩrefl : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → fun (flipΩIso {A = A} (suc n)) refl ≡ refl -flipΩrefl {A = A} n j = - transp (λ i₁ → fst (Ω (flipΩPath {A = A} n ((i₁ ∨ j))))) - j (snd (Ω (flipΩPath n j))) - - --- Solves some termination issues -private - +nInd : ∀ {ℓ} {P : ℕ → Type ℓ} - → P 0 - → P 1 - → ((n : ℕ) → P (suc n) → P (suc (suc n))) - → (n : ℕ) → P n - +nInd 0c 1c indc zero = 0c - +nInd 0c 1c indc (suc zero) = 1c - +nInd {P = P} 0c 1c indc (suc (suc n)) = - indc n (+nInd {P = P} 0c 1c indc (suc n)) - --- move to loop -Ω^→ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) - → (A →∙ B) → ((Ω^ n) A →∙ (Ω^ n) B) -Ω^→ zero f = f -Ω^→ (suc n) f = Ω→ (Ω^→ n f) - --- move to pointed -post∘∙ : ∀ {ℓX ℓ ℓ'} (X : Pointed ℓX) {A : Pointed ℓ} {B : Pointed ℓ'} - → (A →∙ B) → ((X →∙ A ∙) →∙ (X →∙ B ∙)) -post∘∙ X f .fst g = f ∘∙ g -post∘∙ X f .snd = - ΣPathP - ( (funExt λ _ → f .snd) - , (sym (lUnit (f .snd)) ◁ λ i j → f .snd (i ∨ j)) - ) - -suspMap : ∀ {ℓ} {A : Pointed ℓ}(n : ℕ) - → S₊∙ (suc n) →∙ A - → S₊∙ (suc (suc n)) →∙ Susp∙ (typ A) -fst (suspMap n f) north = north -fst (suspMap n f) south = north -fst (suspMap {A = A} n f) (merid a i) = - (merid (fst f a) ∙ sym (merid (pt A))) i -snd (suspMap n f) = refl - -lMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} - → typ ((Ω^ n) A) → (S₊∙ n →∙ A) -lMapId : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (a : _) - → lMap n {A = A} (pt ((Ω^ n) A)) .fst a ≡ pt A -lMapId2 : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} - → lMap n {A = A} (pt ((Ω^ n) A)) ≡ ((λ _ → pt A) , refl) -fst (lMap zero {A = A} a) false = a -fst (lMap zero {A = A} a) true = pt A -snd (lMap zero {A = A} a) = refl -fst (lMap (suc zero) {A = A} p) base = pt A -fst (lMap (suc zero) {A = A} p) (loop i) = p i -snd (lMap (suc zero) {A = A} p) = refl -fst (lMap (suc (suc n)) {A = A} p) north = pt A -fst (lMap (suc (suc n)) {A = A} p) south = pt A -fst (lMap (suc (suc n)) {A = A} p) (merid a i) = - (sym (lMapId (suc n) a) - ∙∙ (λ i → lMap (suc n) (p i) .fst a) - ∙∙ lMapId (suc n) a) i -snd (lMap (suc (suc n)) {A = A} p) = refl -lMapId zero false = refl -lMapId zero true = refl -lMapId (suc zero) base = refl -lMapId (suc zero) (loop i) = refl -lMapId (suc (suc n)) north = refl -lMapId (suc (suc n)) south = refl -lMapId (suc (suc n)) {A = A} (merid a i) j = - ∙∙lCancel (lMapId (suc n) {A = A} a) j i -lMapId2 zero {A = A} = - ΣPathP ((funExt (λ { false → refl - ; true → refl})) - , refl) -lMapId2 (suc zero) {A = A} = - ΣPathP ((funExt (λ { base → refl - ; (loop i) → refl})) - , refl) -lMapId2 (suc (suc n)) {A = A} = - ΣPathP ((funExt (λ { north → refl - ; south → refl - ; (merid a i) j - → ∙∙lCancel (lMapId (suc n) {A = A} a) j i})) - , refl) - -lMap∙ : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} - → ((Ω^ n) A) →∙ (S₊∙ n →∙ A ∙) -lMap∙ n .fst = lMap n -lMap∙ n .snd = lMapId2 n - --- We define the following maps which will be used to --- show that lMap is an equivalence -lMapSplit₁ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → typ ((Ω^ (suc n)) A) - → typ (Ω (S₊∙ n →∙ A ∙)) -lMapSplit₁ n = Ω→ (lMap∙ n) .fst - -ΩSphereMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → typ (Ω (S₊∙ n →∙ A ∙)) - → (S₊∙ (suc n) →∙ A) -fst (ΩSphereMap {A = A} zero p) base = p i0 .fst false -fst (ΩSphereMap {A = A} zero p) (loop i) = p i .fst false -snd (ΩSphereMap {A = A} zero p) = refl -fst (ΩSphereMap {A = A} (suc n) p) north = pt A -fst (ΩSphereMap {A = A} (suc n) p) south = pt A -fst (ΩSphereMap {A = A} (suc n) p) (merid a i) = p i .fst a -snd (ΩSphereMap {A = A} (suc n) p) = refl - -SphereMapΩ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → (S₊∙ (suc n) →∙ A) - → typ (Ω (S₊∙ n →∙ A ∙)) -SphereMapΩ {A = A} zero (f , p) = - ΣPathP ((funExt λ { false → sym p ∙∙ cong f loop ∙∙ p - ; true → refl}) - , refl) -SphereMapΩ {A = A} (suc n) (f , p) = - ΣPathP (funExt (λ x → sym p ∙∙ cong f (merid x ∙ sym (merid (ptSn _))) ∙∙ p) - , flipSquare (cong (sym p ∙∙_∙∙ p) - (cong (cong f) (rCancel (merid (ptSn _)))) - ∙ ∙∙lCancel p)) - -SphereMapΩIso : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → Iso (S₊∙ (suc n) →∙ A) - (typ (Ω (S₊∙ n →∙ A ∙))) -fun (SphereMapΩIso {A = A} n) = SphereMapΩ n -inv (SphereMapΩIso {A = A} n) = ΩSphereMap n -fst (rightInv (SphereMapΩIso {A = A} zero) p k i) false = - rUnit (funExt⁻ (cong fst p) false) (~ k) i -fst (rightInv (SphereMapΩIso {A = A} zero) p k i) true = snd (p i) (~ k) -snd (rightInv (SphereMapΩIso {A = A} zero) p k i) l = snd (p i) (~ k ∨ l) -fst (rightInv (SphereMapΩIso {A = A} (suc n)) p k i) x = help k i - where - help : cong (fst (ΩSphereMap {A = A} (suc n) p)) (merid x ∙ sym (merid (ptSn _))) ∙ refl - ≡ funExt⁻ (cong fst p) x - help = (cong (refl ∙∙_∙∙ refl) (cong-∙ (fst (ΩSphereMap {A = A} (suc n) p)) - (merid x) (sym (merid (ptSn _)))) - ∙ cong (refl ∙∙_∙∙ refl) (cong (funExt⁻ (cong fst p) x ∙_) - (cong sym (flipSquare (cong snd p))) - ∙ sym (rUnit _))) - ∙ sym (rUnit _) -snd (rightInv (SphereMapΩIso {A = A} (suc n)) p k i) j = - hcomp (λ r → - λ { (i = i0) → pt A - ; (i = i1) → pt A - ; (j = i0) → ((wrap-refl (cong-∙ Ωp (merid (ptSn (suc n))) - (sym (merid (ptSn _)))) - ∙ wrap-refl (cong (p-refl ∙_) - (cong sym (flipSquare (cong snd p))) - ∙ sym (rUnit _))) ∙ sym (rUnit _)) k i - ; (j = i1) → rUnit (λ _ → pt A) (~ r ∧ ~ k) i - ; (k = i0) → compPath-filler (wrap-refl (cong (cong Ωp) - (rCancel (merid (ptSn _))))) - (sym (rUnit refl)) r j i - ; (k = i1) → snd (p i) j}) - (hcomp (λ r → - λ { (i = i0) → rUnit (λ _ → pt A) (~ k ∨ r) i - ; (i = i1) → rUnit (λ _ → pt A) (~ k ∨ r) i - ; (j = i0) → ((cong (λ x → rUnit x r) - (cong-∙ Ωp (merid (ptSn (suc n))) - (sym (merid (ptSn _)))) - ∙ cong (λ x → rUnit x r) - (cong (p-refl ∙_) - (cong sym (flipSquare (cong snd p))) - ∙ sym (rUnit _))) - ∙ λ i → rUnit p-refl (~ i ∧ r)) k i - ; (j = i1) → rUnit (λ _ → pt A) (~ k ∧ r) i - ; (k = i0) → rUnit (cong Ωp (rCancel (merid (ptSn _)) j)) r i - ; (k = i1) → snd (p i) j}) - (hcomp (λ r → - λ { (i = i0) → pt A - ; (i = i1) → pt A - ; (j = i0) → ((rUnit (compPath-filler' - (cong-∙ Ωp (merid (ptSn (suc n))) - (sym (merid (ptSn _)))) - ((cong (p-refl ∙_) - (cong sym (flipSquare (cong snd p))) - ∙ sym (rUnit _))) r) r) k i) - ; (j = i1) → pt A - ; (k = i0) → help Ωp (merid (ptSn (suc n))) r j i - ; (k = i1) → snd (p i) j}) - (lem p-refl (sym (flipSquare (cong snd p))) k j i))) - where - Ωp = (fst (ΩSphereMap {A = A} (suc n) p)) - wrap-refl : ∀ {ℓ} {A : Type ℓ} {x : A} {r s : x ≡ x} (p : r ≡ s) → _ - wrap-refl p = cong (refl ∙∙_∙∙ refl) p - p-refl = funExt⁻ (cong fst p) (ptSn (suc n)) - - cong-∙∙-filler : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y z w : A} - (f : A → B) (p : w ≡ x) (q : x ≡ y) (r : y ≡ z) → I → I → I → B - cong-∙∙-filler f p q r k j i = - hfill (λ k → λ { (j = i0) → f (doubleCompPath-filler p q r k i) - ; (i = i0) → f (p (~ k)) - ; (i = i1) → f (r k) }) - (inS (f (q i))) k - - lem : ∀ {ℓ} {A : Pointed ℓ} (p : Path (typ A) (pt A) (pt A)) (q : refl ≡ p) - → PathP (λ i → (cong (p ∙_) (cong sym (sym q)) ∙ sym (rUnit p)) i ≡ refl) - (rCancel p) (sym q) - lem {A = A} p = - J (λ p q → - PathP (λ i → (cong (p ∙_) (cong sym (sym q)) ∙ sym (rUnit p)) i ≡ refl) - (rCancel p) (sym q)) - (flipSquare (sym (lUnit (sym (rUnit refl))) - ◁ λ k i → rCancel (λ _ → pt A) (k ∨ i))) - - help : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} {x y : A} (f : A → B) (p : x ≡ y) - → PathP (λ k → cong-∙ f p (sym p) (~ k) ≡ refl) - (rCancel (cong f p)) (cong (cong f) (rCancel p)) - help f p i j k = - hcomp (λ r → λ { (i = i0) → rCancel-filler (cong f p) r j k - ; (i = i1) → f (rCancel-filler p r j k) - ; (j = i0) → cong-∙∙-filler f refl p (sym p) r (~ i) k - ; (j = i1) → f (p i0) - ; (k = i0) → f (p i0) - ; (k = i1) → f (p (~ r ∧ ~ j))}) - (f (p (~ j ∧ k))) - -leftInv (SphereMapΩIso {A = A} zero) (f , p) = - ΣPathP ((funExt (λ { base → sym p - ; (loop i) j - → doubleCompPath-filler (sym p) (cong f loop) p (~ j) i})) - , λ i j → p (j ∨ ~ i)) -leftInv (SphereMapΩIso {A = A} (suc n)) (f , p) = - ΣPathP (funExt (λ { north → sym p - ; south → sym p ∙ cong f (merid (ptSn (suc n))) - ; (merid a i) j → help a j i}) - , λ i j → p (j ∨ ~ i)) - where - help : (a : S₊ (suc n)) → PathP (λ i → p (~ i) ≡ (sym p ∙ cong f (merid (ptSn (suc n)))) i) - (sym p ∙∙ cong f (merid a ∙ sym (merid (ptSn _))) ∙∙ p) - (cong f (merid a)) - help a i j = - hcomp (λ k → - λ { (i = i0) → doubleCompPath-filler (sym p) - (cong f (merid a ∙ sym (merid (ptSn _)))) p k j - ; (i = i1) → f (merid a j) - ; (j = i0) → p (~ i ∧ k) - ; (j = i1) → compPath-filler' (sym p) - (cong f (merid (ptSn (suc n)))) k i}) - (f (compPath-filler (merid a) - (sym (merid (ptSn _))) (~ i) j)) - -{- -In order to show that lMap is an equivalence, we show that it factors - - SphereMapΩ lMapSplit₁ -Ωⁿ⁺¹A ----------------> Ω (Sⁿ →∙ A) -----------> (Sⁿ⁺¹ →∙ A) --} - - -lMap-split : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p : typ ((Ω^ (suc n)) A)) - → lMap (suc n) p ≡ ΩSphereMap n (lMapSplit₁ n p) -lMap-split {A = A} zero p = - ΣPathP ((funExt (λ { base → refl - ; (loop i) j → lem (~ j) i})) - , refl) - where - lem : funExt⁻ (cong fst (lMapSplit₁ zero p)) false ≡ p - lem = (λ i → funExt⁻ (cong-∙∙ fst (sym (lMapId2 zero)) - (cong (lMap zero) p) - (lMapId2 zero) i) false) - ∙ sym (rUnit _) -lMap-split {A = A} (suc n) p = - ΣPathP ((funExt (λ { north → refl - ; south → refl - ; (merid a i) j → lem₂ a j i})) - , refl) - where - lem : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (a : S₊ (suc n)) - → lMapId (suc n) {A = A} a - ≡ (λ i → fst (lMapId2 (suc n) {A = A} i) a) - lem zero base = refl - lem zero (loop i) = refl - lem (suc n) north = refl - lem (suc n) south = refl - lem (suc n) (merid a i) = refl - - lem₂ : (a : S₊ (suc n)) - → ((λ i₁ → lMapId (suc n) {A = A} a (~ i₁)) - ∙∙ (λ i₁ → lMap (suc n) (p i₁) .fst a) - ∙∙ lMapId (suc n) a) - ≡ (λ i → lMapSplit₁ (suc n) p i .fst a) - lem₂ a = cong (λ x → sym x - ∙∙ funExt⁻ (cong fst (λ i → lMap (suc n) (p i))) a - ∙∙ x) - (lem n a) - ∙∙ sym (cong-∙∙ (λ x → x a) - (cong fst (λ i → lMapId2 (suc n) (~ i))) - (cong fst (λ i → lMap (suc n) (p i))) - (cong fst (lMapId2 (suc n)))) - ∙∙ (λ i → funExt⁻ (cong-∙∙ fst (sym (lMapId2 (suc n))) - (cong (lMap (suc n)) p) - (lMapId2 (suc n)) (~ i)) a) - -isEquiv-lMap : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} - → isEquiv (lMap (suc n) {A = A}) -isEquiv-lMap zero {A = A} = - isoToIsEquiv (iso _ invFun sec λ p → sym (rUnit p)) - where - invFun : S₊∙ 1 →∙ A → typ (Ω A) - invFun (f , p) = sym p ∙∙ cong f loop ∙∙ p - - sec : section (lMap 1) invFun - sec (f , p) = - ΣPathP ((funExt (λ { base → sym p - ; (loop i) j → doubleCompPath-filler - (sym p) (cong f loop) p (~ j) i})) - , λ i j → p (~ i ∨ j)) - -isEquiv-lMap (suc n) = - subst isEquiv (sym (funExt (lMap-split (suc n)))) - (snd (compEquiv - ((lMapSplit₁ (suc n)) , - (isEquivΩ→ (lMap (suc n) , lMapId2 (suc n)) - (isEquiv-lMap n))) - (invEquiv (isoToEquiv (SphereMapΩIso (suc n)))))) - - --- The homogeneity assumption is not necessary but simplifying -module _ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') (homogB : isHomogeneous B) (f : A →∙ B) - where - - isNaturalΩSphereMap : (n : ℕ) - → ∀ g → f ∘∙ ΩSphereMap n g ≡ ΩSphereMap n (Ω→ (post∘∙ (S₊∙ n) f) .fst g) - isNaturalΩSphereMap 0 g = - →∙Homogeneous≡ homogB (funExt lem) - where - lem : ∀ x → f .fst (ΩSphereMap 0 g .fst x) ≡ ΩSphereMap 0 (Ω→ (post∘∙ (S₊∙ 0) f) .fst g) .fst x - lem base = f .snd - lem (loop i) j = - hfill - (λ j → λ - { (i = i0) → post∘∙ _ f .snd j - ; (i = i1) → post∘∙ _ f .snd j - }) - (inS (f ∘∙ g i)) - j .fst false - isNaturalΩSphereMap (n@(suc _)) g = - →∙Homogeneous≡ homogB (funExt lem) - where - lem : ∀ x → f .fst (ΩSphereMap n g .fst x) ≡ ΩSphereMap n (Ω→ (post∘∙ (S₊∙ n) f) .fst g) .fst x - lem north = f .snd - lem south = f .snd - lem (merid a i) j = - hfill - (λ j → λ - { (i = i0) → post∘∙ _ f .snd j - ; (i = i1) → post∘∙ _ f .snd j - }) - (inS (f ∘∙ g i)) - j .fst a - - mutual - isNatural-lMap : ∀ n p → f ∘∙ lMap n p ≡ lMap n (Ω^→ n f .fst p) - isNatural-lMap 0 p = - →∙Homogeneous≡ homogB (funExt λ {true → f .snd; false → refl}) - isNatural-lMap (n@(suc n')) p = - cong (f ∘∙_) (lMap-split n' p) - ∙ isNaturalΩSphereMap n' (lMapSplit₁ n' p) - ∙ cong (ΩSphereMap n') inner - ∙ sym (lMap-split n' (Ω^→ n f .fst p)) - where - inner : Ω→ (post∘∙ (S₊∙ n') f) .fst (Ω→ (lMap∙ n') .fst p) ≡ Ω→ (lMap∙ n') .fst (Ω^→ (suc n') f .fst p) - inner = - sym (Ω→∘ (post∘∙ (S₊∙ n') f) (lMap∙ n') p) - ∙ cong (λ g∙ → Ω→ g∙ .fst p) (isNatural-lMap∙ n') - ∙ Ω→∘ (lMap∙ n') (Ω^→ n' f) p - - isNatural-lMap∙ : ∀ n → post∘∙ (S₊∙ n) f ∘∙ (lMap∙ n) ≡ (lMap∙ n {A = B} ∘∙ Ω^→ n f) - isNatural-lMap∙ n = →∙Homogeneous≡ (isHomogeneous→∙ homogB) (funExt (isNatural-lMap n)) - - -σ : ∀ {ℓ} (A : Pointed ℓ) (a : typ A) → Path (Susp (typ A)) north north -σ A a = merid a ∙ sym (merid (pt A)) - -σ∙ : ∀ {ℓ} (A : Pointed ℓ) → A →∙ Ω (Susp∙ (typ A)) -fst (σ∙ A) = σ A -snd (σ∙ A) = rCancel _ - - -top□ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → Ω^→ (suc n) (σ∙ A) - ≡ (((Iso.fun (flipΩIso (suc n))) , flipΩrefl n) - ∘∙ suspMapΩ∙ (suc n)) -top□ {A = A} zero = - →∙Homogeneous≡ (isHomogeneousPath _ _) - (funExt λ p → sym (transportRefl _)) -top□ {A = A} (suc n) = - cong Ω→ (top□ {A = A} n) - ∙ →∙Homogeneous≡ - (isHomogeneousPath _ _) - (funExt λ x - → Ω→∘ (fun (flipΩIso (suc n)) , flipΩrefl n) (suspMapΩ∙ (suc n)) x) - -{- suspMapΩ - Ωⁿ A --------------------> Ω¹⁺ⁿ (Susp A) - | | - | = | ≃ flipΩ - | Ωⁿ→ σ v -Ωⁿ A -------------------> Ωⁿ (Ω (Susp A)) - | | - | | - | ≃ lMap | ≃ lMap - | | - v post∘∙ . σ v - (Sⁿ →∙ A) -------------- > (Sⁿ →∙ Ω (Susp A)) - | | - | = | ≃ botᵣ - | | - v suspMap v - (Sⁿ →∙ A) -------------- > (Sⁿ⁺¹→∙ Susp A) - -} - - -mid□ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → (p : typ ((Ω^ (suc n)) A)) - → fst (post∘∙ (S₊∙ (suc n)) (σ∙ A)) (lMap (suc n) p) - ≡ lMap (suc n) (fst (Ω^→ (suc n) (σ∙ A)) p) -mid□ {A = A} n p = - funExt⁻ - (cong fst - (isNatural-lMap∙ - A (Ω (Susp∙ (typ A))) - (isHomogeneousPath _ _) - (σ∙ A) (suc n))) p - -botᵣ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → (S₊∙ n →∙ Ω (Susp∙ (typ A))) - → S₊∙ (suc n) →∙ Susp∙ (typ A) -fst (botᵣ zero (f , p)) base = north -fst (botᵣ zero (f , p)) (loop i) = f false i -snd (botᵣ zero (f , p)) = refl -fst (botᵣ (suc n) (f , p)) north = north -fst (botᵣ (suc n) (f , p)) south = north -fst (botᵣ (suc n) (f , p)) (merid a i) = f a i -snd (botᵣ (suc n) (f , p)) = refl - - -botᵣ⁻' : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → S₊∙ (suc n) →∙ Susp∙ (typ A) → (S₊ n → typ (Ω (Susp∙ (typ A)))) -botᵣ⁻' zero f false = - sym (snd f) ∙∙ cong (fst f) loop ∙∙ snd f -botᵣ⁻' zero f true = refl -botᵣ⁻' (suc n) f x = - sym (snd f) - ∙∙ cong (fst f) (merid x ∙ sym (merid (ptSn (suc n)))) - ∙∙ snd f - -botᵣ⁻ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → S₊∙ (suc n) →∙ Susp∙ (typ A) - → (S₊∙ n →∙ Ω (Susp∙ (typ A))) -fst (botᵣ⁻ {A = A} n f) = botᵣ⁻' {A = A} n f -snd (botᵣ⁻ {A = A} zero f) = refl -snd (botᵣ⁻ {A = A} (suc n) f) = - cong (sym (snd f) ∙∙_∙∙ snd f) - (cong (cong (fst f)) (rCancel (merid (ptSn _)))) - ∙ ∙∙lCancel (snd f) - --- botᵣ is an Iso -botᵣIso : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → Iso (S₊∙ n →∙ Ω (Susp∙ (typ A))) - (S₊∙ (suc n) →∙ Susp∙ (typ A)) -botᵣIso {A = A} n = (iso (botᵣ {A = A} n) (botᵣ⁻ {A = A} n) (h n) (retr n)) - where - h : (n : ℕ) → section (botᵣ {A = A} n) (botᵣ⁻ {A = A} n) - h zero (f , p) = - ΣPathP (funExt (λ { base → sym p - ; (loop i) j → doubleCompPath-filler - (sym p) (cong f loop) p (~ j) i}) - , λ i j → p (~ i ∨ j)) - h (suc n) (f , p) = - ΣPathP (funExt (λ { north → sym p - ; south → sym p ∙ cong f (merid (ptSn _)) - ; (merid a i) j - → hcomp (λ k - → λ { (i = i0) → p (~ j ∧ k) - ; (i = i1) → compPath-filler' - (sym p) (cong f (merid (ptSn _))) k j - ; (j = i1) → f (merid a i)}) - (f (compPath-filler - (merid a) (sym (merid (ptSn _))) (~ j) i))}) - , λ i j → p (~ i ∨ j)) - - retr : (n : ℕ) → retract (botᵣ {A = A} n) (botᵣ⁻ {A = A} n) - retr zero (f , p) = - ΣPathP ((funExt (λ { false → sym (rUnit _) ; true → sym p})) - , λ i j → p (~ i ∨ j)) - retr (suc n) (f , p) = - →∙Homogeneous≡ (isHomogeneousPath _ _) - (funExt λ x → (λ i - → rUnit (cong-∙ (fst ((botᵣ {A = A}(suc n) (f , p)))) - (merid x) - (sym (merid (ptSn (suc n)))) i) (~ i)) - ∙∙ (λ i → f x ∙ sym (p i) ) - ∙∙ sym (rUnit (f x))) - - -bot□ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (f : (S₊∙ (suc n) →∙ A)) - → suspMap n f ≡ botᵣ {A = A} (suc n) (post∘∙ (S₊∙ (suc n)) (σ∙ A) .fst f) -bot□ {A = A} n f = - ΣPathP ((funExt (λ { north → refl - ; south → refl - ; (merid a i) → refl})) - , refl) - - -IsoΩSphereMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → Iso (typ ((Ω^ n) A)) ((S₊∙ n →∙ A)) -fun (IsoΩSphereMap zero) = lMap zero -inv (IsoΩSphereMap zero) f = fst f false -rightInv (IsoΩSphereMap zero) f = - ΣPathP ((funExt (λ { false → refl - ; true → sym (snd f)})) - , λ i j → snd f (~ i ∨ j)) -leftInv (IsoΩSphereMap zero) p = refl -IsoΩSphereMap (suc n) = equivToIso (_ , isEquiv-lMap n) - -IsoΩSphereMapᵣ : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → Iso (typ ((Ω^ (suc n)) (Susp∙ (typ A)))) ((S₊∙ (suc n) →∙ Susp∙ (typ A))) -IsoΩSphereMapᵣ {A = A} n = - compIso (flipΩIso n) - (compIso (IsoΩSphereMap n) (botᵣIso {A = A} n)) - -suspMapPathP : ∀ {ℓ} (A : Pointed ℓ) (n : ℕ) - → (typ ((Ω^ (suc n)) A) → (typ ((Ω^ (suc (suc n))) (Susp∙ (typ A))))) - ≡ ((S₊∙ (suc n) →∙ A) → S₊∙ (suc (suc n)) →∙ (Susp∙ (typ A))) -suspMapPathP A n i = - isoToPath (IsoΩSphereMap {A = A} (suc n)) i - → isoToPath (IsoΩSphereMapᵣ {A = A} (suc n)) i - -Ωσ→suspMap : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) - → PathP (λ i → suspMapPathP A n i) - (suspMapΩ (suc n)) - (suspMap n) -Ωσ→suspMap {A = A} n = - toPathP (funExt (λ p → - (λ i → transportRefl - (Iso.fun (IsoΩSphereMapᵣ {A = A} (suc n)) - (suspMapΩ {A = A} (suc n) - (Iso.inv (IsoΩSphereMap {A = A} (suc n)) - (transportRefl p i)))) i) - ∙∙ cong (botᵣ {A = A} (suc n)) - (cong (lMap (suc n) {A = Ω (Susp∙ (typ A)) }) - ((sym (funExt⁻ (cong fst (top□ {A = A} n)) - (invEq (lMap (suc n) , isEquiv-lMap n) p)))) - ∙ (sym (mid□ n (invEq (lMap (suc n) , isEquiv-lMap n) p)) - ∙ cong (σ∙ (fst A , snd A) ∘∙_) - (secEq (lMap (suc n) , isEquiv-lMap n) p))) - ∙∙ sym (bot□ n p))) diff --git a/Cubical/Homotopy/HopfInvariant/Base.agda b/Cubical/Homotopy/HopfInvariant/Base.agda index b5a09a59d..1efa57d9a 100644 --- a/Cubical/Homotopy/HopfInvariant/Base.agda +++ b/Cubical/Homotopy/HopfInvariant/Base.agda @@ -3,7 +3,7 @@ module Cubical.Homotopy.HopfInvariant.Base where open import Cubical.Homotopy.Connected open import Cubical.Homotopy.Group.Base -open import Cubical.Homotopy.Group.Properties +open import Cubical.Homotopy.Group.SuspensionMap open import Cubical.Relation.Nullary diff --git a/Cubical/Homotopy/HopfInvariant/Homomorphism.agda b/Cubical/Homotopy/HopfInvariant/Homomorphism.agda index 99a8c6fec..528626755 100644 --- a/Cubical/Homotopy/HopfInvariant/Homomorphism.agda +++ b/Cubical/Homotopy/HopfInvariant/Homomorphism.agda @@ -3,7 +3,7 @@ module Cubical.Homotopy.HopfInvariant.Homomorphism where open import Cubical.Homotopy.HopfInvariant.Base open import Cubical.Homotopy.Group.Base -open import Cubical.Homotopy.Group.Properties +open import Cubical.Homotopy.Group.SuspensionMap open import Cubical.ZCohomology.Base open import Cubical.ZCohomology.GroupStructure diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda index 93681bd77..153c9634d 100644 --- a/Cubical/Homotopy/Loopspace.agda +++ b/Cubical/Homotopy/Loopspace.agda @@ -13,11 +13,15 @@ open import Cubical.Foundations.GroupoidLaws open import Cubical.HITs.SetTruncation open import Cubical.HITs.Truncation hiding (elim2) renaming (rec to trRec) open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Transport open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.Function open import Cubical.Foundations.Path open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.Equiv +open import Cubical.Functions.Morphism +open import Cubical.Data.Sigma +-- open import Cubical.HITs.Sn open Iso {- loop space of a pointed type -} @@ -35,6 +39,11 @@ open Iso fst (Ω→ {A = A} {B = B} (f , p)) q = sym p ∙∙ cong f q ∙∙ p snd (Ω→ {A = A} {B = B} (f , p)) = ∙∙lCancel p +Ω^→ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) + → (A →∙ B) → ((Ω^ n) A →∙ (Ω^ n) B) +Ω^→ zero f = f +Ω^→ (suc n) f = Ω→ (Ω^→ n f) + {- loop space map functoriality (missing pointedness proof) -} Ω→∘ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} (g : B →∙ C) (f : A →∙ B) @@ -42,9 +51,7 @@ snd (Ω→ {A = A} {B = B} (f , p)) = ∙∙lCancel p Ω→∘ g f p k i = hcomp (λ j → λ - { (k = i0) → doubleCompPath-filler (sym (snd (g ∘∙ f))) (cong (fst (g ∘∙ f)) p) (snd (g ∘∙ f)) j i - ; (k = i1) → doubleCompPath-filler (sym (snd g)) (cong (fst g) (fst (Ω→ f) p)) (snd g) j i - ; (i = i0) → compPath-filler' (cong (g .fst) (f .snd)) (g .snd) (~ k) j + { (i = i0) → compPath-filler' (cong (g .fst) (f .snd)) (g .snd) (~ k) j ; (i = i1) → compPath-filler' (cong (g .fst) (f .snd)) (g .snd) (~ k) j }) (g .fst (doubleCompPath-filler (sym (f .snd)) (cong (f .fst) p) (f .snd) k i)) @@ -59,7 +66,8 @@ isEquivΩ→ {B = (B , b)} = λ eqf → subst isEquiv (funExt (rUnit ∘ cong f)) (isoToIsEquiv (congIso (equivToIso (f , eqf)))) -ΩfunExtIso : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') → Iso (typ (Ω (A →∙ B ∙))) (A →∙ Ω B) +ΩfunExtIso : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') + → Iso (typ (Ω (A →∙ B ∙))) (A →∙ Ω B) fst (fun (ΩfunExtIso A B) p) x = funExt⁻ (cong fst p) x snd (fun (ΩfunExtIso A B) p) i j = snd (p j) i fst (inv (ΩfunExtIso A B) (f , p) i) x = f x i @@ -77,10 +85,13 @@ private ≡ (λ i → refl ∙ β i) ∙ (λ i → α i ∙ refl) mainPath n α β i = (λ j → α (j ∧ ~ i) ∙ β (j ∧ i)) ∙ λ j → α (~ i ∨ j) ∙ β (i ∨ j) -EH-filler : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → typ ((Ω^ (2 + n)) A) → typ ((Ω^ (2 + n)) A) → I → I → I → _ +EH-filler : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → typ ((Ω^ (2 + n)) A) + → typ ((Ω^ (2 + n)) A) → I → I → I → _ EH-filler {A = A} n α β i j z = - hfill (λ k → λ { (i = i0) → ((cong (λ x → rUnit x (~ k)) α) ∙ cong (λ x → lUnit x (~ k)) β) j - ; (i = i1) → ((cong (λ x → lUnit x (~ k)) β) ∙ cong (λ x → rUnit x (~ k)) α) j + hfill (λ k → λ { (i = i0) → ((cong (λ x → rUnit x (~ k)) α) + ∙ cong (λ x → lUnit x (~ k)) β) j + ; (i = i1) → ((cong (λ x → lUnit x (~ k)) β) + ∙ cong (λ x → rUnit x (~ k)) α) j ; (j = i0) → rUnit refl (~ k) ; (j = i1) → rUnit refl (~ k)}) (inS (mainPath n α β i j)) z @@ -225,18 +236,67 @@ syllepsis {A = A} n α β k i j = ; (k = i1) → rUnit guy (j ∧ ~ r)}) (rUnit guy (~ r ∧ ~ k)) +------ Ωⁿ⁺¹ A ≃ Ωⁿ(Ω A) ------ +flipΩPath : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → ((Ω^ (suc n)) A) ≡ (Ω^ n) (Ω A) +flipΩPath {A = A} zero = refl +flipΩPath {A = A} (suc n) = cong Ω (flipΩPath {A = A} n) + +flipΩIso : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → Iso (fst ((Ω^ (suc n)) A)) (fst ((Ω^ n) (Ω A))) +flipΩIso {A = A} n = pathToIso (cong fst (flipΩPath n)) + +flipΩIso⁻pres· : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (f g : fst ((Ω^ (suc n)) (Ω A))) + → inv (flipΩIso (suc n)) (f ∙ g) + ≡ (inv (flipΩIso (suc n)) f) + ∙ (inv (flipΩIso (suc n)) g) +flipΩIso⁻pres· {A = A} n f g i = + transp (λ j → flipΩPath {A = A} n (~ i ∧ ~ j) .snd + ≡ flipΩPath n (~ i ∧ ~ j) .snd) i + (transp (λ j → flipΩPath {A = A} n (~ i ∨ ~ j) .snd + ≡ flipΩPath n (~ i ∨ ~ j) .snd) (~ i) f + ∙ transp (λ j → flipΩPath {A = A} n (~ i ∨ ~ j) .snd + ≡ flipΩPath n (~ i ∨ ~ j) .snd) (~ i) g) + +flipΩIsopres· : {ℓ : Level} {A : Pointed ℓ} (n : ℕ) + → (f g : fst (Ω ((Ω^ (suc n)) A))) + → fun (flipΩIso (suc n)) (f ∙ g) + ≡ (fun (flipΩIso (suc n)) f) + ∙ (fun (flipΩIso (suc n)) g) +flipΩIsopres· n = + morphLemmas.isMorphInv _∙_ _∙_ + (inv (flipΩIso (suc n))) + (flipΩIso⁻pres· n) + (fun (flipΩIso (suc n))) + (Iso.leftInv (flipΩIso (suc n))) + (Iso.rightInv (flipΩIso (suc n))) + +flipΩrefl : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → fun (flipΩIso {A = A} (suc n)) refl ≡ refl +flipΩrefl {A = A} n j = + transp (λ i₁ → fst (Ω (flipΩPath {A = A} n ((i₁ ∨ j))))) + j (snd (Ω (flipΩPath n j))) + +---- Misc. ---- + isCommA→isCommTrunc : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → isComm∙ A → isOfHLevel (suc n) (typ A) → isComm∙ (∥ typ A ∥ (suc n) , ∣ pt A ∣) isCommA→isCommTrunc {A = (A , a)} n comm hlev p q = ((λ i j → (leftInv (truncIdempotentIso (suc n) hlev) ((p ∙ q) j) (~ i))) - ∙∙ (λ i → cong {B = λ _ → ∥ A ∥ (suc n) } (λ x → ∣ x ∣) (cong (trRec hlev (λ x → x)) (p ∙ q))) - ∙∙ (λ i → cong {B = λ _ → ∥ A ∥ (suc n) } (λ x → ∣ x ∣) (congFunct {A = ∥ A ∥ (suc n)} {B = A} (trRec hlev (λ x → x)) p q i))) - ∙ ((λ i → cong {B = λ _ → ∥ A ∥ (suc n) } (λ x → ∣ x ∣) (comm (cong (trRec hlev (λ x → x)) p) (cong (trRec hlev (λ x → x)) q) i)) - ∙∙ (λ i → cong {B = λ _ → ∥ A ∥ (suc n) } (λ x → ∣ x ∣) (congFunct {A = ∥ A ∥ (suc n)} {B = A} (trRec hlev (λ x → x)) q p (~ i))) + ∙∙ (λ i → cong {B = λ _ → ∥ A ∥ (suc n) } (λ x → ∣ x ∣) + (cong (trRec hlev (λ x → x)) (p ∙ q))) + ∙∙ (λ i → cong {B = λ _ → ∥ A ∥ (suc n) } (λ x → ∣ x ∣) + (congFunct {A = ∥ A ∥ (suc n)} {B = A} (trRec hlev (λ x → x)) p q i))) + ∙ ((λ i → cong {B = λ _ → ∥ A ∥ (suc n) } (λ x → ∣ x ∣) + (comm (cong (trRec hlev (λ x → x)) p) (cong (trRec hlev (λ x → x)) q) i)) + ∙∙ (λ i → cong {B = λ _ → ∥ A ∥ (suc n) } (λ x → ∣ x ∣) + (congFunct {A = ∥ A ∥ (suc n)} {B = A} (trRec hlev (λ x → x)) q p (~ i))) ∙∙ (λ i j → (leftInv (truncIdempotentIso (suc n) hlev) ((q ∙ p) j) i))) -ptdIso→comm : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Type ℓ'} (e : Iso (typ A) B) → isComm∙ A → isComm∙ (B , fun e (pt A)) +ptdIso→comm : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Type ℓ'} (e : Iso (typ A) B) + → isComm∙ A → isComm∙ (B , fun e (pt A)) ptdIso→comm {A = (A , a)} {B = B} e comm p q = sym (rightInv (congIso e) (p ∙ q)) ∙∙ (cong (fun (congIso e)) ((invCongFunct e p q) From ab506c92748dccce228a0a97b4dfc9b61a5d17c4 Mon Sep 17 00:00:00 2001 From: aljungstrom Date: Thu, 11 Nov 2021 00:20:16 +0100 Subject: [PATCH 89/89] minor --- Cubical/HITs/Susp/Properties.agda | 19 ++++++++++--------- Cubical/Homotopy/Loopspace.agda | 1 - 2 files changed, 10 insertions(+), 10 deletions(-) diff --git a/Cubical/HITs/Susp/Properties.agda b/Cubical/HITs/Susp/Properties.agda index bc017d4a1..2315c9106 100644 --- a/Cubical/HITs/Susp/Properties.agda +++ b/Cubical/HITs/Susp/Properties.agda @@ -154,15 +154,16 @@ module _ {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} where fun ΩSuspAdjointIso = toΩ→fromSusp inv ΩSuspAdjointIso = fromSusp→toΩ rightInv ΩSuspAdjointIso f = - ΣPathP (funExt (λ { north → sym (snd f) - ; south → sym (snd f) ∙ cong (fst f) (merid (pt A)) - ; (merid a i) j → - hcomp (λ k → λ { (i = i0) → snd f (~ j ∧ k) - ; (i = i1) → compPath-filler' - (sym (snd f)) - (cong (fst f) (merid (pt A))) k j - ; (j = i1) → fst f (merid a i)}) - (fst f (compPath-filler (merid a) (sym (merid (pt A))) (~ j) i))}) + ΣPathP (funExt + (λ { north → sym (snd f) + ; south → sym (snd f) ∙ cong (fst f) (merid (pt A)) + ; (merid a i) j → + hcomp (λ k → λ { (i = i0) → snd f (~ j ∧ k) + ; (i = i1) → compPath-filler' + (sym (snd f)) + (cong (fst f) (merid (pt A))) k j + ; (j = i1) → fst f (merid a i)}) + (fst f (compPath-filler (merid a) (sym (merid (pt A))) (~ j) i))}) , λ i j → snd f (~ i ∨ j)) leftInv ΩSuspAdjointIso f = →∙Homogeneous≡ (isHomogeneousPath _ _) diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda index 153c9634d..9c3ef34ee 100644 --- a/Cubical/Homotopy/Loopspace.agda +++ b/Cubical/Homotopy/Loopspace.agda @@ -21,7 +21,6 @@ open import Cubical.Foundations.Equiv.HalfAdjoint open import Cubical.Foundations.Equiv open import Cubical.Functions.Morphism open import Cubical.Data.Sigma --- open import Cubical.HITs.Sn open Iso {- loop space of a pointed type -}