diff --git a/Cubical/Algebra/AbGroup/FinitePresentation.agda b/Cubical/Algebra/AbGroup/FinitePresentation.agda new file mode 100644 index 0000000000..06cfc2879e --- /dev/null +++ b/Cubical/Algebra/AbGroup/FinitePresentation.agda @@ -0,0 +1,40 @@ +module Cubical.Algebra.AbGroup.FinitePresentation where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Nat +open import Cubical.Data.Int + +open import Cubical.HITs.PropositionalTruncation + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup + +open import Cubical.Algebra.Group.QuotientGroup +open import Cubical.Algebra.Group.Subgroup + +private + variable + ℓ : Level + +record FinitePresentation (A : AbGroup ℓ) : Type ℓ where + field + nGens : ℕ + nRels : ℕ + rels : AbGroupHom ℤ[Fin nRels ] ℤ[Fin nGens ] + fpiso : AbGroupIso A (ℤ[Fin nGens ] /Im rels) + +isFinitelyPresented : AbGroup ℓ → Type ℓ +isFinitelyPresented G = ∥ FinitePresentation G ∥₁ + +open FinitePresentation +GrIsoPresFinitePresentation : ∀ {ℓ ℓ'} {A : AbGroup ℓ} {B : AbGroup ℓ'} + → AbGroupIso A B → FinitePresentation A → FinitePresentation B +nGens (GrIsoPresFinitePresentation abG fpA) = nGens fpA +nRels (GrIsoPresFinitePresentation abG fpA) = nRels fpA +rels (GrIsoPresFinitePresentation abG fpA) = rels fpA +fpiso (GrIsoPresFinitePresentation abG fpA) = + compGroupIso (invGroupIso abG) (fpiso fpA) diff --git a/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda b/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda index 14cb90e4e8..a72ba69860 100644 --- a/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda +++ b/Cubical/Algebra/AbGroup/Instances/FreeAbGroup.agda @@ -5,6 +5,7 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv open import Cubical.Data.Sigma open import Cubical.Data.Nat hiding (_·_) renaming (_+_ to _+ℕ_) @@ -15,6 +16,8 @@ open import Cubical.Data.Fin.Inductive open import Cubical.Data.Empty as ⊥ open import Cubical.HITs.FreeAbGroup +open import Cubical.HITs.FreeGroup as FG hiding (rec) +open import Cubical.HITs.SetQuotients as SQ hiding (_/_ ; rec) open import Cubical.Algebra.AbGroup open import Cubical.Algebra.AbGroup.Instances.Pi @@ -23,6 +26,10 @@ open import Cubical.Algebra.AbGroup.Instances.DirectProduct open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.Abelianization.Base +open import Cubical.Algebra.Group.Abelianization.Properties as Abi hiding (rec) +open import Cubical.Algebra.Group.Subgroup +open import Cubical.Algebra.Group.QuotientGroup private variable @@ -33,11 +40,66 @@ module _ {A : Type ℓ} where FAGAbGroup : AbGroup ℓ FAGAbGroup = makeAbGroup {G = FreeAbGroup A} ε _·_ _⁻¹ trunc assoc identityᵣ invᵣ comm +FAGAbGroup→AbGroupHom : ∀ {ℓ ℓ'} {A : Type ℓ} {G : AbGroup ℓ'} + → (A → fst G) → AbGroupHom (FAGAbGroup {A = A}) G +fst (FAGAbGroup→AbGroupHom {G = G} f) = + Rec.f (AbGroupStr.is-set (snd G)) f + (AbGroupStr.0g (snd G)) (AbGroupStr._+_ (snd G)) (AbGroupStr.-_ (snd G)) + (AbGroupStr.+Assoc (snd G)) (AbGroupStr.+Comm (snd G)) + (AbGroupStr.+IdR (snd G)) (AbGroupStr.+InvR (snd G)) +snd (FAGAbGroup→AbGroupHom {G = G} f) = makeIsGroupHom λ x y → refl + +FAGAbGroupGroupHom≡ : ∀ {ℓ ℓ'} {A : Type ℓ} {G : AbGroup ℓ'} + (f g : AbGroupHom (FAGAbGroup {A = A}) G) + → (∀ a → (fst f) (⟦ a ⟧) ≡ (fst g) (⟦ a ⟧)) → f ≡ g +FAGAbGroupGroupHom≡ {G = G} f g p = + GroupHom≡ (funExt (ElimProp.f (AbGroupStr.is-set (snd G) _ _) + p (IsGroupHom.pres1 (snd f) ∙ sym (IsGroupHom.pres1 (snd g))) + (λ p q → IsGroupHom.pres· (snd f) _ _ + ∙ cong₂ (AbGroupStr._+_ (snd G)) p q + ∙ sym (IsGroupHom.pres· (snd g) _ _)) + λ p → IsGroupHom.presinv (snd f) _ + ∙ cong (AbGroupStr.-_ (snd G)) p + ∙ sym (IsGroupHom.presinv (snd g) _))) + +module _ {A : Type ℓ} where + freeGroup→freeAbGroup : GroupHom (freeGroupGroup A) + (AbGroup→Group (FAGAbGroup {A = A})) + freeGroup→freeAbGroup = FG.rec {Group = AbGroup→Group (FAGAbGroup {A = A})} ⟦_⟧ + + AbelienizeFreeGroup→FreeAbGroup : + AbGroupHom (AbelianizationAbGroup (freeGroupGroup A)) (FAGAbGroup {A = A}) + AbelienizeFreeGroup→FreeAbGroup = + fromAbelianization FAGAbGroup freeGroup→freeAbGroup + + FreeAbGroup→AbelienizeFreeGroup : + AbGroupHom (FAGAbGroup {A = A}) (AbelianizationAbGroup (freeGroupGroup A)) + FreeAbGroup→AbelienizeFreeGroup = FAGAbGroup→AbGroupHom λ a → η (η a) + + GroupIso-AbelienizeFreeGroup→FreeAbGroup : + AbGroupIso (AbelianizationAbGroup (freeGroupGroup A)) (FAGAbGroup {A = A}) + Iso.fun (fst GroupIso-AbelienizeFreeGroup→FreeAbGroup) = + AbelienizeFreeGroup→FreeAbGroup .fst + Iso.inv (fst GroupIso-AbelienizeFreeGroup→FreeAbGroup) = + FreeAbGroup→AbelienizeFreeGroup .fst + Iso.rightInv (fst GroupIso-AbelienizeFreeGroup→FreeAbGroup) x i = + FAGAbGroupGroupHom≡ + (compGroupHom FreeAbGroup→AbelienizeFreeGroup + AbelienizeFreeGroup→FreeAbGroup) + idGroupHom (λ _ → refl) i .fst x + Iso.leftInv (fst GroupIso-AbelienizeFreeGroup→FreeAbGroup) = + Abi.elimProp _ (λ _ → isset _ _) + (funExt⁻ (cong fst (freeGroupHom≡ + {f = compGroupHom freeGroup→freeAbGroup FreeAbGroup→AbelienizeFreeGroup} + {g = AbelianizationGroupStructure.ηAsGroupHom (freeGroupGroup A)} + λ _ → refl))) + snd GroupIso-AbelienizeFreeGroup→FreeAbGroup = + AbelienizeFreeGroup→FreeAbGroup .snd + -- Alternative definition of case when A = Fin n ℤ[Fin_] : (n : ℕ) → AbGroup ℓ-zero ℤ[Fin n ] = ΠℤAbGroup (Fin n) - -- generator of ℤ[Fin_] ℤFinGenerator : {n : ℕ} (k : Fin n) → ℤ[Fin n ] .fst ℤFinGenerator {n = n} k s with (fst k ≟ᵗ fst s) @@ -398,21 +460,21 @@ snd (ℤFinFunct {n = n} {m} f) = ... | gt _ = refl -- Homs are equal if they agree on generators -agreeOnℤFinGenerator→≡ : ∀ {n m : ℕ} - → {ϕ ψ : AbGroupHom (ℤ[Fin n ]) (ℤ[Fin m ])} +agreeOnℤFinGenerator→≡ : ∀ {ℓ} {n : ℕ} {G : Group ℓ} + → {ϕ ψ : GroupHom (AbGroup→Group (ℤ[Fin n ])) G} → ((x : _) → fst ϕ (ℤFinGenerator x) ≡ fst ψ (ℤFinGenerator x)) → ϕ ≡ ψ -agreeOnℤFinGenerator→≡ {n} {m} {ϕ} {ψ} idr = +agreeOnℤFinGenerator→≡ {G = G} {ϕ} {ψ} w = Σ≡Prop (λ _ → isPropIsGroupHom _ _) (funExt - (elimPropℤFin _ _ (λ _ → isOfHLevelPath' 1 (isSetΠ (λ _ → isSetℤ)) _ _) + (elimPropℤFin _ _ (λ _ → isOfHLevelPath' 1 (GroupStr.is-set (snd G)) _ _) (IsGroupHom.pres1 (snd ϕ) ∙ sym (IsGroupHom.pres1 (snd ψ))) - idr + w (λ f g p q → IsGroupHom.pres· (snd ϕ) f g - ∙∙ (λ i x → p i x + q i x) + ∙∙ (λ i → GroupStr._·_ (snd G) (p i) (q i)) ∙∙ sym (IsGroupHom.pres· (snd ψ) f g )) λ f p → IsGroupHom.presinv (snd ϕ) f - ∙∙ (λ i x → -ℤ (p i x)) + ∙∙ cong (GroupStr.inv (G .snd) ) p ∙∙ sym (IsGroupHom.presinv (snd ψ) f))) -- @@ -460,3 +522,47 @@ snd (sumCoefficients n) = makeIsGroupHom (λ x y → funExt λ _ → sumFinℤHo ℤFinProduct : (n m : ℕ) → AbGroupIso ℤ[Fin (n +ℕ m) ] (AbDirProd ℤ[Fin n ] ℤ[Fin m ]) fst (ℤFinProduct n m) = ℤFinProductIso n m snd (ℤFinProduct n m) = makeIsGroupHom (λ x y → refl) + +-- lemmas about quotients of Free abelian groups +ℤ[]/-GroupHom≡ : ∀ {ℓ} {n : ℕ} (G : Group ℓ) + {Q : NormalSubgroup (AbGroup→Group ℤ[Fin n ])} + (ϕ ψ : GroupHom (AbGroup→Group (ℤ[Fin n ]) / Q ) G) + → ((k : _) → fst ϕ [ ℤFinGenerator k ] ≡ fst ψ [ ℤFinGenerator k ]) + → ϕ ≡ ψ +ℤ[]/-GroupHom≡ G ϕ ψ s = Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt (SQ.elimProp (λ _ → GroupStr.is-set (snd G) _ _) + λ x → funExt⁻ (cong fst (agreeOnℤFinGenerator→≡ + {ϕ = compGroupHom ([_] , makeIsGroupHom λ f g → refl) ϕ} + {ψ = compGroupHom ([_] , makeIsGroupHom λ f g → refl) ψ} + s)) x)) + +makeℤ[]/Equiv : ∀ {ℓ ℓ'} {G : Group ℓ} {H : Group ℓ'} {n : ℕ} + {T : NormalSubgroup (AbGroup→Group ℤ[Fin n ])} + (ϕ : GroupEquiv (AbGroup→Group ℤ[Fin n ] / T) G) + (ψ : GroupEquiv (AbGroup→Group ℤ[Fin n ] / T) H) + (m : GroupHom G H) + → ((k : _) → fst m (fst (fst ϕ) [ ℤFinGenerator k ]) + ≡ fst (fst ψ) [ ℤFinGenerator k ]) + → isEquiv (fst m) +makeℤ[]/Equiv {n = n} {T = T} ϕ ψ m ind = + subst isEquiv (cong fst lem) + (compEquiv (invEquiv (fst ϕ)) (fst ψ) .snd) + where + ξ : GroupHom (AbGroup→Group ℤ[Fin n ] / T) (AbGroup→Group ℤ[Fin n ] / T) + ξ = compGroupHom (GroupEquiv→GroupHom ϕ) + (compGroupHom m (GroupEquiv→GroupHom (invGroupEquiv ψ))) + + ξ≡id : ξ ≡ idGroupHom + ξ≡id = ℤ[]/-GroupHom≡ _ _ _ + λ w → cong (invEq (fst ψ)) (ind w) + ∙ retEq (fst ψ) [ ℤFinGenerator w ] + + lem : compGroupHom (GroupEquiv→GroupHom (invGroupEquiv ϕ)) + (GroupEquiv→GroupHom ψ) + ≡ m + lem = Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt λ x → (sym (funExt⁻ + (cong fst (cong (compGroupHom (GroupEquiv→GroupHom (invGroupEquiv ϕ))) + (cong (λ X → compGroupHom X (GroupEquiv→GroupHom ψ)) ξ≡id))) x)) + ∙ secEq (fst ψ) _ + ∙ cong (fst m) (secEq (fst ϕ) x)) diff --git a/Cubical/Algebra/AbGroup/Properties.agda b/Cubical/Algebra/AbGroup/Properties.agda index 26e09ae8ee..6feb7ca2c4 100644 --- a/Cubical/Algebra/AbGroup/Properties.agda +++ b/Cubical/Algebra/AbGroup/Properties.agda @@ -60,6 +60,15 @@ snd (negGroupHom A B ϕ) = subtrGroupHom : (A : AbGroup ℓ) (B : AbGroup ℓ') (ϕ ψ : AbGroupHom A B) → AbGroupHom A B subtrGroupHom A B ϕ ψ = addGroupHom A B ϕ (negGroupHom A B ψ) +-- Abelian groups quotiented by image of a map +_/Im_ : {H : Group ℓ} (G : AbGroup ℓ) (ϕ : GroupHom H (AbGroup→Group G)) → AbGroup ℓ +G /Im ϕ = + Group→AbGroup (G /' ϕ) + (elimProp2 (λ _ _ → squash/ _ _) λ a b → cong [_] (AbGroupStr.+Comm (snd G) _ _)) + where + _/'_ : {H : Group ℓ} (G : AbGroup ℓ) (ϕ : GroupHom H (AbGroup→Group G)) → Group ℓ + G /' ϕ = AbGroup→Group G + / (imSubgroup ϕ , isNormalIm ϕ λ _ _ → AbGroupStr.+Comm (snd G) _ _) -- ℤ-multiplication defines a homomorphism for abelian groups @@ -104,13 +113,7 @@ snd (multₗHom G n) = makeIsGroupHom (ℤ·isHom n G) -- Abelian groups quotiented by a natural number _/^_ : (G : AbGroup ℓ) (n : ℕ) → AbGroup ℓ -G /^ n = - Group→AbGroup - ((AbGroup→Group G) - / (imSubgroup (multₗHom G (pos n)) - , isNormalIm (multₗHom G (pos n)) (AbGroupStr.+Comm (snd G)))) - (SQ.elimProp2 (λ _ _ → squash/ _ _) - λ a b → cong [_] (AbGroupStr.+Comm (snd G) a b)) +G /^ n = G /Im multₗHom G (pos n) -- Torsion subgrous _[_]ₜ : (G : AbGroup ℓ) (n : ℕ) → AbGroup ℓ diff --git a/Cubical/Algebra/Group/Abelianization/Properties.agda b/Cubical/Algebra/Group/Abelianization/Properties.agda index ca620b1023..75b7dcbef2 100644 --- a/Cubical/Algebra/Group/Abelianization/Properties.agda +++ b/Cubical/Algebra/Group/Abelianization/Properties.agda @@ -12,6 +12,8 @@ module Cubical.Algebra.Group.Abelianization.Properties where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Structure +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv open import Cubical.Data.Sigma @@ -19,15 +21,17 @@ open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Group.Properties open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties - using (isPropIsGroupHom; compGroupHom; idGroupHom) - + using (isPropIsGroupHom; compGroupHom; idGroupHom + ; makeIsGroupHom ; GroupEquiv→GroupHom ; invGroupEquiv) +open import Cubical.Algebra.Monoid.Base +open import Cubical.Algebra.Semigroup open import Cubical.Algebra.AbGroup.Base open import Cubical.Algebra.Group.Abelianization.Base private variable - ℓ : Level + ℓ ℓ' : Level module _ (G : Group ℓ) where open GroupStr {{...}} @@ -37,7 +41,7 @@ module _ (G : Group ℓ) where _ = snd G -- Some helpful lemmas, similar to those in Cubical/HITs/SetQuotients/Properties.agda - elimProp : {B : Abelianization G → Type ℓ} + elimProp : {B : Abelianization G → Type ℓ'} → (Bprop : (x : Abelianization G) → isProp (B x)) → (f : (g : fst G) → B (η g)) → (x : Abelianization G) @@ -51,7 +55,7 @@ module _ (G : Group ℓ) where where g = elimProp Bprop f - elimProp2 : {C : Abelianization G → Abelianization G → Type ℓ} + elimProp2 : {C : Abelianization G → Abelianization G → Type ℓ'} → (Cprop : (x y : Abelianization G) → isProp (C x y)) → (f : (a b : fst G) → C (η a) (η b)) → (x y : Abelianization G) @@ -59,7 +63,7 @@ module _ (G : Group ℓ) where elimProp2 Cprop f = elimProp (λ x → isPropΠ (λ y → Cprop x y)) (λ x → elimProp (λ y → Cprop (η x) y) (f x)) - elimProp3 : {D : Abelianization G → Abelianization G → Abelianization G → Type ℓ} + elimProp3 : {D : Abelianization G → Abelianization G → Abelianization G → Type ℓ'} → (Dprop : (x y z : Abelianization G) → isProp (D x y z)) → ((a b c : fst G) → D (η a) (η b) (η c)) → (x y z : Abelianization G) @@ -67,21 +71,21 @@ module _ (G : Group ℓ) where elimProp3 Dprop f = elimProp (λ x → isPropΠ2 (λ y z → Dprop x y z)) (λ x → elimProp2 (λ y z → Dprop (η x) y z) (f x)) - elimContr : {B : Abelianization G → Type ℓ} + elimContr : {B : Abelianization G → Type ℓ'} → (Bcontr : ∀ (a : fst G) → isContr (B (η a))) → (x : Abelianization G) → B x elimContr Bcontr = elimProp (elimProp (λ _ → isPropIsProp) λ _ → isContr→isProp (Bcontr _)) λ _ → Bcontr _ .fst - elimContr2 : {C : Abelianization G → Abelianization G → Type ℓ} + elimContr2 : {C : Abelianization G → Abelianization G → Type ℓ'} → (Ccontr : ∀ (a b : fst G) → isContr (C (η a) (η b))) → (x y : Abelianization G) → C x y elimContr2 Ccontr = elimContr λ _ → isOfHLevelΠ 0 (elimContr λ _ → inhProp→isContr (Ccontr _ _) isPropIsContr) - rec : {M : Type ℓ} + rec : {M : Type ℓ'} (Mset : isSet M) (f : fst G → M) (fcomm : (a b c : fst G) → f (a · (b · c)) ≡ f (a · (c · b))) @@ -92,7 +96,7 @@ module _ (G : Group ℓ) where where g = rec Mset f fcomm - rec2 : {M : Type ℓ} + rec2 : {M : Type ℓ'} (Mset : isSet M) (f : fst G → fst G → M) (fcomml : (a b c d : fst G) → f (a · (b · c)) d ≡ f (a · (c · b)) d) @@ -194,9 +198,24 @@ module AbelianizationGroupStructure (G : Group ℓ) where η (y · x) ≡⟨ refl ⟩ (η y) ·Ab (η x) ∎) + open AbGroupStr + open IsAbGroup + open IsGroup + open IsSemigroup + open IsMonoid -- The proof that the abelianization is in fact an abelian group. asAbelianGroup : AbGroup ℓ - asAbelianGroup = makeAbGroup 1Ab _·Ab_ invAb isset assocAb ridAb rinvAb commAb + fst asAbelianGroup = Abelianization G + 0g (snd asAbelianGroup) = 1Ab + _+_ (snd asAbelianGroup) = _·Ab_ + - snd asAbelianGroup = invAb + is-set (isSemigroup (isMonoid (isGroup (isAbGroup (snd asAbelianGroup))))) = isset + ·Assoc (isSemigroup (isMonoid (isGroup (isAbGroup (snd asAbelianGroup))))) = assocAb + ·IdR (isMonoid (isGroup (isAbGroup (snd asAbelianGroup)))) = ridAb + ·IdL (isMonoid (isGroup (isAbGroup (snd asAbelianGroup)))) x = commAb 1Ab x ∙ ridAb x + ·InvR (isGroup (isAbGroup (snd asAbelianGroup))) = rinvAb + ·InvL (isGroup (isAbGroup (snd asAbelianGroup))) x = commAb (invAb x) x ∙ rinvAb x + +Comm (isAbGroup (snd asAbelianGroup)) = commAb -- The proof that η can be seen as a group homomorphism ηAsGroupHom : GroupHom G (AbGroup→Group asAbelianGroup) @@ -211,6 +230,9 @@ module AbelianizationGroupStructure (G : Group ℓ) where AbelianizationAbGroup : (G : Group ℓ) → AbGroup ℓ AbelianizationAbGroup G = AbelianizationGroupStructure.asAbelianGroup G +AbelianizationGroup : (G : Group ℓ) → Group ℓ +AbelianizationGroup G = AbGroup→Group (AbelianizationAbGroup G) + AbelianizationHom : (G : Group ℓ) → GroupHom G (AbGroup→Group (AbelianizationAbGroup G)) AbelianizationHom G = AbelianizationGroupStructure.ηAsGroupHom G @@ -221,80 +243,135 @@ module UniversalProperty (G : Group ℓ) where private instance _ = snd G - abstract - {- The proof of the universal property of the abelianization. - - G --η--> abelianization - \ . - \ . - f ∃! inducedHom - \ . - \ . - H - commuting diagram - -} - inducedHom : (H : AbGroup ℓ) - → (f : GroupHom G (AbGroup→Group H)) - → AbGroupHom asAbelianGroup H - inducedHom H f = g , gIsHom - where open IsGroupHom - instance - _ : GroupStr (fst H) - _ = snd (AbGroup→Group H) - f' : fst G → fst H - f' = fst f - g : Abelianization G → fst H - g = (rec G) - is-set - (λ x → (f') x) - (λ a b c → f' (a · b · c) ≡⟨ (snd f).pres· a (b · c) ⟩ - (f' a) · (f' (b · c)) ≡⟨ cong - (λ x → (f' a) · x) - ((snd f).pres· b c) ⟩ - (f' a) · (f' b) · (f' c) ≡⟨ cong - (λ x → (f' a) · x) - ((snd H).AbGroupStr.+Comm (f' b) (f' c)) ⟩ - (f' a) · (f' c) · (f' b) ≡⟨ cong - (λ x → (f' a) · x) - (sym ((snd f).pres· c b)) ⟩ - (f' a) · (f' (c · b)) ≡⟨ sym ((snd f).pres· a (c · b)) ⟩ - f' (a · c · b) ∎) - gIsHom : IsGroupHom (snd (AbGroup→Group asAbelianGroup)) g (snd (AbGroup→Group H)) - pres· gIsHom = - (elimProp2 G) - (λ x y → is-set _ _) - ((snd f).pres·) - pres1 gIsHom = (snd f).pres1 - presinv gIsHom = - (elimProp G) - (λ x → is-set _ _) - ((snd f).presinv) - - commutativity : (H : AbGroup ℓ) - → (f : GroupHom G (AbGroup→Group H)) - → (compGroupHom ηAsGroupHom (inducedHom H f) ≡ f) - commutativity H f = - Σ≡Prop - (λ _ → isPropIsGroupHom _ _) - (λ i x → q x i) - where q : (x : fst G) - → fst (compGroupHom ηAsGroupHom (inducedHom H f)) x ≡ fst f x - q = (λ x → refl) - - uniqueness : (H : AbGroup ℓ) - → (f : GroupHom G (AbGroup→Group H)) - → (g : AbGroupHom asAbelianGroup H) - → (p : compGroupHom ηAsGroupHom g ≡ f) - → (g ≡ inducedHom H f) - uniqueness H f g p = - Σ≡Prop - (λ _ → isPropIsGroupHom _ _) - (λ i x → q x i) - where - module H = AbGroupStr (str H) - q : (x : Abelianization G) → fst g x ≡ fst (inducedHom H f) x - q = (elimProp G) - (λ _ → H.is-set _ _) - (λ x → fst g (η x) ≡⟨ cong (λ f → f x) (cong fst p) ⟩ - (fst f) x ≡⟨ refl ⟩ - fst (inducedHom H f) (η x)∎) + + {- The proof of the universal property of the abelianization. + + G --η--> abelianization + \ . + \ . + f ∃! inducedHom + \ . + \ . + H + commuting diagram + -} + inducedHom : (H : AbGroup ℓ) + → (f : GroupHom G (AbGroup→Group H)) + → AbGroupHom asAbelianGroup H + inducedHom H f = g , gIsHom + where open IsGroupHom + instance + _ : GroupStr (fst H) + _ = snd (AbGroup→Group H) + f' : fst G → fst H + f' = fst f + g : Abelianization G → fst H + g = (rec G) + is-set + (λ x → (f') x) + (λ a b c → f' (a · b · c) ≡⟨ (snd f).pres· a (b · c) ⟩ + (f' a) · (f' (b · c)) ≡⟨ cong + (λ x → (f' a) · x) + ((snd f).pres· b c) ⟩ + (f' a) · (f' b) · (f' c) ≡⟨ cong + (λ x → (f' a) · x) + ((snd H).AbGroupStr.+Comm (f' b) (f' c)) ⟩ + (f' a) · (f' c) · (f' b) ≡⟨ cong + (λ x → (f' a) · x) + (sym ((snd f).pres· c b)) ⟩ + (f' a) · (f' (c · b)) ≡⟨ sym ((snd f).pres· a (c · b)) ⟩ + f' (a · c · b) ∎) + gIsHom : IsGroupHom (snd (AbGroup→Group asAbelianGroup)) g (snd (AbGroup→Group H)) + pres· gIsHom = + (elimProp2 G) + (λ x y → is-set _ _) + ((snd f).pres·) + pres1 gIsHom = (snd f).pres1 + presinv gIsHom = + (elimProp G) + (λ x → is-set _ _) + ((snd f).presinv) + + commutativity : (H : AbGroup ℓ) + → (f : GroupHom G (AbGroup→Group H)) + → (compGroupHom ηAsGroupHom (inducedHom H f) ≡ f) + commutativity H f = + Σ≡Prop + (λ _ → isPropIsGroupHom _ _) + (λ i x → q x i) + where q : (x : fst G) + → fst (compGroupHom ηAsGroupHom (inducedHom H f)) x ≡ fst f x + q = (λ x → refl) + + uniqueness : (H : AbGroup ℓ) + → (f : GroupHom G (AbGroup→Group H)) + → (g : AbGroupHom asAbelianGroup H) + → (p : compGroupHom ηAsGroupHom g ≡ f) + → (g ≡ inducedHom H f) + uniqueness H f g p = + Σ≡Prop + (λ _ → isPropIsGroupHom _ _) + (λ i x → q x i) + where + module H = AbGroupStr (str H) + q : (x : Abelianization G) → fst g x ≡ fst (inducedHom H f) x + q = (elimProp G) + (λ _ → H.is-set _ _) + (λ x → fst g (η x) ≡⟨ cong (λ f → f x) (cong fst p) ⟩ + (fst f) x ≡⟨ refl ⟩ + fst (inducedHom H f) (η x)∎) + + +module _ {ℓ} {G : Group ℓ} (H : AbGroup ℓ) (ϕ : GroupHom G (AbGroup→Group H)) where + fromAbelianization : AbGroupHom (AbelianizationAbGroup G) H + fst fromAbelianization = rec G (AbGroupStr.is-set (snd H)) (fst ϕ) + λ x y z → IsGroupHom.pres· (snd ϕ) _ _ + ∙ cong₂ (AbGroupStr._+_ (snd H)) refl + (IsGroupHom.pres· (snd ϕ) _ _ + ∙ AbGroupStr.+Comm (snd H) _ _ + ∙ sym (IsGroupHom.pres· (snd ϕ) _ _)) + ∙ sym (IsGroupHom.pres· (snd ϕ) _ _) + snd fromAbelianization = + makeIsGroupHom (elimProp2 _ + (λ _ _ → AbGroupStr.is-set (snd H) _ _) + λ x y → IsGroupHom.pres· (snd ϕ) x y) + +AbelianizationFun : ∀ {ℓ} {G : Group ℓ} {H : Group ℓ} + → GroupHom G H → AbGroupHom (AbelianizationAbGroup G) (AbelianizationAbGroup H) +fst (AbelianizationFun {G = G} {H} ϕ) = rec _ isset (λ x → η (fst ϕ x)) λ a b c + → cong η (IsGroupHom.pres· (snd ϕ) a _ + ∙ cong₂ (GroupStr._·_ (snd H)) refl (IsGroupHom.pres· (snd ϕ) b c)) + ∙ comm _ _ _ + ∙ sym (cong η (IsGroupHom.pres· (snd ϕ) a _ + ∙ cong₂ (GroupStr._·_ (snd H)) refl (IsGroupHom.pres· (snd ϕ) c b))) +snd (AbelianizationFun {G = G} {H} ϕ) = makeIsGroupHom + (elimProp2 _ (λ _ _ → isset _ _) + λ a b → cong η (IsGroupHom.pres· (snd ϕ) a b)) + +AbelianizationEquiv : ∀ {ℓ} {G : Group ℓ} {H : Group ℓ} + → GroupEquiv G H + → AbGroupEquiv (AbelianizationAbGroup G) (AbelianizationAbGroup H) +fst (AbelianizationEquiv {G = G} {H} ϕ) = isoToEquiv main + where + main : Iso _ _ + Iso.fun main = fst (AbelianizationFun (GroupEquiv→GroupHom ϕ)) + Iso.inv main = fst (AbelianizationFun (GroupEquiv→GroupHom (invGroupEquiv ϕ))) + Iso.rightInv main = + elimProp _ (λ _ → isset _ _) λ g → cong η (secEq (fst ϕ) g) + Iso.leftInv main = + elimProp _ (λ _ → isset _ _) λ g → cong η (retEq (fst ϕ) g) +snd (AbelianizationEquiv {G = G} {H} ϕ) = + snd (AbelianizationFun (fst (fst ϕ) , snd ϕ)) + +AbelianizationIdempotent : ∀ {ℓ} (G : AbGroup ℓ) + → AbGroupIso G (AbelianizationAbGroup (AbGroup→Group G)) +Iso.fun (fst (AbelianizationIdempotent G)) = η +Iso.inv (fst (AbelianizationIdempotent G)) = + rec _ (AbGroupStr.is-set (snd G)) + (λ x → x) + λ a b c → cong (AbGroupStr._+_ (snd G) a) (AbGroupStr.+Comm (snd G) _ _) +Iso.rightInv (fst (AbelianizationIdempotent G)) = + elimProp _ (λ _ → isset _ _) (λ _ → refl) +Iso.leftInv (fst (AbelianizationIdempotent G)) x = refl +snd (AbelianizationIdempotent G) = + snd (AbelianizationGroupStructure.ηAsGroupHom _) diff --git a/Cubical/Algebra/Group/Instances/Pi.agda b/Cubical/Algebra/Group/Instances/Pi.agda index e8b1c166e4..8c592a0463 100644 --- a/Cubical/Algebra/Group/Instances/Pi.agda +++ b/Cubical/Algebra/Group/Instances/Pi.agda @@ -2,9 +2,12 @@ module Cubical.Algebra.Group.Instances.Pi where open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism open import Cubical.Algebra.Group.Base open import Cubical.Algebra.Semigroup open import Cubical.Algebra.Monoid +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.Morphisms open IsGroup open GroupStr @@ -29,3 +32,21 @@ module _ {ℓ ℓ' : Level} {X : Type ℓ} (G : X → Group ℓ') where funExt λ x → ·InvR (isGroup (snd (G x))) (f x) ·InvL (isGroup (snd ΠGroup)) f = funExt λ x → ·InvL (isGroup (snd (G x))) (f x) + +ΠGroupHom : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {G : A → Group ℓ'} {H : A → Group ℓ''} + → ((a : _) → GroupHom (G a) (H a)) + → GroupHom (ΠGroup G) (ΠGroup H) +fst (ΠGroupHom fam) f a = fst (fam a) (f a) +snd (ΠGroupHom fam) = + makeIsGroupHom λ f g + → funExt λ a → IsGroupHom.pres· (snd (fam a)) _ _ + +ΠGroupIso : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {G : A → Group ℓ'} {H : A → Group ℓ''} + → ((a : _) → GroupIso (G a) (H a)) + → GroupIso (ΠGroup G) (ΠGroup H) +Iso.fun (fst (ΠGroupIso fam)) = fst (ΠGroupHom λ a → GroupIso→GroupHom (fam a)) +Iso.inv (fst (ΠGroupIso fam)) = + fst (ΠGroupHom λ a → GroupIso→GroupHom (invGroupIso (fam a))) +Iso.rightInv (fst (ΠGroupIso fam)) f = funExt λ x → Iso.rightInv (fst (fam x)) _ +Iso.leftInv (fst (ΠGroupIso fam)) f = funExt λ x → Iso.leftInv (fst (fam x)) _ +snd (ΠGroupIso fam) = snd (ΠGroupHom λ a → GroupIso→GroupHom (fam a)) diff --git a/Cubical/Algebra/Group/IsomorphismTheorems.agda b/Cubical/Algebra/Group/IsomorphismTheorems.agda index 9c0863a249..1f5bb7278d 100644 --- a/Cubical/Algebra/Group/IsomorphismTheorems.agda +++ b/Cubical/Algebra/Group/IsomorphismTheorems.agda @@ -43,12 +43,6 @@ module _ {G H : Group ℓ} (ϕ : GroupHom G H) where imϕ : Group ℓ imϕ = imGroup ϕ - -- for completeness: - imNormalSubgroup : ((x y : ⟨ H ⟩) - → GroupStr._·_ (snd H) x y ≡ GroupStr._·_ (snd H) y x) - → NormalSubgroup H - imNormalSubgroup comm = imSubgroup ϕ , isNormalIm ϕ comm - private module G = GroupStr (snd G) module H = GroupStr (snd H) diff --git a/Cubical/Algebra/Group/Properties.agda b/Cubical/Algebra/Group/Properties.agda index 2f837738e2..c2cb2a51f8 100644 --- a/Cubical/Algebra/Group/Properties.agda +++ b/Cubical/Algebra/Group/Properties.agda @@ -4,6 +4,7 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function open import Cubical.Foundations.HLevels open import Cubical.Foundations.Structure +open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.GroupoidLaws hiding (assoc) open import Cubical.Data.Sigma @@ -101,3 +102,27 @@ congIdLeft≡congIdRight _·G_ -G_ 0G rUnitG lUnitG r≡l p = ∙∙ cong₂ (λ x y → x ∙∙ p ∙∙ y) (sym r≡l) (cong sym (sym r≡l)) ∙∙ λ i → (λ j → rUnitG 0G (~ i ∧ j)) ∙∙ cong (λ x → rUnitG x (~ i)) p ∙∙ λ j → rUnitG 0G (~ i ∧ ~ j)) ∙∙ sym (rUnit (congL _·G_ p)) + +·GroupAutomorphismL : ∀ {ℓ} (G : Group ℓ) (g : fst G) → Iso (fst G) (fst G) +Iso.fun (·GroupAutomorphismL G g) = GroupStr._·_ (snd G) g +Iso.inv (·GroupAutomorphismL G g) = GroupStr._·_ (snd G) (GroupStr.inv (snd G) g) +Iso.rightInv (·GroupAutomorphismL G g) h = + GroupStr.·Assoc (snd G) _ _ _ + ∙ cong₂ (GroupStr._·_ (snd G)) (GroupStr.·InvR (snd G) g) refl + ∙ GroupStr.·IdL (snd G) h +Iso.leftInv (·GroupAutomorphismL G g) h = + GroupStr.·Assoc (snd G) _ _ _ + ∙ cong₂ (GroupStr._·_ (snd G)) (GroupStr.·InvL (snd G) g) refl + ∙ GroupStr.·IdL (snd G) h + +·GroupAutomorphismR : ∀ {ℓ} (G : Group ℓ) (g : fst G) → Iso (fst G) (fst G) +Iso.fun (·GroupAutomorphismR G g) x = GroupStr._·_ (snd G) x g +Iso.inv (·GroupAutomorphismR G g) x = GroupStr._·_ (snd G) x (GroupStr.inv (snd G) g) +Iso.rightInv (·GroupAutomorphismR G g) h = + sym (GroupStr.·Assoc (snd G) _ _ _) + ∙ cong₂ (GroupStr._·_ (snd G)) refl (GroupStr.·InvL (snd G) g) -- + ∙ GroupStr.·IdR (snd G) h +Iso.leftInv (·GroupAutomorphismR G g) h = + sym (GroupStr.·Assoc (snd G) _ _ _) + ∙ cong₂ (GroupStr._·_ (snd G)) refl (GroupStr.·InvR (snd G) g) -- + ∙ GroupStr.·IdR (snd G) h diff --git a/Cubical/Algebra/Group/QuotientGroup.agda b/Cubical/Algebra/Group/QuotientGroup.agda index 3c5632f94f..5a1c9d9061 100644 --- a/Cubical/Algebra/Group/QuotientGroup.agda +++ b/Cubical/Algebra/Group/QuotientGroup.agda @@ -11,9 +11,11 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Structure open import Cubical.Foundations.Powerset open import Cubical.Foundations.GroupoidLaws hiding (assoc) +open import Cubical.Foundations.Function open import Cubical.Data.Sigma open import Cubical.HITs.SetQuotients.Base renaming (_/_ to _/s_) open import Cubical.HITs.SetQuotients.Properties +open import Cubical.HITs.PropositionalTruncation as PT hiding (rec; elim) open import Cubical.Relation.Binary.Base open import Cubical.Algebra.Group.Base @@ -94,8 +96,16 @@ module _ (G' : Group ℓ) (H' : Subgroup G') (Hnormal : isNormal H') where ·/H-invr = elimProp (λ x → squash/ _ _) λ x → cong [_] (·InvR x) asGroup : Group ℓ - asGroup = makeGroup-right 1/H _·/H_ inv/H squash/ ·/H-assoc ·/H-rid ·/H-invr - + fst asGroup = G/H + GroupStr.1g (snd asGroup) = [ 1g ] + GroupStr._·_ (snd asGroup) = _·/H_ + GroupStr.inv (snd asGroup) = inv/H + GroupStr.isGroup (snd asGroup) = isGrp + where + opaque + isGrp : IsGroup [ 1g ] _·/H_ inv/H + isGrp = GroupStr.isGroup (makeGroup-right 1/H _·/H_ inv/H squash/ + ·/H-assoc ·/H-rid ·/H-invr .snd) _/_ : (G : Group ℓ) → (H : NormalSubgroup G) → Group ℓ G / H = asGroup G (H .fst) (H .snd) @@ -134,3 +144,68 @@ module _ {G' : Group ℓ} (H' : NormalSubgroup G') Iso.leftInv (fst trivialRelIso) _ = refl snd trivialRelIso = makeIsGroupHom λ _ _ → refl + +Hom/ : ∀ {ℓ ℓ'} {G : Group ℓ} {H : Group ℓ'} + {G' : NormalSubgroup G} {H' : NormalSubgroup H} + → (ϕ : GroupHom G H) + → (ϕ' : (a b : _) (r : (G ~ G' .fst) (G' .snd) a b) + → (H ~ H' .fst) (H' .snd) (fst ϕ a) (fst ϕ b)) + → GroupHom (G / G') (H / H') +fst (Hom/ ϕ ϕ') = + elim (λ _ → squash/) (λ x → [ fst ϕ x ]) λ a b r → eq/ _ _ (ϕ' a b r) +snd (Hom/ ϕ ϕ') = + makeIsGroupHom (elimProp2 (λ _ _ → squash/ _ _) + λ a b → cong [_] (IsGroupHom.pres· (snd ϕ) a b)) + +Hom/Iso : ∀ {ℓ ℓ'} {G : Group ℓ} {H : Group ℓ'} + {G' : NormalSubgroup G} {H' : NormalSubgroup H} + → (ϕ : GroupIso G H) + → (ϕ' : (a b : _) (r : (G ~ G' .fst) (G' .snd) a b) + → (H ~ H' .fst) (H' .snd) (Iso.fun (fst ϕ) a) (Iso.fun (fst ϕ) b)) + (ψ' : (a b : _) (r : (H ~ H' .fst) (H' .snd) a b) + → (G ~ G' .fst) (G' .snd) (Iso.inv (fst ϕ) a) (Iso.inv (fst ϕ) b)) + → GroupIso (G / G') (H / H') +Iso.fun (fst (Hom/Iso {G' = G'} {H' = H'} ϕ ϕ' ψ')) = + fst (Hom/ {G' = G'} {H' = H'} (GroupIso→GroupHom ϕ) ϕ') +Iso.inv (fst (Hom/Iso {G' = G'} {H' = H'} ϕ ϕ' ψ')) = + fst (Hom/ {G' = H'} {H' = G'} (GroupIso→GroupHom (invGroupIso ϕ)) ψ') +Iso.rightInv (fst (Hom/Iso ϕ ϕ' ψ')) = + elimProp (λ _ → squash/ _ _) λ a → cong [_] (Iso.rightInv (fst ϕ) a) +Iso.leftInv (fst (Hom/Iso ϕ ϕ' ψ')) = + elimProp (λ _ → squash/ _ _) λ a → cong [_] (Iso.leftInv (fst ϕ) a) +snd (Hom/Iso {G' = G'} {H' = H'} ϕ ϕ' ψ') = + makeIsGroupHom λ x y + → IsGroupHom.pres· (snd (Hom/ {G' = G'} {H' = H'} + (GroupIso→GroupHom ϕ) ϕ')) x y + +Hom/ImIso : ∀ {ℓ ℓ'} {G H : Group ℓ'} (ϕ : GroupHom G H) + {G' H' : Group ℓ} (ψ : GroupHom G' H') + {ϕ' : isNormal (imSubgroup ϕ)} {ψ' : isNormal (imSubgroup ψ)} + (eG : GroupIso G G') (eH : GroupIso H H') + (e~ : (g : fst G) + → fst ψ (Iso.fun (fst eG) g) ≡ + Iso.fun (fst eH) (ϕ .fst g)) + → GroupIso (H / (imSubgroup ϕ , ϕ')) (H' / (imSubgroup ψ , ψ')) +Hom/ImIso {G = G} {H} ϕ {G'} {H'} ψ {ϕ'} {ψ'} eG eH e∼ = + (fst main) + , makeIsGroupHom λ x y → IsGroupHom.pres· (snd main) x y + where + -- Faster type checking this way... + main = Hom/Iso {G = H} {H = H'} + {G' = (imSubgroup ϕ , ϕ')} {H' = (imSubgroup ψ , ψ')} eH + (λ a b → PT.map (uncurry λ x p + → (Iso.fun (fst eG) x) , e∼ x + ∙ cong (Iso.fun (fst eH)) p + ∙ IsGroupHom.pres· (snd eH) _ _ + ∙ cong₂ (GroupStr._·_ (snd H')) + refl (IsGroupHom.presinv (snd eH) _))) + λ a b → PT.map (uncurry λ x p + → (Iso.inv (fst eG) x) + , (sym ((cong₂ (GroupStr._·_ (snd H)) + refl (sym (IsGroupHom.presinv (snd (invGroupIso eH)) b)) + ∙ sym (IsGroupHom.pres· (snd (invGroupIso eH)) + a (GroupStr.inv (H' .snd) b)) + ∙ cong (Iso.inv (fst eH)) (sym p) + ∙ cong (Iso.inv (fst eH) ∘ fst ψ) (sym (Iso.rightInv (fst eG) x))) + ∙ cong (Iso.inv (fst eH)) (e∼ (Iso.inv (fst eG) x)) + ∙ Iso.leftInv (fst eH) _))) diff --git a/Cubical/Algebra/Group/Subgroup.agda b/Cubical/Algebra/Group/Subgroup.agda index bb44242727..69349c8d4c 100644 --- a/Cubical/Algebra/Group/Subgroup.agda +++ b/Cubical/Algebra/Group/Subgroup.agda @@ -185,6 +185,10 @@ module _ {G H : Group ℓ} (ϕ : GroupHom G H) where ((x H.· y) H.· H.inv x) ≡⟨ sym (H.·Assoc x y (H.inv x)) ⟩ x H.· y H.· H.inv x ∎ )} + imNormalSubgroup : ((x y : ⟨ H ⟩) → x H.· y ≡ y H.· x) → NormalSubgroup H + fst (imNormalSubgroup _) = imSubgroup + snd (imNormalSubgroup comm) = isNormalIm comm + kerSubset : ℙ ⟨ G ⟩ kerSubset x = isInKer ϕ x , isPropIsInKer ϕ x diff --git a/Cubical/Algebra/Group/ZAction.agda b/Cubical/Algebra/Group/ZAction.agda index 684bf7c77a..79d069cb75 100644 --- a/Cubical/Algebra/Group/ZAction.agda +++ b/Cubical/Algebra/Group/ZAction.agda @@ -642,7 +642,7 @@ GroupEquivℤ/abs-gen G H L = λ f g ex → GroupIso→GroupEquiv (GroupIsoℤ/abs f L g ex)) -- for type checking reasons, let's also do it with an abstract type -abstract +opaque abstractℤGroup/_ : ℕ → Group₀ abstractℤGroup/_ n = ℤGroup/ n @@ -662,7 +662,7 @@ GroupEquiv-abstractℤ/abs-gen : (G H L : Group₀) → GroupEquiv (abstractℤGroup/_ n) L GroupEquiv-abstractℤ/abs-gen G H L e r f g ex n p = main where - abstract + opaque main : GroupEquiv (abstractℤGroup/_ n) L main = transport (λ i diff --git a/Cubical/CW/Approximation.agda b/Cubical/CW/Approximation.agda index a331b1eefb..c4b819a35e 100644 --- a/Cubical/CW/Approximation.agda +++ b/Cubical/CW/Approximation.agda @@ -18,6 +18,7 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Function open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Nat.Order open import Cubical.Data.Fin.Inductive.Base open import Cubical.Data.Fin.Inductive.Properties open import Cubical.Data.Sigma @@ -26,6 +27,7 @@ open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Sequence open import Cubical.Data.FinSequence open import Cubical.Data.Nat.Order.Inductive +open import Cubical.Data.Sum as ⊎ open import Cubical.HITs.SequentialColimit open import Cubical.HITs.PropositionalTruncation as PT hiding (elimFin) @@ -146,7 +148,7 @@ module _ (C : CWskel ℓ) (D : CWskel ℓ') (f : realise C → realise D) where invEq propTrunc≃Trunc1 (invEq (_ , InductiveFinSatAC 1 (CWskel-fields.card C (suc m)) _) λ a → fst propTrunc≃Trunc1 - (sphereToTrunc m λ y → + (sphereToTrunc (suc m) λ y → TR.map fst (isConnectedCong _ _ (isConnected-CW↪∞ (suc (suc m)) D) (sym (push _) ∙ (fh (CWskel-fields.α C (suc m) (a , y)) @@ -190,7 +192,7 @@ module _ (C : CWskel ℓ) (D : CWskel ℓ') (f : realise C → realise D) where (fib-f-r x)) ∥₁ mere-fib-f-coh = invEq propTrunc≃Trunc1 (invEq (_ , InductiveFinSatAC 1 (card (suc m)) _) - λ a → fst propTrunc≃Trunc1 (sphereToTrunc m + λ a → fst propTrunc≃Trunc1 (sphereToTrunc (suc m) (sphereElim' m (λ x → isOfHLevelRetractFromIso m (invIso (PathPIdTruncIso (suc m))) @@ -269,6 +271,129 @@ CWmap→finCellMap C D f m = , →FinSeqColimHomotopy _ _ (g flast .snd)}) (approx C D f m) +-- Version for finite CW complexes +finCWmap→CellMap : ∀ {ℓ ℓ'} (n : ℕ) (C : finCWskel ℓ n) (D : CWskel ℓ') + (f : realise (finCWskel→CWskel n C) → realise D) + → ∃[ ϕ ∈ cellMap (finCWskel→CWskel n C) D ] + realiseCellMap ϕ ≡ f +finCWmap→CellMap n C D f = + PT.map (λ {(ϕ , p) → ψ ϕ (funExt⁻ p) + , funExt λ x + → subst (λ x → realiseCellMap (ψ ϕ (funExt⁻ p)) x ≡ f x) + (Iso.rightInv (converges→ColimIso + {seq = realiseSeq (finCWskel→CWskel n C)} n (C .snd .snd)) x) + (cong (incl {n = n}) + (silly ϕ (funExt⁻ p) _) + ∙ funExt⁻ p (fincl (n , (<ᵗsucm {n})) + (Iso.inv (converges→ColimIso n (C .snd .snd)) x)))}) + (CWmap→finCellMap + (finCWskel→CWskel n C) D f n) + where + open SequenceMap renaming (map to smap) + open FinSequenceMap + module _ (ϕ : finCellMap n (finCWskel→CWskel n C) D) + (ϕid : (x : FinSeqColim n + (realiseSeq (finCWskel→CWskel n C))) + → FinSeqColim→Colim n + (finCellMap→FinSeqColim (finCWskel→CWskel n C) D ϕ x) + ≡ f (FinSeqColim→Colim n x)) where + -- ψm' : (m : ℕ) → fst (finCWskel→CWskel n C) m → fst D m + -- ψm' m with (m ≟ suc n) + -- ... | lt q = fmap ϕ (m , q) + -- ... | eq q = fmap ϕ (m , {!!}) + -- ... | gt q = {!!} + + C≃ : (k : ℕ) → fst C n ≃ fst C (k +ℕ n) + C≃ zero = idEquiv _ + C≃ (suc k) = compEquiv (C≃ k) (_ , snd (snd C) k) + + C→D : (k : ℕ) → fst C (k +ℕ n) → fst D (k +ℕ n) + C→D k = CW↪Iterate D n k + ∘ fmap ϕ (n , <ᵗsucm {n}) + ∘ invEq (C≃ k) + + C→D-cellular : (k : ℕ) (x : fst (finCWskel→CWskel n C) (k +ℕ n)) + → C→D (suc k) (CW↪ (finCWskel→CWskel n C) (k +ℕ n) x) + ≡ CW↪ D (k +ℕ n) (C→D k x) + C→D-cellular k x = + cong (CW↪ D (k +ℕ n) ∘ CW↪Iterate D n k ∘ fmap ϕ (n , <ᵗsucm)) + (cong (invEq (C≃ k)) (retEq (_ , snd (snd C) k) x)) + + mainlem : ∀ {ℓ ℓ'} (D : ℕ → Type ℓ) (C : ℕ → Type ℓ') (n : ℕ) + → (C→D : (k : ℕ) → C (k +ℕ n) → D (k +ℕ n)) + → (D↑ : (n : ℕ) → D n → D (suc n)) + → (C↑ : (n : ℕ) → C n → C (suc n)) + → (cm : (k : ℕ) (x : _) + → C→D (suc k) (C↑ (k +ℕ n) x) + ≡ D↑ (k +ℕ n) (C→D k x)) + (r : ℕ) (k : ℕ) (p4 : r +ℕ n ≡ k) + (x : C k) (z : ℕ) (p0 : suc r ≡ z) (p1 : suc k ≡ z +ℕ n) + (m : ℕ) (p3 : r +ℕ n ≡ m) + (p2 : z +ℕ n ≡ suc m) + → D↑ m (subst D p3 (C→D r (subst C (sym p4) x))) + ≡ subst D p2 (C→D z (subst C p1 (C↑ k x))) + mainlem C D n C→D D↑ C↑ cm r = + J> λ x → J> λ p1 → J> λ p2 + → cong (D↑ (r +ℕ n)) + (transportRefl _ ∙ cong (C→D r) (transportRefl _)) + ∙ sym ((λ i → subst C (isSetℕ _ _ p2 refl i) + (C→D (suc r) (subst D (isSetℕ _ _ p1 refl i) + (C↑ (r +ℕ n) x)))) + ∙ transportRefl _ + ∙ cong (C→D (suc r)) (transportRefl _) + ∙ cm r x) + + ψm : (m : ℕ) → fst (finCWskel→CWskel n C) m → fst D m + ψm m with (Dichotomyℕ n m) + ... | inl q = subst (fst D) (snd q) + ∘ C→D (fst q) + ∘ subst (fst C) (sym (snd q)) + ... | inr q = fmap ϕ (m , <ᵗ-trans (<→<ᵗ q) (<ᵗsucm {n})) + + ψ : cellMap (finCWskel→CWskel n C) D + smap ψ = ψm + comm ψ m x with (Dichotomyℕ n m) | Dichotomyℕ n (suc m) + ... | inl n≤m | inl n≤sucm = + mainlem (fst D) (fst C) n C→D (CW↪ D) (CW↪ (finCWskel→CWskel n C)) + C→D-cellular _ _ _ _ _ + (inj-+m {m = n} (cong suc (snd n≤m) + ∙ sym (cong predℕ (cong suc (snd n≤sucm))))) _ _ _ _ + ... | inl n≤m | inr n>sucm = + ⊥.rec (<-asym (<≤-trans n>sucm n≤m) (1 , refl)) + ... | inr n>m | inl (zero , n≤sucm) = + (cong (CW↪ D m) + (cong (λ p → fmap ϕ (m , p) x) (isProp<ᵗ _ _)) + ∙ fcomm ϕ (m , <→<ᵗ n>m) x) + ∙ cong (λ p → fmap ϕ (suc m , p) c) (isProp<ᵗ _ _) + ∙ λ j → transp (λ i → fst D (n≤sucm (i ∨ ~ j))) (~ j) + (fmap ϕ (n≤sucm (~ j) , + isProp→PathP {B = λ j → n≤sucm (~ j) <ᵗ suc n} + (λ _ → isProp<ᵗ) (<→<ᵗ n>m) <ᵗsucm j) + (transp (λ i → fst C (n≤sucm (~ i ∨ ~ j))) (~ j) + c)) + where + c = CW↪ (finCWskel→CWskel n C) m x + lem : n>m ≡ (0 , sym n≤sucm) + lem = isProp≤ _ _ + ... | inr n>m | inl (suc diff , n≤sucm) = + ⊥.rec (<-asym (<≤-trans (diff , +-suc diff n ∙ n≤sucm) n>m) (0 , refl)) + ... | inr n>m | inr n>sucm = + cong (CW↪ D m) + (funExt⁻ (cong (fmap ϕ) (ΣPathP (refl , (isProp<ᵗ _ _)))) x) + ∙ fcomm ϕ (m , <→<ᵗ n>m) x + ∙ funExt⁻ (cong (fmap ϕ) (ΣPathP (refl , (isProp<ᵗ _ _)))) _ + + silly : (x : _) → smap ψ n x ≡ fmap ϕ (n , <ᵗsucm {n}) x + silly x with (Dichotomyℕ n n) + ... | inl (zero , p) = + cong (λ p → subst (fst D) p (C→D zero (subst (fst C) (sym p) x))) + (isSetℕ _ _ p refl) + ∙ transportRefl _ + ∙ cong (C→D zero) (transportRefl x) + ... | inl (suc diff , p) = + ⊥.rec (¬m<ᵗm {n} (<→<ᵗ (diff , (+-suc diff n ∙ p)))) + ... | inr p = ⊥.rec (¬m<ᵗm (<→<ᵗ p)) + ---- Part 3. The (finite) cellular approximation theorem for cellular homotopies: ----- -- Given two (m)-finite cellular maps fₙ, gₙ : Cₙ → Dₙ agreeing on -- Dₘ ↪ D∞, there is a finite cellular homotopy fₙ ∼ₘ gₙ. @@ -520,7 +645,7 @@ module approx {C : CWskel ℓ} {D : CWskel ℓ'} Iso.inv propTruncTrunc1Iso (invEq (_ , InductiveFinSatAC 1 _ _) λ x → Iso.fun propTruncTrunc1Iso - (sphereToTrunc n' (fiber-cong²-hₙ₊₁-push∞ x))) + (sphereToTrunc (suc n') (fiber-cong²-hₙ₊₁-push∞ x))) module _ (q : (x : Fin (fst (snd C) (suc n'))) (y : S₊ n') → fiber (cong² x y) (hₙ₊₁-push∞ x y)) where diff --git a/Cubical/CW/Base.agda b/Cubical/CW/Base.agda index 52c4f73f53..21ca55fb88 100644 --- a/Cubical/CW/Base.agda +++ b/Cubical/CW/Base.agda @@ -71,6 +71,9 @@ yieldsFinCWskel n X = finCWskel : (ℓ : Level) → (n : ℕ) → Type (ℓ-suc ℓ) finCWskel ℓ n = Σ[ C ∈ (ℕ → Type ℓ) ] (yieldsFinCWskel n C) +isFinCWskel : ∀ {ℓ} (C : CWskel ℓ) → Type ℓ +isFinCWskel C = Σ[ m ∈ ℕ ] ((k : ℕ) → isEquiv (CW↪ C (k +ℕ m))) + finCWskel→CWskel : (n : ℕ) → finCWskel ℓ n → CWskel ℓ finCWskel→CWskel n C = fst C , fst (snd C) @@ -86,35 +89,66 @@ realise : CWskel ℓ → Type ℓ realise C = SeqColim (realiseSeq C) -- Finally: definition of CW complexes +hasCWskel : (X : Type ℓ) → Type (ℓ-suc ℓ) +hasCWskel {ℓ = ℓ} X = Σ[ X' ∈ CWskel ℓ ] X ≃ realise X' + isCW : (X : Type ℓ) → Type (ℓ-suc ℓ) -isCW {ℓ = ℓ} X = Σ[ X' ∈ CWskel ℓ ] X ≃ realise X' +isCW X = ∥ hasCWskel X ∥₁ CW : (ℓ : Level) → Type (ℓ-suc ℓ) -CW ℓ = Σ[ A ∈ Type ℓ ] ∥ isCW A ∥₁ +CW ℓ = Σ[ A ∈ Type ℓ ] (isCW A) CWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) -CWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isCW A) +CWexplicit ℓ = Σ[ A ∈ Type ℓ ] (hasCWskel A) + +CWexplicit→CWskel : ∀ {ℓ} → CWexplicit ℓ → CWskel ℓ +CWexplicit→CWskel C = fst (snd C) + +CWexplicit→CW : ∀ {ℓ} → CWexplicit ℓ → CW ℓ +CWexplicit→CW C = fst C , ∣ snd C ∣₁ -- Finite CW complexes -isFinCW : (X : Type ℓ) → Type (ℓ-suc ℓ) -isFinCW {ℓ = ℓ} X = +isFinIsCW : {X : Type ℓ} → hasCWskel X → Type ℓ +isFinIsCW X = Σ[ n ∈ ℕ ] (((k : ℕ) → isEquiv (CW↪ (X .fst) (k +ℕ n)))) + +hasFinCWskel : (X : Type ℓ) → Type (ℓ-suc ℓ) +hasFinCWskel {ℓ = ℓ} X = Σ[ m ∈ ℕ ] (Σ[ X' ∈ finCWskel ℓ m ] X ≃ realise (finCWskel→CWskel m X')) +isFinCW : (X : Type ℓ) → Type (ℓ-suc ℓ) +isFinCW {ℓ = ℓ} X = ∥ hasFinCWskel X ∥₁ + finCW : (ℓ : Level) → Type (ℓ-suc ℓ) -finCW ℓ = Σ[ A ∈ Type ℓ ] ∥ isFinCW A ∥₁ +finCW ℓ = Σ[ A ∈ Type ℓ ] (isFinCW A) finCW∙ : (ℓ : Level) → Type (ℓ-suc ℓ) -finCW∙ ℓ = Σ[ A ∈ Pointed ℓ ] ∥ isFinCW (fst A) ∥₁ +finCW∙ ℓ = Σ[ A ∈ Pointed ℓ ] (isFinCW (fst A)) finCWexplicit : (ℓ : Level) → Type (ℓ-suc ℓ) -finCWexplicit ℓ = Σ[ A ∈ Type ℓ ] (isFinCW A) +finCWexplicit ℓ = Σ[ A ∈ Type ℓ ] (hasFinCWskel A) -isFinCW→isCW : (X : Type ℓ) → isFinCW X → isCW X -isFinCW→isCW X (n , X' , str) = (finCWskel→CWskel n X') , str +hasFinCWskel→hasCWskel : (X : Type ℓ) → hasFinCWskel X → hasCWskel X +hasFinCWskel→hasCWskel X (n , X' , str) = (finCWskel→CWskel n X') , str finCW→CW : finCW ℓ → CW ℓ -finCW→CW (X , p) = X , PT.map (isFinCW→isCW X) p +finCW→CW (X , p) = X , PT.map (hasFinCWskel→hasCWskel X) p + +-- Pointed complexes (with basepoint in X₀) +CWskel∙ : ∀ {ℓ} (X : CWskel ℓ) → fst X 1 → (n : ℕ) → fst X (suc n) +CWskel∙ X x zero = x +CWskel∙ X x (suc n) = CW↪ X (suc n) (CWskel∙ X x n) + +CWskel∞∙ : ∀ {ℓ} (X : CWskel ℓ) → fst X 1 → (n : ℕ) → realise X +CWskel∞∙ X x₀ n = incl (CWskel∙ X x₀ n) + +CWskel∞∙Id : ∀ {ℓ} (X : CWskel ℓ) (x₀ : fst X 1) (n : ℕ) → CWskel∞∙ X x₀ n ≡ incl x₀ +CWskel∞∙Id X x₀ zero = refl +CWskel∞∙Id X x₀ (suc n) = sym (push (CWskel∙ X x₀ n)) ∙ CWskel∞∙Id X x₀ n +incl∙ : ∀ {ℓ} (X : CWskel ℓ) (x₀ : fst X 1) {n : ℕ} + → (fst X (suc n) , CWskel∙ X x₀ n) →∙ (realise X , incl x₀) +fst (incl∙ X x₀ {n = n}) = incl +snd (incl∙ X x₀ {n = n}) = CWskel∞∙Id X x₀ n -- morphisms _→ᶜʷ_ : CW ℓ → CW ℓ' → Type (ℓ-max ℓ ℓ') diff --git a/Cubical/CW/Connected.agda b/Cubical/CW/Connected.agda index 34bfaf5485..c85d2c85fb 100644 --- a/Cubical/CW/Connected.agda +++ b/Cubical/CW/Connected.agda @@ -8,12 +8,17 @@ is defined by saying that it has non-trivial cells only in dimension The main result is packaged up in makeConnectedCW. This says that the usual notion of connectedness in terms of truncations (merely) coincides with the above definition for CW complexes. + +It also contains a proof that of πₙ₊₂X is finitely presented for X an +(n+1)-connected CW complex -} module Cubical.CW.Connected where open import Cubical.CW.Base open import Cubical.CW.Properties +open import Cubical.CW.Instances.SphereBouquetMap +open import Cubical.CW.Subcomplex open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv @@ -51,6 +56,14 @@ open import Cubical.Axiom.Choice open import Cubical.Relation.Nullary open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.Group.Properties +open import Cubical.Homotopy.Group.PiCofibFinSphereBouquetMap + +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.AbGroup.Base +open import Cubical.Algebra.AbGroup.FinitePresentation open Sequence @@ -73,10 +86,20 @@ yieldsCombinatorialConnectedCWskel A n = connectedCWskel : (ℓ : Level) (n : ℕ) → Type (ℓ-suc ℓ) connectedCWskel ℓ n = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsConnectedCWskel X n) +connectedCWskel→CWskel : ∀ {ℓ} {n : ℕ} + → connectedCWskel ℓ n → CWskel ℓ +fst (connectedCWskel→CWskel sk) = fst sk +snd (connectedCWskel→CWskel sk) = fst (snd sk) + combinatorialConnectedCWskel : (ℓ : Level) (n : ℕ) → Type (ℓ-suc ℓ) combinatorialConnectedCWskel ℓ n = Σ[ X ∈ (ℕ → Type ℓ) ] (yieldsCombinatorialConnectedCWskel X n) +combinatorialConnectedCWskel→CWskel : ∀ {ℓ} {n : ℕ} + → combinatorialConnectedCWskel ℓ n → CWskel ℓ +fst (combinatorialConnectedCWskel→CWskel sk) = fst sk +snd (combinatorialConnectedCWskel→CWskel sk) = fst (snd sk) + isConnectedCW : ∀ {ℓ} (n : ℕ) → Type ℓ → Type (ℓ-suc ℓ) isConnectedCW {ℓ = ℓ} n A = Σ[ sk ∈ connectedCWskel ℓ n ] realise (_ , (snd sk .fst)) ≃ A @@ -85,6 +108,18 @@ isConnectedCW' : ∀ {ℓ} (n : ℕ) → Type ℓ → Type (ℓ-suc ℓ) isConnectedCW' {ℓ = ℓ} n A = Σ[ sk ∈ combinatorialConnectedCWskel ℓ n ] realise (_ , (snd sk .fst)) ≃ A +ConnectedCW : (ℓ : Level) (n : ℕ) → Type (ℓ-suc ℓ) +ConnectedCW ℓ n = Σ[ X ∈ Type ℓ ] isConnectedCW n X + +ConnectedCW→CWexplicit : ∀ {ℓ} {n : ℕ} → ConnectedCW ℓ n → CWexplicit ℓ +fst (ConnectedCW→CWexplicit (X , p , con)) = X +fst (fst (snd (ConnectedCW→CWexplicit (X , (Xsk , _ , _) , con)))) = Xsk +snd (fst (snd (ConnectedCW→CWexplicit (X , (Xsk , p , _) , con)))) = p +snd (snd (ConnectedCW→CWexplicit (X , (Xsk , _ , _) , con))) = invEquiv con + +ConnectedCW→CW : ∀ {ℓ} {n : ℕ} → ConnectedCW ℓ n → CW ℓ +ConnectedCW→CW X = CWexplicit→CW (ConnectedCW→CWexplicit X) + --- Goal: show that these two definitions coincide (note that indexing is off by 2) --- -- For the base case, we need to analyse α₀ : Fin n × S⁰ → X₁ (recall, -- indexing of skeleta is shifted by 1). In partiuclar, we will need @@ -622,7 +657,7 @@ isConnectedCW→Contr (suc n) sk (suc x , p) ind) makeConnectedCW : ∀ {ℓ} (n : ℕ) {C : Type ℓ} - → isCW C + → hasCWskel C → isConnected (suc (suc n)) C → ∥ isConnectedCW n C ∥₁ makeConnectedCW zero {C = C} (cwsk , e) cA = @@ -903,7 +938,7 @@ makeConnectedCW {ℓ = ℓ} (suc n) {C = C} (cwsk , eqv) cA with sphereVanish : (f : S₊ (suc n) → C4+n) → ∥ ((x : S₊ (suc n)) → f x ≡ C⋆) ∥₁ sphereVanish f = - sphereToTrunc (suc n) + sphereToTrunc (suc (suc n)) λ x → isConnectedPath 2+n isConnectedC4+n _ _ .fst pted = Iso.fun (PathIdTruncIso 2+n) @@ -1041,7 +1076,7 @@ makeConnectedCW {ℓ = ℓ} (suc n) {C = C} (cwsk , eqv) cA with invEq propTrunc≃Trunc1 (invEq (_ , InductiveFinSatAC _ _ _) λ x → fst propTrunc≃Trunc1 - (sphereToTrunc 2+n + (sphereToTrunc (suc 2+n) λ y → isConnectedInr-cofib∘inr (inl (inr (x , y))) .fst)) where isConnectedInr-cofibβ : @@ -1321,3 +1356,195 @@ makeConnectedCW {ℓ = ℓ} (suc n) {C = C} (cwsk , eqv) cA with ((λ m → C'-realise m ((m +ℕ 4+n) ≟ᵗ 3+n)) , (λ n → C'-realise-coh n _ _))) (invIso (SeqColimIso _ (4 +ℕ n)))))) (ind .snd) + +-- As a consequence, we can compute Xₘ for m small enough when X is an +-- n-connected CW complex. This is done in the following three theorems +open CWskel-fields +connectedCWContr : {ℓ : Level} (n m : ℕ) (l : m <ᵗ suc n) (X : Type ℓ) + (cwX : isConnectedCW n X) → isContr (fst (fst cwX) (suc m)) +connectedCWContr n zero l X cwX = + subst isContr + (cong (Lift ∘ Fin) (sym ((snd (fst cwX)) .snd .fst)) + ∙ sym (ua (compEquiv (CW₁-discrete (connectedCWskel→CWskel (fst cwX))) + LiftEquiv))) + (isOfHLevelLift 0 (inhProp→isContr fzero isPropFin1)) +connectedCWContr n (suc m) l X cwX = + subst isContr + (ua (compEquiv (isoToEquiv (PushoutEmptyFam + (¬Fin0 ∘ subst Fin cardₘ₊₁≡0 ∘ fst) + (¬Fin0 ∘ subst Fin cardₘ₊₁≡0))) + (invEquiv (e (connectedCWskel→CWskel (fst cwX)) (suc m))) + )) + (connectedCWContr n m (<ᵗ-trans l <ᵗsucm) X cwX) + where + cardₘ₊₁≡0 = snd (snd (snd (fst cwX))) m l + +connectedCW≃SphereBouquet : {ℓ : Level} (n : ℕ) (X : Type ℓ) (cwX : isConnectedCW n X) + → fst (fst cwX) (suc (suc n)) + ≃ SphereBouquet (suc n) (Fin (card (connectedCWskel→CWskel (fst cwX)) (suc n))) +connectedCW≃SphereBouquet n X cwX = + compEquiv + (e (connectedCWskel→CWskel (fst cwX)) (suc n)) + (compEquiv + (pushoutEquiv _ _ _ fst + (idEquiv _) + (isContr→≃Unit (connectedCWContr n n <ᵗsucm X cwX)) + (idEquiv _) + (λ _ _ → tt) + (λ i x → fst x)) + (compEquiv (isoToEquiv (Iso-cofibFst-⋁ (λ A → S₊∙ n))) + (pushoutEquiv _ _ _ _ (idEquiv _) (idEquiv _) + (Σ-cong-equiv-snd (λ _ → isoToEquiv (invIso (IsoSucSphereSusp n)))) + (λ _ _ → tt) (λ i x → x , IsoSucSphereSusp∙ n i)))) + + +open import Cubical.CW.Instances.Lift +module _ {ℓ : Level} (n : ℕ) (X : Type ℓ) (cwX : isConnectedCW n X) + (str : isCW (fst cwX .fst (suc (suc (suc n))))) where + + private + X₃₊ₙ = fst (fst cwX) (suc (suc (suc n))) + X₂₊ₙ = fst (fst cwX) (suc (suc n)) + Xˢᵏᵉˡ = connectedCWskel→CWskel (fst cwX) + + X₃₊ₙᶜʷ : CW ℓ + X₃₊ₙᶜʷ = X₃₊ₙ , str + + connectedCW≃CofibFinSphereBouquetMap : + ∃[ α ∈ FinSphereBouquetMap∙ + (card Xˢᵏᵉˡ (suc (suc n))) (card Xˢᵏᵉˡ (suc n)) n ] + (X₃₊ₙᶜʷ ≡ CWLift ℓ (SphereBouquet/ᶜʷ (fst α))) + connectedCW≃CofibFinSphereBouquetMap = + PT.rec squash₁ + (λ { (x , ptz , t) → + ∣ (≃∘α' x ptz t) + , (Σ≡Prop (λ _ → squash₁) + (isoToPath (compIso + (connectedCW≅CofibFinSphereBouquetMap x ptz t) + LiftIso))) ∣₁}) + lem + where + isConX₂₊ₙ : isConnected 2 X₂₊ₙ + isConX₂₊ₙ = + isConnectedRetractFromIso 2 + (equivToIso (connectedCW≃SphereBouquet n X cwX)) + (isConnectedSubtr' n 2 isConnectedSphereBouquet') + + lem : ∃[ x ∈ X₂₊ₙ ] + (((a : Fin (card Xˢᵏᵉˡ (suc (suc n)))) + → x ≡ α Xˢᵏᵉˡ (suc (suc n)) (a , ptSn (suc n))) + × (fst (connectedCW≃SphereBouquet n X cwX) x ≡ inl tt)) + lem = TR.rec (isProp→isSet squash₁) + (λ x₀ → TR.rec squash₁ + (λ pts → TR.rec squash₁ (λ F → ∣ x₀ , F , pts ∣₁) + (invEq (_ , InductiveFinSatAC 1 (card Xˢᵏᵉˡ (suc (suc n))) _) + λ a → isConnectedPath 1 isConX₂₊ₙ _ _ .fst)) + (isConnectedPath 1 (isConnectedSubtr' n 2 isConnectedSphereBouquet') + (fst (connectedCW≃SphereBouquet n X cwX) x₀) (inl tt) .fst)) + (fst isConX₂₊ₙ) + + module _ (x : X₂₊ₙ) + (pts : ((a : Fin (card Xˢᵏᵉˡ (suc (suc n)))) + → x ≡ α Xˢᵏᵉˡ (suc (suc n)) (a , ptSn (suc n)))) + (ptd : fst (connectedCW≃SphereBouquet n X cwX) x ≡ inl tt) where + α' : SphereBouquet (suc n) (Fin (card Xˢᵏᵉˡ (suc (suc n)))) → X₂₊ₙ + α' (inl tt) = x + α' (inr x) = α Xˢᵏᵉˡ (suc (suc n)) x + α' (push a i) = pts a i + + ≃∘α' : SphereBouquet∙ (suc n) (Fin (card Xˢᵏᵉˡ (suc (suc n)))) + →∙ SphereBouquet∙ (suc n) (Fin (card Xˢᵏᵉˡ (suc n))) + fst ≃∘α' = fst (connectedCW≃SphereBouquet n X cwX) ∘ α' + snd ≃∘α' = ptd + + connectedCW≅CofibFinSphereBouquetMap : + Iso X₃₊ₙ (cofib (fst ≃∘α')) + connectedCW≅CofibFinSphereBouquetMap = + compIso (equivToIso (compEquiv + (e Xˢᵏᵉˡ (suc (suc n))) + (pushoutEquiv _ _ _ _ + (idEquiv _) (connectedCW≃SphereBouquet n X cwX) (idEquiv _) + (λ i x → fst ≃∘α' (inr x)) + (λ i x → fst x)))) + (⋁-cofib-Iso (SphereBouquet∙ (suc n) (Fin (card Xˢᵏᵉˡ (suc n)))) ≃∘α') + +-- Proof that πₙ₊₂(X) is FP when X is (n+1)-connected +-- first: a proof of the result with some additional explicit assumptions +-- (which we later get for free up to propositional truncation) +module isFinitelyPresented-π'connectedCW-lemmas {ℓ : Level} + (X : Pointed ℓ) (n : ℕ) + (X' : isConnectedCW (1 +ℕ n) (typ X)) + (isConX' : isConnected 2 (X' .fst .fst (4 +ℕ n))) + (x : X' .fst .fst (4 +ℕ n)) + (x≡ : X' .snd .fst (CW↪∞ (connectedCWskel→CWskel (fst X')) (4 +ℕ n) x) + ≡ snd X) + where + private + Xˢᵏᵉˡ : CWskel _ + Xˢᵏᵉˡ = connectedCWskel→CWskel (fst X') + + e∞ = X' .snd + + X₄₊ₙ∙ : Pointed _ + fst X₄₊ₙ∙ = X' .fst .fst (4 +ℕ n) + snd X₄₊ₙ∙ = x + + firstEquiv : GroupEquiv (π'Gr (suc n) X₄₊ₙ∙) (π'Gr (suc n) X) + firstEquiv = + (connectedFun→π'Equiv (suc n) + (fst e∞ ∘ CW↪∞ Xˢᵏᵉˡ (4 +ℕ n) , x≡) + (isConnectedComp (fst e∞) (CW↪∞ Xˢᵏᵉˡ (4 +ℕ n)) _ + (isEquiv→isConnected _ (snd e∞) (4 +ℕ n)) + (isConnected-CW↪∞ (4 +ℕ n) Xˢᵏᵉˡ))) + + isFP-π'X₄₊ₙ : isFinitelyPresented (Group→AbGroup (π'Gr (suc n) X₄₊ₙ∙) + (π'-comm n)) + isFP-π'X₄₊ₙ = PT.rec squash₁ + (λ {(α , e) → PT.map + (λ pp → subst FinitePresentation + (cong (λ X → Group→AbGroup (π'Gr (suc n) X) (π'-comm n)) + (ΣPathP ((sym (cong fst e)) , pp))) + (GrIsoPresFinitePresentation + (invGroupIso (π'GrLiftIso ℓ (suc n))) + (hasFPπ'CofibFinSphereBouquetMap α))) + (lem α (cong fst e))}) + (connectedCW≃CofibFinSphereBouquetMap (suc n) (fst X) + X' (subCW (4 +ℕ n) (fst X , Xˢᵏᵉˡ , invEquiv e∞) .snd)) + where + ll = π'GrIso + lem : (α : FinSphereBouquetMap∙ + (card Xˢᵏᵉˡ (suc (suc (suc n)))) (card Xˢᵏᵉˡ (suc (suc n))) + (suc n)) + (e : fst X₄₊ₙ∙ ≡ Lift (cofib (fst α))) + → ∥ PathP (λ i → e (~ i)) (lift (inl tt)) x ∥₁ + lem α e = TR.rec squash₁ ∣_∣₁ (isConnectedPathP _ isConX' _ _ .fst) + + isFPX : isFinitelyPresented (Group→AbGroup (π'Gr (suc n) X) (π'-comm n)) + isFPX = + PT.map (λ fp → subst FinitePresentation (AbGroupPath _ _ .fst firstEquiv) fp) + isFP-π'X₄₊ₙ + +-- Main result +isFinitelyPresented-π'connectedCW : ∀ {ℓ} (X : Pointed ℓ) + (cwX : isCW (fst X)) (n : ℕ) (cX : isConnected (3 +ℕ n) (typ X)) + → isFinitelyPresented (Group→AbGroup (π'Gr (suc n) X) (π'-comm n)) +isFinitelyPresented-π'connectedCW X = + PT.rec (isPropΠ2 (λ _ _ → squash₁)) λ cwX n cX → + PT.rec squash₁ (λ a → + PT.rec squash₁ (λ b → + PT.rec squash₁ (λ c → + PT.rec squash₁ (isFPX X n a b c) + (TR.rec (isProp→isOfHLevelSuc (suc n) squash₁) ∣_∣₁ + (isConnectedPath _ cX _ _ .fst))) + (TR.rec (isOfHLevelSuc 1 squash₁) ∣_∣₁ (b .fst))) + ∣ connectedFunPresConnected 2 + {f = fst (snd a) ∘ CW↪∞ (connectedCWskel→CWskel (fst a)) (4 +ℕ n)} + (isConnectedSubtr' (suc n) 2 cX) + (isConnectedComp (fst (snd a)) _ _ + (isEquiv→isConnected _ (snd (snd a)) 2) + λ b → isConnectedSubtr' (suc (suc n)) 2 + ((isConnected-CW↪∞ (4 +ℕ n) + (connectedCWskel→CWskel (fst a))) b)) ∣₁) + (makeConnectedCW (1 +ℕ n) cwX cX) + where + open isFinitelyPresented-π'connectedCW-lemmas diff --git a/Cubical/CW/Homology.agda b/Cubical/CW/Homology/Base.agda similarity index 86% rename from Cubical/CW/Homology.agda rename to Cubical/CW/Homology/Base.agda index 0c2bfd7c43..e2ee70456f 100644 --- a/Cubical/CW/Homology.agda +++ b/Cubical/CW/Homology/Base.agda @@ -5,7 +5,7 @@ This file contains: 2. Construction of cellular homology, including funtoriality -} -module Cubical.CW.Homology where +module Cubical.CW.Homology.Base where open import Cubical.CW.Base open import Cubical.CW.Properties @@ -22,15 +22,19 @@ open import Cubical.Foundations.Function open import Cubical.Data.Nat open import Cubical.Data.Fin.Inductive.Base +open import Cubical.Data.Sigma +open import Cubical.Data.Int open import Cubical.Algebra.Group open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.Group.MorphismProperties open import Cubical.Algebra.Group.GroupPath +open import Cubical.Algebra.AbGroup.Base open import Cubical.Algebra.ChainComplex open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.HITs.SequentialColimit +open import Cubical.HITs.SetQuotients as SQ private variable @@ -138,7 +142,7 @@ H̃ᶜʷ (C , CWstr) n = coh CWstr where - coh : (Cskel Dskel Eskel : isCW C) (t : fst (H̃ˢᵏᵉˡ (Cskel .fst) n)) + coh : (Cskel Dskel Eskel : hasCWskel C) (t : fst (H̃ˢᵏᵉˡ (Cskel .fst) n)) → fst (fst (H̃ˢᵏᵉˡ→Equiv n (compEquiv (invEquiv (snd Dskel)) (snd Eskel)))) (fst (fst (H̃ˢᵏᵉˡ→Equiv n (compEquiv (invEquiv (snd Cskel)) (snd Dskel)))) t) ≡ fst (fst (H̃ˢᵏᵉˡ→Equiv n (compEquiv (invEquiv (snd Cskel)) (snd Eskel)))) t @@ -151,7 +155,7 @@ H̃ᶜʷ (C , CWstr) n = -- lemmas for functoriality private module _ {C : Type ℓ} {D : Type ℓ'} (f : C → D) (n : ℕ) where - right : (cwC : isCW C) (cwD1 : isCW D) (cwD2 : isCW D) + right : (cwC : hasCWskel C) (cwD1 : hasCWskel D) (cwD2 : hasCWskel D) → PathP (λ i → GroupHom (H̃ᶜʷ (C , ∣ cwC ∣₁) n) (H̃ᶜʷ (D , squash₁ ∣ cwD1 ∣₁ ∣ cwD2 ∣₁ i) n)) (H̃ˢᵏᵉˡ→ n (λ x → fst (snd cwD1) (f (invEq (snd cwC) x)))) @@ -162,7 +166,7 @@ private → cong (fst (snd cwD2)) (sym (retEq (snd cwD1) _)))) ∙ H̃ˢᵏᵉˡ→comp n _ _) - left : (cwC1 : isCW C) (cwC2 : isCW C) (cwD : isCW D) + left : (cwC1 : hasCWskel C) (cwC2 : hasCWskel C) (cwD : hasCWskel D) → PathP (λ i → GroupHom (H̃ᶜʷ (C , squash₁ ∣ cwC1 ∣₁ ∣ cwC2 ∣₁ i) n) (H̃ᶜʷ (D , ∣ cwD ∣₁) n)) (H̃ˢᵏᵉˡ→ n (λ x → fst (snd cwD) (f (invEq (snd cwC1) x)))) @@ -175,7 +179,7 @@ private → cong (fst (snd cwD) ∘ f) (retEq (snd cwC2) _)))) - left-right : (x y : isCW C) (v w : isCW D) → + left-right : (x y : hasCWskel C) (v w : hasCWskel D) → SquareP (λ i j → GroupHom (H̃ᶜʷ (C , squash₁ ∣ x ∣₁ ∣ y ∣₁ i) n) @@ -184,7 +188,7 @@ private left-right _ _ _ _ = isSet→SquareP (λ _ _ → isSetGroupHom) _ _ _ _ - H̃ᶜʷ→pre : (cwC : ∥ isCW C ∥₁) (cwD : ∥ isCW D ∥₁) + H̃ᶜʷ→pre : (cwC : isCW C) (cwD : isCW D) → GroupHom (H̃ᶜʷ (C , cwC) n) (H̃ᶜʷ (D , cwD) n) H̃ᶜʷ→pre = elim2→Set (λ _ _ → isSetGroupHom) @@ -242,3 +246,25 @@ snd (H̃ᶜʷ→Iso m e) = snd (H̃ᶜʷ→ m (fst e)) H̃ᶜʷ→Equiv : {C : CW ℓ} {D : CW ℓ'} (m : ℕ) (e : fst C ≃ fst D) → GroupEquiv (H̃ᶜʷ C m) (H̃ᶜʷ D m) H̃ᶜʷ→Equiv m e = GroupIso→GroupEquiv (H̃ᶜʷ→Iso m e) + +-- As abelian groups +H̃ˢᵏᵉˡ-comm : ∀ {ℓ} {n : ℕ} {X : CWskel ℓ} (x y : H̃ˢᵏᵉˡ X (suc n) .fst) + → GroupStr._·_ (H̃ˢᵏᵉˡ X (suc n) .snd) x y + ≡ GroupStr._·_ (H̃ˢᵏᵉˡ X (suc n) .snd) y x +H̃ˢᵏᵉˡ-comm = SQ.elimProp2 (λ _ _ → GroupStr.is-set (H̃ˢᵏᵉˡ _ _ .snd) _ _) + λ a b → cong [_] (Σ≡Prop (λ _ → isSetΠ (λ _ → isSetℤ) _ _) + (funExt λ _ → +Comm _ _)) + +H̃ᶜʷ-comm : ∀ {ℓ} {n : ℕ} (X : CW ℓ) (x y : H̃ᶜʷ X (suc n) .fst) + → GroupStr._·_ (H̃ᶜʷ X (suc n) .snd) x y + ≡ GroupStr._·_ (H̃ᶜʷ X (suc n) .snd) y x +H̃ᶜʷ-comm {n = n} = uncurry λ X + → PT.elim (λ _ → isPropΠ2 λ _ _ + → GroupStr.is-set (H̃ᶜʷ (X , _) (suc n) .snd) _ _) + λ x → H̃ˢᵏᵉˡ-comm + +H̃ˢᵏᵉˡAb : ∀ {ℓ} → CWskel ℓ → ℕ → AbGroup ℓ-zero +H̃ˢᵏᵉˡAb X n = Group→AbGroup (H̃ˢᵏᵉˡ X (suc n)) H̃ˢᵏᵉˡ-comm + +H̃ᶜʷAb : ∀ {ℓ} → CW ℓ → ℕ → AbGroup ℓ-zero +H̃ᶜʷAb X n = Group→AbGroup (H̃ᶜʷ X (suc n)) (H̃ᶜʷ-comm X) diff --git a/Cubical/CW/Homology/Groups/CofibFinSphereBouquetMap.agda b/Cubical/CW/Homology/Groups/CofibFinSphereBouquetMap.agda new file mode 100644 index 0000000000..8f2420a317 --- /dev/null +++ b/Cubical/CW/Homology/Groups/CofibFinSphereBouquetMap.agda @@ -0,0 +1,386 @@ +{-# OPTIONS --lossy-unification #-} +{- +File contains : a proof that Hₙ₊₁(cofib α) ≅ ℤ[e]/Im(deg α) for +α : ⋁ₐ Sⁿ → ⋁ₑ Sⁿ (with a and e finite sets) +-} +module Cubical.CW.Homology.Groups.CofibFinSphereBouquetMap where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.CW.Base +open import Cubical.CW.ChainComplex +open import Cubical.CW.Instances.SphereBouquetMap + +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Int +open import Cubical.Data.Bool +open import Cubical.Data.Fin.Inductive +open import Cubical.Data.Nat.Order.Inductive +open import Cubical.Data.Sigma +open import Cubical.Data.Empty as ⊥ + +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.Pushout +open import Cubical.HITs.SphereBouquet +open import Cubical.HITs.SphereBouquet.Degree +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.Wedge +open import Cubical.HITs.SetQuotients.Base renaming (_/_ to _/s_) +open import Cubical.HITs.SetQuotients.Properties as SQ + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup +open import Cubical.Algebra.Group.Subgroup +open import Cubical.Algebra.Group.QuotientGroup + +open import Cubical.Relation.Nullary + +open Iso + +module _ {c1 c2 : ℕ} {n : ℕ} (α : FinSphereBouquetMap c1 c2 n) where + private + α∙ : ∥ α (inl tt) ≡ inl tt ∥₁ + α∙ = isConnectedSphereBouquet _ + + ℤ[]/ImSphereMap : Group₀ + ℤ[]/ImSphereMap = (AbGroup→Group ℤ[Fin c2 ]) + / (imSubgroup (bouquetDegree α) + , isNormalIm (bouquetDegree α) + λ _ _ → AbGroupStr.+Comm (snd ℤ[Fin c2 ]) _ _) + + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-inv : + (p : Trichotomyᵗ (suc n) (suc n)) (q : Trichotomyᵗ (suc n) (suc (suc n))) + → ℤ[Fin c2 ] .fst → (Fin (SphereBouquet/CardGen c1 c2 (suc n) p q) → ℤ) + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-inv (lt x) q = λ _ _ → 0 + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-inv (eq x) (lt y) = λ f → f + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-inv (eq x) (eq y) = λ f → f + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-inv (eq x) (gt y) = λ f → f + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-inv (gt x) q = λ _ _ → 0 + + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-fun : + (p : Trichotomyᵗ (suc n) (suc n)) (q : Trichotomyᵗ (suc n) (suc (suc n))) + → (Fin (SphereBouquet/CardGen c1 c2 (suc n) p q) → ℤ) → ℤ[Fin c2 ] .fst + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-fun (lt x) q = λ _ _ → 0 + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-fun (eq x) (lt y) = λ f → f + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-fun (eq x) (eq y) = λ f → f + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-fun (eq x) (gt y) = λ f → f + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-fun (gt x) q = λ _ _ → 0 + + Iso-HₙSphereBouquetⁿ/-ℤ[]/ImSphereMap-pre : + (p : Trichotomyᵗ (suc n) (suc n)) (q : Trichotomyᵗ (suc n) (suc (suc n))) + → Iso (Fin (SphereBouquet/CardGen c1 c2 (suc n) p q) → ℤ) (ℤ[Fin c2 ] .fst) + fun (Iso-HₙSphereBouquetⁿ/-ℤ[]/ImSphereMap-pre p q) = + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-fun p q + inv (Iso-HₙSphereBouquetⁿ/-ℤ[]/ImSphereMap-pre p q) = + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-inv p q + rightInv (Iso-HₙSphereBouquetⁿ/-ℤ[]/ImSphereMap-pre (lt x) q) f = + ⊥.rec (¬m<ᵗm x) + rightInv (Iso-HₙSphereBouquetⁿ/-ℤ[]/ImSphereMap-pre (eq x) (lt y)) f = refl + rightInv (Iso-HₙSphereBouquetⁿ/-ℤ[]/ImSphereMap-pre (eq x) (eq y)) f = refl + rightInv (Iso-HₙSphereBouquetⁿ/-ℤ[]/ImSphereMap-pre (eq x) (gt y)) f = refl + rightInv (Iso-HₙSphereBouquetⁿ/-ℤ[]/ImSphereMap-pre (gt x) q) f = + ⊥.rec (¬m<ᵗm x) + leftInv (Iso-HₙSphereBouquetⁿ/-ℤ[]/ImSphereMap-pre (lt x) q) f = + ⊥.rec (¬m<ᵗm x) + leftInv (Iso-HₙSphereBouquetⁿ/-ℤ[]/ImSphereMap-pre (eq x) (lt y)) f = refl + leftInv (Iso-HₙSphereBouquetⁿ/-ℤ[]/ImSphereMap-pre (eq x) (eq y)) f = refl + leftInv (Iso-HₙSphereBouquetⁿ/-ℤ[]/ImSphereMap-pre (eq x) (gt y)) f = refl + leftInv (Iso-HₙSphereBouquetⁿ/-ℤ[]/ImSphereMap-pre (gt x) q) f = + ⊥.rec (¬m<ᵗm x) + + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-fun-hom : + (p : Trichotomyᵗ (suc n) (suc n)) (q : Trichotomyᵗ (suc n) (suc (suc n))) + (x y : Fin (SphereBouquet/CardGen c1 c2 (suc n) p q) → ℤ) + → HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-fun p q (λ z → x z + y z) + ≡ λ z → HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-fun p q x z + + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-fun p q y z + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-fun-hom (lt _) _ _ _ = refl + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-fun-hom (eq _) (lt _) _ _ = refl + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-fun-hom (eq _) (eq _) _ _ = refl + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-fun-hom (eq _) (gt _) _ _ = refl + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-fun-hom (gt _) _ _ _ = refl + + private + module _ (x : n ≡ n) where + EQ = compEquiv (SphereBouquet/EqBottomMainGen c1 c2 α) + (pathToEquiv (λ i → cofib {A = Fin c2 × S₊ (x (~ i))} fst)) + + FIs : Iso _ _ + FIs = BouquetIso-gen (suc n) c2 (λ _ → tt) EQ + + F1 = fun (sphereBouquetSuspIso {A = Fin c2} {n = (suc n)}) + F2' = fun FIs + F2 = suspFun F2' + F3 : Susp (SphereBouquet (suc n) (Fin c2)) → Susp (cofib (invEq EQ ∘ inl)) + F3 = suspFun inr + F4 = δ-pre (invEq (SphereBouquet/EqTopGen' c1 c2 (suc n) α (cong suc x)) + ∘ inl) + F5 = F1 ∘ F2 ∘ F3 ∘ F4 + F5' = F1 ∘ F2 ∘ F3 + + bouquetSusp→Charac : (a b : Fin c2 → ℤ) (t : α (inl tt) ≡ inl tt) + (x : _) → refl ≡ x + → F1 x ∘ F2 x ∘ F3 x ∘ F4 x + ∘ inv (BouquetIso-gen (suc (suc n)) c1 + (λ w → α (inr (fst w , subst (S₊ ∘ suc) x (snd w)))) + (SphereBouquet/EqTopGen' c1 c2 (suc n) α (cong suc x))) + ≡ bouquetSusp→ α + bouquetSusp→Charac a b αpt = + J> funExt λ { (inl x) → refl + ; (inr (x , north)) → refl + ; (inr (x , south)) → refl + ; (inr (x , merid a i)) j → main αpt x a j i + ; (push a i) → refl} + where + F2'≡id : (a : _) → F2' refl (inr a) ≡ a + F2'≡id a = + cong (Pushout→Bouquet (suc n) c2 (λ _ → tt) + (compEquiv (SphereBouquet/EqBottomMainGen c1 c2 α) + (pathToEquiv (λ i → cofib {A = Fin c2 × S₊ n} fst)))) + (transportRefl (fst (SphereBouquet/EqBottomMainGen c1 c2 α) a)) + ∙ SphereBouquet/EqBottomMainGenLem n α {e = EQ refl} a + + main : (α (inl tt) ≡ inl tt) → (x : Fin c1) (a : S₊ (suc n)) + → cong (F5 refl) + (push (α (inr (x , transport refl a))) + ∙∙ (λ i → inr (invEq (SphereBouquet/EqTopGen' c1 c2 (suc n) α refl) + ((push (x , a) ∙ sym (push (x , ptSn (suc n)))) i))) + ∙∙ sym (push (α (inr (x , transport refl (ptSn (suc n))))))) + ≡ Bouquet→ΩBouquetSusp (Fin c2) (λ _ → S₊∙ (suc n)) (α (inr (x , a))) + main apt x a = + cong-∙∙ (F5 refl) _ _ _ + ∙ cong₃ _∙∙_∙∙_ + (λ i → Bouquet→ΩBouquetSusp (Fin c2) (λ _ → S₊∙ (suc n)) + (F2' refl (inr (α (inr (x , transportRefl a i)))))) + refl + ((λ j i + → F5 refl (push ((cong α + ((λ j → inr (x , transportRefl (ptSn (suc n)) j)) + ∙ sym (push x)) ∙ apt) j) (~ i)))) + ∙ (sym (compPath≡compPath' _ _) ∙ sym (rUnit _)) + ∙ cong (Bouquet→ΩBouquetSusp (Fin c2) (λ _ → S₊∙ (suc n))) + (F2'≡id (α (inr (x , a)))) + + module _ + (a b : Fin (SphereBouquet/CardGen c1 c2 (suc n) + (suc n ≟ᵗ suc n) (suc n ≟ᵗ suc (suc n))) → ℤ) + (ak : ∂ (SphereBouquet/ˢᵏᵉˡ α) n .fst a ≡ (λ _ → 0)) + (bk : ∂ (SphereBouquet/ˢᵏᵉˡ α) n .fst b ≡ (λ _ → 0)) + (r : Σ[ t ∈ (Fin (SphereBouquet/CardGen c1 c2 (suc (suc n)) + (suc (suc n) ≟ᵗ suc n) (suc (suc n) ≟ᵗ suc (suc n))) → ℤ) ] + ∂ (SphereBouquet/ˢᵏᵉˡ α) (suc n) .fst t ≡ λ q → a q - b q) where + + pathlemma : Path (ℤ[]/ImSphereMap .fst) + [ HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-fun + (suc n ≟ᵗ suc n) (suc n ≟ᵗ suc (suc n)) a ] + [ HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-fun + (suc n ≟ᵗ suc n) (suc n ≟ᵗ suc (suc n)) b ] + pathlemma with (n ≟ᵗ n) | (n ≟ᵗ suc n) | (suc n ≟ᵗ n) + ... | lt x | st | ah = ⊥.rec (¬m<ᵗm x) + ... | eq x | lt y | lt z = ⊥.rec (¬-suc-n<ᵗn z) + ... | eq x | lt y | eq z = ⊥.rec (falseDichotomies.eq-eq (x , sym z)) + ... | eq x | lt y | gt z = + PT.rec (squash/ _ _) (λ apt → + eq/ _ _ ∣ (fst r) + , ((λ i → bouquetDegree α .fst (fst r)) + ∙ funExt⁻ (cong fst (bouquetDegreeSusp α)) (fst r) + ∙ λ i → bouquetDegree (bouquetSusp→Charac a b apt x + (isSetℕ _ _ refl x) (~ i)) .fst (fst r)) + ∙ snd r ∣₁) α∙ + ... | eq x | eq y | ah = ⊥.rec (falseDichotomies.eq-eq (x , y)) + ... | eq x | gt y | ah = ⊥.rec (¬-suc-n<ᵗn y) + ... | gt x | st | ah = refl + + + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap : + H̃ˢᵏᵉˡ (SphereBouquet/ˢᵏᵉˡ α) (suc n) .fst → ℤ[]/ImSphereMap .fst + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap = + SQ.rec squash/ + (λ f → [ HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-fun + (suc n ≟ᵗ suc n) (suc n ≟ᵗ suc (suc n)) (fst f) ]) + λ {(a , ak) (b , bk) → PT.elim (λ _ → squash/ _ _) + λ {(t , s) → pathlemma a b ak bk (t , cong fst s)}} + + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-hom : + (x y : H̃ˢᵏᵉˡ (SphereBouquet/ˢᵏᵉˡ α) (suc n) .fst) + → HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap + (GroupStr._·_ (H̃ˢᵏᵉˡ (SphereBouquet/ˢᵏᵉˡ α) (suc n) .snd) x y) + ≡ GroupStr._·_ (ℤ[]/ImSphereMap .snd) + (HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap x) + (HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap y) + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-hom = SQ.elimProp2 + (λ _ _ → squash/ _ _) + λ f g → cong [_] (HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-fun-hom + _ _ (fst f) (fst g)) + + ℤ[]/ImSphereMap→HₙSphereBouquetⁿ/ : + ℤ[]/ImSphereMap .fst → H̃ˢᵏᵉˡ (SphereBouquet/ˢᵏᵉˡ α) (suc n) .fst + ℤ[]/ImSphereMap→HₙSphereBouquetⁿ/ = + SQ.rec squash/ + (λ a → [ HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-inv + (suc n ≟ᵗ suc n) (suc n ≟ᵗ suc (suc n)) a + , cancel a ]) + λ {a b → PT.elim (λ _ → squash/ _ _) + λ {(r , k) → PT.rec (squash/ _ _) + (λ apt → eq/ _ _ + ∣ lem apt a b r k .fst + , Σ≡Prop (λ _ → isOfHLevelPath' 1 (isSetΠ (λ _ → isSetℤ)) _ _) + (lem apt a b r k .snd ) ∣₁) α∙}} + where + lem : (apt : α (inl tt) ≡ inl tt) (a b : Fin c2 → ℤ) (r : Fin c1 → ℤ) + → bouquetDegree α .fst r ≡ (λ x → a x + - b x) + → Σ[ w ∈ (Fin (SphereBouquet/CardGen c1 c2 (suc (suc n)) + (Trichotomyᵗ-suc (suc n ≟ᵗ n)) + ((Trichotomyᵗ-suc (Trichotomyᵗ-suc (n ≟ᵗ n))))) → ℤ) ] + ∂ (SphereBouquet/ˢᵏᵉˡ α) (suc n) .fst w + ≡ λ x → HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-inv + (suc n ≟ᵗ suc n) (suc n ≟ᵗ suc (suc n)) a x + - HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-inv + (suc n ≟ᵗ suc n) (suc n ≟ᵗ suc (suc n)) b x + lem apt a b r q with (n ≟ᵗ n) | (n ≟ᵗ suc n) | (suc n ≟ᵗ n) + ... | lt x | d | e = ⊥.rec (¬m<ᵗm x) + ... | eq x | lt y | lt z = ⊥.rec (¬-suc-n<ᵗn z) + ... | eq x | lt y | eq z = ⊥.rec (falseDichotomies.eq-eq (refl , sym z)) + ... | eq x | lt y | gt z = r + , ((funExt⁻ (cong fst + (cong bouquetDegree (bouquetSusp→Charac a b apt x (isSetℕ _ _ refl x)) + ∙ sym (bouquetDegreeSusp α))) r) ∙ q) + ... | eq x | eq y | e = ⊥.rec (falseDichotomies.eq-eq (refl , y)) + ... | eq x | gt y | e = ⊥.rec (⊥.rec (¬-suc-n<ᵗn y)) + ... | gt x | d | e = ⊥.rec (¬m<ᵗm x) + + pre∂-vanish : (x : _) + → preboundary.pre∂ (SphereBouquet/ˢᵏᵉˡ α) zero x ≡ inl tt + pre∂-vanish (inl x) = refl + pre∂-vanish (inr (x , base)) = refl + pre∂-vanish (inr (x , loop i)) j = help j i + where + help : + cong (preboundary.isoSuspBouquet (SphereBouquet/ˢᵏᵉˡ α) zero) + (cong (suspFun (preboundary.isoCofBouquet (SphereBouquet/ˢᵏᵉˡ α) zero)) + (cong (suspFun inr) (cong (δ 1 (SphereBouquet/ˢᵏᵉˡ α)) + (cong (preboundary.isoCofBouquetInv↑ (SphereBouquet/ˢᵏᵉˡ α) zero) + (λ i → inr (x , loop i)))))) + ≡ refl {x = inl tt} + help = cong-∙∙ + (preboundary.isoSuspBouquet (SphereBouquet/ˢᵏᵉˡ α) zero + ∘ suspFun (preboundary.isoCofBouquet (SphereBouquet/ˢᵏᵉˡ α) zero) + ∘ suspFun inr ∘ δ 1 (SphereBouquet/ˢᵏᵉˡ α)) _ _ _ + ∙ ∙∙lCancel (sym + (cong (preboundary.isoSuspBouquet (SphereBouquet/ˢᵏᵉˡ α) zero) + (merid (inr (fzero , false))))) + pre∂-vanish (push a i) = refl + + ∂-zero : ∂ (SphereBouquet/ˢᵏᵉˡ α) zero ≡ trivGroupHom + ∂-zero = cong bouquetDegree (funExt pre∂-vanish) ∙ bouquetDegreeConst _ _ _ + + ∂-vanish : (m : ℕ) → m <ᵗ suc n → ∂ (SphereBouquet/ˢᵏᵉˡ α) m ≡ trivGroupHom + ∂-vanish zero p = ∂-zero + ∂-vanish (suc m) p = Σ≡Prop (λ _ → isPropIsGroupHom _ _) + (funExt λ x → funExt λ y → ⊥.rec (¬bottom _ _ y)) + where + ¬bottom : (p : _)(q : _) → ¬ Fin (SphereBouquet/CardGen c1 c2 (suc m) p q) + ¬bottom (lt x) q = ¬Fin0 + ¬bottom (eq x) q = ⊥.rec (¬m<ᵗm (subst (_<ᵗ n) (cong predℕ x) p)) + ¬bottom (gt x) q = ⊥.rec (¬m<ᵗm (<ᵗ-trans x p)) + + cancel : (a : Fin c2 → ℤ) + → ∂ (SphereBouquet/ˢᵏᵉˡ α) n .fst + ((HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-inv (suc n ≟ᵗ suc n) + (suc n ≟ᵗ suc (suc n)) a)) + ≡ (λ _ → 0) + cancel a = funExt⁻ (cong fst (∂-vanish n <ᵗsucm)) + (HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-inv + (suc n ≟ᵗ suc n) (suc n ≟ᵗ suc (suc n)) a) + + -- Main result + GroupIso-Hₙ₊₁SphereBouquetⁿ/→ℤ[]/ImSphereMap : + GroupIso (H̃ˢᵏᵉˡ (SphereBouquet/ˢᵏᵉˡ α) (suc n)) ℤ[]/ImSphereMap + fun (fst GroupIso-Hₙ₊₁SphereBouquetⁿ/→ℤ[]/ImSphereMap) = + HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap + inv (fst GroupIso-Hₙ₊₁SphereBouquetⁿ/→ℤ[]/ImSphereMap) = + ℤ[]/ImSphereMap→HₙSphereBouquetⁿ/ + rightInv (fst GroupIso-Hₙ₊₁SphereBouquetⁿ/→ℤ[]/ImSphereMap) = + SQ.elimProp (λ _ → squash/ _ _) + λ f → cong [_] (rightInv (Iso-HₙSphereBouquetⁿ/-ℤ[]/ImSphereMap-pre + (suc n ≟ᵗ suc n) (suc n ≟ᵗ suc (suc n))) f) + leftInv (fst GroupIso-Hₙ₊₁SphereBouquetⁿ/→ℤ[]/ImSphereMap) = + SQ.elimProp (λ _ → squash/ _ _) + λ f → cong [_] (Σ≡Prop (λ _ → isSetΠ (λ _ → isSetℤ) _ _) + (leftInv (Iso-HₙSphereBouquetⁿ/-ℤ[]/ImSphereMap-pre + (suc n ≟ᵗ suc n) (suc n ≟ᵗ suc (suc n))) (fst f))) + snd GroupIso-Hₙ₊₁SphereBouquetⁿ/→ℤ[]/ImSphereMap = + makeIsGroupHom HₙSphereBouquetⁿ/→ℤ[]/ImSphereMap-hom + + fin→SphereBouquet/Cell : {n : ℕ} (p : _) (q : _) + → Fin c2 → Fin (SphereBouquet/CardGen c1 c2 {n = n} (suc n) p q) + fin→SphereBouquet/Cell (lt y) q x = ⊥.rec (¬m<ᵗm y) + fin→SphereBouquet/Cell (eq y) q x = x + fin→SphereBouquet/Cell (gt y) q x = ⊥.rec (¬m<ᵗm y) + +-- Description of genrators +opaque + inKerGenerator : {c1 c2 : ℕ} {n : ℕ} + (α : FinSphereBouquetMap c1 c2 n) (k : Fin c2) + → bouquetDegree (preboundary.pre∂ (SphereBouquet/ˢᵏᵉˡ α) n) .fst + (ℤFinGenerator (fin→SphereBouquet/Cell α + (suc n ≟ᵗ suc n) (suc n ≟ᵗ suc (suc n)) k)) + ≡ λ _ → 0 + inKerGenerator {n = n} α k = inKerAll α k _ + where + inKerAll : {c1 c2 : ℕ} {n : ℕ} + (α : FinSphereBouquetMap c1 c2 n) (k : Fin c2) (t : _) + → bouquetDegree (preboundary.pre∂ (SphereBouquet/ˢᵏᵉˡ α) n) .fst t + ≡ (λ _ → 0) + inKerAll {c1 = c1} {c2 = c2} {n = n} α k = + funExt⁻ (cong fst (agreeOnℤFinGenerator→≡ + {ϕ = bouquetDegree (preboundary.pre∂ (SphereBouquet/ˢᵏᵉˡ α) n)} + {trivGroupHom} + (BT n α))) + where + BT : (n : ℕ) (α : FinSphereBouquetMap c1 c2 n) + (x : Fin (preboundary.An+1 (SphereBouquet/ˢᵏᵉˡ α) n)) → + fst (bouquetDegree (preboundary.pre∂ (SphereBouquet/ˢᵏᵉˡ α) n)) + (ℤFinGenerator x) + ≡ λ _ → pos zero + BT n α with (n ≟ᵗ n) | (n ≟ᵗ suc n) + ... | lt x | q = ⊥.rec (¬m<ᵗm x) + BT zero α | eq x | lt y = + λ q → funExt λ { (zero , snd₁) + → sumFinℤId _ (λ r → ·Comm (ℤFinGenerator q r) (pos zero)) + ∙ sumFinℤ0 _} + BT (suc n) α | eq x | lt y = λ q → funExt λ x → ⊥.rec (snd x) + ... | eq x | eq y = ⊥.rec (¬m<ᵗm (subst (_<ᵗ suc n) y <ᵗsucm)) + ... | eq x | gt y = ⊥.rec (¬-suc-n<ᵗn y) + ... | gt x | q = ⊥.rec (¬m<ᵗm x) + +genH̃ˢᵏᵉˡSphereBouquet/ˢᵏᵉˡ : {c1 c2 n : ℕ} + (α : FinSphereBouquetMap c1 c2 n) (k : Fin c2) + → H̃ˢᵏᵉˡ (SphereBouquet/ˢᵏᵉˡ α) (suc n) .fst +genH̃ˢᵏᵉˡSphereBouquet/ˢᵏᵉˡ α k = + [ ℤFinGenerator (fin→SphereBouquet/Cell α _ _ k) + , inKerGenerator α k ] + +isGen-genH̃ˢᵏᵉˡSphereBouquet/ˢᵏᵉˡ : {c1 c2 n : ℕ} + (α : FinSphereBouquetMap c1 c2 n) (k : Fin c2) + → fun (fst (GroupIso-Hₙ₊₁SphereBouquetⁿ/→ℤ[]/ImSphereMap α)) + (genH̃ˢᵏᵉˡSphereBouquet/ˢᵏᵉˡ α k) + ≡ [ ℤFinGenerator k ] +isGen-genH̃ˢᵏᵉˡSphereBouquet/ˢᵏᵉˡ {n = n} α k + with (suc n ≟ᵗ suc (suc n)) | (n ≟ᵗ n) +... | lt x | lt y = ⊥.rec (¬m<ᵗm y) +... | lt x | eq y = refl +... | lt x | gt y = ⊥.rec (¬m<ᵗm y) +... | eq x | q = ⊥.rec (¬m<ᵗm (subst (suc n <ᵗ_) (sym x) <ᵗsucm)) +... | gt x | q = ⊥.rec (¬-suc-n<ᵗn x) diff --git a/Cubical/CW/Homology/Groups/Sn.agda b/Cubical/CW/Homology/Groups/Sn.agda new file mode 100644 index 0000000000..4f4ec9f954 --- /dev/null +++ b/Cubical/CW/Homology/Groups/Sn.agda @@ -0,0 +1,91 @@ +{-# OPTIONS --lossy-unification #-} +module Cubical.CW.Homology.Groups.Sn where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Nat.Order.Inductive +open import Cubical.Data.Int +open import Cubical.Data.Fin.Inductive +open import Cubical.Data.Sigma +open import Cubical.Data.Empty as ⊥ + +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.SetQuotients.Base renaming (_/_ to _/s_) +open import Cubical.HITs.SetQuotients.Properties as SQ + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.Instances.Int + +open import Cubical.CW.ChainComplex +open import Cubical.CW.Instances.Sn + +open Iso + +module _ (n : ℕ) where + ScardDiag : isContr (Fin (Scard (suc n) (suc n))) + ScardDiag with (n ≟ᵗ n) + ... | lt x = ⊥.rec (¬m<ᵗm x) + ... | eq x = inhProp→isContr fzero isPropFin1 + ... | gt x = ⊥.rec (¬m<ᵗm x) + + HₙSⁿ→ℤ-fun : (a : Fin (Scard (suc n) (suc n)) → ℤ) → ℤ + HₙSⁿ→ℤ-fun a = a (ScardDiag .fst) + + HₙSⁿ→ℤ-coh : + (a b : Fin (Scard (suc n) (suc n)) → ℤ) + → (aker : ∂ (Sˢᵏᵉˡ (suc n)) n .fst a ≡ λ _ → 0) + → (bker : ∂ (Sˢᵏᵉˡ (suc n)) n .fst b ≡ λ _ → 0) + → (r : Σ[ t ∈ (Fin (Scard (suc n) (suc (suc n))) → ℤ) ] + ∂ (Sˢᵏᵉˡ (suc n)) (suc n) .fst t ≡ λ q → a q - b q) + → HₙSⁿ→ℤ-fun a ≡ HₙSⁿ→ℤ-fun b + HₙSⁿ→ℤ-coh a b aker bker r with (n ≟ᵗ n) | (suc n ≟ᵗ n) + ... | lt x | t = ⊥.rec (¬m<ᵗm x) + ... | eq x | lt x₁ = ⊥.rec (¬-suc-n<ᵗn x₁) + ... | eq x | eq x₁ = ⊥.rec (sucn≠n x₁) + ... | eq x | gt x₁ = sym (+Comm (b fzero) 0 + ∙ cong (_+ b fzero) (funExt⁻ (snd r) fzero) + ∙ sym (+Assoc (a fzero) (- b fzero) (b fzero)) + ∙ cong (a fzero +_) (-Cancel' (b fzero))) + ... | gt x | t = ⊥.rec (¬m<ᵗm x) + + HₙSⁿ→ℤ : H̃ˢᵏᵉˡ (Sˢᵏᵉˡ (suc n)) (suc n) . fst → ℤ + HₙSⁿ→ℤ = SQ.elim (λ _ → isSetℤ) (λ a → HₙSⁿ→ℤ-fun (fst a)) + λ a b → PT.elim (λ _ → isSetℤ _ _) + λ x → HₙSⁿ→ℤ-coh (fst a) (fst b) (snd a) (snd b) (fst x , cong fst (snd x)) + + ∂VanishS : (n : ℕ) (t : _) → ∂ (Sˢᵏᵉˡ (suc n)) n .fst t ≡ λ _ → pos 0 + ∂VanishS zero t = funExt λ { (zero , p) → ·Comm (t fzero) (pos 0)} + ∂VanishS (suc n) t = funExt λ y → ⊥.rec (¬Scard' n y) + + ℤ→HₙSⁿ-fun : ℤ → H̃ˢᵏᵉˡ (Sˢᵏᵉˡ (suc n)) (suc n) . fst + ℤ→HₙSⁿ-fun z = [ (λ _ → z) , ∂VanishS n (λ _ → z) ] + + ℤ→HₙSⁿ : GroupHom ℤGroup (H̃ˢᵏᵉˡ (Sˢᵏᵉˡ (suc n)) (suc n)) + fst (ℤ→HₙSⁿ) = ℤ→HₙSⁿ-fun + snd (ℤ→HₙSⁿ) = makeIsGroupHom λ x y + → cong [_] (Σ≡Prop (λ _ → isOfHLevelPath' 1 (isSetΠ (λ _ → isSetℤ)) _ _) + refl) + + HₙSⁿ→ℤ→HₙSⁿ : (x : _) → ℤ→HₙSⁿ-fun (HₙSⁿ→ℤ x) ≡ x + HₙSⁿ→ℤ→HₙSⁿ = + SQ.elimProp (λ _ → GroupStr.is-set (snd (H̃ˢᵏᵉˡ (Sˢᵏᵉˡ (suc n)) (suc n))) _ _) + λ {(a , p) → cong [_] (Σ≡Prop (λ _ → isSetΠ (λ _ → isSetℤ) _ _) + (funExt λ t → cong a (ScardDiag .snd t)))} + + ℤ≅HₙSⁿ : GroupIso ℤGroup (H̃ˢᵏᵉˡ (Sˢᵏᵉˡ (suc n)) (suc n)) + fun (fst ℤ≅HₙSⁿ) = ℤ→HₙSⁿ .fst + inv (fst ℤ≅HₙSⁿ) = HₙSⁿ→ℤ + rightInv (fst ℤ≅HₙSⁿ) = HₙSⁿ→ℤ→HₙSⁿ + leftInv (fst ℤ≅HₙSⁿ) _ = refl + snd ℤ≅HₙSⁿ = ℤ→HₙSⁿ .snd + + HₙSⁿ≅ℤ : GroupIso (H̃ˢᵏᵉˡ (Sˢᵏᵉˡ (suc n)) (suc n)) ℤGroup + HₙSⁿ≅ℤ = invGroupIso ℤ≅HₙSⁿ + + genHₙSⁿ : H̃ˢᵏᵉˡ (Sˢᵏᵉˡ (suc n)) (suc n) .fst + genHₙSⁿ = [ (λ _ → 1) , (∂VanishS n (λ _ → 1)) ] diff --git a/Cubical/CW/Homology/Groups/Subcomplex.agda b/Cubical/CW/Homology/Groups/Subcomplex.agda new file mode 100644 index 0000000000..6f575c55f2 --- /dev/null +++ b/Cubical/CW/Homology/Groups/Subcomplex.agda @@ -0,0 +1,338 @@ +{-# OPTIONS --lossy-unification #-} + +{- +This file contains a definition of a notion of (strict) +subcomplexes of a CW complex. Here, a subcomplex of a complex +C = colim ( C₀ → C₁ → ...) +is simply the complex +C⁽ⁿ⁾ := colim (C₀ → ... → Cₙ = Cₙ = ...) + +This file contains. +1. Definition of above +2. A `strictification' of finite CW complexes in terms of above +3. An elimination principle for finite CW complexes +4. A proof that C and C⁽ⁿ⁾ has the same homology in appropriate dimensions. +-} + +module Cubical.CW.Homology.Groups.Subcomplex where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport +open import Cubical.Foundations.GroupoidLaws + + +open import Cubical.Data.Nat +open import Cubical.Data.Fin.Inductive.Base +open import Cubical.Data.Fin.Inductive.Properties +open import Cubical.Data.Sigma +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Nat.Order.Inductive +open import Cubical.Data.Sequence +open import Cubical.Data.FinSequence + +open import Cubical.CW.Base +open import Cubical.CW.Properties +open import Cubical.CW.ChainComplex +open import Cubical.CW.Approximation +open import Cubical.CW.Map +open import Cubical.CW.Homology.Base +open import Cubical.CW.Subcomplex + +open import Cubical.CW.ChainComplex +open import Cubical.HITs.SetQuotients.Base renaming (_/_ to _/s_) +open import Cubical.HITs.SetQuotients.Properties as SQ +open import Cubical.Data.Int hiding (_+_) +open import Cubical.Algebra.Group.QuotientGroup + +open import Cubical.HITs.SequentialColimit +open import Cubical.HITs.Sn +open import Cubical.HITs.Pushout +open import Cubical.HITs.SphereBouquet +open import Cubical.HITs.SphereBouquet.Degree +open import Cubical.HITs.PropositionalTruncation as PT + +open import Cubical.Relation.Nullary + +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.ChainComplex +open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup + + +private + variable + ℓ ℓ' ℓ'' : Level + +-- Goal: Show that a cell complex C and its subcomplex C⁽ⁿ⁾ share +-- homology in low enough dimensions +module _ (C : CWskel ℓ) where + ChainSubComplex : (n : ℕ) → snd C .fst n ≡ snd (subComplex C (suc n)) .fst n + ChainSubComplex n with (n ≟ᵗ suc n) + ... | lt x = refl + ... | eq x = ⊥.rec (¬m<ᵗm (subst (n <ᵗ_) (sym x) <ᵗsucm)) + ... | gt x = ⊥.rec (¬-suc-n<ᵗn x) + + ≤ChainSubComplex : (n : ℕ) (m : Fin n) + → snd C .fst (fst m) ≡ snd (subComplex C n) .fst (fst m) + ≤ChainSubComplex n (m , p) with (m ≟ᵗ n) + ... | lt x = refl + ... | eq x = ⊥.rec (¬m<ᵗm (subst (_<ᵗ n) x p)) + ... | gt x = ⊥.rec (¬m<ᵗm (<ᵗ-trans x p)) + +subChainIsoGen : (C : CWskel ℓ) (n : ℕ) (m : Fin n) (p : Trichotomyᵗ (fst m) n) + → AbGroupIso (CWskel-fields.ℤ[A_] C (fst m)) + ℤ[Fin SubComplexGen.subComplexCard C n (fst m) p ] +subChainIsoGen C n (m , p) (lt x) = idGroupIso +subChainIsoGen C n (m , p) (eq x) = ⊥.rec (¬m<ᵗm (subst (m <ᵗ_) (sym x) p)) +subChainIsoGen C n (m , p) (gt x) = ⊥.rec (¬m<ᵗm (<ᵗ-trans x p)) + +subChainIso : (C : CWskel ℓ) (n : ℕ) (m : Fin n) + → AbGroupIso (CWskel-fields.ℤ[A_] C (fst m)) + (CWskel-fields.ℤ[A_] (subComplex C n) (fst m)) +subChainIso C n (m , p) = subChainIsoGen C n (m , p) _ + +-- main result +subComplexHomology : (C : CWskel ℓ) (n m : ℕ) (p : suc (suc m) <ᵗ n) + → GroupIso (H̃ˢᵏᵉˡ C (suc m)) (H̃ˢᵏᵉˡ (subComplex C n) (suc m)) +subComplexHomology C n m p = + homologyIso (suc m) _ _ + (subChainIso C n (suc (suc m) , p)) + (subChainIso C n (suc m , <ᵗ-trans <ᵗsucm p)) + (subChainIso C n (m , <ᵗ-trans <ᵗsucm (<ᵗ-trans <ᵗsucm p))) + lem₁ + lem₂ + where + lem₁ : {q : _} {r : _} + → Iso.fun (subChainIso C n (m , q) .fst) ∘ ∂ C m .fst + ≡ ∂ (subComplex C n) m .fst + ∘ Iso.fun (subChainIso C n (suc m , r) .fst) + lem₁ {q} {r} with (m ≟ᵗ n) | (suc m ≟ᵗ n) | (suc (suc m) ≟ᵗ n) + ... | lt x | lt x₁ | lt x₂ = refl + ... | lt x | lt x₁ | eq x₂ = refl + ... | lt x | lt x₁ | gt x₂ = ⊥.rec (¬squeeze (x₁ , x₂)) + ... | lt x | eq x₁ | t = ⊥.rec (¬m<ᵗm (subst (_<ᵗ n) x₁ r)) + ... | lt x | gt x₁ | t = ⊥.rec (¬m<ᵗm (<ᵗ-trans x₁ r)) + ... | eq x | s | t = ⊥.rec (¬-suc-n<ᵗn (subst (suc m <ᵗ_) (sym x) r)) + ... | gt x | s | t = ⊥.rec (¬-suc-n<ᵗn (<ᵗ-trans r x)) + + lem₂ : {q : suc m <ᵗ n} {r : (suc (suc m)) <ᵗ n} + → Iso.fun (subChainIso C n (suc m , q) .fst) + ∘ ∂ C (suc m) .fst + ≡ ∂ (subComplex C n) (suc m) .fst + ∘ Iso.fun (subChainIso C n (suc (suc m) , r) .fst) + lem₂ {q} with (suc m ≟ᵗ n) | (suc (suc m) ≟ᵗ n) | (suc (suc (suc m)) ≟ᵗ n) + ... | lt x | lt x₁ | lt x₂ = refl + ... | lt x | lt x₁ | eq x₂ = refl + ... | lt x | lt x₁ | gt x₂ = ⊥.rec (¬squeeze (x₁ , x₂)) + ... | lt x | eq x₁ | t = ⊥.rec (¬m<ᵗm (subst (_<ᵗ n) x₁ p)) + ... | lt x | gt x₁ | t = ⊥.rec (¬m<ᵗm (<ᵗ-trans x₁ p)) + ... | eq x | s | t = ⊥.rec (¬m<ᵗm (subst (suc m <ᵗ_) (sym x) q)) + ... | gt x | s | t = ⊥.rec (¬-suc-n<ᵗn (<ᵗ-trans p x)) + + +subComplexHomologyEquiv≡ : ∀ {ℓ} (C : CWskel ℓ) (m n : ℕ) (q : suc (suc n) <ᵗ m) + → Iso.inv (fst (subComplexHomology C m n q)) + ≡ H̃ˢᵏᵉˡ→ (suc n) (incl ∘ Iso.inv (realiseSubComplex m C)) .fst +subComplexHomologyEquiv≡ C m n q = + funExt (SQ.elimProp (λ _ → squash/ _ _) + λ a → cong [_] (Σ≡Prop (λ _ → isSetΠ (λ _ → isSetℤ) _ _) + (mainGen (suc n ≟ᵗ m) (suc (suc n) ≟ᵗ m) (fst a) + ∙ (λ i → bouquetDegree ((CTB* ∘ f1/f2≡ i ∘ + BouquetFuns.BTC (suc n) + (G.subComplexCard C m (suc n) (suc n ≟ᵗ m)) + (G.subComplexα C m (suc n) (suc n ≟ᵗ m)) + (G.subComplexFam≃Pushout C m (suc n) (suc n ≟ᵗ m) + (suc (suc n) ≟ᵗ m)))) .fst (fst a))))) + ∙ cong fst (sym (H̃ˢᵏᵉˡ→β (subComplex C m) C (suc n) + {f = (incl ∘ Iso.inv (realiseSubComplex m C))} + (help q))) + where + open CWskel-fields C + + help' : (m : _) (k : _) (q : _) + → finCellApprox (subComplex C m) C + (λ x → incl (Iso.inv (realiseSubComplex m C) x)) k + fst (help' m k q) = subComplex→ C m k + snd (help' m k q) = →FinSeqColimHomotopy _ _ + λ x → CW↑Gen≡ C k (suc m) (suc m ≟ᵗ suc k) q _ + ∙ cong (incl {n = suc m}) + (funExt⁻ (CW↑GenComm C k (suc m) m (suc m ≟ᵗ suc k) q) x + ∙ funExt⁻ (subComplex→map'Charac C m (suc m ≟ᵗ m) (m ≟ᵗ m)) + (CW↑Gen (subComplex C m) k (suc m) (suc m ≟ᵗ suc k) q x) + ∙ cong (CW↪ C m) (sym (Iso.leftInv ( (realiseSubComplex m C) ) _) + ∙ cong (Iso.inv (realiseSubComplex m C)) + ((push _ ∙ cong (incl {n = suc m}) + (cong (CW↪ (subComplex C m) m) + (secEq (complex≃subcomplex' C m m <ᵗsucm (m ≟ᵗ m)) _) + ∙ CW↪subComplexFam↓ C m (m ≟ᵗ m) (suc m ≟ᵗ m) _)) + ∙ sym (CW↑Gen≡ (subComplex C m) k (suc m) (suc m ≟ᵗ suc k) q x)))) + ∙ sym (push _) + + help : (q : suc (suc n) <ᵗ m) + → finCellApprox (subComplex C m) C + (λ x → incl (Iso.inv (realiseSubComplex m C) x)) + (suc (suc (suc (suc n)))) + fst (help q) = subComplex→ C m (suc (suc (suc (suc n)))) + snd (help q) with (suc (suc (suc n)) ≟ᵗ m) + ... | lt x = help' m (suc (suc (suc (suc n)))) x .snd + ... | eq x = funExt (subst (λ m → (x : _) + → FinSeqColim→Colim (suc (suc (suc (suc n)))) + (finCellMap→FinSeqColim (subComplex C m) C + (subComplex→ C m (suc (suc (suc (suc n))))) x) ≡ incl + (Iso.inv (realiseSubComplex m C) + (FinSeqColim→Colim (suc (suc (suc (suc n)))) x))) x + (funExt⁻ (→FinSeqColimHomotopy _ _ λ w → + (cong (incl {n = suc (suc (suc (suc n)))}) + (cong (subComplex→map C + (suc (suc (suc n))) (suc (suc (suc (suc n))))) + (sym (secEq (_ , subComplexFin C (suc (suc (suc n))) + (suc (suc (suc n)) , <ᵗsucm)) w))) + ∙ (cong (incl {n = suc (suc (suc (suc n)))}) + (CW↪Eq (3 + n) ((4 + n) ≟ᵗ (3 + n)) ((3 + n) ≟ᵗ (3 + n)) + (invEq + (CW↪ (subComplex C (suc (suc (suc n)))) (suc (suc (suc n))) + , subComplexFin C (suc (suc (suc n))) (suc (suc (suc n)) + , <ᵗsucm)) w)) + ∙ sym (push (FinSequenceMap.fmap + (fst (help' (suc (suc (suc n))) (suc (suc (suc n))) <ᵗsucm)) + (suc (suc (suc n)) , <ᵗsucm) + (invEq + (CW↪ (subComplex C (suc (suc (suc n)))) (suc (suc (suc n))) + , subComplexFin C (suc (suc (suc n))) (suc (suc (suc n)) + , <ᵗsucm)) w))) + ∙ funExt⁻ (help' (suc (suc (suc n))) (suc (suc (suc n))) <ᵗsucm .snd) + (fincl (suc (suc (suc n)) , <ᵗsucm) _))) + ∙ cong (incl {n = suc (suc (suc n))}) + (cong (Iso.inv (realiseSubComplex (suc (suc (suc n))) C)) + (push _ + ∙ cong (incl {n = suc (suc (suc (suc n)))}) + (secEq (_ , subComplexFin C (suc (suc (suc n))) (suc (suc (suc n)) + , <ᵗsucm)) w)))))) + where + CW↑GenEq : (n : ℕ) (x₂ : _) (x : _) + → CW↑Gen C n (suc n) (eq (λ _ → suc n)) x₂ x ≡ invEq (e n) (inl x) + CW↑GenEq zero x₂ x = ⊥.rec (C .snd .snd .snd .fst x) + CW↑GenEq (suc n) x₂ x = cong (CW↪ C (suc n)) (transportRefl x) + + CW↪Eq : (n : ℕ) (P : _) (Q : _) (x : _) + → subComplexMapGen.subComplex→map' C n (suc n) P + (invEq (G.subComplexFam≃Pushout C n n Q P) (inl x)) + ≡ CW↪ C n (subComplexMapGen.subComplex→map' C n n Q x) + CW↪Eq n P (lt y) x = ⊥.rec (¬m<ᵗm y) + CW↪Eq n (lt x₂) (eq y) x = ⊥.rec (¬-suc-n<ᵗn x₂) + CW↪Eq n (eq x₂) (eq y) x = ⊥.rec (¬m<ᵗm (subst (_<ᵗ suc n) (sym x₂) <ᵗsucm)) + CW↪Eq n (gt x₂) (eq y) x = ah + where + ah : CW↑Gen C n (suc n) (Trichotomyᵗ-suc (n ≟ᵗ n)) x₂ + (invEq (G.subComplexFam≃Pushout C n n (eq y) (gt x₂)) (inl x)) + ≡ CW↪ C n (idfun (fst C n) x) + ah with (n ≟ᵗ n) + ... | lt x = ⊥.rec (¬m<ᵗm x) + ... | eq q = cong₂ (λ p r → CW↑Gen C n (suc n) (eq p) x₂ + (transport refl (subst (fst C) r x))) + (isSetℕ _ _ _ refl) (isSetℕ _ _ _ refl) + ∙ cong (CW↑Gen C n (suc n) (eq (λ _ → suc n)) x₂) + (transportRefl _ ∙ transportRefl x) + ∙ CW↑GenEq n x₂ x + ... | gt x = ⊥.rec (¬m<ᵗm x) + CW↪Eq n P (gt y) x = ⊥.rec (¬m<ᵗm y) + ... | gt x = ⊥.rec (¬squeeze (q , x)) + + [3+n] : Fin (suc (suc (suc n))) + fst [3+n] = suc n + snd [3+n] = <ᵗ-trans {n = suc n} {m = suc (suc n)} {k = suc (suc (suc n))} + <ᵗsucm <ᵗsucm + + CTB* = BouquetFuns.CTB (suc n) (card (suc n)) (α (suc n)) (e (suc n)) + + f1/f2gen : (q1 : _) (q2 : _) + → cofib (invEq (G.subComplexFam≃Pushout C m (suc n) q1 q2) ∘ inl) + → Pushout (λ _ → tt) (invEq (e (suc n)) ∘ inl) + f1/f2gen q1 q2 (inl x) = inl x + f1/f2gen q1 q2 (inr x) = + inr (subComplexMapGen.subComplex→map' C m (suc (suc n)) q2 x) + f1/f2gen q1 q2 (push a i) = + (push (subComplexMapGen.subComplex→map' C m (suc n) q1 a) + ∙ λ i → inr (subComplex→comm C (suc n) m q1 q2 a i)) i + + f1/f2≡ : f1/f2gen (suc n ≟ᵗ m) (suc (suc n) ≟ᵗ m) + ≡ prefunctoriality.fn+1/fn (suc (suc (suc (suc n)))) + (subComplex→ C m (suc (suc (suc (suc n))))) + ((suc n , <ᵗ-trans-suc (<ᵗ-trans (snd flast) <ᵗsucm))) + f1/f2≡ = funExt λ { (inl x) → refl ; (inr x) → refl ; (push a i) → refl} + + f1/f2genId : (q1 : _) (q2 : _) → f1/f2gen (lt q1) (lt q2) ≡ idfun _ + f1/f2genId q1 q2 = + funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a i) j + → ((λ i → push a ∙ (λ j → inr (lem m q1 q2 a i j))) + ∙ sym (rUnit (push a))) j i} + where + lem : (m : ℕ) (q1 : _) (q2 : _) (a : _) + → subComplex→comm C (suc n) m (lt q1) (lt q2) a ≡ refl + lem (suc m) q1 q2 a = refl + + mainGen : (q1 : _) (q2 : _) (a : _) + → Iso.inv (fst (subChainIsoGen C m (suc n , <ᵗ-trans <ᵗsucm q) q1)) a + ≡ bouquetDegree + (CTB* ∘ f1/f2gen q1 q2 + ∘ (BouquetFuns.BTC (suc n) + (G.subComplexCard C m (suc n) q1) + (G.subComplexα C m (suc n) q1) + (G.subComplexFam≃Pushout C m (suc n) q1 q2))) .fst a + mainGen (lt x) (lt y) a = + funExt⁻ (sym (cong fst (bouquetDegreeId))) a + ∙ λ i → bouquetDegree (lem (~ i)) .fst a + where + lem : CTB* ∘ f1/f2gen (lt x) (lt y) + ∘ BouquetFuns.BTC (suc n) (G.subComplexCard C m (suc n) (lt x)) + (G.subComplexα C m (suc n) (lt x)) + (G.subComplexFam≃Pushout C m (suc n) (lt x) (lt y)) + ≡ idfun _ + lem = funExt λ a → cong CTB* + (funExt⁻ (f1/f2genId x y) _) + ∙ CTB-BTC-cancel (suc n) (card (suc n)) + (α (suc n)) (e (suc n)) .fst _ + mainGen (lt x) (eq y) a = ⊥.rec (¬m<ᵗm (subst (suc (suc n) <ᵗ_) (sym y) q)) + mainGen (lt x) (gt y) a = ⊥.rec (¬squeeze (x , y)) + mainGen (eq x) q2 a = + ⊥.rec (¬m<ᵗm (subst (_<ᵗ_ (suc n)) (sym x) (<ᵗ-trans <ᵗsucm q))) + mainGen (gt x) q2 a = + ⊥.rec (¬m<ᵗm (<ᵗ-trans x (<ᵗ-trans <ᵗsucm q))) + +subComplexHomologyEquiv : ∀ {ℓ} (C : CWskel ℓ) (n : ℕ) (m : ℕ) + (p : suc (suc n) <ᵗ m) + → GroupEquiv (H̃ˢᵏᵉˡ (subComplex C m) (suc n)) + (H̃ˢᵏᵉˡ C (suc n)) +fst (fst (subComplexHomologyEquiv C n m p)) = + H̃ˢᵏᵉˡ→ (suc n) (incl ∘ Iso.inv (realiseSubComplex m C)) .fst +snd (fst (subComplexHomologyEquiv C n m p)) = + subst isEquiv (subComplexHomologyEquiv≡ C m n p) + (isoToIsEquiv (invIso (fst (subComplexHomology C m n p)))) +snd (subComplexHomologyEquiv C n m p) = + H̃ˢᵏᵉˡ→ (suc n) (incl ∘ Iso.inv (realiseSubComplex m C)) .snd + +-- H̃ᶜʷₙ Cₘ ≅ H̃ᶜʷₙ C∞ for n sufficiently small +subComplexHomologyᶜʷEquiv : ∀ {ℓ} (C : CWexplicit ℓ) (n : ℕ) (m : ℕ) + (p : suc (suc n) <ᵗ m) + → GroupEquiv (H̃ᶜʷ ((fst (fst (snd C))) m + , ∣ subComplex (snd C .fst) m + , isoToEquiv (realiseSubComplex m (snd C .fst)) ∣₁) + (suc n)) + (H̃ᶜʷ (realise (snd C .fst) + , ∣ fst (snd C) + , idEquiv _ ∣₁) + (suc n)) +fst (subComplexHomologyᶜʷEquiv C n m p) = + H̃ᶜʷ→ (suc n) (incl {n = m}) .fst + , subComplexHomologyEquiv (snd C .fst) n m p .fst .snd +snd (subComplexHomologyᶜʷEquiv C n m p) = + H̃ᶜʷ→ (suc n) (incl {n = m}) .snd diff --git a/Cubical/CW/HurewiczTheorem.agda b/Cubical/CW/HurewiczTheorem.agda new file mode 100644 index 0000000000..031468afa5 --- /dev/null +++ b/Cubical/CW/HurewiczTheorem.agda @@ -0,0 +1,1294 @@ +{-# OPTIONS --lossy-unification #-} +{- +This file contains +1. a construction of the Hurewicz map πₙᵃᵇ(X) → H̃ᶜʷₙ(X), +2. a proof that it's a homomorphism, and +3. the fact that it is an equivalence when X in (n-1)-connected +-} +module Cubical.CW.HurewiczTheorem where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.GroupoidLaws as GL +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Path + +open import Cubical.CW.Base +open import Cubical.CW.Properties +open import Cubical.CW.Map +open import Cubical.CW.Connected +open import Cubical.CW.Subcomplex +open import Cubical.CW.Approximation +open import Cubical.CW.ChainComplex +open import Cubical.CW.Instances.Sn +open import Cubical.CW.Instances.SphereBouquetMap +open import Cubical.CW.Homology.Base +open import Cubical.CW.Homology.Groups.Sn +open import Cubical.CW.Homology.Groups.CofibFinSphereBouquetMap +open import Cubical.CW.Homology.Groups.Subcomplex +open import Cubical.CW.Instances.Lift + +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Bool +open import Cubical.Data.Fin.Inductive +open import Cubical.Data.Nat.Order.Inductive +open import Cubical.Data.Sigma +open import Cubical.Data.FinSequence +open import Cubical.Data.Int + +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.Sn.Degree +open import Cubical.HITs.SphereBouquet.Degree +open import Cubical.HITs.Susp +open import Cubical.HITs.Pushout +open import Cubical.HITs.SequentialColimit +open import Cubical.HITs.SphereBouquet +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.Truncation as TR +open import Cubical.HITs.SetQuotients as SQ renaming (_/_ to _/s_) +open import Cubical.HITs.Wedge + +open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.Group.Properties +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Group.PiAbCofibFinSphereBouquetMap + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.Abelianization.Base +open import Cubical.Algebra.Group.Abelianization.Properties as Abi +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup + +open import Cubical.Relation.Nullary + +open Iso renaming (inv to inv') +open FinSequenceMap + +private + variable + ℓ ℓ' : Level + +-- Generators of ℤ[k] with locked definitions for faster type checking +private + module _ {k : ℕ} (w : Fin k) (x : _) where + ℤFinGenerator* : lockUnit {ℓ} → ℤ + ℤFinGenerator* unlock = ℤFinGenerator w x + + ℤFinGenerator*₀ : (l : _) → ¬ (fst w ≡ fst x) → ℤFinGenerator* {ℓ} l ≡ pos 0 + ℤFinGenerator*₀ unlock nope with (fst w ≟ᵗ fst x) + ... | (lt ineq) = refl + ... | (eq p) = ⊥.rec (nope p) + ... | (gt ineq) = refl + + ℤFinGenerator*₁ : (l : _) → (fst w ≡ fst x) → ℤFinGenerator* {ℓ} l ≡ pos 1 + ℤFinGenerator*₁ unlock aye with (fst w ≟ᵗ fst x) + ... | (lt ineq) = ⊥.rec (¬m<ᵗm (subst (_<ᵗ fst x) aye ineq)) + ... | (eq p) = refl + ... | (gt ineq) = ⊥.rec (¬m<ᵗm (subst (fst x <ᵗ_) aye ineq)) + +--- Parts 1 and 2: construction of map and homomorphism proof --- + +-- Construction of the Hurewicz map πₙ(X) → H̃ᶜʷₙ(X) +module _ where + -- variations of the map + preHurewiczMap : {n : ℕ} (X : CW ℓ) (x : fst X) + (f : S₊∙ (suc n) →∙ (fst X , x)) + → GroupHom (H̃ᶜʷ (Sᶜʷ (suc n)) (suc n)) (H̃ᶜʷ X (suc n)) + preHurewiczMap {n = n} X x f = H̃ᶜʷ→ {C = Sᶜʷ (suc n)} {D = X} (suc n) (fst f) + + HurewiczMapUntrunc : {n : ℕ} (X : CW ℓ) (x : fst X) + (f : S₊∙ (suc n) →∙ (fst X , x)) → H̃ᶜʷ X (suc n) .fst + HurewiczMapUntrunc {n = n} X x f = preHurewiczMap X x f .fst (genHₙSⁿ n) + + HurewiczMap : {n : ℕ} (X : CW ℓ) (x : fst X) + → π' (suc n) (fst X , x) + → fst (H̃ᶜʷ X (suc n)) + HurewiczMap X x = + ST.rec (GroupStr.is-set (snd (H̃ᶜʷ X _))) (HurewiczMapUntrunc X x) + + -- proof that it is a homomorphism + HurewiczMapHom : {n : ℕ} (X : CW ℓ) (x : fst X) + (f g : π' (suc n) (fst X , x)) → isConnected 2 (fst X) + → HurewiczMap X x (·π' n f g) + ≡ GroupStr._·_ (snd (H̃ᶜʷ X (suc n))) + (HurewiczMap X x f) (HurewiczMap X x g) + HurewiczMapHom {n = n} = uncurry λ X → PT.elim + (λ x → isPropΠ4 λ _ _ _ _ + → GroupStr.is-set (snd (H̃ᶜʷ (X , x) (suc n))) _ _) + (uncurry λ Xsk → EquivJ (λ X y → (x : X) + (f g : π' (suc n) (X , x)) → isConnected 2 X + → HurewiczMap (X , ∣ Xsk , y ∣₁) x (·π' n f g) + ≡ GroupStr._·_ (snd (H̃ᶜʷ (X , ∣ Xsk , y ∣₁) (suc n))) + (HurewiczMap (X , ∣ Xsk , y ∣₁) x f) + (HurewiczMap (X , ∣ Xsk , y ∣₁) x g)) + (λ x → TR.rec (isPropΠ3 (λ _ _ _ → squash/ _ _)) + (uncurry λ x₀ → main Xsk x x₀ x) + (isConnected-CW↪∞ 1 Xsk x .fst))) + where + module _ (Xsk : CWskel ℓ) (x : realise Xsk) where + ∥x₀∥ : hLevelTrunc 1 (Xsk .fst 1) + ∥x₀∥ = TR.map fst (isConnected-CW↪∞ 1 Xsk x .fst) + + X' : CW ℓ + X' = realise Xsk , ∣ Xsk , idEquiv (realise Xsk) ∣₁ + + main : (x₁ : fst Xsk 1) (x : realise Xsk) (y : CW↪∞ Xsk 1 x₁ ≡ x) + (f g : π' (suc n) (realise Xsk , x)) + → isConnected 2 (realise Xsk) + → HurewiczMap X' x (·π' n f g) + ≡ GroupStr._·_ (snd (H̃ᶜʷ X' (suc n))) + (HurewiczMap X' x f) (HurewiczMap X' x g) + main x₀ = J> ST.elim2 (λ _ _ → isSetΠ λ _ → isProp→isSet (squash/ _ _)) + λ f g t → PT.rec2 (squash/ _ _) + (λ {(f' , fp) (g' , gp) → goal f' g' f fp g gp}) + (approxSphereMap∙ Xsk x₀ n f) + (approxSphereMap∙ Xsk x₀ n g) + where + X∙ : Pointed _ + X∙ = fst Xsk (suc (suc n)) , CW↪ Xsk (suc n) (CWskel∙ Xsk x₀ n) + + X* : (n : ℕ) → Pointed _ + X* n = fst Xsk (suc (suc n)) , CW↪ Xsk (suc n) (CWskel∙ Xsk x₀ n) + + goalTy : (f g : S₊∙ (suc n) →∙ (realise Xsk , CW↪∞ Xsk 1 x₀)) → Type _ + goalTy f g = + HurewiczMap X' (CW↪∞ Xsk 1 x₀) (·π' n ∣ f ∣₂ ∣ g ∣₂) + ≡ GroupStr._·_ (snd (H̃ᶜʷ X' (suc n))) + (HurewiczMap X' (CW↪∞ Xsk 1 x₀) ∣ f ∣₂) + (HurewiczMap X' (CW↪∞ Xsk 1 x₀) ∣ g ∣₂) + + module _ (f' g' : S₊∙ (suc n) →∙ X∙) where + open CWskel-fields Xsk + finCellApprox∙Π : finCellApprox (Sˢᵏᵉˡ (suc n)) Xsk + (fst (∙Π (incl∙ Xsk x₀ ∘∙ f') (incl∙ Xsk x₀ ∘∙ g')) ∘ + invEq (hasCWskelSphere (suc n) .snd)) (suc (suc (suc (suc n)))) + finCellApprox∙Π = + finCellApproxSˢᵏᵉˡImproved Xsk (suc n) x₀ + (∙Π f' g') (∙Π (incl∙ Xsk x₀ ∘∙ f') (incl∙ Xsk x₀ ∘∙ g')) + (λ x → funExt⁻ (cong fst (∙Π∘∙ n f' g' (incl∙ Xsk x₀))) x ∙ refl) + (suc (suc (suc (suc n)))) + + CTB→ : (n : ℕ) → _ + CTB→ n = BouquetFuns.CTB (suc n) + (card (suc n)) (α (suc n)) (e (suc n)) + + cofib→cofibCW : (n : ℕ) (f : S₊∙ (suc n) →∙ X* n) (p : _) (q : _) → + cofib (invEq (SαEqGen (suc n) (suc n) p q) ∘ inl) → cofibCW (suc n) Xsk + cofib→cofibCW n f p q (inl x) = inl x + cofib→cofibCW n f (lt _) q (inr x) = inl tt + cofib→cofibCW n f (eq _) p (inr x) = inr (f .fst x) + cofib→cofibCW n f (gt _) q (inr x) = inl tt + cofib→cofibCW n f (lt x) q (push a i) = inl tt + cofib→cofibCW n f (eq x) q (push a i) = + (push (CWskel∙ Xsk x₀ n) ∙ λ i → inr (f'∘SαEqGen⁻¹≡ q x a (~ i))) i + where + f'∘SαEqGen⁻¹≡ : (q : _) (x : _) (a : _) + → f .fst ((invEq (SαEqGen (suc n) (suc n) (eq x) q) ∘ inl) a) + ≡ CWskel∙ Xsk x₀ (suc n) + f'∘SαEqGen⁻¹≡ (lt _) x a = snd f + f'∘SαEqGen⁻¹≡ (eq p) x a = + ⊥.rec (¬m<ᵗm (subst (_<ᵗ_ (suc n)) ((sym p) ∙ cong predℕ x) <ᵗsucm)) + f'∘SαEqGen⁻¹≡ (gt p) x a = + ⊥.rec (¬-suc-n<ᵗn (subst (_<ᵗ_ (suc (suc n))) (λ i → predℕ (x i)) p)) + + cofib→cofibCW n f (gt x) q (push a i) = inl tt + + CTB∘cofib→cofibCW∘BTC : (n : ℕ) (f : S₊∙ (suc n) →∙ X* n) + (p : _) (q : _) (x : _) → _ + CTB∘cofib→cofibCW∘BTC n f' p q x = + CTB→ n (cofib→cofibCW n f' p q (BouquetFuns.BTC (suc n) + (ScardGen (suc n) (suc n) p) + (SαGen (suc n) (suc n) p q) + (SαEqGen (suc n) (suc n) p q) x)) + + module _ (f' : S₊∙ (suc n) →∙ X∙) (Q : _) where + private + f = finCellApproxSˢᵏᵉˡImproved Xsk (suc n) x₀ f' + (incl∙ Xsk x₀ ∘∙ f') Q (suc (suc (suc (suc n)))) + + cofib→cofibCW≡inr : (x : _) + → prefunctoriality.fn+1/fn (suc (suc (suc (suc n)))) (f .fst) + (suc n , <ᵗ-trans-suc (<ᵗ-trans <ᵗsucm <ᵗsucm)) (inr x) + ≡ cofib→cofibCW n f' (suc (suc n) ≟ᵗ suc (suc n)) + (suc n ≟ᵗ suc (suc n)) (inr x) + cofib→cofibCW≡inr x with (n ≟ᵗ n) + ... | lt p = ⊥.rec (¬m<ᵗm p) + ... | eq q = λ i → inr ((cong (λ p → subst (fst Xsk) p (fst f' x)) + (cong sym (isSetℕ _ _ (cong suc (cong suc q)) refl)) + ∙ transportRefl (fst f' x)) i) + ... | gt p = ⊥.rec (¬m<ᵗm p) + + cofib→cofibCW≡push : (a : _) + → Square refl (cofib→cofibCW≡inr (CW↪ (Sˢᵏᵉˡ (suc n)) (suc n) a)) + (push (cellMapSˢᵏᵉˡFunGenGen Xsk (suc n) x₀ + (fst f') (snd f') (suc n) (suc n ≟ᵗ suc (suc n)) a) + ∙ (λ i → inr (cellMapSˢᵏᵉˡFunGenComm Xsk (suc n) x₀ + (fst f') (snd f') (suc n) + (suc (suc n) ≟ᵗ suc (suc n)) + (suc n ≟ᵗ suc (suc n)) a (~ i)))) + (cong (cofib→cofibCW n f' + (suc (suc n) ≟ᵗ suc (suc n)) (suc n ≟ᵗ suc (suc n))) (push a)) + cofib→cofibCW≡push a with (n ≟ᵗ n) + ... | lt x = ⊥.rec (¬m<ᵗm x) + ... | eq x = + flipSquare (help (cong suc (cong suc x)) (sym (isSetℕ _ _ _ _))) + where + cool : cellMapSˢᵏᵉˡFunGenGen∙ Xsk _ x₀ + (fst f') (snd f') (suc n) (eq refl) + ≡ transportRefl _ ∙ snd f' + cool = + cong₂ _∙_ + (λ j i → subst (fst Xsk) + (isSet→isGroupoid isSetℕ _ _ _ _ + (isSetℕ (suc (suc n)) _ refl refl) refl j i) + (snd f' i)) + (transportRefl _) + ∙ λ i → (λ j → transportRefl (snd f' (j ∧ ~ i)) (j ∧ i)) + ∙ λ j → transportRefl (snd f' (~ i ∨ j)) (i ∨ j) + + help : (w : suc (suc n) ≡ suc (suc n)) (t : refl ≡ w) + → Square + ((push (cellMapSˢᵏᵉˡFunGenGen Xsk (suc n) x₀ (fst f') + (snd f') (suc n) (suc n ≟ᵗ suc (suc n)) a) + ∙ (λ i → inr (cellMapSˢᵏᵉˡFunGenComm Xsk (suc n) x₀ + (fst f') (snd f') (suc n) + (eq w) (suc n ≟ᵗ suc (suc n)) a (~ i))))) + (λ i → cofib→cofibCW n f' (eq w) (suc n ≟ᵗ suc (suc n)) + (push a i)) + (λ _ → inl tt) + (λ i → inr ((cong (λ p → subst (fst Xsk) p + (fst f' (invEq (SαEqGen (suc n) (suc n) (eq w) + (suc n ≟ᵗ suc (suc n))) (inl a)))) + (sym (cong sym t)) ∙ transportRefl _) i)) + help with (n ≟ᵗ suc n) + ... | lt w = + J> (cong₂ _∙_ + refl + ((λ j i → inr ((lUnit (cool j) (~ j)) (~ i))) + ∙ cong sym (cong-∙ inr (transportRefl _) + (snd f')) + ∙ symDistr _ _) + ∙ GL.assoc _ _ _) + ◁ flipSquare (flipSquare (symP (compPath-filler + (push (CWskel∙ Xsk x₀ n) + ∙ (λ i₁ → inr (snd f' (~ i₁)))) + (sym (transportRefl (inr (f' .snd i0)))))) + ▷ λ j i → inr (lUnit (transportRefl (fst f' (ptSn (suc n)))) j i)) + ... | eq x = ⊥.rec (¬m<ᵗm (subst (_<ᵗ suc n) x <ᵗsucm)) + ... | gt x = ⊥.rec (¬-suc-n<ᵗn x) + ... | gt x = ⊥.rec (¬m<ᵗm x) + + cofib→cofibCW≡ : (x : _) + → prefunctoriality.fn+1/fn (suc (suc (suc (suc n)))) (f .fst) + (suc n , <ᵗ-trans-suc (<ᵗ-trans <ᵗsucm <ᵗsucm)) x + ≡ cofib→cofibCW n f' + (suc (suc n) ≟ᵗ suc (suc n)) (suc n ≟ᵗ suc (suc n)) x + cofib→cofibCW≡ (inl x) = refl + cofib→cofibCW≡ (inr x) = cofib→cofibCW≡inr x + cofib→cofibCW≡ (push a i) = cofib→cofibCW≡push a i + + bouquetFunct≡ : + prefunctoriality.bouquetFunct (suc (suc (suc (suc n)))) + (f .fst) (suc n , <ᵗ-trans-suc (<ᵗ-trans <ᵗsucm <ᵗsucm)) + ≡ CTB∘cofib→cofibCW∘BTC n f' + (suc (suc n) ≟ᵗ suc (suc n)) (suc n ≟ᵗ suc (suc n)) + bouquetFunct≡ = funExt (λ x → cong (CTB→ n) (cofib→cofibCW≡ _)) + + f = finCellApproxSˢᵏᵉˡImproved Xsk (suc n) x₀ f' + (incl∙ Xsk x₀ ∘∙ f') (λ _ → refl) (suc (suc (suc (suc n)))) + g = finCellApproxSˢᵏᵉˡImproved Xsk (suc n) x₀ g' + (incl∙ Xsk x₀ ∘∙ g') (λ _ → refl) (suc (suc (suc (suc n)))) + + wraplem : ∀ {ℓ} {A : Type ℓ} {x : A} (y : A) (p : x ≡ y) (l1 l2 : y ≡ y) + → p ∙∙ (l1 ∙ l2) ∙∙ sym p + ≡ (p ∙∙ l1 ∙∙ sym p) ∙ (p ∙∙ l2 ∙∙ sym p) + wraplem = J> λ l1 l2 → sym (rUnit _) ∙ cong₂ _∙_ (rUnit l1) (rUnit l2) + + CTB∘cofib→cofibCW∘BTC-Hom : (n : ℕ) (f' g' : _) (p : _) (q : _) (x : _) + → CTB∘cofib→cofibCW∘BTC n (∙Π f' g') p q x + ≡ SphereBouquet∙Π (CTB∘cofib→cofibCW∘BTC n f' p q , refl) + (CTB∘cofib→cofibCW∘BTC n g' p q , refl) .fst x + CTB∘cofib→cofibCW∘BTC-Hom n f' g' (lt s) q x = ⊥.rec (¬m<ᵗm s) + CTB∘cofib→cofibCW∘BTC-Hom n f' g' (eq _) (lt _) (inl _) = refl + CTB∘cofib→cofibCW∘BTC-Hom zero f' g' (eq s) (lt d) + (inr (t , base)) = refl + CTB∘cofib→cofibCW∘BTC-Hom zero f' g' (eq s) (lt d) + (inr ((zero , tt) , loop i)) j = CTB∘cofib→cofibCW∘BTC-Hom₀loop j i + where + p : I → _ + p i = subst S₊ (isSetℕ _ _ (cong (predℕ ∘ predℕ) s) refl i) + + q = cong (CTB∘cofib→cofibCW∘BTC zero f' (eq s) (lt d)) + (sym (push fzero)) ∙ refl + + lem : (h : S₊∙ 1 →∙ X* zero) + → cong (cofib→cofibCW zero h (eq s) (lt d) + ∘ inr ∘ invEq (SαEqGen 1 1 (eq s) (lt d))) + (push (fzero , false) ∙ sym (push (fzero , true))) + ≡ (λ i → inr (fst h (loop i))) + lem h = cong-∙ (cofib→cofibCW zero h (eq s) (lt d) + ∘ inr ∘ invEq (SαEqGen 1 1 (eq s) (lt d))) + (push (fzero , false)) (sym (push (fzero , true))) + ∙ cong₂ _∙_ + (λ i j → inr (h .fst (SuspBool→S¹ (merid (p i false) j)))) + (λ i j → inr (h .fst (SuspBool→S¹ (merid (p i true) (~ j))))) + ∙ sym (rUnit _) + + CTB∘cofib→cofibCW∘BTC-Hom₀loop : + cong (CTB∘cofib→cofibCW∘BTC zero (∙Π f' g') (eq s) (lt d)) + (λ i → inr (fzero , loop i)) + ≡ (sym q ∙∙ cong (CTB∘cofib→cofibCW∘BTC zero f' (eq s) (lt d)) + (λ i → inr (fzero , loop i)) ∙∙ q) + ∙ (sym q ∙∙ cong (CTB∘cofib→cofibCW∘BTC zero g' (eq s) (lt d)) + (λ i → inr (fzero , loop i)) ∙∙ q) + CTB∘cofib→cofibCW∘BTC-Hom₀loop = + cong (cong (CTB→ zero)) + (cong-∙∙ (cofib→cofibCW zero (∙Π f' g') (eq s) (lt d)) _ _ _ + ∙ cong₃ _∙∙_∙∙_ (sym (rUnit (push x₀))) (lem (∙Π f' g') + ∙ cong-∙ inr _ _) + (cong sym (sym (rUnit (push x₀)))) + ∙ wraplem _ _ _ _) + ∙ (cong-∙ (CTB→ zero) _ _ + ∙ cong₂ _∙_ (cong (cong (CTB→ zero)) + λ i → compPath-filler (push x₀) (λ t → inr (sym (snd f') t)) i + ∙∙ (λ j → inr (doubleCompPath-filler (sym (snd f')) + (cong (fst f') loop) (snd f') (~ i) j)) + ∙∙ sym (compPath-filler (push x₀) + (λ t → inr (sym (snd f') t)) i)) + λ i → (cong (CTB→ zero)) + (compPath-filler (push x₀) (λ t → inr (sym (snd g') t)) i + ∙∙ (λ j → inr (doubleCompPath-filler (sym (snd g')) + (cong (fst g') loop) (snd g') (~ i) j)) + ∙∙ sym (compPath-filler (push x₀) + (λ t → inr (sym (snd g') t)) i))) + ∙ cong₂ _∙_ + (sym (cong (cong (CTB→ zero)) + (cong-∙∙ (cofib→cofibCW zero f' (eq s) (lt d)) _ _ _ + ∙ cong₃ _∙∙_∙∙_ refl (lem f') refl)) + ∙ rUnit (cong (CTB∘cofib→cofibCW∘BTC zero f' (eq s) (lt d)) + (λ i → inr (fzero , loop i))) + ∙ cong₃ _∙∙_∙∙_ (rUnit refl) refl (rUnit refl)) + (sym (cong (cong (CTB→ zero)) + (cong-∙∙ (cofib→cofibCW zero g' (eq s) (lt d)) _ _ _ + ∙ cong₃ _∙∙_∙∙_ refl (lem g') refl)) + ∙ rUnit (cong (CTB∘cofib→cofibCW∘BTC zero g' (eq s) (lt d)) + (λ i → inr (fzero , loop i))) + ∙ cong₃ _∙∙_∙∙_ (rUnit refl) refl (rUnit refl)) + CTB∘cofib→cofibCW∘BTC-Hom (suc n) f' g' (eq s) (lt d) + (inr (t , north)) = refl + CTB∘cofib→cofibCW∘BTC-Hom (suc n) f' g' (eq s) (lt d) + (inr (t , south)) = refl + CTB∘cofib→cofibCW∘BTC-Hom (suc n) f' g' (eq s) (lt d) + (inr ((zero , tt) , merid a i)) j = CTB∘cofib→cofibCW∘BTC-Hom₀merid j i + where + p : (x : _) → transport (λ i₂ → S₊ (predℕ (predℕ (s i₂)))) x ≡ x + p x = cong (λ p → transport p x) (cong (cong S₊) (isSetℕ _ _ _ refl)) + ∙ transportRefl x + + q = cong (CTB∘cofib→cofibCW∘BTC (suc n) f' (eq s) (lt d)) + (sym (push fzero)) ∙ refl + + cong-h-σ : (h : S₊∙ (suc (suc n)) →∙ X* (suc n)) (a : _) + → _ + cong-h-σ h a = (push (CWskel∙ Xsk x₀ (suc n)) + ∙∙ (λ i → inr ((sym (snd h) ∙∙ cong (fst h) (σS a) ∙∙ snd h) i)) + ∙∙ sym (push (CWskel∙ Xsk x₀ (suc n)))) + + cong-CTB→∘h-σ≡ : (h : S₊∙ (suc (suc n)) →∙ X* (suc n)) (a : _) + → cong (CTB∘cofib→cofibCW∘BTC (suc n) h (eq s) (lt d)) + (λ i → inr (fzero , merid a i)) + ≡ cong (CTB→ (suc n)) (cong-h-σ h a) + cong-CTB→∘h-σ≡ h a = cong (cong (CTB→ (suc n))) + (cong-∙∙ (cofib→cofibCW (suc n) h (eq s) (lt d)) _ _ _ + ∙ cong₃ _∙∙_∙∙_ refl + (cong-∙ (cofib→cofibCW (suc n) h (eq s) (lt d) ∘ inr + ∘ invEq (SαEqGen (suc (suc n)) (suc (suc n)) (eq s) (lt d))) + (push (fzero , a)) (sym (push (fzero , ptSn (suc n)))) + ∙ cong₂ _∙_ + (λ j i → inr (h .fst (merid (p a j) i))) + (λ j i → inr (h .fst (merid (p (ptSn (suc n)) j) (~ i)))) + ∙ sym (cong-∙ (λ x → inr (h .fst x)) + (merid a) (sym (merid (ptSn (suc n)))))) refl + ∙ λ i → compPath-filler (push (CW↪ Xsk (suc n) (CWskel∙ Xsk x₀ n))) + (λ i → inr (snd h (~ i))) (~ i) + ∙∙ (λ j → inr (doubleCompPath-filler (sym (snd h)) + (λ i → fst h (σS a i)) + (snd h) i j)) + ∙∙ sym (compPath-filler (push (CW↪ Xsk (suc n) (CWskel∙ Xsk x₀ n))) + (λ i → inr (snd h (~ i))) (~ i))) + + CTB∘cofib→cofibCW∘BTC-merid≡ : (h : S₊∙ (suc (suc n)) →∙ X* (suc n)) + → cong (CTB∘cofib→cofibCW∘BTC (suc n) h (eq s) (lt d)) + (λ i → inr (fzero , merid (ptSn (suc n)) i)) ≡ refl + CTB∘cofib→cofibCW∘BTC-merid≡ h = + cong-CTB→∘h-σ≡ h (ptSn (suc n)) + ∙ cong (cong (CTB→ (suc n))) + (cong₃ _∙∙_∙∙_ refl + ((λ j i → inr ((cong₃ _∙∙_∙∙_ + refl (cong (cong (fst h)) σS∙) refl + ∙ ∙∙lCancel (snd h)) j i))) refl + ∙ ∙∙lCancel _) + + CTB∘cofib→cofibCW∘BTC-merid≡' : + (h : S₊∙ (suc (suc n)) →∙ X* (suc n)) (a : _) + → cong (CTB∘cofib→cofibCW∘BTC (suc n) h (eq s) (lt d)) + (λ i → inr (fzero , σS a i)) + ≡ cong-CTB→∘h-σ≡ h a i1 + CTB∘cofib→cofibCW∘BTC-merid≡' h a = + cong-∙ (λ q → CTB∘cofib→cofibCW∘BTC (suc n) h (eq s) (lt d) + (inr (fzero , q))) + (merid a) (sym (merid (ptSn (suc n)))) + ∙ cong₂ _∙_ (cong-CTB→∘h-σ≡ h a) + (cong sym (CTB∘cofib→cofibCW∘BTC-merid≡ h)) + ∙ sym (rUnit (cong-CTB→∘h-σ≡ h a i1)) + + cong-h-σ-Hom : (a : _) → cong-h-σ (·Susp (S₊∙ (suc n)) f' g') a + ≡ cong-h-σ f' a ∙ cong-h-σ g' a + cong-h-σ-Hom a = + cong₃ _∙∙_∙∙_ refl + (λ i j → inr ((sym (rUnit (cong (fst (·Susp (S₊∙ (suc n)) f' g')) + (σS a))) + ∙ cong-∙ (fst (·Susp (S₊∙ (suc n)) f' g')) + (merid a) (sym (merid (ptSn (suc n)))) + ∙ cong₂ _∙_ refl (cong sym + (cong₂ _∙_ + (cong₃ _∙∙_∙∙_ refl + (cong (cong (fst f')) + (rCancel (merid (ptSn (suc n))))) refl + ∙ ∙∙lCancel (snd f')) + (cong₃ _∙∙_∙∙_ refl + (cong (cong (fst g')) + (rCancel (merid (ptSn (suc n))))) refl + ∙ ∙∙lCancel (snd g')) + ∙ sym (rUnit refl))) + ∙ sym (rUnit _)) i j)) refl + ∙ cong₃ _∙∙_∙∙_ refl (cong-∙ inr _ _) refl + ∙ wraplem _ _ _ _ + + CTB∘cofib→cofibCW∘BTC-Hom₀merid : + cong (CTB∘cofib→cofibCW∘BTC (suc n) (∙Π f' g') (eq s) (lt d)) + (λ i → inr (fzero , merid a i)) + ≡ (sym q ∙∙ cong (CTB∘cofib→cofibCW∘BTC (suc n) f' (eq s) (lt d)) + (λ i → inr (fzero , σS a i)) ∙∙ q) + ∙ (sym q ∙∙ cong (CTB∘cofib→cofibCW∘BTC (suc n) g' (eq s) (lt d)) + (λ i → inr (fzero , σS a i)) ∙∙ q) + CTB∘cofib→cofibCW∘BTC-Hom₀merid = cong-CTB→∘h-σ≡ (∙Π f' g') a + ∙ cong (cong (CTB→ (suc n))) (cong-h-σ-Hom a) + ∙ cong-∙ (CTB→ (suc n)) (cong-h-σ f' a) (cong-h-σ g' a) + ∙ cong₂ _∙_ + (rUnit _ ∙ cong₃ _∙∙_∙∙_ (rUnit refl) + (sym (CTB∘cofib→cofibCW∘BTC-merid≡' f' a)) (rUnit refl)) + (rUnit _ ∙ cong₃ _∙∙_∙∙_ (rUnit refl) + (sym (CTB∘cofib→cofibCW∘BTC-merid≡' g' a)) (rUnit refl)) + + CTB∘cofib→cofibCW∘BTC-Hom zero f' g' (eq s) (lt d) (push a i) = refl + CTB∘cofib→cofibCW∘BTC-Hom (suc n) f' g' (eq s) (lt d) (push a i) = refl + CTB∘cofib→cofibCW∘BTC-Hom n f' g' (eq s) (eq d) x = + ⊥.rec (¬m<ᵗm (subst (_<ᵗ suc (suc n)) d <ᵗsucm)) + CTB∘cofib→cofibCW∘BTC-Hom n f' g' (eq s) (gt d) x = + ⊥.rec (¬-suc-n<ᵗn d) + CTB∘cofib→cofibCW∘BTC-Hom n f' g' (gt s) q x = ⊥.rec (¬m<ᵗm s) + + CTB∘cofib→cofibCW∘BTC-Hom' : (x : _) + → prefunctoriality.bouquetFunct (suc (suc (suc (suc n)))) + (finCellApprox∙Π .fst) + (suc n , <ᵗ-trans-suc (<ᵗ-trans <ᵗsucm <ᵗsucm)) x + ≡ SphereBouquet∙Π + (prefunctoriality.bouquetFunct (suc (suc (suc (suc n)))) + (f .fst) (suc n , <ᵗ-trans-suc (<ᵗ-trans <ᵗsucm <ᵗsucm)) , refl) + (prefunctoriality.bouquetFunct (suc (suc (suc (suc n)))) + (g .fst) (suc n , <ᵗ-trans-suc (<ᵗ-trans <ᵗsucm <ᵗsucm)) , refl) + .fst x + CTB∘cofib→cofibCW∘BTC-Hom' x = + funExt⁻ (bouquetFunct≡ (∙Π f' g') λ _ → refl) x + ∙ CTB∘cofib→cofibCW∘BTC-Hom n f' g' _ _ x + ∙ λ i → SphereBouquet∙Π + (bouquetFunct≡ f' (λ _ → refl) (~ i) , (λ _ → inl tt)) + (bouquetFunct≡ g' (λ _ → refl) (~ i) , (λ _ → inl tt)) .fst x + + goal' : goalTy (incl∙ Xsk x₀ ∘∙ f') (incl∙ Xsk x₀ ∘∙ g') + goal' = + funExt⁻ (cong fst (H̃ˢᵏᵉˡ→β (Sˢᵏᵉˡ (suc n)) Xsk (suc n) + {f = (fst (∙Π (incl∙ Xsk x₀ ∘∙ f') (incl∙ Xsk x₀ ∘∙ g')) ∘ + invEq (hasCWskelSphere (suc n) .snd))} finCellApprox∙Π)) (genHₙSⁿ n) + ∙ cong [_] (Σ≡Prop (λ _ → isSetΠ (λ _ → isSetℤ) _ _) + ((λ i → bouquetDegree (λ x → CTB∘cofib→cofibCW∘BTC-Hom' x i) + .fst (λ _ → pos 1)) + ∙ funExt⁻ (cong fst (bouquetDegree∙Π _ _ _ + (prefunctoriality.bouquetFunct (suc (suc (suc (suc n)))) + (f .fst) (suc n , <ᵗ-trans-suc (<ᵗ-trans <ᵗsucm <ᵗsucm)) , refl) + (prefunctoriality.bouquetFunct (suc (suc (suc (suc n)))) + (g .fst) (suc n , <ᵗ-trans-suc (<ᵗ-trans <ᵗsucm <ᵗsucm)) , refl))) + λ _ → pos 1)) + ∙ cong₂ (GroupStr._·_ (snd (H̃ᶜʷ X' (suc n)))) + (funExt⁻ (cong fst (sym (H̃ˢᵏᵉˡ→β (Sˢᵏᵉˡ (suc n)) Xsk (suc n) + {f = incl ∘ fst f' ∘ invEq (hasCWskelSphere (suc n) .snd)} f))) + (genHₙSⁿ n)) + ((funExt⁻ (cong fst (sym (H̃ˢᵏᵉˡ→β (Sˢᵏᵉˡ (suc n)) Xsk (suc n) + {f = incl ∘ fst g' ∘ invEq (hasCWskelSphere (suc n) .snd)} g))) + (genHₙSⁿ n))) + + goal : (f : _) (fp : incl∙ Xsk x₀ ∘∙ f' ≡ f) + (g : _) (gp : incl∙ Xsk x₀ ∘∙ g' ≡ g) → goalTy f g + goal = J> (J> goal') + + -- THe Hurewicz map as a group homomorphism + HurewiczHom : {n : ℕ} (X : CW ℓ) (x : fst X) (conX : isConnected 2 (fst X)) + → GroupHom (π'Gr n (fst X , x)) (H̃ᶜʷ X (suc n)) + fst (HurewiczHom {n = n} X x conX) = HurewiczMap X x + snd (HurewiczHom {n = n} X x conX) = + makeIsGroupHom λ f g → HurewiczMapHom X x f g conX + +-- Naturality of the Hurewicz map +HurewiczMapNat : {n : ℕ} (X Y : CW ℓ) (x : fst X) (y : fst Y) + (g : (fst X , x) →∙ (fst Y , y)) + → H̃ᶜʷ→ {C = X} {D = Y} (suc n) (fst g) .fst ∘ HurewiczMap X x + ≡ HurewiczMap Y y ∘ π'∘∙Hom n g .fst +HurewiczMapNat {n = n} X Y x y g = funExt + (ST.elim (λ _ → isOfHLevelPath 2 (GroupStr.is-set (H̃ᶜʷ Y (suc n) .snd)) _ _) + λ f → funExt⁻ (sym (cong fst (H̃ᶜʷ→comp + {C = Sᶜʷ (suc n)} {D = X} {E = Y} (suc n) (fst g) (fst f)))) + (genHₙSⁿ n)) + +-- The Hurewicz map on abelisanised homotopy groups +HurewiczHomAb : (X : CW ℓ) (x : fst X) (isC : isConnected 2 (fst X)) + (n : ℕ) → AbGroupHom (AbelianizationAbGroup (π'Gr n (fst X , x))) (H̃ᶜʷAb X n) +fst (HurewiczHomAb X x isC n) = + Abi.rec _ (AbGroupStr.is-set (snd (H̃ᶜʷAb X n))) + (HurewiczHom X x isC .fst) + lem + where + lem : (a b c : _) → _ + lem a b c = IsGroupHom.pres· (HurewiczHom X x isC .snd) _ _ + ∙ cong₂ (GroupStr._·_ (H̃ᶜʷ X (suc n) .snd)) refl + (IsGroupHom.pres· (HurewiczHom X x isC .snd) _ _ + ∙ AbGroupStr.+Comm (snd (H̃ᶜʷAb X n)) _ _ + ∙ sym (IsGroupHom.pres· (HurewiczHom X x isC .snd) _ _)) + ∙ sym (IsGroupHom.pres· (HurewiczHom X x isC .snd) _ _) +snd (HurewiczHomAb X x isC n) = makeIsGroupHom + (Abi.elimProp2 _ (λ _ _ → GroupStr.is-set (snd (H̃ᶜʷ X (suc n))) _ _) + (IsGroupHom.pres· (HurewiczHom X x isC .snd))) + + +--- Part 3: proof of the Hurewicz theorem --- + +-- The Hurewicz map is an equivalence on cofibres of sphere bouquets +private + makeHurewiczMapCofibEquiv : + {n m k : ℕ} (α : FinSphereBouquetMap∙ m k n) + (ϕ : GroupHom (AbelianizationGroup (π'Gr n (cofib∙ (fst α)))) + (H̃ˢᵏᵉˡ (SphereBouquet/ˢᵏᵉˡ (fst α)) (suc n))) + → ((k : _) → fst ϕ (πᵃᵇFinSphereBouquetMapGenerator α k) + ≡ genH̃ˢᵏᵉˡSphereBouquet/ˢᵏᵉˡ (fst α) k) + → isEquiv (fst ϕ) + makeHurewiczMapCofibEquiv α ϕ hyp = + makeℤ[]/Equiv + (GroupIso→GroupEquiv + (invGroupIso (π'ᵃᵇCofibFinSphereBouquetMap≅ℤ[]/BouquetDegree α))) + (GroupIso→GroupEquiv + (invGroupIso (GroupIso-Hₙ₊₁SphereBouquetⁿ/→ℤ[]/ImSphereMap (fst α)))) ϕ + λ k → cong (fst ϕ) + (sym (cong (inv' (fst (π'ᵃᵇCofibFinSphereBouquetMap≅ℤ[]/BouquetDegree α))) + (π'ᵃᵇCofibFinSphereBouquetMap≅ℤ[]/BouquetDegreePresGens α k)) + ∙ leftInv (fst (π'ᵃᵇCofibFinSphereBouquetMap≅ℤ[]/BouquetDegree α)) _) + ∙ hyp k + ∙ sym (leftInv (fst (GroupIso-Hₙ₊₁SphereBouquetⁿ/→ℤ[]/ImSphereMap (fst α))) + (genH̃ˢᵏᵉˡSphereBouquet/ˢᵏᵉˡ (fst α) k)) + ∙ cong (ℤ[]/ImSphereMap→HₙSphereBouquetⁿ/ (fst α)) + (isGen-genH̃ˢᵏᵉˡSphereBouquet/ˢᵏᵉˡ (fst α) k) + +HurewiczMapCofibEquiv : ∀ {n m k : ℕ} + → (α : FinSphereBouquetMap∙ m k n) + → (isC : isConnected 2 (cofib (fst α))) + → isEquiv (fst (HurewiczHomAb (SphereBouquet/ᶜʷ (fst α)) (inl tt) isC n)) +HurewiczMapCofibEquiv {n = n} {m} {k} α isC = makeHurewiczMapCofibEquiv α + (HurewiczHomAb (SphereBouquet/ᶜʷ (fst α)) (inl tt) isC n) λ w → + funExt⁻ (cong fst (H̃ˢᵏᵉˡ→β (Sˢᵏᵉˡ (suc n)) (SphereBouquet/ˢᵏᵉˡ (fst α)) + (suc n) {f = realiseInr w} (finCellApproxInr' w))) + (genHₙSⁿ n) + ∙ cong [_] (Σ≡Prop (λ _ → isOfHLevelPathP' 1 (isSetΠ (λ _ → isSetℤ)) _ _) + ((λ i x → sumFinℤ (λ a → degree (suc n) λ t + → pickPetal x (CTB' n α trich₁ trich₂ + (cofib→cofib≡' w (suc n , <ᵗ-trans <ᵗsucm <ᵗsucm) + (preBTC' n α trich₁ trich₂ a .fst t) (~ i))))) + ∙ funExt λ x → sumFin-choose _+_ 0 (λ _ → refl) +Comm + (λ a → degree (suc n) + λ s → pickPetal x (CTB' n α trich₁ trich₂ + (cofib→cofib n α w (suc n , <ᵗ-trans <ᵗsucm <ᵗsucm) + trich₁ trich₂ (preBTC' n α trich₁ trich₂ a .fst s )))) + (ℤFinGenerator (fin→SphereBouquet/Cell (fst α) trich₁ trich₂ w) x) + (pickCell n trich₁) + (nonVanish n α _ _ x w) + (λ x' q → ⊥.rec (≠pickCell→Empty trich₁ x' q)))) + where + pickCell : (n : ℕ) (t : _) + → Fin (ScardGen (suc n) (suc n) (Trichotomyᵗ-suc t)) + pickCell n (lt s) = ⊥.rec (¬m<ᵗm s) + pickCell n (eq s) = fzero + pickCell n (gt s) = ⊥.rec (¬m<ᵗm s) + + ≠pickCell→Empty : (t : _) + (s : Fin (ScardGen (suc n) (suc n) (Trichotomyᵗ-suc t))) + → ¬ s ≡ pickCell n t → ⊥ + ≠pickCell→Empty (eq x) (zero , tt) s = s refl + + -- some abbreviations + module _ (n : ℕ) + (α : FinSphereBouquetMap∙ m k n) + (trich₁ : _) (trich₂ : _) where + + CTB' = BouquetFuns.CTB (suc n) + (SphereBouquet/CardGen m k (suc n) trich₁ trich₂) + (SphereBouquet/αGen m k (fst α) (suc n) trich₁ trich₂) + (SphereBouquet/EqGen m k (suc n) (fst α) (Trichotomyᵗ-suc trich₁) + trich₁ trich₂) + + Pushout→Bouquet' = Pushout→Bouquet (suc n) + (SphereBouquet/CardGen m k (suc n) trich₁ trich₂) + (SphereBouquet/αGen m k (fst α) (suc n) trich₁ trich₂) + (SphereBouquet/EqGen m k (suc n) (fst α) (Trichotomyᵗ-suc trich₁) + trich₁ trich₂) + + Pushout→Bouquet'∘EqGen = Pushout→Bouquet' + ∘ fst (SphereBouquet/EqGen m k (suc n) (fst α) + (Trichotomyᵗ-suc trich₁) trich₁ trich₂) + + preBTC' = preBTC (suc n) + (ScardGen (suc n) (suc n) (Trichotomyᵗ-suc trich₁)) + (λ t → Sgen.Sfam∙ (suc n) n trich₂) + (SαEqGen (suc n) (suc n) (Trichotomyᵗ-suc trich₁) trich₂) + + trich₁ = suc n ≟ᵗ suc n + trich₂ = suc n ≟ᵗ suc (suc n) + + realiseInr : (w : Fin k) + → realise (Sˢᵏᵉˡ (suc n)) → realise (SphereBouquet/ˢᵏᵉˡ (fst α)) + realiseInr w = fst (hasCWskelSphereBouquet/ (fst α) .snd) + ∘ preπ'FinSphereBouquetMapGenerator α w .fst + ∘ invEq (hasCWskelSphere (suc n) .snd) + + Sⁿ→cofib : {n : ℕ} (m k : _) (α : FinSphereBouquetMap∙ m k n) + (w : Fin k) (r : _) (P : _) + → Sgen.Sfam (suc n) r P → SphereBouquet/FamGen m k (fst α) r P + Sⁿ→cofib m k α w (suc r) (lt x) s = tt + Sⁿ→cofib m k α w (suc r) (eq x) s = inr (w , s) + Sⁿ→cofib m k α w (suc r) (gt x) s = inr (inr (w , s)) + + Sⁿ→cofib≡ : {n : ℕ} (m k : _) (α : FinSphereBouquetMap∙ m k n) + (w : Fin k) (r : _) (P : _) (Q : _) (x : Sgen.Sfam (suc n) r Q) + → invEq (SphereBouquet/EqGen m k r (fst α) + (Trichotomyᵗ-suc P) P Q) + (inl (Sⁿ→cofib m k α w r Q x)) + ≡ Sⁿ→cofib m k α w (suc r) (Trichotomyᵗ-suc P) + (invEq (SαEqGen (suc n) r (Trichotomyᵗ-suc P) Q) (inl x)) + Sⁿ→cofib≡ m k α w (suc r) (lt s) Q x = refl + Sⁿ→cofib≡ m k α w (suc r) (eq s) (lt t) x = push w + Sⁿ→cofib≡ m k α w (suc r) (eq s) (eq t) x = + ⊥.rec (falseDichotomies.eq-eq (s , t)) + Sⁿ→cofib≡ m k α w (suc r) (eq s) (gt t) x = + ⊥.rec (¬-suc-n<ᵗn (transport (λ i → s (~ i) <ᵗ r) t)) + Sⁿ→cofib≡ m k α w (suc r) (gt s) (lt t) x = ⊥.rec (¬squeeze (t , s)) + Sⁿ→cofib≡ m k α w (suc r) (gt s) (eq t) x = refl + Sⁿ→cofib≡ m k α w (suc r) (gt s) (gt t) x = refl + + module _ (n : ℕ) (α : FinSphereBouquetMap∙ m k n) where + cofib→cofib : (w : _) (n' : Fin (suc (suc (suc n)))) (P : _) (Q : _) + → cofib (invEq (SαEqGen (suc n) (fst n') (Trichotomyᵗ-suc P) Q) ∘ inl) + → cofib (invEq (SphereBouquet/EqGen m k (fst n') (fst α) + (Trichotomyᵗ-suc P) P Q) ∘ inl) + cofib→cofib w n' P Q (inl x) = inl tt + cofib→cofib w n' P Q (inr x) = + inr (Sⁿ→cofib m k α w (suc (fst n')) (Trichotomyᵗ-suc P) x) + cofib→cofib w n' P Q (push a i) = + (push (Sⁿ→cofib m k α w (fst n') Q a) + ∙ λ i → inr (Sⁿ→cofib≡ m k α w (fst n') P Q a i)) i + + module _ (n : ℕ) (w : Fin k) (x : _) + (p : Path (S₊ (suc n)) (ptSn (suc n)) (ptSn (suc n))) where + teGen : ¬ (fst w ≡ fst x) + → (cong (pickPetal x) (push w) ∙∙ + (λ i → pickPetal x (inr (w , p i))) ∙∙ + cong (pickPetal x) (sym (push w))) ≡ refl + teGen nope with (fst x ≟ᵗ fst w) + ... | lt x = sym (rUnit refl) + ... | eq x = ⊥.rec (nope (sym x)) + ... | gt x = sym (rUnit refl) + + teGen' : (fst w ≡ fst x) + → (cong (pickPetal x) (push w) + ∙∙ (λ i → pickPetal x (inr (w , p i))) + ∙∙ cong (pickPetal x) (sym (push w))) ≡ p + teGen' aye with (fst x ≟ᵗ fst w) + ... | lt ine = ⊥.rec (¬m<ᵗm (subst (fst x <ᵗ_) aye ine)) + ... | eq x = sym (rUnit p) + ... | gt ine = ⊥.rec (¬m<ᵗm (subst (_<ᵗ fst x) aye ine)) + + -- key lemma: roughly, Hurewicz map preserves generators + nonVanish : (n : ℕ) (α : _) (trich₁ : _) (trich₂ : _) (x : Fin _) (w : _) + → degree (suc n) (λ s → + pickPetal x + (CTB' n α trich₁ trich₂ + (cofib→cofib n α w (suc n , <ᵗ-trans <ᵗsucm <ᵗsucm) trich₁ trich₂ + (preBTC' n α trich₁ trich₂ (pickCell n trich₁) .fst s)))) + ≡ ℤFinGenerator (fin→SphereBouquet/Cell (fst α) trich₁ trich₂ w) x + nonVanish zero α (eq s) (lt t) x w = + cong (degree (suc zero)) + (funExt (λ a → cong (pickPetal x) + λ i → CTB' zero α (eq (isSetℕ _ _ s refl i)) (lt t) + (cofib→cofib zero α w (1 , tt) (eq (isSetℕ _ _ s refl i)) (lt t) + (preBTC' zero α (eq (isSetℕ _ _ s refl i)) (lt t) fzero .fst a)))) + ∙ main pick∘CTB'∘cofib→cofib∘BTC' refl + where + pick∘CTB'∘cofib→cofib∘BTC' = pickPetal x + ∘ CTB' zero α (eq refl) (lt t) + ∘ cofib→cofib zero α w (1 , tt) (eq refl) (lt t) + ∘ preBTC' zero α (eq refl) (lt t) fzero .fst + + CTB'∘cofib→cofib = + CTB' zero α (eq refl) (lt t) + ∘ cofib→cofib zero α w (1 , <ᵗ-trans <ᵗsucm <ᵗsucm) (eq refl) (lt t) + + lem : cong pick∘CTB'∘cofib→cofib∘BTC' loop + ≡ cong (pickPetal x) (push w) + ∙∙ (λ i → pickPetal x (inr (w , σSn 0 false i))) + ∙∙ cong (pickPetal x) (sym (push w)) + lem = cong (cong (pickPetal x ∘ CTB'∘cofib→cofib)) lem1 + ∙ cong-∙∙ (pickPetal x ∘ CTB'∘cofib→cofib) + (push tt) (λ i₁ → inr (loop i₁)) (sym (push tt)) + ∙ cong₃ _∙∙_∙∙_ refl + lem3 + refl + ∙ comm∙∙S¹ _ _ + where + lem1 : cong (preBTC' zero α (eq refl) (lt t) fzero .fst) loop + ≡ push tt ∙∙ (λ i → inr (loop i)) ∙∙ sym (push tt) + lem1 = cong (λ p → push tt ∙∙ p ∙∙ sym (push tt)) + ((λ j i → inr (cong-∙ (invEq (SαEqGen 1 1 + (Trichotomyᵗ-suc (eq refl)) (lt t))) + (push (fzero , false)) (sym (push (fzero , true))) j i)) + ∙ λ j i → inr (rUnit loop (~ j) i)) + + lem2 : cong (fst (SphereBouquet/EqGen m k (suc zero) (fst α) + (Trichotomyᵗ-suc (eq refl)) (eq refl) (lt t))) + (λ i → inr (w , loop i)) + ≡ (push (w , false) ∙ sym (push (w , true))) + lem2 = (λ j i → transportRefl (SphereBouquet/EqBottomMainGen m k (fst α) + .fst (inr (w , loop i))) j) + ∙ cong-∙ (λ K → ⋁→cofibFst (λ _ → Bool , true) (inr (w , K))) + (merid false) (sym (merid true)) + + lem3 : cong (pickPetal x ∘ CTB' zero α (eq refl) (lt t)) + (λ i → inr (inr (w , loop i))) + ≡ (cong (pickPetal x) (push w ) + ∙∙ (λ i → pickPetal x (inr (w , σSn 0 false i))) + ∙∙ cong (pickPetal x) (sym (push w))) + lem3 = + cong (cong (pickPetal x)) + ((λ j i → Pushout→Bouquet' zero α (eq refl) (lt t) (lem2 j i)) + ∙ cong-∙ (Pushout→Bouquet' zero α (eq refl) (lt t)) + (push (w , false)) (sym (push (w , true))) + ∙ cong₂ _∙_ refl (sym (rUnit _)) + ∙ sym (GL.assoc _ _ _) ∙ sym (doubleCompPath≡compPath _ _ _)) + ∙ cong-∙∙ (pickPetal x) + (push w) (λ i → (inr (w , σSn 0 false i))) (sym (push w)) + + comm∙∙S¹ : ∀ (p q : ΩS¹) → p ∙∙ q ∙∙ sym p ≡ q + comm∙∙S¹ p q = doubleCompPath≡compPath p q (sym p) + ∙ comm-ΩS¹ p _ + ∙ sym (GL.assoc _ _ _) + ∙ cong (q ∙_) (lCancel p) + ∙ sym (rUnit q) + + pick∘CTB'∘cofib→cofib∘BTC'-const : ¬ (fst w ≡ fst x) → (r : _) + → pick∘CTB'∘cofib→cofib∘BTC' r ≡ base + pick∘CTB'∘cofib→cofib∘BTC'-const nope base = refl + pick∘CTB'∘cofib→cofib∘BTC'-const nope (loop i) j = + (lem ∙ teGen _ α zero w x (σS false) nope) j i + + pick∘CTB'∘cofib→cofib∘BTC'-id : (fst w ≡ fst x) → (r : _) + → pick∘CTB'∘cofib→cofib∘BTC' r ≡ r + pick∘CTB'∘cofib→cofib∘BTC'-id aye base = refl + pick∘CTB'∘cofib→cofib∘BTC'-id aye (loop i) j = + (lem ∙ teGen' _ α zero w x (σS false) aye) j i + + main : (pick∘CTB'∘cofib→cofib∘BTC* : _) + → pick∘CTB'∘cofib→cofib∘BTC' ≡ pick∘CTB'∘cofib→cofib∘BTC* + → degree 1 pick∘CTB'∘cofib→cofib∘BTC* ≡ ℤFinGenerator w x + main _ p with (fst w ≟ᵗ fst x) + ... | lt wa = + cong (degree (suc zero)) + (sym p ∙ funExt (λ d → pick∘CTB'∘cofib→cofib∘BTC'-const + (λ s → ¬m<ᵗm (subst (_<ᵗ fst x) s wa)) d)) + ∙ degreeConst (suc zero) + ... | eq x = (cong (degree (suc zero)) (sym p) + ∙ cong (degree 1) (funExt (pick∘CTB'∘cofib→cofib∘BTC'-id x))) + ∙ degreeIdfun (suc zero) + ... | gt wa = + cong (degree (suc zero)) + (sym p ∙ funExt (λ d → pick∘CTB'∘cofib→cofib∘BTC'-const + (λ s → ¬m<ᵗm (subst (fst x <ᵗ_) s wa)) d)) + ∙ degreeConst (suc zero) + + nonVanish (suc n) α (eq s) (lt t) x w = + cong (degree (suc (suc n))) + (funExt (λ asd → cong (pickPetal x) + λ i → CTB' (suc n) α (eq (isSetℕ _ _ s refl i)) (lt t) + (cofib→cofib (suc n) α w (suc (suc n) , <ᵗ-trans <ᵗsucm <ᵗsucm) + (eq (isSetℕ _ _ s refl i)) (lt t) + (preBTC' (suc n) α (eq (isSetℕ _ _ s refl i)) (lt t) fzero .fst asd)))) + ∙ TR.rec (isProp→isOfHLevelSuc n (isSetℤ _ _)) + (λ hyp → main hyp (discreteℕ _ _) unlock) + (isConnectedPath _ + (isConnectedPath _ (sphereConnected (suc (suc n))) _ _) + (cong (λ x₃ → pickPetal x (CTB'∘cofib→cofib x₃)) (push tt)) + refl .fst) + where + pick∘CTB'∘cofib→cofib∘BTC' = pickPetal x + ∘ CTB' (suc n) α (eq refl) (lt t) + ∘ cofib→cofib (suc n) α w (suc (suc n) , <ᵗ-trans <ᵗsucm <ᵗsucm) + (eq refl) (lt t) + ∘ preBTC' (suc n) α (eq refl) (lt t) fzero .fst + + CTB'∘cofib→cofib = CTB' (suc n) α (eq refl) (lt t) + ∘ cofib→cofib (suc n) α w (suc (suc n) , <ᵗ-trans <ᵗsucm <ᵗsucm) + (eq refl) (lt t) + + module _ (hyp : cong (λ w → pickPetal x (CTB'∘cofib→cofib w)) (push tt) + ≡ refl) where + lem : (a : _) → cong pick∘CTB'∘cofib→cofib∘BTC' (merid a) + ≡ cong (pickPetal x) (push w) + ∙∙ cong (pickPetal x) (λ i → inr (w , σS a i)) + ∙∙ cong (pickPetal x) (sym (push w)) + lem a = cong (cong (pickPetal x ∘ CTB'∘cofib→cofib)) (lem1 a) + ∙ cong-∙∙ (pickPetal x ∘ CTB'∘cofib→cofib) _ _ _ + ∙ cong₃ _∙∙_∙∙_ + hyp + (cong (cong (pickPetal x)) + (cong-∙ (λ z → CTB'∘cofib→cofib (inr z)) + (merid a) (sym (merid (ptSn (suc n)))) + ∙ cong₂ _∙_ (lem2 a) (cong sym (lem2 (ptSn (suc n)))) + ∙ pathLem _ (push w) + (λ i → inr (w , σS a i)) _ + (λ i → inr (transportRefl w (~ i) , north)) + (λ i → inr (w , σS (ptSn (suc n)) i)) + λ j i → inr (w , rCancel (merid (ptSn (suc n))) (~ j) i)) + ∙ cong-∙∙ (pickPetal x) _ _ _) + (cong sym hyp) + ∙ sym (rUnit _) + where + lem1 : (a : _) + → cong (preBTC' (suc n) α (eq refl) (lt t) fzero .fst) (merid a) + ≡ push tt ∙∙ (λ i → inr (σS a i)) ∙∙ sym (push tt) + lem1 a = cong (λ p → push tt ∙∙ p ∙∙ sym (push tt)) + λ j i → inr ((cong-∙ (invEq + (SαEqGen (suc (suc n)) (suc (suc n)) (Trichotomyᵗ-suc (eq refl)) + (lt t))) (push (fzero , a)) (sym (push (fzero , ptSn (suc n)))) + ∙ cong₂ _∙_ (cong merid (transportRefl a)) + (cong (sym ∘ merid) (transportRefl (ptSn (suc n))))) j i) + + transportRefl-transportRefl : ∀ {ℓ} {A : Type ℓ} (x : A) + → Square {A = A} (λ i → transportRefl (transportRefl x i) (~ i)) + refl refl refl + transportRefl-transportRefl x = + (gen _ _ _ _ _ _ _ _ + (λ i j → transportRefl (transportRefl x i) j) + ∙ rCancel _) + where + gen : ∀ {ℓ} {A : Type ℓ} (x y : A) (p : x ≡ y) + (z : A) (l : x ≡ z) (w : A) (r : y ≡ w) (q : z ≡ w) + (P : Square p q l r) → (λ i → P i (~ i)) ≡ r ∙ sym q + gen x = J> (J> (J> (J> rUnit refl))) + + lem2 : (a : _) + → cong (Pushout→Bouquet'∘EqGen (suc n) α (eq refl) (lt t)) + (λ i → inr (w , merid a i)) + ≡ push w + ∙∙ (λ i → (inr (w , σS a i))) + ∙∙ λ i → (inr (transportRefl w (~ i) , north)) + lem2 a = + (cong (cong (Pushout→Bouquet' (suc n) α (eq refl) (lt t))) + (λ j i → transport refl (push (w , a) i)) + ∙ cong (cong (Pushout→Bouquet' (suc n) α (eq refl) (lt t))) + (cong₂ _∙_ refl refl) + ∙ cong-∙ (Pushout→Bouquet' (suc n) α (eq refl) (lt t)) _ _ + ∙ cong₃ _∙∙_∙∙_ refl + (cong₂ _∙_ refl refl) + refl + ∙ cong₂ _∙_ refl + (λ j i → inr (transportRefl-transportRefl w j i , north)) + ∙ sym (GL.assoc _ _ _) + ∙ (λ j → push (transportRefl w j) + ∙ (λ i → inr (transportRefl w j + , σS (transportRefl a j) i)) + ∙ λ i → inr (transp (λ i₁ → Fin k) (j ∧ ~ i) w , north)) + ∙ sym (doubleCompPath≡compPath _ _ _)) + + pathLem : ∀ {ℓ} {A : Type ℓ} {x : A} + (y : A) (p : x ≡ y) (q : y ≡ y) (z : A) (r : y ≡ z) + (q2 : y ≡ y) → refl ≡ q2 + → (p ∙∙ q ∙∙ r) ∙ (sym r ∙∙ sym q2 ∙∙ sym p) + ≡ (p ∙∙ q ∙∙ sym p) + pathLem = + J> λ q → J> (J> cong₂ _∙_ (sym (rUnit q)) (sym (rUnit refl))) + + pick∘CTB'∘cofib→cofib∘BTC'-const : ¬ (fst w ≡ fst x) → (r : _) + → pick∘CTB'∘cofib→cofib∘BTC' r ≡ north + pick∘CTB'∘cofib→cofib∘BTC'-const nope north = refl + pick∘CTB'∘cofib→cofib∘BTC'-const nope south = refl + pick∘CTB'∘cofib→cofib∘BTC'-const nope (merid a i) j = + (lem a ∙ teGen _ α (suc n) w x (σS a) nope) j i + + pick∘CTB'∘cofib→cofib∘BTC'-id : (fst w ≡ fst x) → (r : _) + → pick∘CTB'∘cofib→cofib∘BTC' r ≡ r + pick∘CTB'∘cofib→cofib∘BTC'-id aye north = refl + pick∘CTB'∘cofib→cofib∘BTC'-id aye south = merid (ptSn (suc n)) + pick∘CTB'∘cofib→cofib∘BTC'-id aye (merid a i) j = + ((lem a ∙ teGen' _ α (suc n) w x (σS a) aye) + ◁ symP (compPath-filler (merid a) (sym (merid (ptSn (suc n)))))) j i + + main : Dec (fst w ≡ fst x) → (l : _) + → degree (suc (suc n)) pick∘CTB'∘cofib→cofib∘BTC' + ≡ ℤFinGenerator* w x {ℓ-zero} l + main (yes p) l = + cong (degree (suc (suc n))) (funExt (pick∘CTB'∘cofib→cofib∘BTC'-id p)) + ∙ degreeIdfun (suc (suc n)) + ∙ sym (ℤFinGenerator*₁ w x l p) + main (no q) l = + cong (degree (suc (suc n))) (funExt (pick∘CTB'∘cofib→cofib∘BTC'-const q)) + ∙ degreeConst (suc (suc n)) + ∙ sym (ℤFinGenerator*₀ w x l q) + + nonVanish n α (eq s) (eq t) x w = + ⊥.rec (¬m<ᵗm (subst (_<ᵗ suc (suc n)) t <ᵗsucm)) + nonVanish n α (eq s) (gt t) x w = + ⊥.rec (¬-suc-n<ᵗn t) + nonVanish n α (gt s) trich₂ x w = ⊥.rec (¬m<ᵗm s) + + finCellApproxInr : (w : _) + → finCellApprox (Sˢᵏᵉˡ (suc n)) (SphereBouquet/ˢᵏᵉˡ (fst α)) + (realiseInr w) (suc (suc (suc n))) + fmap (fst (finCellApproxInr w)) (r , s) = + Sⁿ→cofib m k α w r (r ≟ᵗ suc (suc n)) + fcomm (fst (finCellApproxInr w)) (r , s) = + Sⁿ→cofib≡ m k α w r _ _ + snd (finCellApproxInr w) = + →FinSeqColimHomotopy _ _ λ s → + cong (incl {n = suc (suc (suc n))}) + (lem1 _ s + ∙ cong (SphereBouquet/FamTopElementGen m k (suc (suc (suc n))) (fst α) + <ᵗsucm (suc (suc (suc n)) ≟ᵗ suc (suc n)) .fst + ∘ preπ'FinSphereBouquetMapGenerator α w .fst) + (sym (lem2 s))) + where + lem1 : (P : _) (s : _) + → Sⁿ→cofib m k α w (suc (suc (suc n))) P s + ≡ SphereBouquet/FamTopElementGen m k (suc (suc (suc n))) + (fst α) <ᵗsucm P .fst + (preπ'FinSphereBouquetMapGenerator α w .fst + (invEq + (SfamGenTopElement (suc n) (suc (suc (suc n))) + (<ᵗ-trans <ᵗsucm <ᵗsucm) P) s)) + lem1 (lt x) = λ _ → refl + lem1 (eq x) = + ⊥.rec (¬m<ᵗm (subst (_<ᵗ suc (suc n)) (λ i → predℕ (x (~ i))) <ᵗsucm)) + lem1 (gt x) = λ _ → refl + + lem2 : (x : Sfam (suc n) (suc (suc (suc n)))) + → invEq (hasCWskelSphere (suc n) .snd) (incl x) + ≡ invEq (SfamGenTopElement (suc n) (suc (suc (suc n))) + (<ᵗ-trans {n = n} {m = suc n} {k = suc (suc n)} <ᵗsucm <ᵗsucm) + (suc (suc (suc n)) ≟ᵗ suc (suc n))) x + lem2 x = cong (invEq (hasCWskelSphere (suc n) .snd)) genLem + ∙ retEq (hasCWskelSphere (suc n) .snd) _ + where + gen : (P : _) (Q : _) (x : Sgen.Sfam (suc n) (suc (suc (suc n))) P) + → x ≡ invEq (SαEqGen (suc n) (suc (suc n)) P Q) + (inl (fst (SfamGenTopElement (suc n) (suc (suc n)) <ᵗsucm Q) + (invEq (SfamGenTopElement (suc n) (suc (suc (suc n))) + (<ᵗ-trans <ᵗsucm <ᵗsucm) P) x))) + gen P (lt s) x = ⊥.rec (¬squeeze (s , <ᵗsucm)) + gen (lt t) (eq s) x = refl + gen (eq t) (eq s) x = + ⊥.rec (¬m<ᵗm (subst (_<ᵗ suc n) (cong (predℕ ∘ predℕ) (sym t)) <ᵗsucm)) + gen (gt t) (eq s) x = refl + gen P (gt s) x = ⊥.rec (¬m<ᵗm s) + + genLem : incl x + ≡ fst (hasCWskelSphere (suc n) .snd) + (invEq (SfamGenTopElement (suc n) (suc (suc (suc n))) + (<ᵗ-trans {n = n} {m = suc n} {k = suc (suc n)} + <ᵗsucm <ᵗsucm) + (suc (suc (suc n)) ≟ᵗ suc (suc n))) x) + genLem = cong incl (gen _ _ x) ∙ sym (push _) + + + cofib→cofib≡ : (w : _) (n' : Fin (suc (suc (suc n)))) (x : _) + → cofib→cofib n α w n' (fst n' ≟ᵗ suc n) (fst n' ≟ᵗ suc (suc n)) x + ≡ prefunctoriality.fn+1/fn (suc (suc (suc n))) + (fst (finCellApproxInr w)) n' x + cofib→cofib≡ w n' (inl x) = refl + cofib→cofib≡ w n' (inr x) = refl + cofib→cofib≡ w n' (push a i) = refl + + finCellApproxInr' : (w : _) + → finCellApprox (Sˢᵏᵉˡ (suc n)) + (SphereBouquet/ˢᵏᵉˡ (fst α)) (realiseInr w) (suc (suc (suc (suc n)))) + fmap (fst (finCellApproxInr' w)) m' x = + Sⁿ→cofib m k α w (fst m') (fst m' ≟ᵗ suc (suc n)) x + fcomm (fst (finCellApproxInr' w)) n x = + Sⁿ→cofib≡ m k α w (fst n) _ _ x + snd (finCellApproxInr' w) = →FinSeqColimHomotopy _ _ λ s → + cong (incl {n = suc (suc (suc (suc n)))}) + (cong (Sⁿ→cofib m k α w (suc (suc (suc (suc n)))) + ((4 +ℕ n) ≟ᵗ (2 +ℕ n))) + (sym (secEq (_ , SˢᵏᵉˡConverges (suc n) 1) s)) + ∙ mainLemma ((4 +ℕ n) ≟ᵗ (2 +ℕ n)) + (suc (suc (suc n)) ≟ᵗ suc n) + (suc (suc (suc n)) ≟ᵗ suc (suc n)) + (invEq (invEq (SαEq (suc n) (1 +ℕ suc (suc n))) + ∘ inl , SˢᵏᵉˡConverges (suc n) 1) s)) + ∙ sym (push _) + ∙ funExt⁻ (snd (finCellApproxInr w)) + (fincl (suc (suc (suc n)) , <ᵗsucm) + (invEq (invEq (SαEq (suc n) (1 +ℕ suc (suc n))) + ∘ inl , SˢᵏᵉˡConverges (suc n) 1) s)) + ∙ cong (realiseInr w) + (push (invEq (invEq (SαEq (suc n) (1 +ℕ suc (suc n))) + ∘ inl , SˢᵏᵉˡConverges (suc n) 1) s)) + ∙ λ i → realiseInr w (incl {n = suc (suc (suc (suc n)))} + (secEq (_ , SˢᵏᵉˡConverges (suc n) 1) s i)) + where + mainLemma : (P : _) (Q : _) (R : _) (s : _) + → Sⁿ→cofib m k α w (suc (suc (suc (suc n)))) P + (invEq (SαEqGen (suc n) (suc (suc (suc n))) P R) (inl s)) + ≡ invEq (SphereBouquet/EqGen m k (suc (suc (suc n))) (fst α) P Q R) + (inl (Sⁿ→cofib m k α w (suc (suc (suc n))) R s)) + mainLemma (lt x) Q R _ = ⊥.rec (¬-suc-n<ᵗn (<ᵗ-trans x <ᵗsucm)) + mainLemma (eq x) Q R _ = + ⊥.rec (¬m<ᵗm (subst (_<ᵗ suc (suc n)) (sym (cong (predℕ ∘ predℕ) x)) + (<ᵗ-trans <ᵗsucm <ᵗsucm))) + mainLemma (gt x) Q (lt s) _ = ⊥.rec (¬-suc-n<ᵗn s) + mainLemma (gt x) Q (eq s) _ = + ⊥.rec (¬m<ᵗm (subst (_<ᵗ suc (suc n)) (sym (cong predℕ s)) <ᵗsucm)) + mainLemma (gt x) (lt t) (gt s) _ = ⊥.rec (¬m<ᵗm (<ᵗ-trans x t)) + mainLemma (gt x) (eq t) (gt s) _ = + ⊥.rec (¬m<ᵗm (subst (_<ᵗ suc (suc (suc n))) (λ i → t (~ i)) x)) + mainLemma (gt x) (gt t) (gt s) _ = refl + + cofib→cofib≡' : (w : _) (n' : Fin (suc (suc (suc n)))) (x : _) + → cofib→cofib n α w n' (fst n' ≟ᵗ suc n) (fst n' ≟ᵗ suc (suc n)) x + ≡ prefunctoriality.fn+1/fn (suc (suc (suc (suc n)))) + (fst (finCellApproxInr' w)) (injectSuc n') x + cofib→cofib≡' w n' (inl x) = refl + cofib→cofib≡' w n' (inr x) = refl + cofib→cofib≡' w n' (push a i) = refl + +private + preHurewiczLemma : (n : ℕ) (X : CWexplicit ℓ) (conX : isConnected 2 (fst X)) + → ((l : _) (str : _) (t : _) + → isEquiv (HurewiczHomAb (CWexplicit→CWskel X .fst (suc (suc (suc n))) + , str) l t n .fst)) + → (x : fst X) → isEquiv (HurewiczHomAb (CWexplicit→CW X) x conX n .fst) + preHurewiczLemma {ℓ = ℓ} n X conX ind' x = + TR.rec (isPropIsEquiv _) + (λ t → subst isEquiv (funExt (altEquiv≡ t)) (altEquiv t .fst .snd)) + (isConnected-CW↪∞ (suc zero) Xˢᵏᵉˡ x₀ .fst) + where + e∞ = snd (snd X) + e∞⃗ = fst e∞ + e∞⃖ = invEq e∞ + x₀ = e∞⃗ x + Xˢᵏᵉˡ = CWexplicit→CWskel X + + X∙ : Pointed _ + fst X∙ = fst X + snd X∙ = x + + X₃₊ₙ = fst Xˢᵏᵉˡ (suc (suc (suc n))) + X₃₊ₙ≅X∞ = realiseSubComplex (suc (suc (suc n))) Xˢᵏᵉˡ + + Xˢᵘᵇ : CW ℓ + Xˢᵘᵇ = realise (subComplex Xˢᵏᵉˡ (suc (suc (suc n)))) + , ∣ (subComplex Xˢᵏᵉˡ (suc (suc (suc n)))) + , idEquiv _ ∣₁ + + module _ (t : fiber (CW↪∞ Xˢᵏᵉˡ (suc zero)) x₀) where + xˢᵘᵇ : fst Xˢᵘᵇ + xˢᵘᵇ = fun X₃₊ₙ≅X∞ (CWskel∙ Xˢᵏᵉˡ (fst t) (suc (suc n))) + + Xˢᵘᵇ∙ : Pointed _ + fst Xˢᵘᵇ∙ = fst Xˢᵘᵇ + snd Xˢᵘᵇ∙ = xˢᵘᵇ + + Xˢᵘᵇ→∙X₃₊ₙ : Xˢᵘᵇ∙ →∙ (X₃₊ₙ , CWskel∙ Xˢᵏᵉˡ (fst t) (suc (suc n))) + fst Xˢᵘᵇ→∙X₃₊ₙ = inv' X₃₊ₙ≅X∞ + snd Xˢᵘᵇ→∙X₃₊ₙ = leftInv X₃₊ₙ≅X∞ _ + + Xˢᵘᵇ→∙X : Xˢᵘᵇ∙ →∙ X∙ + Xˢᵘᵇ→∙X = (e∞⃖ , e∞⃖-incl) + ∘∙ (incl∙ Xˢᵏᵉˡ (fst t) + ∘∙ Xˢᵘᵇ→∙X₃₊ₙ) + where + e∞⃖-incl : e∞⃖ (incl (fst t)) ≡ x + e∞⃖-incl = cong e∞⃖ (snd t) ∙ retEq e∞ x + + isConn₃₊ₙXˢᵘᵇ→∙X : isConnectedFun (suc (suc (suc n))) (fst Xˢᵘᵇ→∙X) + isConn₃₊ₙXˢᵘᵇ→∙X = + isConnectedComp (invEq (snd (snd X))) _ (suc (suc (suc n))) + (isEquiv→isConnected _ (snd (invEquiv e∞)) _) + (isConnectedComp incl (Xˢᵘᵇ→∙X₃₊ₙ .fst) (suc (suc (suc n))) + (isConnected-CW↪∞ (suc (suc (suc n))) _) + (isEquiv→isConnected _ + (isoToIsEquiv (invIso X₃₊ₙ≅X∞)) _)) + + isConnX₃₊ₙ : isConnected 2 (fst Xˢᵏᵉˡ (suc (suc (suc n)))) + isConnX₃₊ₙ = + connectedFunPresConnected 2 + (subst (isConnected 2) (ua e∞) conX) (λ b → + isConnectedSubtr' (suc n) 2 + (isConnected-CW↪∞ (suc (suc (suc n))) Xˢᵏᵉˡ b)) + + isConnXˢᵘᵇ : isConnected 2 (fst Xˢᵘᵇ) + isConnXˢᵘᵇ = subst (isConnected 2) (isoToPath X₃₊ₙ≅X∞) isConnX₃₊ₙ + + HurewiczInst = HurewiczHomAb Xˢᵘᵇ xˢᵘᵇ isConnXˢᵘᵇ n + + isEquivHurewiczInst : isEquiv (fst HurewiczInst) + isEquivHurewiczInst = + transport (λ i → isEquiv (fst (HurewiczHomAb' i))) + (ind' (CWskel∙ Xˢᵏᵉˡ (fst t) (suc (suc n))) + ∣ subComplex Xˢᵏᵉˡ (suc (suc (suc n))) + , isoToEquiv X₃₊ₙ≅X∞ ∣₁ + isConnX₃₊ₙ) + where + lem : Path (Σ[ X ∈ CW ℓ ] (Σ[ x ∈ fst X ] isConnected 2 (fst X))) + ((_ , ∣ subComplex Xˢᵏᵉˡ (suc (suc (suc n))) + , isoToEquiv X₃₊ₙ≅X∞ ∣₁) + , (CWskel∙ Xˢᵏᵉˡ (fst t) (suc (suc n))) , isConnX₃₊ₙ) + (Xˢᵘᵇ , xˢᵘᵇ , isConnXˢᵘᵇ) + lem = ΣPathP ((Σ≡Prop (λ _ → squash₁) + (isoToPath X₃₊ₙ≅X∞)) + , (ΣPathPProp (λ _ → isPropIsContr) + (toPathP (cong incl (transportRefl _))))) + + HurewiczHomAb' : (i : I) → _ + HurewiczHomAb' i = + HurewiczHomAb (lem i .fst) (lem i .snd .fst) (lem i .snd .snd) n + + altEquiv : AbGroupEquiv (AbelianizationAbGroup (π'Gr n (fst X , x))) + (H̃ᶜʷAb (fst X , ∣ snd X ∣₁) n) + altEquiv = + compGroupEquiv + (invGroupEquiv (connected→Abπ'Equiv n Xˢᵘᵇ→∙X isConn₃₊ₙXˢᵘᵇ→∙X)) + (compGroupEquiv ((fst HurewiczInst , isEquivHurewiczInst) + , snd HurewiczInst) + (subComplexHomologyEquiv Xˢᵏᵉˡ n (suc (suc (suc n))) <ᵗsucm)) + + altEquiv≡ : (a : _) → altEquiv .fst .fst a + ≡ HurewiczHomAb (fst X , ∣ snd X ∣₁) x conX n .fst a + altEquiv≡ = Abi.elimProp _ (λ _ → squash/ _ _) + (λ f → H̃ᶜʷ→≡Hurewicz _ + ∙ cong (HurewiczHomAb (fst X , ∣ snd X ∣₁) x conX n .fst) + (secEq (fst (connected→Abπ'Equiv n Xˢᵘᵇ→∙X isConn₃₊ₙXˢᵘᵇ→∙X)) + (η f))) + where + H̃ᶜʷ→compInst = H̃ᶜʷ→comp {C = Sᶜʷ (suc n)} {D = Xˢᵘᵇ} + {E = realise Xˢᵏᵉˡ , ∣ Xˢᵏᵉˡ , (idEquiv _) ∣₁} (suc n) + (incl ∘ inv' X₃₊ₙ≅X∞) + + H̃ᶜʷ→≡Hurewicz : (t : _) + → H̃ᶜʷ→ {C = subCW (suc (suc (suc n))) X} + {D = realise Xˢᵏᵉˡ , ∣ Xˢᵏᵉˡ , (idEquiv _) ∣₁} (suc n) incl + .fst (HurewiczHomAb Xˢᵘᵇ xˢᵘᵇ isConnXˢᵘᵇ n .fst (η t)) + ≡ HurewiczHomAb (fst X , ∣ snd X ∣₁) x conX n .fst + (fst (fst (connected→Abπ'Equiv n Xˢᵘᵇ→∙X isConn₃₊ₙXˢᵘᵇ→∙X)) (η t)) + H̃ᶜʷ→≡Hurewicz = + ST.elim (λ _ → isProp→isSet (squash/ _ _)) λ g → + funExt⁻ (cong fst (sym (H̃ᶜʷ→compInst (fst g)))) (genHₙSⁿ n) + ∙ λ i → H̃ᶜʷ→ {C = Sᶜʷ (suc n)} + {D = realise Xˢᵏᵉˡ , ∣ Xˢᵏᵉˡ , (idEquiv _) ∣₁} (suc n) + (λ z → secEq e∞ (incl (Xˢᵘᵇ→∙X₃₊ₙ .fst (fst g z))) (~ i)) + .fst (genHₙSⁿ n) + +-- Finally, the main theorem +HurewiczTheorem : (n : ℕ) (X : CW ℓ-zero) + (conX : isConnected (suc (suc n)) (fst X)) (x : _) + → isEquiv (HurewiczHomAb X x (isConnectedSubtr' n 2 conX) n .fst) +HurewiczTheorem n = + uncurry λ X → PT.elim (λ _ → isPropΠ2 λ _ _ → isPropIsEquiv _) + λ cw isc → PT.rec (isPropΠ (λ _ → isPropIsEquiv _)) + (λ cw' → main X cw cw' isc) + (makeConnectedCW n cw isc) + where + isEqTransport : (cw1 cw2 : CW ℓ) (P : cw1 ≡ cw2) + (con1 : isConnected 2 (fst cw1)) (con2 : isConnected 2 (fst cw2)) + (x' : fst cw1) (x : fst cw2) (PP : PathP (λ i → fst (P i)) x' x) + → isEquiv (HurewiczHomAb cw1 x' con1 n .fst) + → isEquiv (HurewiczHomAb cw2 x con2 n .fst) + isEqTransport cw1 = + J> λ con1 con2 x' → + J> (subst (λ c → isEquiv (HurewiczHomAb cw1 x' c n .fst)) + (isPropIsContr _ _)) + + module _ (X : Type) (cw : hasCWskel X) (cw' : isConnectedCW n X) + (isc : isConnected (suc (suc n)) X) (x : X) where + Xˢᵏᵉˡ Xˢᵏᵉˡ' : CWskel ℓ-zero + Xˢᵏᵉˡ = fst cw + Xˢᵏᵉˡ' = connectedCWskel→CWskel (fst cw') + + Xᶜʷ Xᶜʷ' : CWexplicit ℓ-zero + Xᶜʷ = X , cw + Xᶜʷ' = X , Xˢᵏᵉˡ' , (invEquiv (snd cw')) + + liftLem : (A : CW ℓ-zero) (a : fst A) (e : isConnected 2 (Lift (fst A))) + → Path (Σ[ A ∈ CW ℓ-zero ] (Σ[ a ∈ fst A ] isConnected 2 (fst A))) + (A , a , subst (isConnected 2) (ua (invEquiv LiftEquiv)) e) + ((CWLift ℓ-zero A) , (lift a , e)) + liftLem A a e = + ΣPathP ((Σ≡Prop (λ _ → squash₁) (ua LiftEquiv)) + , (ΣPathPProp (λ _ → isPropIsContr) + λ i → ua-gluePt LiftEquiv i a)) + + main : isEquiv (HurewiczHomAb (X , ∣ cw ∣₁) x + (isConnectedSubtr' n 2 isc) n .fst) + main = + isEqTransport (CWexplicit→CW Xᶜʷ') (CWexplicit→CW Xᶜʷ) + (Σ≡Prop (λ _ → squash₁) refl) + (isConnectedSubtr' n 2 isc) + (isConnectedSubtr' n 2 isc) x x refl + (preHurewiczLemma n Xᶜʷ' + (isConnectedSubtr' n 2 isc) + (λ l str con' → PT.rec (isPropIsEquiv _) + (λ {(α , e) → TR.rec (isPropIsEquiv _) + (λ linl → isEqTransport _ (_ , str) (sym e) + (subst (isConnected 2) (cong fst e) con') + con' + (lift (inl tt)) + l + (toPathP (sym linl)) + (transport (λ i → isEquiv + (HurewiczHomAb + (liftLem (SphereBouquet/ᶜʷ (fst α)) (inl tt) + (subst (isConnected 2) + (λ i → fst (e i)) con') i .fst) + (liftLem (SphereBouquet/ᶜʷ (fst α)) (inl tt) + (subst (isConnected 2) + (λ i → fst (e i)) con') i .snd .fst) + (liftLem (SphereBouquet/ᶜʷ (fst α)) (inl tt) + (subst (isConnected 2) + (λ i → fst (e i)) con') i .snd .snd) + n .fst)) + (HurewiczMapCofibEquiv α _))) + (isConnectedPath 1 con' l + (transport (sym (cong fst e)) (lift (inl tt))) + .fst)}) + (connectedCW≃CofibFinSphereBouquetMap n X cw' str)) x) diff --git a/Cubical/CW/Instances/Empty.agda b/Cubical/CW/Instances/Empty.agda new file mode 100644 index 0000000000..a8ab3319f4 --- /dev/null +++ b/Cubical/CW/Instances/Empty.agda @@ -0,0 +1,75 @@ +{- +The empty type as a CW complex +-} +module Cubical.CW.Instances.Empty where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Fin.Inductive +open import Cubical.Data.Empty + +open import Cubical.HITs.Pushout +open import Cubical.HITs.SequentialColimit +open import Cubical.HITs.PropositionalTruncation + +open import Cubical.CW.Base + +module _ (ℓ : Level) where + finCWskel⊥* : finCWskel ℓ 0 + fst finCWskel⊥* _ = ⊥* + fst (fst (snd finCWskel⊥*)) _ = 0 + fst (snd (fst (snd finCWskel⊥*))) _ (x , _) = lift (¬Fin0 x) + fst (snd (snd (fst (snd finCWskel⊥*)))) () + snd (snd (snd (fst (snd finCWskel⊥*)))) n = + uninhabEquiv (λ()) + λ x → invEq LiftEquiv + ((Iso.inv (PushoutEmptyFam (λ x → ¬Fin0 (fst x)) ¬Fin0)) x) + snd (snd finCWskel⊥*) _ = uninhabEquiv _ (λ{()}) .snd + + CWskel⊥* : CWskel ℓ + CWskel⊥* = finCWskel→CWskel 0 finCWskel⊥* + + hasFinCWskel⊥* : hasFinCWskel ⊥* + fst hasFinCWskel⊥* = 0 + fst (snd hasFinCWskel⊥*) = finCWskel⊥* + snd (snd hasFinCWskel⊥*) = + uninhabEquiv (λ{()}) λ{ (incl ()) ; (push () i)} + + hasCWskel⊥* : hasCWskel ⊥* + fst hasCWskel⊥* = CWskel⊥* + snd hasCWskel⊥* = + uninhabEquiv (λ{()}) λ{ (incl ()) ; (push () i)} + + finCW⊥* : finCW ℓ + fst finCW⊥* = ⊥* + snd finCW⊥* = ∣ hasFinCWskel⊥* ∣₁ + +finCWskel⊥ : finCWskel ℓ-zero 0 +fst finCWskel⊥ _ = ⊥ +fst (fst (snd finCWskel⊥)) _ = 0 +fst (snd (fst (snd finCWskel⊥))) _ (x , _) = ¬Fin0 x +fst (snd (snd (fst (snd finCWskel⊥)))) () +snd (snd (snd (fst (snd finCWskel⊥)))) n = + uninhabEquiv (λ()) + (Iso.inv (PushoutEmptyFam (λ x → ¬Fin0 (fst x)) ¬Fin0)) +snd (snd finCWskel⊥) _ = uninhabEquiv _ (λ{()}) .snd + +CWskel⊥ : CWskel ℓ-zero +CWskel⊥ = finCWskel→CWskel 0 finCWskel⊥ + +hasFinCWskel⊥ : hasFinCWskel ⊥ +fst hasFinCWskel⊥ = 0 +fst (snd hasFinCWskel⊥) = finCWskel⊥ +snd (snd hasFinCWskel⊥) = + uninhabEquiv (λ{()}) λ{ (incl ()) ; (push () i)} + +hasCWskel⊥ : hasCWskel ⊥ +fst hasCWskel⊥ = CWskel⊥ +snd hasCWskel⊥ = + uninhabEquiv (λ{()}) λ{ (incl ()) ; (push () i)} + +finCW⊥ : finCW ℓ-zero +fst finCW⊥ = ⊥ +snd finCW⊥ = ∣ hasFinCWskel⊥ ∣₁ diff --git a/Cubical/CW/Instances/Join.agda b/Cubical/CW/Instances/Join.agda new file mode 100644 index 0000000000..6d93c2df2d --- /dev/null +++ b/Cubical/CW/Instances/Join.agda @@ -0,0 +1,24 @@ +{- +CW structure on joins +-} + +module Cubical.CW.Instances.Join where + +open import Cubical.Foundations.Prelude + +open import Cubical.CW.Base +open import Cubical.CW.Instances.Pushout +open import Cubical.CW.Instances.Sigma + +open import Cubical.HITs.Join + +-- joins as finite CW complexes +isFinCWJoinPushout : {ℓ : Level} (A B : finCW ℓ) + → isFinCW (joinPushout (fst A) (fst B)) +isFinCWJoinPushout A B = isFinCWPushout (_ , (isFinCW× A B)) A B fst snd + +isFinCWJoin : {ℓ : Level} (A B : finCW ℓ) → isFinCW (join (fst A) (fst B)) +isFinCWJoin A B = + subst isFinCW (joinPushout≡join (fst A) (fst B)) (isFinCWJoinPushout A B) + +-- todo: joins of arbitrary complexes diff --git a/Cubical/CW/Instances/Lift.agda b/Cubical/CW/Instances/Lift.agda new file mode 100644 index 0000000000..b009b478ef --- /dev/null +++ b/Cubical/CW/Instances/Lift.agda @@ -0,0 +1,63 @@ +{- +Universe lifts of CW complexes +-} +module Cubical.CW.Instances.Lift where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Sequence + +open import Cubical.HITs.Pushout +open import Cubical.HITs.SequentialColimit +open import Cubical.HITs.PropositionalTruncation as PT + +open import Cubical.CW.Base + +private + variable + ℓA : Level + +module _ (ℓ : Level) where + CWskelLift : CWskel ℓA → CWskel (ℓ-max ℓA ℓ) + fst (CWskelLift sk) n = Lift {j = ℓ} (fst sk n) + fst (snd (CWskelLift sk)) = CWskel-fields.card sk + fst (snd (snd (CWskelLift sk))) n (x , a) = + lift (CWskel-fields.α sk n (x , a)) + fst (snd (snd (snd (CWskelLift sk)))) (lift x) = + fst (snd (snd (snd sk))) x + snd (snd (snd (snd (CWskelLift sk)))) n = + compEquiv (invEquiv LiftEquiv) + (compEquiv (CWskel-fields.e sk n) + (pushoutEquiv _ _ _ _ (idEquiv _) + LiftEquiv (idEquiv _) refl refl)) + + hasCWskelLift : {A : Type ℓA} → hasCWskel A → hasCWskel (Lift {j = ℓ} A) + fst (hasCWskelLift (sk , e)) = CWskelLift sk + snd (hasCWskelLift (sk , e)) = + compEquiv (invEquiv LiftEquiv) + (compEquiv e + (invEquiv (isoToEquiv (SeqColimLift {ℓ = ℓ} _)))) + + CWLift : CW ℓA → CW (ℓ-max ℓA ℓ) + fst (CWLift (A , sk)) = Lift {j = ℓ} A + snd (CWLift (A , sk)) = PT.map hasCWskelLift sk + + finCWskelLift : ∀ {n} → finCWskel ℓA n → finCWskel (ℓ-max ℓA ℓ) n + fst (finCWskelLift sk) = CWskelLift (finCWskel→CWskel _ sk) .fst + fst (snd (finCWskelLift sk)) = CWskelLift (finCWskel→CWskel _ sk) .snd + snd (snd (finCWskelLift sk)) k = + compEquiv (invEquiv LiftEquiv) + (compEquiv (_ , snd (snd sk) k) LiftEquiv) .snd + + hasFinCWskelLift : {A : Type ℓA} + → hasFinCWskel A → hasFinCWskel (Lift {j = ℓ} A) + fst (hasFinCWskelLift (dim , sk , e)) = dim + fst (snd (hasFinCWskelLift (dim , sk , e))) = finCWskelLift sk + snd (snd (hasFinCWskelLift c)) = + hasCWskelLift (hasFinCWskel→hasCWskel _ c) .snd + + finCWLift : finCW ℓA → finCW (ℓ-max ℓA ℓ) + fst (finCWLift (A , sk)) = Lift {j = ℓ} A + snd (finCWLift (A , sk)) = PT.map hasFinCWskelLift sk diff --git a/Cubical/CW/Instances/Pushout.agda b/Cubical/CW/Instances/Pushout.agda new file mode 100644 index 0000000000..0a0005680b --- /dev/null +++ b/Cubical/CW/Instances/Pushout.agda @@ -0,0 +1,830 @@ +module Cubical.CW.Instances.Pushout where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Path +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Equiv + +open import Cubical.CW.Base +open import Cubical.CW.Properties +open import Cubical.CW.Map + +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.NatMinusOne +open import Cubical.Data.Nat.Order +open import Cubical.Data.Bool +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Fin.Inductive +open import Cubical.Data.Sigma +open import Cubical.Data.Sequence +open import Cubical.Data.FinSequence + +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.Pushout +open import Cubical.HITs.SequentialColimit +open import Cubical.HITs.SphereBouquet +open import Cubical.HITs.PropositionalTruncation as PT + +open import Cubical.CW.Strictification renaming (strictCWskel to str) +open import Cubical.CW.Approximation + +-- Lemmas +module pushoutCWLemmas where + IsoFinSplit3 : ∀ n m l + → Iso (Fin ((n +ℕ m) +ℕ l)) (((Fin n) ⊎ (Fin m)) ⊎ (Fin l)) + IsoFinSplit3 n m l = + compIso (invIso (Iso-Fin⊎Fin-Fin+ {n +ℕ m}{l})) + (⊎Iso {A = Fin (n +ℕ m)} + {C = (Fin n) ⊎ (Fin m)} + {B = Fin l} + {D = Fin l} + (invIso (Iso-Fin⊎Fin-Fin+ {n}{m})) idIso) + + finSplit3 : ∀ n m l → Fin ((n +ℕ m) +ℕ l) + ≃ ((Fin n) ⊎ (Fin m)) ⊎ (Fin l) + finSplit3 n m l = isoToEquiv (IsoFinSplit3 n m l) + + module _ {ℓ : Level} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : x ≡ z) where + invSides-hfill : I → I → I → A + invSides-hfill i j = + hfill (λ k → λ { (i = i0) → p (k ∧ j) + ; (i = i1) → q (~ j ∧ k) + ; (j = i0) → q (i ∧ k) + ; (j = i1) → p (~ i ∧ k)}) + (inS x) + + invSides-hfill1 : I → I → I → A + invSides-hfill1 i j = + hfill (λ k → λ { (i = i0) → p j + ; (i = i1) → q (~ j ∧ k) + ; (j = i0) → q (i ∧ k) + ; (j = i1) → p (~ i)}) + (inS (p ((~ i) ∧ j))) + + invSides-hfill2 : I → I → I → A + invSides-hfill2 i j = + hfill (λ k → λ { (i = i0) → p (k ∧ j) + ; (i = i1) → q (~ j) + ; (j = i0) → q (i) + ; (j = i1) → p (~ i ∧ k)}) + (inS (q (i ∧ (~ j)))) + +open pushoutCWLemmas + +-- This module defines a CW structure when B, C, D are strict CW skeleta +-- The non-strict version can be derived from there using the fact that every +-- CW skeleton is isomorphic to a strict one +module CWPushout (ℓ : Level) (Bʷ Cʷ Dʷ : CWskel ℓ) + (fʷ : cellMap (str Bʷ) (str Cʷ)) + (gʷ : cellMap (str Bʷ) (str Dʷ)) where + + B = str Bʷ + C = str Cʷ + D = str Dʷ + + f = strictCwMap fʷ + g = strictCwMap gʷ + + open CWskel-fields + open SequenceMap renaming (map to ∣_∣) + + pushoutA : ℕ → Type ℓ + pushoutA zero = B .fst zero + pushoutA (suc n) = + Pushout {A = B .fst n} {B = C .fst (suc n)} {C = D .fst (suc n)} + (inl ∘ ∣ f ∣ n) (inl ∘ ∣ g ∣ n) + + pushoutCells : ℕ → ℕ + pushoutCells zero = (card C zero) +ℕ (card D zero) + pushoutCells (suc n) = (card C (suc n)) +ℕ (card B n) +ℕ (card D (suc n)) + + pushoutMap₀ : (Fin (pushoutCells zero)) × (S⁻ 0) → pushoutA 0 + pushoutMap₀ () + + pushoutMapₛ : (n : ℕ) + → (((A C (suc n)) ⊎ (A B n)) ⊎ (A D (suc n))) × (Susp (S⁻ n)) + → pushoutA (suc n) + pushoutMapₛ n (inl (inl c) , x) = + inl (e C n .fst (α C (suc n) (c , (EquivSphereSusp n) .fst x))) + pushoutMapₛ n (inl (inr b) , north) = inl (∣ f ∣ (suc n) (inr b)) + pushoutMapₛ n (inl (inr b) , south) = inr (∣ g ∣ (suc n) (inr b)) + pushoutMapₛ n (inl (inr b) , merid x i) = + ((λ i → inl (∣ f ∣ (suc n) (push (b , x) (~ i)))) + ∙∙ (push (α B n (b , x))) + ∙∙ (λ i → inr (∣ g ∣ (suc n) (push (b , x) i)))) i + pushoutMapₛ n (inr d , x) = + inr (e D n .fst (α D (suc n) (d , (EquivSphereSusp n) .fst x ))) + + pushoutMap : (n : ℕ) → (Fin (pushoutCells n)) × (S⁻ n) → pushoutA n + pushoutMap zero = pushoutMap₀ + pushoutMap (suc n) (a , x) = + pushoutMapₛ n (finSplit3 (card C (suc n)) (card B n) (card D (suc n)) .fst a + , invEq (EquivSphereSusp n) x) + pushoutSpanₛ : (n : ℕ) → 3-span + 3-span.A0 (pushoutSpanₛ n) = pushoutA (suc n) + 3-span.A2 (pushoutSpanₛ n) = + (((A C (suc n)) ⊎ (A B n)) ⊎ (A D (suc n))) × Susp (S⁻ n) + 3-span.A4 (pushoutSpanₛ n) = ((A C (suc n)) ⊎ (A B n)) ⊎ (A D (suc n)) + 3-span.f1 (pushoutSpanₛ n) = pushoutMapₛ n + 3-span.f3 (pushoutSpanₛ n) = fst + + pushoutSpan : (n : ℕ) → 3-span + 3-span.A0 (pushoutSpan n) = pushoutA (suc n) + 3-span.A2 (pushoutSpan n) = Fin (pushoutCells (suc n)) × S⁻ (suc n) + 3-span.A4 (pushoutSpan n) = Fin (pushoutCells (suc n)) + 3-span.f1 (pushoutSpan n) = pushoutMap (suc n) + 3-span.f3 (pushoutSpan n) = fst + + pushoutSpanEquiv : (n : ℕ) → 3-span-equiv (pushoutSpan n) (pushoutSpanₛ n) + 3-span-equiv.e0 (pushoutSpanEquiv n) = idEquiv (pushoutA (suc n)) + 3-span-equiv.e2 (pushoutSpanEquiv n) = + ≃-× (finSplit3 (card C (suc n)) (card B n) (card D (suc n))) + (isoToEquiv (IsoSphereSusp n)) + 3-span-equiv.e4 (pushoutSpanEquiv n) = + finSplit3 (card C (suc n)) (card B n) (card D (suc n)) + 3-span-equiv.H1 (pushoutSpanEquiv n) _ = refl + 3-span-equiv.H3 (pushoutSpanEquiv n) _ = refl + + pushoutₛIso : (n : ℕ) + → Iso (spanPushout (pushoutSpan n)) (spanPushout (pushoutSpanₛ n)) + pushoutₛIso n = pushoutIso _ _ _ _ + (Σ-cong-equiv + (isoToEquiv (IsoFinSplit3 (card C (suc n)) (card B n) (card D (suc n)))) + λ _ → isoToEquiv (IsoSphereSusp n)) + (idEquiv _) + (isoToEquiv (IsoFinSplit3 (card C (suc n)) (card B n) (card D (suc n)))) + refl + refl + + pushoutIso₀-fun : pushoutA (suc zero) → Pushout pushoutMap₀ fst + pushoutIso₀-fun (inl x) = + inr (Iso.fun (Iso-Fin⊎Fin-Fin+ {card C zero} {card D zero}) + (inl (CW₁-discrete C .fst x))) + pushoutIso₀-fun (inr x) = + inr (Iso.fun (Iso-Fin⊎Fin-Fin+ {card C zero} {card D zero}) + (inr (CW₁-discrete D .fst x))) + pushoutIso₀-fun (push a i) with (B .snd .snd .snd .fst a) + pushoutIso₀-fun (push a i) | () + + pushoutIso₀-helper : Fin (card C zero) ⊎ Fin (card D zero) + → pushoutA (suc zero) + pushoutIso₀-helper (inl x) = inl (invEq (CW₁-discrete C) x) + pushoutIso₀-helper (inr x) = inr (invEq (CW₁-discrete D) x) + + pushoutIso₀-inv : Pushout pushoutMap₀ fst → pushoutA (suc zero) + pushoutIso₀-inv (inl x) with (B .snd .snd .snd .fst x) + pushoutIso₀-inv (inl x) | () + pushoutIso₀-inv (inr x) = + pushoutIso₀-helper + (Iso.inv (Iso-Fin⊎Fin-Fin+ {card C zero} {card D zero}) x) + + pushoutIso₀-leftInv : (x : pushoutA (suc zero)) + → pushoutIso₀-inv (pushoutIso₀-fun x) ≡ x + pushoutIso₀-leftInv (inl x) = + (λ i → pushoutIso₀-helper + (Iso.leftInv (Iso-Fin⊎Fin-Fin+ {card C zero} {card D zero}) + (inl (CW₁-discrete C .fst x)) i)) + ∙ λ i → inl (retEq (CW₁-discrete C) x i) + pushoutIso₀-leftInv (inr x) = + (λ i → pushoutIso₀-helper + (Iso.leftInv (Iso-Fin⊎Fin-Fin+ {card C zero} {card D zero}) + (inr (CW₁-discrete D .fst x)) i)) + ∙ λ i → inr (retEq (CW₁-discrete D) x i) + pushoutIso₀-leftInv (push a i) with (B .snd .snd .snd .fst a) + pushoutIso₀-leftInv (push a i) | () + + pushoutIso₀-helper₁ : (x : Fin (card C zero) ⊎ Fin (card D zero)) + → pushoutIso₀-fun (pushoutIso₀-helper x) + ≡ inr (Iso.fun (Iso-Fin⊎Fin-Fin+ {card C zero} {card D zero}) x) + pushoutIso₀-helper₁ (inl x) i = + inr (Iso.fun (Iso-Fin⊎Fin-Fin+ {card C zero} {card D zero}) + (inl (secEq (CW₁-discrete C) x i))) + pushoutIso₀-helper₁ (inr x) i = + inr (Iso.fun (Iso-Fin⊎Fin-Fin+ {card C zero} {card D zero}) + (inr (secEq (CW₁-discrete D) x i))) + + pushoutIso₀-rightInv : (x : Pushout pushoutMap₀ fst) + → pushoutIso₀-fun (pushoutIso₀-inv x) ≡ x + pushoutIso₀-rightInv (inl x) with (B .snd .snd .snd .fst x) + pushoutIso₀-rightInv (inl x) | () + pushoutIso₀-rightInv (inr x) = + pushoutIso₀-helper₁ + (Iso.inv (Iso-Fin⊎Fin-Fin+ {card C zero} {card D zero}) x) + ∙ λ i → inr (Iso.rightInv (Iso-Fin⊎Fin-Fin+ {card C zero} {card D zero}) x i) + + pushoutIso₀ : Iso (pushoutA (suc zero)) (Pushout pushoutMap₀ fst) + pushoutIso₀ = + iso pushoutIso₀-fun pushoutIso₀-inv pushoutIso₀-rightInv pushoutIso₀-leftInv + + -- Technical tool because the definition of S₊ n in the library is messed up + modifySₙ : CWskel ℓ → ℕ → Type ℓ + modifySₙ X zero = X .fst zero + modifySₙ X (suc zero) = X .fst 1 + modifySₙ X (suc (suc n)) = Pushout (λ (a , x) → e X n .fst (α X (suc n) (a , (EquivSphereSusp n) .fst x))) fst + + modifySₙ→id : (X : CWskel ℓ) → (n : ℕ) + → modifySₙ (str X) (suc (suc n)) → (str X) .fst (suc (suc n)) + modifySₙ→id X n (inl x) = inl x + modifySₙ→id X n (inr x) = inr x + modifySₙ→id X n (push (a , x) i) = push (a , (EquivSphereSusp n) .fst x) i + + id→modifySₙ : (X : CWskel ℓ) → (n : ℕ) + → (str X) .fst (suc (suc n)) → modifySₙ (str X) (suc (suc n)) + id→modifySₙ X n (inl x) = inl x + id→modifySₙ X n (inr x) = inr x + id→modifySₙ X n (push (a , x) i) = + ((λ i → inl (e (str X) n .fst (α (str X) (suc n) + (a , secEq (EquivSphereSusp n) x (~ i))))) + ∙∙ (push (a , invEq (EquivSphereSusp n) x)) + ∙∙ refl) i + + id→modifySₙ→id : (X : CWskel ℓ) (n : ℕ) + (x : (str X) .fst (suc (suc n))) → modifySₙ→id X n (id→modifySₙ X n x) ≡ x + id→modifySₙ→id X n (inl x) = refl + id→modifySₙ→id X n (inr x) = refl + id→modifySₙ→id X n (push (a , x) i) j = + hcomp (λ k → + λ{ (i = i0) → inl (strictifyFamα X (suc n) (a , leftInv x (j ∨ k))) + ; (i = i1) → inr a + ; (j = i0) → modifySₙ→id X n + (doubleCompPath-filler + (λ i → inl (e (str X) n .fst + (α (str X) (suc n) (a , leftInv x (~ i))))) + (push (a , fun x)) refl k i) + ; (j = i1) → push (a , x) i}) + (push (a , leftInv x j) i) + where + fun = invEq (EquivSphereSusp n) + inv = (EquivSphereSusp n) .fst + leftInv = secEq (EquivSphereSusp n) + + modifySₙ→id→modifySₙ : (X : CWskel ℓ) (n : ℕ) + (x : modifySₙ (str X) (suc (suc n))) → id→modifySₙ X n (modifySₙ→id X n x) ≡ x + modifySₙ→id→modifySₙ X n (inl x) = refl + modifySₙ→id→modifySₙ X n (inr x) = refl + modifySₙ→id→modifySₙ X n (push (a , x) i) j = + hcomp (λ k → + λ{ (i = i0) → inl (e (str X) n .fst (α (str X) (suc n) + (a , compPath→Square + {p = leftInv (inv x)} {refl} + {λ i → inv (rightInv x i)} {refl} + (cong (λ X → X ∙ refl) + (commPathIsEq (EquivSphereSusp n .snd) x)) k j))) + ; (i = i1) → inr a + ; (j = i0) → doubleCompPath-filler + (λ i → inl (e (str X) n .fst (α (str X) (suc n) + (a , leftInv (inv x) (~ i))))) + (push (a , fun (inv x))) refl k i + ; (j = i1) → push (a , x) i}) + (push (a , rightInv x j) i) + where + fun = invEq (EquivSphereSusp n) + inv = (EquivSphereSusp n) .fst + leftInv = secEq (EquivSphereSusp n) + rightInv = retEq (EquivSphereSusp n) + + modifiedPushout : (n : ℕ) → Type ℓ + modifiedPushout n = Pushout {A = B .fst (suc n)} {B = modifySₙ C (suc (suc n))} + {C = modifySₙ D (suc (suc n))} + (inl ∘ ∣ f ∣ (suc n)) (inl ∘ ∣ g ∣ (suc n)) + + modifiedPushout→Pushout : (n : ℕ) → modifiedPushout n → pushoutA (suc (suc n)) + modifiedPushout→Pushout n (inl x) = inl (modifySₙ→id Cʷ n x) + modifiedPushout→Pushout n (inr x) = inr (modifySₙ→id Dʷ n x) + modifiedPushout→Pushout n (push a i) = push a i + + Pushout→modifiedPushout : (n : ℕ) → pushoutA (suc (suc n)) → modifiedPushout n + Pushout→modifiedPushout n (inl x) = inl (id→modifySₙ Cʷ n x) + Pushout→modifiedPushout n (inr x) = inr (id→modifySₙ Dʷ n x) + Pushout→modifiedPushout n (push a i) = push a i + + Pushout→modP→Pushout : (n : ℕ) → (x : pushoutA (suc (suc n))) + → modifiedPushout→Pushout n (Pushout→modifiedPushout n x) ≡ x + Pushout→modP→Pushout n (inl x) j = inl (id→modifySₙ→id Cʷ n x j) + Pushout→modP→Pushout n (inr x) j = inr (id→modifySₙ→id Dʷ n x j) + Pushout→modP→Pushout n (push a i) j = push a i + + modP→Pushout→modP : (n : ℕ) → (x : modifiedPushout n) + → Pushout→modifiedPushout n (modifiedPushout→Pushout n x) ≡ x + modP→Pushout→modP n (inl x) j = inl (modifySₙ→id→modifySₙ Cʷ n x j) + modP→Pushout→modP n (inr x) j = inr (modifySₙ→id→modifySₙ Dʷ n x j) + modP→Pushout→modP n (push a i) j = push a i + + IsoModifiedPushout : (n : ℕ) → Iso (pushoutA (suc (suc n))) (modifiedPushout n) + Iso.fun (IsoModifiedPushout n) = Pushout→modifiedPushout n + Iso.inv (IsoModifiedPushout n) = modifiedPushout→Pushout n + Iso.rightInv (IsoModifiedPushout n) = modP→Pushout→modP n + Iso.leftInv (IsoModifiedPushout n) = Pushout→modP→Pushout n + + pushoutIsoₛ-filler0 : (n : ℕ) (b : A B n) (x : S⁻ n) → I → I → pushoutA (suc n) + pushoutIsoₛ-filler0 n b x i j = + doubleCompPath-filler + (λ i → inl (∣ f ∣ (suc n) (push (b , x) (~ i)))) + (push (α B n (b , x))) + (λ i → inr (∣ g ∣ (suc n) (push (b , x) i))) i j + + pushoutIsoₛ-filler1 : (n : ℕ) (b : A B n) (x : S⁻ n) + → I → I → I → Pushout (pushoutMapₛ n) fst + pushoutIsoₛ-filler1 n b x i j k = + hfill (λ k → + λ{ (i = i0) → invSides-hfill2 (push (inl (inr b) , north)) + (λ i → inl (inl (∣ f ∣ (suc n) (push (b , x) (~ i))))) + (~ j) (~ k) i1 + ; (i = i1) → invSides-hfill2 (push (inl (inr b) , south)) + (λ i → inl (inr (∣ g ∣ (suc n) (push (b , x) (~ i))))) + (~ j) (~ k) i1 + ; (j = i0) → inl (pushoutIsoₛ-filler0 n b x (~ k) i) + ; (j = i1) → doubleCompPath-filler + (push (inl (inr b) , north)) + (λ _ → inr (inl (inr b))) + (λ i → push (inl (inr b) , south) (~ i)) k i}) + (inS (push (inl (inr b) , merid x i) j)) k + + pushoutIsoₛ-inv↪ : (n : ℕ) → pushoutA (suc n) → modifiedPushout n + pushoutIsoₛ-inv↪ n (inl c) = inl (inl c) + pushoutIsoₛ-inv↪ n (inr d) = inr (inl d) + pushoutIsoₛ-inv↪ n (push b i) = push (inl b) i + + pushoutIsoₛ-filler2 : (n : ℕ) (b : A B n) (x : S⁻ n) + → I → I → I → modifiedPushout n + pushoutIsoₛ-filler2 n b x i j k = + hfill (λ k → + λ{ (i = i0) → pushoutIsoₛ-inv↪ n (pushoutIsoₛ-filler0 n b x k j) + ; (i = i1) → push (inr b) ((~ k) ∧ j) + ; (j = i0) → invSides-hfill1 + (λ i → inl (inl (∣ f ∣ (suc n) (push (b , x) (~ i))))) + (λ _ → push (inr b) i0) i (~ k) i1 + ; (j = i1) → invSides-hfill1 + (λ i → inr (inl (∣ g ∣ (suc n) (push (b , x) (~ i))))) + (λ i → push (inr b) (~ i)) i (~ k) i1}) + (inS (push (push (b , x) i) j)) k + + pushoutIsoₛ-filler3 : (n : ℕ) (b : A B n) + → I → I → I → Pushout (pushoutMapₛ n) fst + pushoutIsoₛ-filler3 n b j k l = + hfill (λ l → + λ{ (j = i0) → inl (inr (∣ gʷ ∣ (suc n) (inr b))) + ; (j = i1) → doubleCompPath-filler + (push (inl (inr b) , north)) + (λ _ → inr (inl (inr b))) + (λ i → push (inl (inr b) , south) (~ i)) (~ k) (~ l) + ; (k = i0) → (push (inl (inr b) , north) + ∙∙ (λ _ → inr (inl (inr b))) + ∙∙ (λ i → push (inl (inr b) , south) (~ i))) (~ l ∨ ~ j) + ; (k = i1) → push (inl (inr b) , south) j}) + (inS (push (inl (inr b) , south) (j ∧ k))) l + + pushoutIsoₛ-filler4 : (n : ℕ) (b : A B n) + → I → I → I → Pushout (pushoutMapₛ n) fst + pushoutIsoₛ-filler4 n b i k j = + hfill (λ j → + λ{ (i = i0) → inl (inr (∣ gʷ ∣ (suc n) (inr b))) + ; (i = i1) → pushoutIsoₛ-filler3 n b j k i1 + ; (k = i0) → (push (inl (inr b) , north) + ∙∙ (λ _ → inr (inl (inr b))) + ∙∙ (λ i → push (inl (inr b) , south) (~ i))) (~ i ∨ ~ j) + ; (k = i1) → push (inl (inr b) , south) (i ∧ j)}) + (inS (push (inl (inr b) , south) i0)) j + + pushoutIsoₛ-fun : (n : ℕ) → modifiedPushout n → Pushout (pushoutMapₛ n) fst + pushoutIsoₛ-fun n (inl (inl c)) = inl (inl c) + pushoutIsoₛ-fun n (inl (inr c)) = inr (inl (inl c)) + pushoutIsoₛ-fun n (inl (push (c , x) i)) = push (inl (inl c) , x) i + pushoutIsoₛ-fun n (inr (inl d)) = inl (inr d) + pushoutIsoₛ-fun n (inr (inr d)) = inr (inr d) + pushoutIsoₛ-fun n (inr (push (d , x) i)) = push (inr d , x) i + pushoutIsoₛ-fun n (push (inl b) i) = inl (push b i) + pushoutIsoₛ-fun n (push (inr b) i) = + (push (inl (inr b) , north) ∙∙ refl + ∙∙ (λ i → push (inl (inr b) , south) (~ i))) i + pushoutIsoₛ-fun n (push (push (b , x) j) i) = pushoutIsoₛ-filler1 n b x i j i1 + + pushoutIsoₛ-inv : (n : ℕ) → Pushout (pushoutMapₛ n) fst → modifiedPushout n + pushoutIsoₛ-inv n (inl x) = pushoutIsoₛ-inv↪ n x + pushoutIsoₛ-inv n (inr (inl (inl c))) = inl (inr c) + pushoutIsoₛ-inv n (inr (inl (inr b))) = push (inr b) i0 + pushoutIsoₛ-inv n (inr (inr d)) = inr (inr d) + pushoutIsoₛ-inv n (push (inl (inl c) , x) i) = inl (push (c , x) i) + pushoutIsoₛ-inv n (push (inl (inr b) , north) i) = push (inr b) i0 + pushoutIsoₛ-inv n (push (inl (inr b) , south) i) = push (inr b) (~ i) + pushoutIsoₛ-inv n (push (inl (inr b) , merid x j) i) = + pushoutIsoₛ-filler2 n b x i j i1 + pushoutIsoₛ-inv n (push (inr d , x) i) = inr (push (d , x) i) + + pushoutIsoₛ-rightInv↪ : (n : ℕ) (x : pushoutA (suc n)) + → pushoutIsoₛ-fun n (pushoutIsoₛ-inv↪ n x) ≡ inl x + pushoutIsoₛ-rightInv↪ n (inl c) = refl + pushoutIsoₛ-rightInv↪ n (inr d) = refl + pushoutIsoₛ-rightInv↪ n (push b i) = refl + + pushoutIsoₛ-rightInv : (n : ℕ) (x : Pushout (pushoutMapₛ n) fst) + → pushoutIsoₛ-fun n (pushoutIsoₛ-inv n x) ≡ x + pushoutIsoₛ-rightInv n (inl x) = pushoutIsoₛ-rightInv↪ n x + pushoutIsoₛ-rightInv n (inr (inl (inl c))) = refl + pushoutIsoₛ-rightInv n (inr (inl (inr b))) = push (inl (inr b) , north) + pushoutIsoₛ-rightInv n (inr (inr d)) = refl + pushoutIsoₛ-rightInv n (push (inl (inl c) , x) i) = refl + pushoutIsoₛ-rightInv n (push (inl (inr b) , north) i) k = push (inl (inr b) , north) (i ∧ k) + pushoutIsoₛ-rightInv n (push (inl (inr b) , south) i) k = pushoutIsoₛ-filler4 n b i k i1 + pushoutIsoₛ-rightInv n (push (inl (inr b) , merid x j) i) k = + hcomp (λ l → + λ{ (i = i0) → i=i0 j k l + ; (i = i1) → doubleCompPath-filler + (push (inl (inr b) , north)) refl + (λ i → push (inl (inr b) , south) (~ i)) (~ k) (~ l ∧ j) + ; (j = i0) → j=i0 i k l + ; (j = i1) → j=i1 i k l + ; (k = i0) → pushoutIsoₛ-fun n (pushoutIsoₛ-filler2 n b x i j l) + ; (k = i1) → push (inl (inr b) , merid x j) i}) + (pushoutIsoₛ-filler1 n b x j i (~ k)) + where + i=i0 j=i0 j=i1 : I → I → I → Pushout (pushoutMapₛ n) fst + i=i0 j k l = hcomp (λ i → + λ {(j = i0) → inl (inl (∣ f ∣ (suc n) (push (b , x) (i ∧ k ∨ l)))) + ; (j = i1) → inl (inr (∣ g ∣ (suc n) (push (b , x) (i ∧ k ∨ l)))) + ; (k = i0) → pushoutIsoₛ-fun n + (pushoutIsoₛ-inv↪ n (pushoutIsoₛ-filler0 n b x (l ∧ i) j)) + ; (k = i1) → inl (pushoutIsoₛ-filler0 n b x i j) + ; (l = i0) → inl (pushoutIsoₛ-filler0 n b x (i ∧ k) j) + ; (l = i1) → pushoutIsoₛ-rightInv↪ n (pushoutIsoₛ-filler0 n b x i j) k}) + (inl (push (α B n (b , x)) j)) + j=i0 i k l = hcomp (λ j → + λ {(i = i0) → inl (inl (∣ f ∣ (suc n) (push (b , x) (k ∨ l)))) + ; (i = i1) → push (inl (inr b) , north) (k ∧ j) + ; (k = i0) → pushoutIsoₛ-fun n (invSides-hfill1 + (λ i → inl (inl (∣ f ∣ (suc n) (push (b , x) (~ i))))) + refl i (~ l) j) + ; (k = i1) → push (inl (inr b) , north) (i ∧ j) + ; (l = i0) → invSides-hfill2 (push (inl (inr b) , north)) + (λ i → inl (inl (∣ f ∣ (suc n) (push (b , x) (~ i))))) + (~ i) ( k) j + ; (l = i1) → push (inl (inr b) , north) (i ∧ k ∧ j)}) + (inl (inl (∣ f ∣ (suc n) (push (b , x) (i ∨ k ∨ l))))) + j=i1 i k l = hcomp (λ j → + λ {(i = i0) → inl (inr (∣ g ∣ (suc n) (push (b , x) (k ∨ l)))) + ; (i = i1) → pushoutIsoₛ-filler3 n b j k l + ; (k = i0) → pushoutIsoₛ-fun n (invSides-hfill1 + (λ i → inr (inl (∣ g ∣ (suc n) (push (b , x) (~ i))))) + (λ i → push (inr b) (~ i)) i (~ l) j) + ; (k = i1) → push (inl (inr b) , south) (i ∧ j) + ; (l = i0) → invSides-hfill2 (push (inl (inr b) , south)) + (λ i → inl (inr (∣ g ∣ (suc n) (push (b , x) (~ i))))) + (~ i) ( k) j + ; (l = i1) → pushoutIsoₛ-filler4 n b i k j}) + (inl (inr (∣ g ∣ (suc n) (push (b , x) (i ∨ k ∨ l))))) + + pushoutIsoₛ-rightInv n (push (inr d , x) i) = refl + + pushoutIsoₛ-leftInv : (n : ℕ) (x : modifiedPushout n) + → pushoutIsoₛ-inv n (pushoutIsoₛ-fun n x) ≡ x + pushoutIsoₛ-leftInv n (inl (inl c)) = refl + pushoutIsoₛ-leftInv n (inl (inr c)) = refl + pushoutIsoₛ-leftInv n (inl (push (c , x) i)) = refl + pushoutIsoₛ-leftInv n (inr (inl d)) = refl + pushoutIsoₛ-leftInv n (inr (inr d)) = refl + pushoutIsoₛ-leftInv n (inr (push (d , x) i)) = refl + pushoutIsoₛ-leftInv n (push (inl b) i) = refl + pushoutIsoₛ-leftInv n (push (inr b) i) k = + hcomp (λ j → λ { (i = i0) → inl (inl (∣ f ∣ (suc n) (inr b))) + ; (i = i1) → push (inr b) (k ∨ j) + ; (k = i0) → pushoutIsoₛ-inv n + (doubleCompPath-filler + (push (inl (inr b) , north)) refl + (λ i → push (inl (inr b) , south) (~ i)) j i) + ; (k = i1) → push (inr b) i }) + (push (inr b) (i ∧ k)) + pushoutIsoₛ-leftInv n (push (push (b , x) j) i) k = + hcomp (λ l → λ {(i = i0) → i=i0 j k l + ; (i = i1) → i=i1 j k l + ; (j = i0) → pushoutIsoₛ-inv↪ n + (pushoutIsoₛ-filler0 n b x ((~ k) ∧ (~ l)) i) + ; (k = i0) → pushoutIsoₛ-inv n + (pushoutIsoₛ-filler1 n b x i j l) + ; (k = i1) → push (push (b , x) j) i}) + (pushoutIsoₛ-filler2 n b x j i (~ k)) + where + i=i0 i=i1 : I → I → I → modifiedPushout n + i=i0 j k l = hcomp (λ i → + λ {(j = i0) → inl (inl (∣ f ∣ (suc n) (push (b , x) ((~ l) ∧ (~ k))))) + ; (j = i1) → inl (inl (e C n .fst (∣ f ∣ (suc n) + (invEq (e B n) (inr b))))) + ; (k = i0) → pushoutIsoₛ-inv n + (invSides-hfill2 (push (inl (inr b) , north)) + (λ i → inl (inl (∣ f ∣ (suc n) (push (b , x) (~ i))))) + (~ j) (~ l) i) + ; (k = i1) → inl (inl (∣ f ∣ (suc n) (push (b , x) j))) + ; (l = i0) → invSides-hfill1 + (λ i → inl (inl (∣ f ∣ (suc n) (push (b , x) (~ i))))) + refl (~ k) (~ j) i + ; (l = i1) → inl (inl (∣ f ∣ (suc n) (push (b , x) j)))}) + (inl (inl (∣ f ∣ (suc n) (push (b , x) (((~ k) ∧ (~ l)) ∨ j))))) + i=i1 j k l = hcomp (λ i → + λ{ (j = i0) → inr (inl (∣ g ∣ (suc n) (push (b , x) ((~ l) ∧ (~ k))))) + ; (j = i1) → push (inr b) (k ∨ ~ i ∨ l) + ; (k = i0) → pushoutIsoₛ-inv n + (invSides-hfill2 (push (inl (inr b) , south)) + (λ i → inl (inr (∣ g ∣ (suc n) (push (b , x) (~ i))))) + (~ j) (~ l) i) + ; (k = i1) → inr (inl (∣ g ∣ (suc n) (push (b , x) j))) + ; (l = i0) → invSides-hfill1 + (λ i → inr (inl (∣ g ∣ (suc n) (push (b , x) (~ i))))) + (λ i → push (inr b) (~ i)) (~ k) (~ j) i + ; (l = i1) → inr (inl (∣ g ∣ (suc n) (push (b , x) j)))}) + (inr (inl (∣ g ∣ (suc n) (push (b , x) (((~ k) ∧ (~ l)) ∨ j))))) + + pushoutIsoₛ : (n : ℕ) → Iso (modifiedPushout n) (Pushout (pushoutMapₛ n) fst) + Iso.fun (pushoutIsoₛ n) = pushoutIsoₛ-fun n + Iso.inv (pushoutIsoₛ n) = pushoutIsoₛ-inv n + Iso.rightInv (pushoutIsoₛ n) = pushoutIsoₛ-rightInv n + Iso.leftInv (pushoutIsoₛ n) = pushoutIsoₛ-leftInv n + + pushoutIsoₜ : (n : ℕ) → Iso (pushoutA (suc n)) (Pushout (pushoutMap n) fst) + pushoutIsoₜ zero = pushoutIso₀ + pushoutIsoₜ (suc n) = compIso (IsoModifiedPushout n) + (compIso (pushoutIsoₛ n) (invIso (pushoutₛIso n))) + + pushoutSkel : CWskel ℓ + fst pushoutSkel = pushoutA + fst (snd pushoutSkel) = pushoutCells + fst (snd (snd pushoutSkel)) = pushoutMap + fst (snd (snd (snd pushoutSkel))) = B .snd .snd .snd .fst + snd (snd (snd (snd pushoutSkel))) n = isoToEquiv (pushoutIsoₜ n) + + f' : SequenceMap (realiseSeq B) (ShiftSeq (realiseSeq C)) + ∣_∣ f' n x = inl (∣ f ∣ n x) + comm f' n x i = inl (comm f n x i) + + g' : SequenceMap (realiseSeq B) (ShiftSeq (realiseSeq D)) + ∣_∣ g' n x = inl (∣ g ∣ n x) + comm g' n x i = inl (comm g n x i) + + PushoutIsoRealise : + Iso (Pushout (realiseSequenceMap f') (realiseSequenceMap g')) + (Pushout (realiseCellMap {C = B} {D = C} f) + (realiseCellMap {C = B} {D = D} g)) + + PushoutIsoRealise = pushoutIso _ _ _ _ (idEquiv _) + (isoToEquiv (invIso (Iso-SeqColim→SeqColimSuc (realiseSeq C)))) + (isoToEquiv (invIso (Iso-SeqColim→SeqColimSuc (realiseSeq D)))) + (λ { i (incl {n = n} x) → push {n = n} (∣ f ∣ n x) (~ i) + ; i (push {n = n} x j) → lem Bʷ Cʷ fʷ n x i j}) + λ { i (incl {n = n} x) → push {n = n} (∣ g ∣ n x) (~ i) + ; i (push {n = n} x j) → lem Bʷ Dʷ gʷ n x i j} + where + module _ (Bʷ Cʷ : CWskel ℓ) (fʷ : cellMap (str Bʷ) (str Cʷ)) where + B' = str Bʷ + C' = str Cʷ + f'' = strictCwMap fʷ -- strMap fʷ + lem : (n : ℕ) (x : fst B' n) + → Square {A = SeqColim (realiseSeq C')} + (cong (Iso.inv (Iso-SeqColim→SeqColimSuc (realiseSeq C'))) + ((push {n = n} (inl (∣ f'' ∣ n x)) + ∙ (λ i → incl {n = suc n} (inl (comm f'' n x i)))))) + (cong (realiseCellMap {C = B'} {D = C'} f'') (push x)) + (sym (push {n = n} (∣ f'' ∣ n x))) + (sym (push {n = suc n} (∣ f'' ∣ (suc n) (inl x)))) + lem n x = + cong-∙ (Iso.inv (Iso-SeqColim→SeqColimSuc (realiseSeq C'))) + (push {n = n} (inl (∣ f'' ∣ n x))) + (λ i → incl {n = suc n} (inl (comm f'' n x i))) + ◁ (sym (rUnit _) + ◁ (invSides-filler {A = SeqColim (realiseSeq C')} + (push {n = suc n} (∣ f'' ∣ (suc n) (inl x))) + (sym (push {n = n} (∣ f'' ∣ n x))) + ▷ rUnit _)) + + SeqIs : SequenceIso (ShiftSeq (realiseSeq pushoutSkel)) + (PushoutSequence f' g') + fst SeqIs n = idIso + snd SeqIs n (inl x) = refl + snd SeqIs n (inr x) = refl + snd SeqIs n (push a i) j = + doubleCompPath-filler + ((λ i₁ → inl (inl (comm f n a i₁)))) + (push (CW↪ B n a)) + (λ i₁ → inr (inl (comm g n a (~ i₁)))) (~ j) i + + realisePushoutSkel : Iso (realise pushoutSkel) + (Pushout (realiseCellMap {C = B} {D = C} f) + (realiseCellMap {C = B} {D = D} g)) + realisePushoutSkel = + compIso (Iso-SeqColim→SeqColimSuc _) + (compIso (sequenceEquiv→ColimIso + (SequenceIso→SequenceEquiv SeqIs)) + (compIso (invIso (Iso-PushoutColim-ColimPushout f' g')) + PushoutIsoRealise)) + + hasCWskelPushout : hasCWskel (Pushout (realiseCellMap {C = B} {D = C} f) + (realiseCellMap {C = B} {D = D} g)) + fst hasCWskelPushout = pushoutSkel + snd hasCWskelPushout = isoToEquiv (invIso realisePushoutSkel) + + hasFinCWskelPushout : isFinCWskel B → isFinCWskel C → isFinCWskel D + → isFinCWskel pushoutSkel + fst (hasFinCWskelPushout (dimB , eB) (dimC , eC) (dimD , eD)) = + suc (dimB +ℕ (dimC +ℕ dimD)) + snd (hasFinCWskelPushout (dimB , eB) (dimC , eC) (dimD , eD)) = + λ k → transport (λ j + → isEquiv (CW↪ pushoutSkel (+-suc k (dimB +ℕ (dimC +ℕ dimD)) (~ j)))) + (subst isEquiv (funExt + (λ { (inl x) → refl + ; (inr x) → refl + ; (push a i) j → rUnit (push (inl a)) (~ j) i})) + (pushoutEquiv _ _ _ _ + (inl , subst (isEquiv ∘ CW↪ B) + (sym (+-assoc k (dimC +ℕ dimD) dimB) + ∙ cong (k +ℕ_) (+-comm (dimC +ℕ dimD) dimB)) + (eB (k +ℕ (dimC +ℕ dimD)))) + (inl , subst (isEquiv ∘ CW↪ C) + (cong suc (sym (+-assoc (k +ℕ dimD) dimB dimC) + ∙ sym (+-assoc k dimD (dimB +ℕ dimC)) + ∙ cong (k +ℕ_ ) (+-comm dimD _ + ∙ sym (+-assoc dimB dimC dimD)))) + (eC (suc (k +ℕ dimD +ℕ dimB)))) + (inl , subst (isEquiv ∘ CW↪ D) + (cong suc (sym (+-assoc (k +ℕ dimC) dimB dimD) + ∙ sym (+-assoc k dimC (dimB +ℕ dimD)) + ∙ cong (k +ℕ_) (+-assoc dimC dimB dimD + ∙ cong (_+ℕ dimD) (+-comm dimC dimB) + ∙ sym (+-assoc dimB dimC dimD)))) + (eD (suc (k +ℕ dimC +ℕ dimB)))) + refl + refl .snd)) + +isFinCWPushout : ∀ {ℓ} (B C D : finCW ℓ) + (f : finCW→CW B →ᶜʷ finCW→CW C) (g : finCW→CW B →ᶜʷ finCW→CW D) + → isFinCW (Pushout f g) +isFinCWPushout {ℓ = ℓ} = + uncurry λ B → PT.elim (λ _ → isPropΠ4 λ _ _ _ _ → squash₁) λ isFinCWB + → uncurry λ C → PT.elim (λ _ → isPropΠ3 λ _ _ _ → squash₁) λ isCWC + → uncurry λ D → PT.elim (λ _ → isPropΠ2 λ _ _ → squash₁) λ isCWD + → subst (λ B∞ → (f : B∞ → C) (g : B∞ → D) + → isFinCW (Pushout f g)) + (sym (ua (isFinCWB .snd .snd))) + (subst2 (λ C∞ D∞ → (f : realise ((isFinCWB .snd .fst .fst) + , (isFinCWB .snd .fst .snd .fst)) → C∞) + (g : realise ((isFinCWB .snd .fst .fst) + , (isFinCWB .snd .fst .snd .fst)) → D∞) + → isFinCW (Pushout f g)) + (sym (ua (isCWC .snd .snd))) + (sym (ua (isCWD .snd .snd))) + λ f g → PT.rec squash₁ (λ {(f' , fid) + → PT.rec squash₁ (λ {(g' , gid) + → subst2 (λ f g → isFinCW (Pushout f g)) + fid gid + ∣ main (hasFinCWskel→hasCWskel _ isFinCWB .fst) + (hasFinCWskel→hasCWskel _ isCWC .fst) + (hasFinCWskel→hasCWskel _ isCWD .fst) + (_ , isFinCWB .snd .fst .snd .snd) + (_ , isCWC .snd .fst .snd .snd) + (_ , isCWD .snd .fst .snd .snd) _ _ ∣₁}) + (finCWmap→CellMap (isFinCWB .fst) (isFinCWB .snd .fst) + (hasFinCWskel→hasCWskel _ isCWD .fst) g)}) + (finCWmap→CellMap (isFinCWB .fst) (isFinCWB .snd .fst) + (hasFinCWskel→hasCWskel _ isCWC .fst) f)) + where + main : ∀ {ℓ} (B C D : CWskel ℓ) + → isFinCWskel B + → isFinCWskel C + → isFinCWskel D + → (f' : cellMap B C) (g' : cellMap B D) + → hasFinCWskel (Pushout (realiseCellMap {C = B} {D = C} f') + (realiseCellMap {C = B} {D = D} g')) + main = elimStrictCW λ B → elimStrictCW λ C → elimStrictCW λ D + → λ fB fC fD f' g' + → subst2 (λ f' g' + → hasFinCWskel (Pushout (realiseCellMap {C = str B} {D = str C} f') + (realiseCellMap {C = str B} {D = str D} g'))) + (strictCwMap≡ f') (strictCwMap≡ g') + ((CWPushout.hasFinCWskelPushout _ B C D f' g' fB fC fD .fst) + , ((CWPushout.hasCWskelPushout _ B C D f' g' .fst .fst + , (CWPushout.hasCWskelPushout _ B C D f' g' .fst .snd + , (CWPushout.hasFinCWskelPushout _ B C D f' g' fB fC fD .snd))) + , CWPushout.hasCWskelPushout _ B C D f' g' .snd)) + +isPushoutᶜʷ : ∀ {ℓ} (B : finCW ℓ) (C D : CW ℓ) + (f : finCW→CW B →ᶜʷ C) (g : finCW→CW B →ᶜʷ D) + → isCW (Pushout f g) +isPushoutᶜʷ {ℓ = ℓ} = uncurry λ B + → PT.elim (λ _ → isPropΠ4 λ _ _ _ _ → squash₁) λ isFinCWB + → uncurry λ C → PT.elim (λ _ → isPropΠ3 λ _ _ _ → squash₁) λ isCWC + → uncurry λ D → PT.elim (λ _ → isPropΠ2 λ _ _ → squash₁) λ isCWD + → subst (λ B∞ → (f : B∞ → C) (g : B∞ → D) + → isCW (Pushout f g)) + (sym (ua (isFinCWB .snd .snd))) + (subst2 (λ C∞ D∞ → (f : realise ((isFinCWB .snd .fst .fst) + , (isFinCWB .snd .fst .snd .fst)) → C∞) + (g : realise ((isFinCWB .snd .fst .fst) + , (isFinCWB .snd .fst .snd .fst)) → D∞) + → isCW (Pushout f g)) + (sym (ua (isCWC .snd))) + (sym (ua (isCWD .snd))) + λ f g → PT.rec squash₁ (λ {(f' , fid) + → PT.rec squash₁ (λ {(g' , gid) + → subst2 (λ f g → isCW (Pushout f g)) + fid gid + ∣ main (isFinCWB .snd .fst .fst , isFinCWB .snd .fst .snd .fst) + (isCWC .fst) (isCWD .fst) f' g' ∣₁}) + (finCWmap→CellMap (isFinCWB .fst) (isFinCWB .snd .fst) + (fst isCWD) g)}) + (finCWmap→CellMap (isFinCWB .fst) (isFinCWB .snd .fst) + (fst isCWC) f)) + where + main : ∀ {ℓ} (B C D : CWskel ℓ) (f' : cellMap B C) (g' : cellMap B D) + → hasCWskel (Pushout (realiseCellMap {C = B} {D = C} f') + (realiseCellMap {C = B} {D = D} g')) + main = elimStrictCW λ B → elimStrictCW λ C → elimStrictCW λ D + → λ f' g' + → subst2 (λ f' g' + → hasCWskel (Pushout (realiseCellMap {C = str B} {D = str C} f') + (realiseCellMap {C = str B} {D = str D} g'))) + (strictCwMap≡ f') (strictCwMap≡ g') + (CWPushout.hasCWskelPushout _ B C D f' g') + +-- Consequence: elimination principle for predicates stable under +-- pushouts (suggested by Reid Barton) +module _ {ℓ ℓ' : Level} (P : Type ℓ → Type ℓ') (P1 : P Unit*) (P0 : P ⊥*) + (Ppush : (A B C : Type ℓ) (f : A → B) (g : A → C) + → P A → P B → P C → P (Pushout f g)) where + private + PFin1 : P (Lift (Fin 1)) + PFin1 = subst P (cong Lift (isoToPath Iso-Unit-Fin1)) P1 + + P⊎ : {B C : Type ℓ} → P B → P C → P (B ⊎ C) + P⊎ {B = B} {C} pB pC = + subst P (isoToPath + (compIso (pushoutIso _ _ _ _ + (invEquiv LiftEquiv) (idEquiv _) (idEquiv _) refl refl) + PushoutEmptyDomainIso)) + (Ppush ⊥* B C (λ ()) (λ ()) P0 pB pC) + + PFin : (n : ℕ) → P (Lift (Fin n)) + PFin zero = + subst P + (cong Lift + (ua (propBiimpl→Equiv isProp⊥ + (λ x → ⊥.rec (¬Fin0 x)) (λ()) ¬Fin0))) + P0 + PFin (suc n) = + subst P (cong Lift (sym (isoToPath Iso-Fin-Unit⊎Fin))) + (subst P (isoToPath (Lift⊎Iso ℓ)) + (P⊎ P1 (PFin n))) + + PS : (n : ℕ) → P (Lift (S⁻ n)) + PS zero = P0 + PS (suc zero) = subst P (cong Lift (isoToPath (invIso Iso-Bool-Fin2))) (PFin 2) + PS (suc (suc n)) = + subst P + (cong Lift (isoToPath (compIso PushoutSuspIsoSusp (invIso (IsoSucSphereSusp n))))) + (subst P (isoToPath (LiftPushoutIso ℓ)) + (Ppush (Lift (S₊ n)) Unit* Unit* (λ _ → tt*) (λ _ → tt*) + (PS (suc n)) P1 P1)) + + PFin×S : (n m : ℕ) → P (Lift (Fin n × S⁻ m)) + PFin×S zero m = + subst P (ua (compEquiv (uninhabEquiv (λ()) λ()) LiftEquiv)) P0 + PFin×S (suc n) m = subst P (isoToPath is) (P⊎ (PS m) (PFin×S n m)) + where + is : Iso (Lift (S⁻ m) ⊎ Lift (Fin n × S⁻ m)) + (Lift (Fin (suc n) × S⁻ m)) + is = + compIso (Lift⊎Iso ℓ) + (compIso (invIso LiftIso) + (compIso + (compIso (compIso (⊎Iso (invIso rUnit×Iso) Σ-swap-Iso) + (compIso (invIso ×DistR⊎Iso) Σ-swap-Iso)) + (Σ-cong-iso-fst (invIso Iso-Fin-Unit⊎Fin))) + LiftIso)) + + PCWskel : (B : CWskel ℓ) → (n : ℕ) → P (fst B n) + PCWskel B zero = subst P (ua (uninhabEquiv (λ()) (CW₀-empty B))) P0 + PCWskel B (suc n) = + subst P + (ua (compEquiv + (pushoutEquiv _ _ _ _ + (invEquiv LiftEquiv) (idEquiv _) (invEquiv LiftEquiv) + refl refl) + (invEquiv (B .snd .snd .snd .snd n)))) + (Ppush _ _ _ + (λ { (lift r) → CWskel-fields.α B n r}) (liftFun fst) + (PFin×S (CWskel-fields.card B n) n) + (PCWskel B n) (PFin (CWskel-fields.card B n))) + + elimPropFinCWFun : (B : Type ℓ) → hasFinCWskel B → P B + elimPropFinCWFun B fCWB = + subst P (ua (compEquiv (isoToEquiv (realiseFin _ (fCWB .snd .fst))) + (invEquiv (fCWB .snd .snd)))) + (PCWskel (finCWskel→CWskel _ (fCWB .snd .fst)) (fCWB .fst)) + + -- main result + elimPropFinCW : (B : finCW ℓ) → isProp (P (fst B)) → P (fst B) + elimPropFinCW (B , Bstr) isp = PT.rec isp (elimPropFinCWFun B) Bstr diff --git a/Cubical/CW/Instances/Sigma.agda b/Cubical/CW/Instances/Sigma.agda new file mode 100644 index 0000000000..86922ac52d --- /dev/null +++ b/Cubical/CW/Instances/Sigma.agda @@ -0,0 +1,63 @@ +{- +CW structure on Σ-types (and binary products) +-} + +module Cubical.CW.Instances.Sigma where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism + +open import Cubical.CW.Base + +open import Cubical.Data.Unit +open import Cubical.Data.Sigma +open import Cubical.Data.Empty + +open import Cubical.HITs.Pushout +open import Cubical.HITs.PropositionalTruncation as PT + +open import Cubical.CW.Instances.Pushout +open import Cubical.CW.Instances.Empty +open import Cubical.CW.Instances.Lift + +isFinCWΣ : ∀ {ℓ ℓ'} (A : finCW ℓ) (B : fst A → finCW ℓ') + → isFinCW (Σ (fst A) (fst ∘ B)) +isFinCWΣ {ℓ} {ℓ'} A B = + subst isFinCW (ua + (isoToEquiv (iso (λ { (lift (lift x , lift y)) → x , y}) + (λ x → lift ((lift (fst x)) , (lift (snd x)))) + (λ _ → refl) (λ _ → refl)))) + (finCWLift ℓ (_ + , main (finCWLift ℓ' A) + λ {(lift x) → finCWLift ℓ (B x)}) .snd) + where + ℓ* = ℓ-max ℓ ℓ' + + main : (A : finCW ℓ*) (B : fst A → finCW ℓ*) + → isFinCW (Σ (fst A) (fst ∘ B)) + main A = elimPropFinCW + (λ A → (B : A → finCW ℓ*) → isFinCW (Σ A (fst ∘ B))) + (λ B → subst isFinCW + (isoToPath (compIso + (invIso lUnit×Iso) + (Σ-cong-iso-fst LiftIso))) + (snd (B tt*))) -- + (λ _ → subst isFinCW (ua (uninhabEquiv {A = ⊥*} (λ()) λ())) + ∣ hasFinCWskel⊥* _ ∣₁) -- + (λ A0 A1 A2 f g p q r B + → subst isFinCW (ua (invEquiv (ΣPushout≃PushoutΣ f g (fst ∘ B)))) + (isFinCWPushout (_ , p (λ x → B (inl (f x)))) + (_ , q λ x → B (inl x)) + (_ , r λ a → B (inr a)) _ _)) + A (isPropΠ λ _ → squash₁) + +-- A × B as as finite CW complex +isFinCW× : ∀ {ℓ ℓ'} (A : finCW ℓ) (B : finCW ℓ') + → isFinCW (fst A × fst B) +isFinCW× A B = isFinCWΣ A (λ _ → B) + +-- Todo: explicit construction of products of arbtirary CW complexes diff --git a/Cubical/CW/Instances/Sn.agda b/Cubical/CW/Instances/Sn.agda new file mode 100644 index 0000000000..77dc9a0093 --- /dev/null +++ b/Cubical/CW/Instances/Sn.agda @@ -0,0 +1,483 @@ +{-# OPTIONS --lossy-unification #-} +{- +CW structure on spheres +-} +module Cubical.CW.Instances.Sn where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.HLevels + +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Bool +open import Cubical.Data.Fin.Inductive +open import Cubical.Data.Nat.Order.Inductive +open import Cubical.Data.Sigma +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Unit +open import Cubical.Data.Sequence +open import Cubical.Data.FinSequence + +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.Pushout +open import Cubical.HITs.SequentialColimit +open import Cubical.HITs.Truncation as TR +open import Cubical.HITs.PropositionalTruncation as PT + +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Group.Base + +open import Cubical.Relation.Nullary + +open import Cubical.CW.Base +open import Cubical.CW.Map +open import Cubical.CW.Properties +open import Cubical.CW.Approximation + +open Iso + +-- Defining family of types +module Sgen (n : ℕ) where + Sfam : (m : ℕ) → Trichotomyᵗ m (suc n) → Type + Sfam zero p = ⊥ + Sfam (suc m) (lt x) = Unit + Sfam (suc m) (eq x) = S₊ n + Sfam (suc m) (gt x) = S₊ n + + Sfam∙ : (m : ℕ) (s : Trichotomyᵗ (suc m) (suc n)) → Sfam (suc m) s + Sfam∙ m (lt x) = tt + Sfam∙ m (eq x) = ptSn n + Sfam∙ m (gt x) = ptSn n + +Sfam : (n : ℕ) → ℕ → Type +Sfam n m = Sgen.Sfam n m (m ≟ᵗ suc n) + +Sfam∙ : (n m : ℕ) → Sfam n (suc m) +Sfam∙ n m = Sgen.Sfam∙ n m (suc m ≟ᵗ suc n) + +Sn→SfamGen : ∀ {n k : ℕ} (p : Trichotomyᵗ k (suc n)) + → 0 <ᵗ k → S₊ n → Sgen.Sfam n k p +Sn→SfamGen {n = n} {suc k} (lt x₁) _ x = tt +Sn→SfamGen {n = n} {suc k} (eq x₁) _ x = x +Sn→SfamGen {n = n} {suc k} (gt x₁) _ x = x + +Sn→SfamGen∙ : ∀ {n k : ℕ} (p : Trichotomyᵗ (suc k) (suc n)) + → Sn→SfamGen p tt (ptSn n) ≡ Sgen.Sfam∙ n k p +Sn→SfamGen∙ (lt x) = refl +Sn→SfamGen∙ (eq x) = refl +Sn→SfamGen∙ (gt x) = refl + +-- Cells of Sⁿ +ScardGen : (n m : ℕ) (s : Trichotomyᵗ (suc m) (suc n)) → ℕ +ScardGen zero zero s = 2 +ScardGen zero (suc m) s = 0 +ScardGen (suc n) zero s = 1 +ScardGen (suc n) (suc m) (lt x) = 0 +ScardGen (suc n) (suc m) (eq x) = 1 +ScardGen (suc n) (suc m) (gt x) = 0 + +Scard : (n : ℕ) → ℕ → ℕ +Scard n m = ScardGen n m (suc m ≟ᵗ suc n) + +-- Attaching maps +SαGen : (n m : ℕ) (s : Trichotomyᵗ (suc m) (suc n)) (q : Trichotomyᵗ m (suc n)) + → Fin (ScardGen n m s) × S⁻ m → Sgen.Sfam n m q +SαGen n (suc m) s q _ = Sgen.Sfam∙ n m q + +Sα : (n m : ℕ) → Fin (Scard n m) × S⁻ m → Sfam n m +Sα n m t = SαGen n m (suc m ≟ᵗ suc n) (m ≟ᵗ suc n) t + +-- Characterisation of trivial cells +¬Scard : (n m : ℕ) → n <ᵗ m → ¬ Fin (Scard n m) +¬Scard n m = ¬ScardGen n m (suc m ≟ᵗ suc n) + where + ¬ScardGen : (n m : ℕ) (p : _) → n <ᵗ m → ¬ Fin (ScardGen n m p) + ¬ScardGen zero (suc m) p q = ¬Fin0 + ¬ScardGen (suc n) (suc m) (lt x) q = snd + ¬ScardGen (suc n) (suc m) (eq x) q = + λ _ → ¬m<ᵗm (subst (n <ᵗ_) (cong (predℕ ∘ predℕ) x) q) + ¬ScardGen (suc n) (suc m) (gt x) q = snd + +¬Scard' : (n : ℕ) → ¬ Fin (Scard (suc (suc n)) (suc n)) +¬Scard' n x with (n ≟ᵗ suc n) +... | eq x₁ = ¬m<ᵗm (subst (n <ᵗ_) (sym x₁) <ᵗsucm) + +-- S⁰ ≃ Bool +Sfam0 : (m : ℕ) (p : _) → Sgen.Sfam zero (suc m) p ≃ Bool +Sfam0 m (eq x) = idEquiv _ +Sfam0 m (gt x) = idEquiv _ + +-- 0-skel of Sⁿ⁺¹ is contractible +SfamContr : (n : ℕ) (p : _) → isContr (Sgen.Sfam (suc n) (suc zero) p) +fst (SfamContr n p) = Sgen.Sfam∙ (suc n) zero p +snd (SfamContr n (lt x)) y = refl +snd (SfamContr n (eq x)) y = ⊥.rec (snotz (sym (cong predℕ x))) + +isContrSfamGen : (n m : ℕ) (s : m <ᵗ n) (q : _) + → isContr (Sgen.Sfam n (suc m) q) +fst (isContrSfamGen n m s q) = Sgen.Sfam∙ n m q +snd (isContrSfamGen n m s (lt x)) y = refl +snd (isContrSfamGen n m s (eq x)) y = + ⊥.rec (¬m<ᵗm (subst (m <ᵗ_) (sym (cong predℕ x)) s)) +snd (isContrSfamGen n m s (gt x)) y = ⊥.rec (¬m<ᵗm (<ᵗ-trans x s)) + +isPushoutSuspSphereIso : (n m : ℕ) (x : n ≡ m) (q : _) + → Iso (Susp (S₊ m)) + (Pushout {A = Fin 1 × S₊ m} (λ _ → Sgen.Sfam∙ (suc n) m q) fst) +fun (isPushoutSuspSphereIso n m e s) north = + inl (Sgen.Sfam∙ (suc n) m s) +fun (isPushoutSuspSphereIso n m e s) south = inr fzero +fun (isPushoutSuspSphereIso n m e s) (merid a i) = push (fzero , a) i +inv (isPushoutSuspSphereIso n m e s) (inl x) = north +inv (isPushoutSuspSphereIso n m e s) (inr x) = south +inv (isPushoutSuspSphereIso n m e s) (push a i) = merid (snd a) i +rightInv (isPushoutSuspSphereIso n m e s) (inl x) i = + inl (isContrSfamGen (suc n) m (subst (_<ᵗ suc n) e <ᵗsucm) s .snd x i) +rightInv (isPushoutSuspSphereIso n m e s) (inr (zero , tt)) j = inr fzero +rightInv (isPushoutSuspSphereIso n m e s) (push ((zero , tt) , a) i) = help i + where + ee = subst (_<ᵗ suc n) e <ᵗsucm + help : Square {A = Pushout {A = Fin 1 × S₊ m} + (λ _ → Sgen.Sfam∙ (suc n) m s) fst} + (λ i₁ → inl (isContrSfamGen (suc n) m ee s .snd + (Sgen.Sfam∙ (suc n) m s) i₁)) + refl + (push (fzero , a)) (push (fzero , a)) + help = (λ i j → inl (isProp→isSet + (isContr→isProp (isContrSfamGen (suc n) m ee s)) _ _ + (isContrSfamGen (suc n) m ee s .snd + (Sgen.Sfam∙ (suc n) m s)) refl i j)) + ◁ λ i j → push (fzero , a) i +leftInv (isPushoutSuspSphereIso n m e s) north = refl +leftInv (isPushoutSuspSphereIso n m e s) south = refl +leftInv (isPushoutSuspSphereIso n m e s) (merid a i) = refl + +SfamGenTopElement : (n m : ℕ) → (n <ᵗ m) → (q : _) → S₊ n ≃ Sgen.Sfam n m q +SfamGenTopElement n (suc m) p (lt x) = ⊥.rec (¬squeeze (x , p)) +SfamGenTopElement n (suc m) p (eq x) = idEquiv _ +SfamGenTopElement n (suc m) p (gt x) = idEquiv _ + +-- improved iso between Susp Sⁿ and Sⁿ⁺¹ used here +SuspSphere→Sphere : (n : ℕ) → Susp (S₊ n) → S₊ (suc n) +SuspSphere→Sphere n north = ptSn (suc n) +SuspSphere→Sphere zero south = base +SuspSphere→Sphere (suc n) south = south +SuspSphere→Sphere zero (merid t i) = SuspBool→S¹ (merid t i) +SuspSphere→Sphere (suc n) (merid a i) = merid a i + +IsoSucSphereSusp' : (n : ℕ) → Iso (S₊ (suc n)) (Susp (S₊ n)) +fun (IsoSucSphereSusp' n) = fun (IsoSucSphereSusp n) +inv (IsoSucSphereSusp' n) = SuspSphere→Sphere n +rightInv (IsoSucSphereSusp' zero) north = refl +rightInv (IsoSucSphereSusp' zero) south = SuspBool→S¹→SuspBool south +rightInv (IsoSucSphereSusp' zero) (merid a i) = + SuspBool→S¹→SuspBool (merid a i) +rightInv (IsoSucSphereSusp' (suc n)) north = refl +rightInv (IsoSucSphereSusp' (suc n)) south = refl +rightInv (IsoSucSphereSusp' (suc n)) (merid a i) = refl +leftInv (IsoSucSphereSusp' zero) base = S¹→SuspBool→S¹ base +leftInv (IsoSucSphereSusp' zero) (loop i) = S¹→SuspBool→S¹ (loop i) +leftInv (IsoSucSphereSusp' (suc n)) north = refl +leftInv (IsoSucSphereSusp' (suc n)) south = refl +leftInv (IsoSucSphereSusp' (suc n)) (merid a i) = refl + +-- Our family Sfam satisfies the pushout property (i.e. is a CW complex) +SαEqGen : (n m : ℕ) (p : Trichotomyᵗ (suc m) (suc n)) (q : _) + → (Sgen.Sfam n (suc m) p) ≃ Pushout (SαGen n m p q) fst +SαEqGen zero zero p q = + compEquiv (Sfam0 0 p) + (compEquiv (isoToEquiv Iso-Bool-Fin2) + (compEquiv (isoToEquiv (PushoutEmptyFam (λ()) λ())) (symPushout _ _))) +SαEqGen (suc n) zero p q = + compEquiv (isContr→Equiv (SfamContr n p) (flast , (λ {(zero , tt) → refl}))) + (compEquiv (isoToEquiv (PushoutEmptyFam (λ()) λ())) (symPushout _ _)) +SαEqGen (suc n) (suc m) (lt x) q = invEquiv + (isContr→≃Unit ((inl (isContrSfamGen (suc n) m (<ᵗ-trans x <ᵗsucm) q .fst)) + , λ { (inl t) i → inl (isContrSfamGen (suc n) m (<ᵗ-trans x <ᵗsucm) q .snd + t i)})) +SαEqGen zero (suc m) (eq x) q = ⊥.rec (snotz (cong predℕ x)) +SαEqGen (suc n) (suc m) (eq x) q = + isoToEquiv (compIso (IsoSucSphereSusp' n) + (compIso (congSuspIso (substIso S₊ (cong (predℕ ∘ predℕ) (sym x)))) + (isPushoutSuspSphereIso n m (cong (predℕ ∘ predℕ) (sym x)) q))) +SαEqGen zero (suc m) (gt x) (eq x₁) = isoToEquiv (PushoutEmptyFam (λ()) λ()) +SαEqGen zero (suc m) (gt x) (gt x₁) = isoToEquiv (PushoutEmptyFam (λ()) λ()) +SαEqGen (suc n) (suc m) (gt x) q = + compEquiv (SfamGenTopElement (suc n) (suc m) x q) + (isoToEquiv (PushoutEmptyFam (λ()) λ())) + +invEqSαEqGen∙ : (n m : ℕ) (p : Trichotomyᵗ (suc (suc m)) (suc n)) (q : _) + → invEq (SαEqGen n (suc m) p q) (inl (Sgen.Sfam∙ n m q)) + ≡ Sgen.Sfam∙ n (suc m) p +invEqSαEqGen∙ (suc n) m (lt x) (lt x₁) = refl +invEqSαEqGen∙ n m (lt x) (eq x₁) = ⊥.rec (¬-suc-n<ᵗn (subst (_<ᵗ n) x₁ x)) +invEqSαEqGen∙ (suc n) (suc m) (lt x) (gt x₁) = + ⊥.rec (¬-suc-n<ᵗn (<ᵗ-trans x x₁)) +invEqSαEqGen∙ (suc n) m (eq x) (lt x₁) = refl +invEqSαEqGen∙ n m (eq x) (eq x₁) = + ⊥.rec (¬m<ᵗm (subst (_<ᵗ suc (suc m)) (x₁ ∙ sym x) <ᵗsucm)) +invEqSαEqGen∙ n m (eq x) (gt x₁) = + ⊥.rec (¬-suc-n<ᵗn (subst (_<ᵗ suc m) (sym x) x₁)) +invEqSαEqGen∙ (suc n) m (gt x) (lt x₁) = ⊥.rec (¬squeeze (x , x₁)) +invEqSαEqGen∙ zero m (gt x) (eq x₁) = refl +invEqSαEqGen∙ (suc n) m (gt x) (eq x₁) = refl +invEqSαEqGen∙ zero m (gt x) (gt x₁) = refl +invEqSαEqGen∙ (suc n) m (gt x) (gt x₁) = refl + +SαEq : (n m : ℕ) → (Sfam n (suc m)) ≃ Pushout (Sα n m) fst +SαEq n m = SαEqGen n m (suc m ≟ᵗ suc n) (m ≟ᵗ suc n) + +-- Finally: Sⁿ as a CW skeleton +Sˢᵏᵉˡ : (n : ℕ) → CWskel ℓ-zero +fst (Sˢᵏᵉˡ n) = Sfam n +fst (snd (Sˢᵏᵉˡ n)) = Scard n +fst (snd (snd (Sˢᵏᵉˡ n))) = Sα n +fst (snd (snd (snd (Sˢᵏᵉˡ n)))) = λ{()} +snd (snd (snd (snd (Sˢᵏᵉˡ n)))) = SαEq n + +-- Proof that our CW structure converges to Sⁿ as a plain type +SfamTopElement : (n : ℕ) → S₊ n ≃ (Sfam n (suc n)) +SfamTopElement n = SfamGenTopElement n (suc n) <ᵗsucm (suc n ≟ᵗ suc n) + +SˢᵏᵉˡConverges : (n : ℕ) (k : ℕ) + → isEquiv (invEq (SαEq n (k +ℕ suc n)) ∘ inl) +SˢᵏᵉˡConverges n k = + compEquiv (inl , h n _ (<→<ᵗ (k , refl))) + (invEquiv (SαEq n (k +ℕ suc n))) .snd + where + h : (n m : ℕ) → n <ᵗ m + → isEquiv {A = Sfam n m} {B = Pushout (Sα n m) fst} inl + h n (suc m) p with (m ≟ᵗ n) + ... | lt x = ⊥.rec (¬squeeze (x , p)) + ... | eq x = isoToIsEquiv (PushoutEmptyFam (¬Scard n (suc m) p ∘ fst) + (¬Scard n (suc m) p)) + ... | gt x = isoToIsEquiv (PushoutEmptyFam (¬Scard n (suc m) p ∘ fst) + (¬Scard n (suc m) p)) + +hasCWskelSphere : (n : ℕ) → hasCWskel (S₊ n) +fst (hasCWskelSphere n) = Sˢᵏᵉˡ n +snd (hasCWskelSphere n) = + compEquiv (SfamTopElement n) + (isoToEquiv (converges→ColimIso (suc n) (SˢᵏᵉˡConverges n))) + +hasFinCWskelSphere : (n : ℕ) → hasFinCWskel (S₊ n) +fst (hasFinCWskelSphere n) = suc n +fst (fst (snd (hasFinCWskelSphere n))) = Sˢᵏᵉˡ n .fst +fst (snd (fst (snd (hasFinCWskelSphere n)))) = Sˢᵏᵉˡ n .snd +snd (snd (fst (snd (hasFinCWskelSphere n)))) = SˢᵏᵉˡConverges n +snd (snd (hasFinCWskelSphere n)) = hasCWskelSphere n .snd + +-- Sⁿ as a CW complex +Sᶜʷ : (n : ℕ) → CW ℓ-zero +fst (Sᶜʷ n) = S₊ n +snd (Sᶜʷ n) = ∣ hasCWskelSphere n ∣₁ + +Sᶠⁱⁿᶜʷ : (n : ℕ) → finCW ℓ-zero +fst (Sᶠⁱⁿᶜʷ n) = S₊ n +snd (Sᶠⁱⁿᶜʷ n) = ∣ hasFinCWskelSphere n ∣₁ + + +--- cellular approximations of maps from spheres into CW complexes +module _ {ℓ} (X : CWskel ℓ) (n : ℕ) (x₀ : fst X 1) + (f : S₊ n → fst X (suc n)) + (f₀ : f (ptSn n) ≡ CWskel∙ X x₀ n) where + private + <ᵗ→0<ᵗ : {n m : ℕ} → n <ᵗ m → 0 <ᵗ m + <ᵗ→0<ᵗ {n} {suc m} _ = tt + + cellMapSˢᵏᵉˡFunGenGen : ∀ k → (p : _) → Sgen.Sfam n k p → fst X k + cellMapSˢᵏᵉˡFunGenGen (suc k) (lt x₁) x = CWskel∙ X x₀ k + cellMapSˢᵏᵉˡFunGenGen (suc k) (eq x₁) x = subst (fst X) (sym x₁) (f x) + cellMapSˢᵏᵉˡFunGenGen (suc k) (gt x₁) x = + CW↪ X k (cellMapSˢᵏᵉˡFunGenGen k (k ≟ᵗ suc n) (Sn→SfamGen _ (<ᵗ→0<ᵗ x₁) x)) + + cellMapSˢᵏᵉˡFunGenGen∙ : ∀ k → (p : _) + → cellMapSˢᵏᵉˡFunGenGen (suc k) p (Sgen.Sfam∙ n k p) ≡ CWskel∙ X x₀ k + cellMapSˢᵏᵉˡFunGenGen∙ k (lt x) = refl + cellMapSˢᵏᵉˡFunGenGen∙ k (eq x) = + cong₂ (λ p q → subst (fst X) p q) (isSetℕ _ _ _ _) f₀ ∙ H _ (cong predℕ x) + where + H : (n : ℕ) (x : k ≡ n) + → subst (fst X) (cong suc (sym x)) (CWskel∙ X x₀ n) ≡ CWskel∙ X x₀ k + H = J> transportRefl _ + cellMapSˢᵏᵉˡFunGenGen∙ (suc k) (gt x) = + cong (CW↪ X (suc k)) + (cong (cellMapSˢᵏᵉˡFunGenGen (suc k) (Trichotomyᵗ-suc (k ≟ᵗ n))) + (Sn→SfamGen∙ (Trichotomyᵗ-suc (k ≟ᵗ n))) + ∙ cellMapSˢᵏᵉˡFunGenGen∙ k (suc k ≟ᵗ suc n)) + + cellMapSˢᵏᵉˡFunGenComm : (k : ℕ) (p : _) (q : _) (x : _) + → cellMapSˢᵏᵉˡFunGenGen (suc k) p (invEq (SαEqGen n k p q) (inl x)) + ≡ CW↪ X k (cellMapSˢᵏᵉˡFunGenGen k q x) + cellMapSˢᵏᵉˡFunGenComm (suc k) (lt x₁) (lt x₂) x = refl + cellMapSˢᵏᵉˡFunGenComm (suc k) (lt x₁) (eq x₂) x = + ⊥.rec (¬-suc-n<ᵗn (subst (suc k <ᵗ_) (cong predℕ (sym x₂)) x₁)) + cellMapSˢᵏᵉˡFunGenComm (suc k) (lt x₁) (gt x₂) x = + ⊥.rec (¬-suc-n<ᵗn (<ᵗ-trans x₁ x₂)) + cellMapSˢᵏᵉˡFunGenComm (suc k) (eq x₁) (lt x₂) x = + cong (subst (fst X) (sym x₁) ∘ f) (invEqSαEqGen∙ n k _ _) + ∙ cellMapSˢᵏᵉˡFunGenGen∙ (suc k) (eq x₁) + cellMapSˢᵏᵉˡFunGenComm k (eq x₁) (eq x₂) x = + ⊥.rec (falseDichotomies.eq-eq (refl , sym (x₁ ∙ sym x₂))) + cellMapSˢᵏᵉˡFunGenComm k (eq x₁) (gt x₂) x = + ⊥.rec (¬-suc-n<ᵗn {n} (subst (suc (suc n) <ᵗ_) x₁ x₂)) + cellMapSˢᵏᵉˡFunGenComm k (gt x₁) (lt x₂) x = ⊥.rec (¬squeeze (x₁ , x₂)) + cellMapSˢᵏᵉˡFunGenComm (suc k) (gt x₁) (eq x₂) x with (k ≟ᵗ n) + ... | lt x₃ = ⊥.rec (¬m<ᵗm (subst (_<ᵗ n) (cong predℕ x₂) x₃)) + ... | eq x₃ = cong (CW↪ X (suc k)) + (cong (subst (fst X) (cong suc (sym x₃))) (cong f (lem n x x₁ x₂)) + ∙ cong (λ p → subst (fst X) p (f x)) + (isSetℕ _ _ (cong suc (sym x₃)) (sym x₂))) + where + lem : (n : ℕ) (x : _) (x₁ : _) (x₂ : _) + → invEq (SαEqGen n (suc k) (gt x₁) (eq x₂)) (inl x) ≡ x + lem zero x x₁ x₂ = refl + lem (suc n) x x₁ x₂ = refl + ... | gt x₃ = ⊥.rec (¬m<ᵗm (subst (n <ᵗ_) (cong predℕ x₂) x₃)) + cellMapSˢᵏᵉˡFunGenComm (suc k) (gt x₁) (gt x₂) x with k ≟ᵗ n + ... | lt x₃ = ⊥.rec (¬m<ᵗm (<ᵗ-trans x₂ x₃)) + ... | eq x₃ = ⊥.rec (¬m<ᵗm (subst (n <ᵗ_) x₃ x₂)) + ... | gt x₃ = + cong (CW↪ X (suc k)) + (cong (CW↪ X k ∘ cellMapSˢᵏᵉˡFunGenGen k (k ≟ᵗ suc n)) + (cong (λ w → Sn→SfamGen (k ≟ᵗ suc n) (<ᵗ→0<ᵗ w) + (invEq (SαEqGen n (suc k) (gt x₁) (gt x₂)) (inl x))) (isProp<ᵗ x₃ x₂) + ∙ cong (Sn→SfamGen (k ≟ᵗ suc n) (<ᵗ→0<ᵗ x₂)) + (lem n k x₁ x₂ x (k ≟ᵗ suc n))) + ∙ cellMapSˢᵏᵉˡFunGenComm k (gt x₂) (k ≟ᵗ suc n) _) + where + lem : (n k : ℕ) (x₁ : _) (x₂ : _) (x : _) (w : _) + → invEq (SαEqGen n (suc k) (gt x₁) (gt x₂)) (inl x) ≡ + invEq (SαEqGen n k (gt x₂) w) + (inl (Sn→SfamGen w (<ᵗ→0<ᵗ x₂) x)) + lem n k x₁ x₂ x (lt x₃) = ⊥.rec (¬squeeze (x₂ , x₃)) + lem zero (suc k) x₁ x₂ x (eq x₃) = refl + lem (suc n) (suc k) x₁ x₂ x (eq x₃) = refl + lem zero (suc k) x₁ x₂ x (gt x₃) = refl + lem (suc n) (suc k) x₁ x₂ x (gt x₃) = refl + + cellMapSˢᵏᵉˡ : cellMap (Sˢᵏᵉˡ n) X + SequenceMap.map cellMapSˢᵏᵉˡ k = cellMapSˢᵏᵉˡFunGenGen k (k ≟ᵗ suc n) + SequenceMap.comm cellMapSˢᵏᵉˡ k x = + sym (cellMapSˢᵏᵉˡFunGenComm k (suc k ≟ᵗ suc n) (k ≟ᵗ suc n) x) + +approxSphereMap∙ : ∀ {ℓ} (Xsk : CWskel ℓ) → (x₀ : fst Xsk (suc zero)) (n : ℕ) + (f : S₊∙ (suc n) →∙ (realise Xsk , incl x₀)) + → ∃[ f' ∈ S₊∙ (suc n) →∙ (fst Xsk (suc (suc n)) , CWskel∙ Xsk x₀ (suc n)) ] + (incl∙ Xsk x₀ ∘∙ f' ≡ f) +approxSphereMap∙ Xsk x₀ n f = + PT.rec squash₁ + (λ F → TR.rec squash₁ + (λ fp → ∣ ((λ x → F x .fst .fst) + , sym (cong fst fp)) + , ΣPathP ((funExt (λ x → F x .fst .snd)) + , (SQ' _ _ _ _ _ _ _ _ (cong snd fp))) ∣₁) + (F (ptSn (suc n)) .snd refl)) + approxSphereMap + where + SQ' : ∀ {ℓ} {A : Type ℓ} (x y : A) (p : x ≡ y) + (z : A) (q : y ≡ z) (w : A) (r : x ≡ w) (t : w ≡ z) + → Square (p ∙ q) t r refl → Square (sym r ∙ p) (sym q) t refl + SQ' x = J> (J> (J> λ t s → sym (rUnit refl) + ◁ λ i j → (rUnit refl ◁ s) (~ j) i)) + + approxSphereMap : ∥ ((x : S₊ (suc n)) + → Σ[ fb ∈ fiber (CW↪∞ Xsk (suc (suc n))) (fst f x) ] + ((p : ptSn (suc n) ≡ x) + → hLevelTrunc 1 (PathP (λ i + → fiber (CW↪∞ Xsk (suc (suc n))) (fst f (p i))) + ((CWskel∙ Xsk x₀ (suc n)) + , (CWskel∞∙Id Xsk x₀ (suc n) ∙ sym (snd f))) fb))) ∥₁ + approxSphereMap = sphereToTrunc (suc (suc n)) + {λ x → Σ[ fb ∈ fiber (CW↪∞ Xsk (suc (suc n))) (fst f x) ] + ((p : ptSn (suc n) ≡ x) + → hLevelTrunc 1 (PathP (λ i → fiber (CW↪∞ Xsk (suc (suc n))) (fst f (p i))) + ((CWskel∙ Xsk x₀ (suc n)) + , (CWskel∞∙Id Xsk x₀ (suc n) ∙ sym (snd f))) fb) )} + λ a → TR.rec (isOfHLevelTrunc (suc (suc n))) + (λ fa → ∣ fa + , (λ p → J (λ a p → (fa : fiber (CW↪∞ Xsk (suc (suc n))) (fst f a)) + → hLevelTrunc 1 + (PathP (λ i → fiber (CW↪∞ Xsk (suc (suc n))) (fst f (p i))) + (CWskel∙ Xsk x₀ (suc n) + , CWskel∞∙Id Xsk x₀ (suc n) ∙ (λ i → snd f (~ i))) fa)) + (λ fa → isConnectedPathP 1 (isConnectedSubtr' n 2 + (isConnected-CW↪∞ (suc (suc n)) + Xsk (fst f (ptSn (suc n))))) _ _ .fst) p fa) ∣ₕ) + (isConnected-CW↪∞ (suc (suc n)) Xsk (fst f a) .fst) + +module _ {ℓ} (X : CWskel ℓ) (n : ℕ) (x₀ : fst X 1) + (faprx : S₊∙ n →∙ (fst X (suc n) , CWskel∙ X x₀ n)) + (f : S₊∙ n →∙ (realise X , incl x₀)) + (faprx≡ : (x : _) → incl {n = suc n} (fst faprx x) ≡ fst f x) where + + cellMapSˢᵏᵉˡImproved : cellMap (Sˢᵏᵉˡ n) X + cellMapSˢᵏᵉˡImproved = cellMapSˢᵏᵉˡ X n x₀ (fst faprx) (snd faprx) + + isApproxCellMapSˢᵏᵉˡImproved : realiseSequenceMap cellMapSˢᵏᵉˡImproved + ≡ fst f ∘ invEq (hasCWskelSphere n .snd) + isApproxCellMapSˢᵏᵉˡImproved = + funExt λ x → cong (realiseSequenceMap cellMapSˢᵏᵉˡImproved) + (sym (secEq (hasCWskelSphere n .snd) x)) + ∙ lem _ + where + lem : (x : _) + → realiseSequenceMap cellMapSˢᵏᵉˡImproved (fst (hasCWskelSphere n .snd) x) + ≡ fst f x + lem x with (n ≟ᵗ n) + ... | lt x₁ = ⊥.rec (¬m<ᵗm x₁) + ... | eq x₁ = cong (incl {n = suc n}) + (cong (λ p → subst (fst X) p (fst faprx x)) + (isSetℕ _ _ _ refl) + ∙ transportRefl (fst faprx x)) ∙ faprx≡ x + ... | gt x₁ = ⊥.rec (¬m<ᵗm x₁) + + finCellApproxSˢᵏᵉˡImproved : (k : ℕ) + → finCellApprox (Sˢᵏᵉˡ n) X (fst f ∘ invEq (hasCWskelSphere n .snd)) k + FinSequenceMap.fmap (fst (finCellApproxSˢᵏᵉˡImproved k)) = + SequenceMap.map cellMapSˢᵏᵉˡImproved ∘ fst + FinSequenceMap.fcomm (fst (finCellApproxSˢᵏᵉˡImproved k)) = + SequenceMap.comm cellMapSˢᵏᵉˡImproved ∘ fst + snd (finCellApproxSˢᵏᵉˡImproved k) = →FinSeqColimHomotopy _ _ + (funExt⁻ isApproxCellMapSˢᵏᵉˡImproved ∘ FinSeqColim→Colim k ∘ (fincl flast)) + +-- cellMapSˢᵏᵉˡ preserves ∙Π (homotopy group addition) +cellMapSˢᵏᵉˡ∙Π : ∀ {ℓ} (X : CWskel ℓ) (n : ℕ) (x₀ : fst X 1) + (faprx gaprx : S₊∙ (suc n) →∙ (fst X (suc (suc n)) , CWskel∙ X x₀ (suc n))) + (f : S₊∙ (suc n) →∙ (realise X , incl x₀)) + (faprx≡ : incl∙ X x₀ ∘∙ faprx ≡ f) + (g : S₊∙ (suc n) →∙ (realise X , incl x₀)) + (gaprx≡ : incl∙ X x₀ ∘∙ gaprx ≡ g) + → realiseCellMap (cellMapSˢᵏᵉˡImproved X (suc n) x₀ (∙Π faprx gaprx) (∙Π f g) + (λ x → funExt⁻ (cong fst (∙Π∘∙ n faprx gaprx (incl∙ X x₀))) x + ∙ λ i → ∙Π (faprx≡ i) (gaprx≡ i) .fst x)) + ≡ (∙Π f g .fst ∘ invEq (hasCWskelSphere (suc n) .snd)) +cellMapSˢᵏᵉˡ∙Π X n x₀ faprx gaprx = + J> (J> funExt λ x → cong h (sym (secEq (hasCWskelSphere (suc n) .snd) x)) + ∙ main _ + ∙ cong (∙Π (incl∙ X x₀ ∘∙ faprx) (incl∙ X x₀ ∘∙ gaprx) .fst) + (retEq (SfamTopElement (suc n)) _)) + where + h = realiseCellMap (cellMapSˢᵏᵉˡImproved X (suc n) x₀ + (∙Π faprx gaprx) (∙Π (incl∙ X x₀ ∘∙ faprx) (incl∙ X x₀ ∘∙ gaprx)) + (λ x → (funExt⁻ (cong fst (∙Π∘∙ n faprx gaprx (incl∙ X x₀))) x) ∙ refl)) + + main : (x : Sgen.Sfam (suc n) (suc (suc n)) (suc (suc n) ≟ᵗ suc (suc n))) + → h (incl {n = suc (suc n)} x) + ≡ ∙Π (incl∙ X x₀ ∘∙ faprx) (incl∙ X x₀ ∘∙ gaprx) .fst + (invEq (SfamTopElement (suc n)) x) + main with (n ≟ᵗ n) + ... | lt x = ⊥.rec (¬m<ᵗm x) + ... | eq x = λ x + → cong (incl {n = suc (suc n)}) + (cong (λ p → subst (fst X) p (fst (∙Π faprx gaprx) x)) (isSetℕ _ _ _ refl) + ∙ transportRefl _) + ∙ funExt⁻ (cong fst (∙Π∘∙ n faprx gaprx (incl∙ X x₀))) x + ... | gt x = ⊥.rec (¬m<ᵗm x) diff --git a/Cubical/CW/Instances/SphereBouquetMap.agda b/Cubical/CW/Instances/SphereBouquetMap.agda new file mode 100644 index 0000000000..d32c8f58e3 --- /dev/null +++ b/Cubical/CW/Instances/SphereBouquetMap.agda @@ -0,0 +1,320 @@ +{- +File contains : a direct description of cell structure for the +cofibre of a map α : ⋁ₐ Sⁿ → ⋁ₑ Sⁿ (with a and e finite sets) +-} + +module Cubical.CW.Instances.SphereBouquetMap where + +open import Cubical.CW.Base + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Bool +open import Cubical.Data.Fin.Inductive +open import Cubical.Data.Nat.Order.Inductive +open import Cubical.Data.Sigma +open import Cubical.Data.Unit + +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp +open import Cubical.HITs.Pushout +open import Cubical.HITs.SequentialColimit +open import Cubical.HITs.SphereBouquet +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.Wedge + +open import Cubical.Relation.Nullary + +-- Explicit definition of CW structure on the cofibre of a map of +-- (finite) sphere bouquets +module _ (c1 c2 : ℕ) {n : ℕ} where + SphereBouquet/FamGen : (α : FinSphereBouquetMap c1 c2 n) + → (m : ℕ) → Trichotomyᵗ m (suc (suc n)) → Type + SphereBouquet/FamGen a zero p = ⊥ + SphereBouquet/FamGen a (suc m) (lt x) = Unit + SphereBouquet/FamGen a (suc m) (eq x) = SphereBouquet (suc n) (Fin c2) + SphereBouquet/FamGen a (suc m) (gt x) = cofib a + + SphereBouquet/CardGen : (m : ℕ) + → Trichotomyᵗ m (suc n) → Trichotomyᵗ m (suc (suc n)) → ℕ + SphereBouquet/CardGen zero p q = 1 + SphereBouquet/CardGen (suc m) (lt x) q = 0 + SphereBouquet/CardGen (suc m) (eq x) q = c2 + SphereBouquet/CardGen (suc m) (gt x) (lt y) = 0 + SphereBouquet/CardGen (suc m) (gt x) (eq y) = c1 + SphereBouquet/CardGen (suc m) (gt x) (gt y) = 0 + + SphereBouquet/αGen : (α : FinSphereBouquetMap c1 c2 n) (m : ℕ) + (p : Trichotomyᵗ m (suc n)) (q : Trichotomyᵗ m (suc (suc n))) + → Fin (SphereBouquet/CardGen m p q) × S⁻ m → SphereBouquet/FamGen α m q + SphereBouquet/αGen a (suc m) p (lt y) x = tt + SphereBouquet/αGen a (suc m) (eq x₂) (eq y) x = inl tt + SphereBouquet/αGen a (suc m) (gt x₂) (eq y) x = + a (inr (fst x , subst S₊ (cong predℕ y) (snd x))) + SphereBouquet/αGen a (suc m) p (gt y) x = inl tt + + SphereBouquet/EqContrGen : (α : FinSphereBouquetMap c1 c2 n) + (m : ℕ) (m< : m <ᵗ suc n) + (p : Trichotomyᵗ m (suc n)) (q : Trichotomyᵗ m (suc (suc n))) + → isContr (Pushout (SphereBouquet/αGen α m p q) fst) + SphereBouquet/EqContrGen a zero m< (lt x) (lt y) = + (inr fzero) , λ { (inr (zero , tt)) → refl} + SphereBouquet/EqContrGen a zero m< (lt x) (eq y) = ⊥.rec (snotz (sym y)) + SphereBouquet/EqContrGen a zero m< (eq x) q = ⊥.rec (snotz (sym x)) + SphereBouquet/EqContrGen a (suc m) m< (lt y) (lt x) = + (inl tt) , (λ {(inl tt) → refl}) + SphereBouquet/EqContrGen a (suc m) m< (eq y) (lt x) = + ⊥.rec (¬m<ᵗm {suc n} (subst (_<ᵗ suc n) y m<)) + SphereBouquet/EqContrGen a (suc m) m< (gt y) (lt x) = + ⊥.rec (¬m<ᵗm {m} (<ᵗ-trans {m} {n} {m} m< y)) + SphereBouquet/EqContrGen a (suc m) m< p (eq x) = + ⊥.rec (falseDichotomies.lt-eq (m< , (cong predℕ x))) + SphereBouquet/EqContrGen a (suc m) m< p (gt x) = + ⊥.rec (¬-suc-n<ᵗn {n} (<ᵗ-trans {suc n} {m} {n} x m<)) + + SphereBouquet/EqBottomMainGen : (α : FinSphereBouquetMap c1 c2 n) + → SphereBouquet (suc n) (Fin c2) + ≃ cofib {A = Fin c2 × S₊ n} {B = Fin c2} fst + SphereBouquet/EqBottomMainGen α = isoToEquiv + (compIso (pushoutIso _ _ _ _ (idEquiv _) (idEquiv Unit) + (Σ-cong-equiv-snd (λ a → isoToEquiv (IsoSucSphereSusp n))) + refl + (funExt (λ a → ΣPathP (refl , IsoSucSphereSusp∙' n)))) + (invIso (Iso-cofibFst-⋁ λ _ → S₊∙ n))) + + SphereBouquet/EqBottomGen : (α : FinSphereBouquetMap c1 c2 n) (m : ℕ) + (r : m ≡ suc n) (p : Trichotomyᵗ m (suc n)) (q : Trichotomyᵗ m (suc (suc n))) + → SphereBouquet (suc n) (Fin c2) ≃ Pushout (SphereBouquet/αGen α m p q) fst + SphereBouquet/EqBottomGen a m m< (lt x) q = + ⊥.rec (¬m<ᵗm {suc n} (subst (_<ᵗ suc n) m< x)) + SphereBouquet/EqBottomGen a zero m< (eq x) (lt y) = ⊥.rec (snotz (sym x)) + SphereBouquet/EqBottomGen a (suc m) m< (eq x) (lt y) = + compEquiv (SphereBouquet/EqBottomMainGen a) + (pathToEquiv λ i → cofib {A = Fin c2 × S₊ (predℕ (m< (~ i)))} + {B = Fin c2} fst) + SphereBouquet/EqBottomGen a m m< (eq x) (eq y) = + ⊥.rec (falseDichotomies.eq-eq (x , y)) + SphereBouquet/EqBottomGen a m m< (eq x) (gt y) = + ⊥.rec (falseDichotomies.eq-gt (x , y)) + SphereBouquet/EqBottomGen a m m< (gt x) q = + ⊥.rec (¬m<ᵗm {suc n} (subst (suc n <ᵗ_) m< x)) + + SphereBouquet/EqTopGen' : (m : ℕ) (α : FinSphereBouquetMap c1 c2 n) + (p : m ≡ suc n) + → cofib α ≃ Pushout (α ∘ (λ x → inr (fst x , subst S₊ p (snd x)))) fst + SphereBouquet/EqTopGen' m a p = + compEquiv (compEquiv (symPushout _ _) + (pushoutEquiv _ _ _ _ (idEquiv _) (idEquiv _) + (invEquiv (isContr→≃Unit (isContrLem c1 n m (sym p)))) + (λ i x → a x) + λ i x → isContrLem c1 n m (sym p) .snd (inl x) i)) + (invEquiv (isoToEquiv + (Iso-PushoutComp-IteratedPushout + (λ x → inr (fst x , subst S₊ p (snd x))) a))) + where + isContrLem : (c1 : ℕ) (n m : ℕ) (x : suc n ≡ m) + → isContr (Pushout {A = Fin c1 × S₊ m} {B = SphereBouquet (suc n) (Fin c1)} + (λ x₂ → inr (fst x₂ , subst S₊ (sym x) (snd x₂))) fst) + isContrLem c1 n = + J> subst isContr + (λ i → Pushout {B = SphereBouquet (suc n) (Fin c1)} + (λ x₂ → inr (fst x₂ , transportRefl (snd x₂) (~ i))) fst) + main + where + main : isContr (Pushout inr fst) + fst main = inl (inl tt) + snd main (inl (inl x)) = refl + snd main (inl (inr x)) = + (λ i → inl (push (fst x) i)) + ∙ push (fst x , ptSn (suc n)) + ∙ sym (push x) + snd main (inl (push a i)) j = lem i j + where + lem : Square refl ((λ i₁ → inl (push a i₁)) + ∙ push (a , ptSn (suc n)) + ∙ sym (push (a , ptSn (suc n)))) + refl λ i → inl (push a i) + lem = (λ i j → inl (push a (i ∧ j))) + ▷ (rUnit _ + ∙ cong ((λ i₁ → inl (push a i₁)) ∙_) + (sym (rCancel (push (a , ptSn (suc n)))))) + snd main (inr x) = (λ i → inl (push x i)) ∙ push (x , ptSn (suc n)) + snd main (push a i) j = + ((λ i₁ → inl (push (fst a) i₁)) + ∙ compPath-filler (push (fst a , ptSn (suc n))) (sym (push a)) (~ i)) j + + SphereBouquet/EqTopGen : (m : ℕ) (α : FinSphereBouquetMap c1 c2 n) + → suc n <ᵗ m → (p : Trichotomyᵗ m (suc n)) (q : Trichotomyᵗ m (suc (suc n))) + → cofib α ≃ Pushout (SphereBouquet/αGen α m p q) fst + SphereBouquet/EqTopGen (suc m) a m< (lt x) q = + ⊥.rec (¬m<ᵗm {n} (<ᵗ-trans {n} {m} {n} m< x)) + SphereBouquet/EqTopGen (suc m) a m< (eq x) q = + ⊥.rec (¬m<ᵗm {suc m} (subst (_<ᵗ suc m) (sym x) m<)) + SphereBouquet/EqTopGen (suc m) a m< (gt x) (lt y) = + ⊥.rec (¬squeeze {m} {suc n} (y , x)) + SphereBouquet/EqTopGen (suc m) a m< (gt x) (eq y) = + SphereBouquet/EqTopGen' m a (cong predℕ y) + SphereBouquet/EqTopGen (suc m) a m< (gt x) (gt y) = + isoToEquiv (PushoutEmptyFam (λ()) λ()) + + SphereBouquet/EqGen : (m : ℕ) (α : FinSphereBouquetMap c1 c2 n) + (p : Trichotomyᵗ (suc m) (suc (suc n))) + (q : Trichotomyᵗ m (suc n)) (q' : Trichotomyᵗ m (suc (suc n))) + → (SphereBouquet/FamGen α (suc m) p) + ≃ Pushout (SphereBouquet/αGen α m q q') fst + SphereBouquet/EqGen m a (lt x) q q' = + invEquiv (isContr→≃Unit (SphereBouquet/EqContrGen a m x q q')) + SphereBouquet/EqGen m a (eq x) q q' = + SphereBouquet/EqBottomGen a m (cong predℕ x) q q' + SphereBouquet/EqGen m a (gt x) q q' = SphereBouquet/EqTopGen m a x q q' + + ¬SphereBouquet/CardGen : (k : ℕ) (ineq : suc (suc n) <ᵗ k) (p : _) (q : _) + → ¬ (Fin (SphereBouquet/CardGen k p q)) + ¬SphereBouquet/CardGen (suc k) ineq (eq x) q c = + falseDichotomies.eq-gt (x , ineq) + ¬SphereBouquet/CardGen (suc k) ineq (gt x) (eq y) c = + ¬m<ᵗm {suc n} (subst (suc n <ᵗ_) (cong predℕ y) ineq) + + SphereBouquet/ˢᵏᵉˡConverges : (k : ℕ) (α : FinSphereBouquetMap c1 c2 n) + → suc (suc n) <ᵗ k + → (p : _) (q : _) + → isEquiv {B = Pushout (SphereBouquet/αGen α k p q) fst} inl + SphereBouquet/ˢᵏᵉˡConverges k a ineq p q = + isoToIsEquiv (PushoutEmptyFam (¬SphereBouquet/CardGen k ineq p q ∘ fst) + (¬SphereBouquet/CardGen k ineq p q)) + + SphereBouquet/FamMidElementGen : + (k : ℕ) (α : FinSphereBouquetMap c1 c2 n) + → suc (suc n) ≡ k → (p : _) + → SphereBouquet (suc n) (Fin c2) ≃ (SphereBouquet/FamGen α k p) + SphereBouquet/FamMidElementGen k q s (lt x) = + ⊥.rec (¬m<ᵗm {n} (subst (_<ᵗ suc (suc n)) (sym s) x)) + SphereBouquet/FamMidElementGen zero q s (eq x) = ⊥.rec (snotz (sym x)) + SphereBouquet/FamMidElementGen (suc k) q s (eq x) = idEquiv _ + SphereBouquet/FamMidElementGen k q s (gt x) = + ⊥.rec (¬m<ᵗm {k} (subst (_<ᵗ k) s x)) + + SphereBouquet/FamTopElementGen : (k : ℕ) (α : FinSphereBouquetMap c1 c2 n) + → suc (suc n) <ᵗ k → (p : _) + → cofib α ≃ (SphereBouquet/FamGen α k p) + SphereBouquet/FamTopElementGen (suc k) α ineq (lt x) = + ⊥.rec (¬m<ᵗm {k} (<ᵗ-trans {k} {suc n} {k} x ineq)) + SphereBouquet/FamTopElementGen (suc k) α ineq (eq x) = + ⊥.rec (¬m<ᵗm {k} (subst (_<ᵗ k) (cong predℕ (sym x)) ineq)) + SphereBouquet/FamTopElementGen (suc k) α ineq (gt x) = idEquiv _ + +SphereBouquet/EqBottomMainGenLem : {C : Type} {c1 c2 : ℕ} (n : ℕ) + (α : FinSphereBouquetMap c1 c2 n) {e : C ≃ _} + → (a : _) → Pushout→Bouquet (suc n) c2 (λ _ → tt) e + (fst (SphereBouquet/EqBottomMainGen c1 c2 α) a) + ≡ a +SphereBouquet/EqBottomMainGenLem n α (inl x) = refl +SphereBouquet/EqBottomMainGenLem zero α (inr (a , base)) = push a +SphereBouquet/EqBottomMainGenLem {c1 = c1} {c2} zero α + {e = e} (inr (a , loop i)) j = main j i + where + main : Square (cong (Pushout→Bouquet 1 c2 (λ _ → tt) e) + λ i → fst (SphereBouquet/EqBottomMainGen c1 c2 α) + (inr (a , loop i))) + (λ i → inr (a , loop i)) + (push a) (push a) + main = (λ j i → Pushout→Bouquet 1 c2 (λ _ → tt) e + (⋁→cofibFst {A = Fin c2} (λ _ → Bool , true) + (inr (a , toSusp (Bool , true) false i)))) + ∙ cong-∙ (λ t → Pushout→Bouquet 1 c2 (λ _ → tt) e + (⋁→cofibFst (λ _ → Bool , true) (inr (a , t)))) + (merid false) + (sym (merid true)) + ∙ cong₂ _∙_ refl (sym (rUnit (sym (push a)))) + ∙ (λ _ → (push a ∙ (λ i₁ → inr (a , loop i₁))) ∙ (λ i₁ → push a (~ i₁))) + ∙ sym (assoc (push a) (λ i → inr (a , loop i)) (sym (push a))) + ∙ sym (doubleCompPath≡compPath + (push a) (λ i → inr (a , loop i)) (sym (push a))) + ◁ symP (doubleCompPath-filler + (push a) (λ i → inr (a , loop i)) (sym (push a))) +SphereBouquet/EqBottomMainGenLem (suc n) α (inr (a , north)) = push a +SphereBouquet/EqBottomMainGenLem (suc n) α (inr (a , south)) = + λ i → inr (a , merid (ptSn (suc n)) i) +SphereBouquet/EqBottomMainGenLem {c1 = c1} {c2} (suc n) α + {e = e} (inr (a , merid x i)) j = main j i + where + main : Square (cong (Pushout→Bouquet (suc (suc n)) c2 (λ _ → tt) e) + λ i → fst (SphereBouquet/EqBottomMainGen c1 c2 α) + (inr (a , merid x i))) + (λ i → inr (a , merid x i)) + (push a) λ i → inr (a , merid (ptSn (suc n)) i) + main = (cong (push a ∙_) + (cong-∙ (inr ∘ (a ,_)) (merid x) (sym (merid (ptSn (suc n))))) + ∙ sym (doubleCompPath≡compPath (push a) + (λ i → inr (a , merid x i)) + (λ i → inr (a , merid (ptSn (suc n)) (~ i))))) + ◁ symP (doubleCompPath-filler (push a) + (λ i → inr (a , merid x i)) + (λ i → inr (a , merid (ptSn (suc n)) (~ i)))) +SphereBouquet/EqBottomMainGenLem {c1 = c1} {c2} zero α + {e = e} (push a i) j = lem j i + where + lem : Square (cong (Pushout→Bouquet (suc zero) c2 (λ _ → tt) e) + (cong (fst (SphereBouquet/EqBottomMainGen c1 c2 α)) + (push a))) + (push a) refl (push a) + lem = (λ j i → Pushout→Bouquet 1 c2 (λ _ → tt) e + (Iso.inv (Iso-cofibFst-⋁ λ _ → S₊∙ zero) + (lUnit (push a) (~ j) i))) + ◁ λ i j → push a (i ∧ j) +SphereBouquet/EqBottomMainGenLem {c1 = c1} {c2} (suc n) α + {e = e} (push a i) j = lem j i + where + lem : Square (cong (Pushout→Bouquet (suc (suc n)) c2 (λ _ → tt) e) + (cong (fst (SphereBouquet/EqBottomMainGen c1 c2 α)) + (push a))) + (push a) refl (push a) + lem = (λ j i → Pushout→Bouquet (suc (suc n)) c2 (λ _ → tt) e + (Iso.inv (Iso-cofibFst-⋁ λ _ → S₊∙ (suc n)) + (lUnit (push a) (~ j) i))) + ◁ λ i j → push a (i ∧ j) + +-- Final product +module _ {c1 c2 : ℕ} {n : ℕ} (α : FinSphereBouquetMap c1 c2 n) where + private + α∙ : ∥ α (inl tt) ≡ inl tt ∥₁ + α∙ = isConnectedSphereBouquet _ + + SphereBouquet/ˢᵏᵉˡ : CWskel ℓ-zero + fst SphereBouquet/ˢᵏᵉˡ m = SphereBouquet/FamGen c1 c2 α m (m ≟ᵗ (suc (suc n))) + fst (snd SphereBouquet/ˢᵏᵉˡ) m = + SphereBouquet/CardGen c1 c2 m (m ≟ᵗ suc n) (m ≟ᵗ suc (suc n)) + fst (snd (snd SphereBouquet/ˢᵏᵉˡ)) m = + SphereBouquet/αGen c1 c2 α m (m ≟ᵗ suc n) (m ≟ᵗ suc (suc n)) + fst (snd (snd (snd SphereBouquet/ˢᵏᵉˡ))) () + snd (snd (snd (snd SphereBouquet/ˢᵏᵉˡ))) m = + SphereBouquet/EqGen c1 c2 m α + (suc m ≟ᵗ suc (suc n)) (m ≟ᵗ suc n) (m ≟ᵗ suc (suc n)) + + hasCWskelSphereBouquet/ : hasCWskel (cofib α) + fst hasCWskelSphereBouquet/ = SphereBouquet/ˢᵏᵉˡ + snd hasCWskelSphereBouquet/ = + compEquiv (SphereBouquet/FamTopElementGen c1 c2 (suc (suc (suc n))) α + (<ᵗsucm {n}) (suc (suc (suc n)) ≟ᵗ suc (suc n))) + (isoToEquiv (converges→ColimIso (suc (suc (suc n))) + λ k → compEquiv (inl + , SphereBouquet/ˢᵏᵉˡConverges c1 c2 (k +ℕ suc (suc (suc n))) α + (<→<ᵗ (k , refl)) + ((k +ℕ suc (suc (suc n))) ≟ᵗ suc n) + ((k +ℕ suc (suc (suc n))) ≟ᵗ suc (suc n))) + (invEquiv (SphereBouquet/EqGen c1 c2 (k +ℕ suc (suc (suc n))) + α + ((suc (k +ℕ suc (suc (suc n)))) ≟ᵗ suc (suc n)) + ((k +ℕ suc (suc (suc n))) ≟ᵗ suc n) _)) .snd)) + + SphereBouquet/ᶜʷ : CW ℓ-zero + fst SphereBouquet/ᶜʷ = cofib α + snd SphereBouquet/ᶜʷ = ∣ hasCWskelSphereBouquet/ ∣₁ diff --git a/Cubical/CW/Instances/Susp.agda b/Cubical/CW/Instances/Susp.agda new file mode 100644 index 0000000000..42ca8693b7 --- /dev/null +++ b/Cubical/CW/Instances/Susp.agda @@ -0,0 +1,33 @@ +{- +CW structure on suspensions +-} + +module Cubical.CW.Instances.Susp where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Nat + +open import Cubical.HITs.Pushout +open import Cubical.HITs.Susp + +open import Cubical.CW.Base +open import Cubical.CW.Instances.Pushout +open import Cubical.CW.Instances.Empty +open import Cubical.CW.Instances.Unit + +-- todo: explicit construction of arbitrary suspensions + +-- Suspensions of finite CW complexes +isFinCWSusp : (A : finCW ℓ-zero) → isFinCW (Susp (fst A)) +isFinCWSusp A = + subst isFinCW PushoutSusp≡Susp + (isFinCWPushout A finCWUnit finCWUnit _ _) + +-- Iterated suspensions of finite CW complexes +isFinCWSusp^' : (A : finCW ℓ-zero) (n : ℕ) → isFinCW (Susp^' n (fst A)) +isFinCWSusp^' A zero = snd A +isFinCWSusp^' A (suc n) = isFinCWSusp (_ , isFinCWSusp^' A n) + +isFinCWSusp^ : (A : finCW ℓ-zero) (n : ℕ) → isFinCW (Susp^ n (fst A)) +isFinCWSusp^ A n = subst isFinCW (Susp^'≡Susp^ n) (isFinCWSusp^' A n) diff --git a/Cubical/CW/Instances/Unit.agda b/Cubical/CW/Instances/Unit.agda new file mode 100644 index 0000000000..5206269abd --- /dev/null +++ b/Cubical/CW/Instances/Unit.agda @@ -0,0 +1,57 @@ +{- +The unit type as a CW complex +-} +module Cubical.CW.Instances.Unit where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_) +open import Cubical.Data.Fin.Inductive +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Unit +open import Cubical.Data.Sequence.Base + +open import Cubical.HITs.Pushout +open import Cubical.HITs.SequentialColimit +open import Cubical.HITs.PropositionalTruncation + +open import Cubical.CW.Base + +CWskelUnit* : ∀ {ℓ} → CWskel ℓ +fst CWskelUnit* zero = ⊥* +fst CWskelUnit* (suc n) = Unit* +fst (snd CWskelUnit*) zero = 1 +fst (snd CWskelUnit*) (suc x) = 0 +fst (snd (snd CWskelUnit*)) zero () +fst (snd (snd CWskelUnit*)) (suc n) _ = tt* +snd (snd (snd (snd CWskelUnit*))) zero = + compEquiv (compEquiv (compEquiv (invEquiv LiftEquiv) + (isoToEquiv Iso-Unit-Fin1)) + (isoToEquiv (PushoutEmptyFam (λ()) λ()))) (symPushout _ _) +snd (snd (snd (snd CWskelUnit*))) (suc n) = + isoToEquiv (PushoutEmptyFam (λ()) λ()) + +convergesCWskelUnit* : ∀ {ℓ} → converges (realiseSeq (CWskelUnit* {ℓ})) 1 +convergesCWskelUnit* zero = idIsEquiv _ +convergesCWskelUnit* (suc k) = idIsEquiv _ + +hasFinCWskelUnit* : ∀ {ℓ} → hasFinCWskel (Unit* {ℓ}) +fst hasFinCWskelUnit* = 1 +fst (fst (snd hasFinCWskelUnit*)) = CWskelUnit* .fst +fst (snd (fst (snd hasFinCWskelUnit*))) = CWskelUnit* .snd +snd (snd (fst (snd hasFinCWskelUnit*))) = convergesCWskelUnit* +fst (snd (snd hasFinCWskelUnit*)) = _ +snd (snd (snd hasFinCWskelUnit*)) = + converges→isEquivIncl {seq = realiseSeq CWskelUnit*} 1 + convergesCWskelUnit* + +finCWUnit* : (ℓ : Level) → finCW ℓ +fst (finCWUnit* ℓ) = Unit* +snd (finCWUnit* ℓ) = ∣ hasFinCWskelUnit* ∣₁ + +finCWUnit : finCW ℓ-zero +fst finCWUnit = Unit +snd finCWUnit = subst isFinCW (sym (ua Unit≃Unit*)) ∣ hasFinCWskelUnit* ∣₁ diff --git a/Cubical/CW/Strictification.agda b/Cubical/CW/Strictification.agda new file mode 100644 index 0000000000..a1e39de8ca --- /dev/null +++ b/Cubical/CW/Strictification.agda @@ -0,0 +1,209 @@ +{-# OPTIONS --lossy-unification #-} +{- +This file contains +1. A 'strictification' functor + strictCWskel : CWskel -> CWskel s.t. strictCWskel C ≡ C + the strictified version of C satisfies the pushout condition + definitionally +2. The same thing for cellular maps +-} +module Cubical.CW.Strictification where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Fin.Inductive.Base +open import Cubical.Data.Sigma +open import Cubical.Data.Sequence +open import Cubical.Data.Nat + +open import Cubical.HITs.Sn +open import Cubical.HITs.Pushout + +open import Cubical.CW.Base +open import Cubical.CW.Map + +open SequenceMap renaming (map to ∣_∣) +open CWskel-fields + + +-- Strictification of CW skels +module _ {ℓ : Level} (B : CWskel ℓ) where + open CWskel-fields + strictifyFam : ℕ → Type ℓ + strictifyFam≡ : (n : ℕ) → strictifyFam n ≡ fst B n + strictifyFame : (n : ℕ) → fst B n ≃ strictifyFam n + strictifyFamα : (n : ℕ) → Fin (fst (B .snd) n) × S⁻ n → strictifyFam n + strictifyFame2 : (n : ℕ) + → (Pushout (α B n) fst) ≃ (Pushout (strictifyFamα n) fst) + strictifyFam zero = fst B 0 + strictifyFam (suc n) = Pushout (strictifyFamα n) fst + strictifyFamα n = fst (strictifyFame n) ∘ α B n + strictifyFame zero = idEquiv _ + strictifyFame (suc n) = + compEquiv (e B n) + (strictifyFame2 n) + strictifyFame2 n = + pushoutEquiv _ _ _ _ (idEquiv _) (strictifyFame n) (idEquiv _) + (λ _ x → fst (strictifyFame n) (α B n x)) + (λ _ x → fst x) + strictifyFam≡ zero = refl + strictifyFam≡ (suc n) = ua (invEquiv (strictifyFame (suc n))) + + strictCWskel : CWskel ℓ + fst strictCWskel = strictifyFam + fst (snd strictCWskel) = card B + fst (snd (snd strictCWskel)) = strictifyFamα + fst (snd (snd (snd strictCWskel))) = fst (snd (snd (snd B))) + snd (snd (snd (snd strictCWskel))) = λ _ → idEquiv _ + + strict≡Gen : ∀ {ℓ ℓ'} {A : Type ℓ} {C D : Type ℓ'} + (α : A → C) (e : C ≃ D) (x : A) + → PathP (λ i → ua (invEquiv e) i) (fst e (α x)) (α x) + strict≡Gen α e x i = + hcomp (λ k → λ {(i = i0) → fst e (α x) + ; (i = i1) → retEq e (α x) k}) + (ua-gluePt (invEquiv e) i (fst e (α x))) + + strict≡GenT' : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C D : Type ℓ''} + {E : Type (ℓ-max ℓ (ℓ-max ℓ' ℓ''))} (g : A → B) + (e : C ≃ D) (α : A → C) (e' : E ≃ Pushout (λ x₁ → α x₁) g) + → PathP (λ k → ua (invEquiv (compEquiv {C = Pushout (fst e ∘ α) g} e' + (pushoutEquiv _ _ _ _ (idEquiv A) e (idEquiv B) + (λ i x → fst e (α x)) λ i x → g x))) k + ≃ Pushout (λ x₁ → strict≡Gen α e x₁ k) g) + (idEquiv _) + e' + strict≡GenT' {A = A} {B} {C} {D} {E} g = + EquivJ (λ C e → (α : A → C) (e' : E ≃ Pushout (λ x₁ → α x₁) g) + → PathP (λ k → ua (invEquiv (compEquiv {C = Pushout (fst e ∘ α) g} e' + (pushoutEquiv _ _ _ _ (idEquiv A) e (idEquiv B) + (λ i x → fst e (α x)) λ i x → g x))) k + ≃ Pushout (λ x₁ → strict≡Gen α e x₁ k) g) + (idEquiv _) e') λ a → + EquivJ (λ E e' + → PathP (λ k → ua (invEquiv (compEquiv e' + (pushoutEquiv a g a g (idEquiv A) (idEquiv D) (idEquiv B) + (λ i x → a x) (λ i → g)))) k + ≃ Pushout (λ x₁ → strict≡Gen a (idEquiv D) x₁ k) g) + (idEquiv (Pushout (λ x → idfun D (a x)) g)) e') + (ΣPathPProp isPropIsEquiv + (transport (λ k + → PathP (λ i → (sym (lem g a) + ∙ sym (cong (ua ∘ invEquiv) (compEquivIdEquiv (pushoutEquiv a g + (λ z → idfun D (a z)) g (idEquiv A) (idEquiv D) + (idEquiv B) (λ i₁ x → idfun D (a x)) (λ i₁ → g))))) k i + → Pushout (λ x₁ → strict≡GenId a x₁ (~ k) i) g) + (idfun _) (idfun _)) + (funExt (main _ _)))) + where + strict≡GenId : ∀ {ℓ ℓ'} {A : Type ℓ} {C : Type ℓ'} (α : A → C)(x : A) + → strict≡Gen α (idEquiv C) x + ≡ λ i → ua-gluePt (invEquiv (idEquiv C)) i (α x) + strict≡GenId {C = C} α x j i = + hcomp (λ k → λ {(i = i0) → α x + ; (i = i1) → α x + ; (j = i1) → ua-gluePt (invEquiv (idEquiv C)) i (α x)}) + (ua-gluePt (invEquiv (idEquiv C)) i (α x)) + + lem : (g : A → B) (α : A → D) + → ua (invEquiv (pushoutEquiv α g α g + (idEquiv A) (idEquiv D) (idEquiv B) refl refl)) + ≡ refl + lem g a = cong ua + (cong invEquiv (Σ≡Prop isPropIsEquiv {v = idEquiv _} + (funExt λ { (inl x) → refl ; (inr x) → refl + ; (push a i) j → rUnit (push a) (~ j) i})) + ∙ invEquivIdEquiv _) ∙ uaIdEquiv + + main : (g : A → B) (α : A → D) (w : _) + → PathP (λ i → Pushout (λ s → ua-gluePt (invEquiv (idEquiv D)) i (α s)) g) + w w + main g α (inl x) i = inl (ua-gluePt (invEquiv (idEquiv D)) i x) + main g α (inr x) i = inr x + main g α (push a j) i = push a j + + strict≡α : (n : ℕ) (x : Fin (card B n)) (y : S⁻ n) + → PathP (λ i → strictifyFam≡ n i) + + (strictifyFamα n (x , y)) + (α B n (x , y)) + strict≡α (suc n) x y = + strict≡Gen (α B (suc n)) (strictifyFame (suc n)) (x , y) + + strict≡e : (n : ℕ) + → PathP (λ i → strictifyFam≡ (suc n) i + ≃ Pushout (λ x → strict≡α n (fst x) (snd x) i) fst) + (idEquiv _) (e B n) + strict≡e zero = + ΣPathPProp (λ _ → isPropIsEquiv _) + (symP (toPathP (funExt + λ { (inl x) → ⊥.rec (B .snd .snd .snd .fst x) + ; (inr x) + → cong (transport (λ i → Pushout (λ s → strict≡α zero + (fst s) (snd s) (~ i)) fst)) + (cong (e B 0 .fst) (transportRefl (invEq (e B 0) (inr x))) + ∙ secEq (e B 0) (inr x)) + ∙ λ j → inr (transportRefl x j)}))) + strict≡e (suc n) = + strict≡GenT' fst (strictifyFame (suc n)) (α B (suc n)) (e B (suc n)) + + strict≡ : strictCWskel ≡ B + fst (strict≡ i) n = strictifyFam≡ n i + fst (snd (strict≡ i)) = card B + fst (snd (snd (strict≡ i))) n (x , y) = strict≡α n x y i + fst (snd (snd (snd (strict≡ i)))) = fst (snd (snd (snd B))) + snd (snd (snd (snd (strict≡ i)))) n = strict≡e n i + +-- Associated elimination principle +module _ {ℓ ℓ'} {P : CWskel ℓ → Type ℓ'} + (e : (B : CWskel ℓ) → P (strictCWskel B)) where + + elimStrictCW : (B : _) → P B + elimStrictCW B = subst P (strict≡ B) (e B) + + elimStrictCWβ : (B : _) → elimStrictCW (strictCWskel B) ≡ e B + elimStrictCWβ B = + (λ j → subst P (λ i → H strictCWskel (funExt (λ x → sym (strict≡ x))) B i j) + (e (strict≡ B j))) + ∙ transportRefl (e B) + where + H : ∀ {ℓ} {A : Type ℓ} (F : A → A) (s : (λ x → x) ≡ F) (a : A) + → Square (λ i → F (s (~ i) a)) refl (λ i → s (~ i) (F a)) refl + H = J> λ _ → refl + +-- Strictification of cellular maps +module _ {ℓ ℓ'} {C : CWskel ℓ} {D : CWskel ℓ'} + (f : cellMap (strictCWskel C) (strictCWskel D)) where + + strictMapFun : (n : ℕ) → strictifyFam C n → strictifyFam D n + strictMapComm : (n : ℕ) (x : strictifyFam C n) → + strictMapFun n x ≡ ∣ f ∣ n x + strictMapFun zero x = ∣ f ∣ 0 x + strictMapFun (suc n) (inl x) = inl (strictMapFun n x) + strictMapFun (suc n) (inr x) = ∣ f ∣ (suc n) (inr x) + strictMapFun (suc (suc n)) (push c i) = + (((λ i → inl (strictMapComm (suc n) (α (strictCWskel C) (suc n) c) i)) + ∙ comm f (suc n) (α (strictCWskel C) (suc n) c)) + ∙ cong (∣ f ∣ (suc (suc n))) (push c)) i + strictMapComm zero x = refl + strictMapComm (suc n) (inl x) = (λ i → inl (strictMapComm n x i)) ∙ comm f n x + strictMapComm (suc n) (inr x) = refl + strictMapComm (suc (suc n)) (push c i) j = + compPath-filler' + ((λ i → inl (strictMapComm (suc n) (α (strictCWskel C) (suc n) c) i)) + ∙ comm f (suc n) (α (strictCWskel C) (suc n) c)) + (cong (∣ f ∣ (suc (suc n))) (push c)) (~ j) i + + strictCwMap : cellMap (strictCWskel C) (strictCWskel D) + SequenceMap.map strictCwMap = strictMapFun + SequenceMap.comm strictCwMap n x = refl + + strictCwMap≡ : strictCwMap ≡ f + ∣_∣ (strictCwMap≡ i) n x = strictMapComm n x i + comm (strictCwMap≡ i) n x j = + compPath-filler ((λ i₁ → inl (strictMapComm n x i₁))) (comm f n x) j i diff --git a/Cubical/CW/Subcomplex.agda b/Cubical/CW/Subcomplex.agda index ba0486f5b2..92ee821bab 100644 --- a/Cubical/CW/Subcomplex.agda +++ b/Cubical/CW/Subcomplex.agda @@ -17,12 +17,12 @@ This file contains. module Cubical.CW.Subcomplex where open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Pointed open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence open import Cubical.Foundations.Function open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport open import Cubical.Data.Nat open import Cubical.Data.Fin.Inductive.Base @@ -30,144 +30,224 @@ open import Cubical.Data.Fin.Inductive.Properties open import Cubical.Data.Sigma open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Nat.Order.Inductive +open import Cubical.Data.Sequence +open import Cubical.Data.FinSequence open import Cubical.CW.Base open import Cubical.CW.Properties -open import Cubical.CW.ChainComplex -open import Cubical.CW.Approximation +open import Cubical.CW.Map +open import Cubical.HITs.SequentialColimit open import Cubical.HITs.Sn open import Cubical.HITs.Pushout open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.Relation.Nullary -open import Cubical.Algebra.AbGroup -open import Cubical.Algebra.Group.Morphisms -open import Cubical.Algebra.Group.MorphismProperties -open import Cubical.Algebra.ChainComplex - private variable ℓ ℓ' ℓ'' : Level -- finite subcomplex C⁽ⁿ⁾ of a cw complex C. -module _ (C : CWskel ℓ) where - -- The underlying family of types is that of C but ending at some - -- fixed n. - subComplexFam : (n : ℕ) (m : ℕ) → Type ℓ - subComplexFam n m with (m ≟ᵗ n) - ... | lt x = fst C m - ... | eq x = fst C m - ... | gt x = fst C n - - -- Similarly for the number of cells - subComplexCard : (n : ℕ) → ℕ → ℕ - subComplexCard n m with (m ≟ᵗ n) - ... | lt x = snd C .fst m - ... | eq x = 0 - ... | gt x = 0 +module SubComplexGen (C : CWskel ℓ) (n : ℕ) where + subComplexFam : (m : ℕ) → Trichotomyᵗ m n → Type ℓ + subComplexFam m (lt x) = fst C m + subComplexFam m (eq x) = fst C m + subComplexFam m (gt x) = fst C n + + subComplexCard : (m : ℕ) → Trichotomyᵗ m n → ℕ + subComplexCard m (lt x) = snd C .fst m + subComplexCard m (eq x) = 0 + subComplexCard m (gt x) = 0 -- Attaching maps - subComplexα : (n m : ℕ) → Fin (subComplexCard n m) × S⁻ m → subComplexFam n m - subComplexα n m with (m ≟ᵗ n) - ... | lt x = snd C .snd .fst m - ... | eq x = λ x → ⊥.rec (¬Fin0 (fst x)) - ... | gt x = λ x → ⊥.rec (¬Fin0 (fst x)) + subComplexα : (m : ℕ) (p : Trichotomyᵗ m n) + → Fin (subComplexCard m p) × S⁻ m → subComplexFam m p + subComplexα m (lt x) = snd C .snd .fst m - subComplex₀-empty : (n : ℕ) → ¬ subComplexFam n zero - subComplex₀-empty n with (zero ≟ᵗ n) - ... | lt x = CW₀-empty C - ... | eq x = CW₀-empty C + subComplex₀-empty : (p : Trichotomyᵗ zero n) → ¬ subComplexFam zero p + subComplex₀-empty (lt x₁) x = CW₀-empty C x + subComplex₀-empty (eq x₁) x = CW₀-empty C x -- Resulting complex has CW structure - subComplexFam≃Pushout : (n m : ℕ) - → subComplexFam n (suc m) ≃ Pushout (subComplexα n m) fst - subComplexFam≃Pushout n m with (m ≟ᵗ n) | ((suc m) ≟ᵗ n) - ... | lt x | lt x₁ = snd C .snd .snd .snd m - ... | lt x | eq x₁ = snd C .snd .snd .snd m - ... | lt x | gt x₁ = ⊥.rec (¬squeeze (x , x₁)) - ... | eq x | lt x₁ = ⊥.rec (¬m<ᵗm (subst (_<ᵗ n) x (<ᵗ-trans <ᵗsucm x₁))) - ... | eq x | eq x₁ = ⊥.rec (¬m<ᵗm (subst (m <ᵗ_) (x₁ ∙ sym x) <ᵗsucm)) - ... | eq x | gt x₁ = + subComplexFam≃Pushout : (m : ℕ) + (p : Trichotomyᵗ m n) (q : Trichotomyᵗ (suc m) n) + → subComplexFam (suc m) q ≃ Pushout (subComplexα m p) fst + subComplexFam≃Pushout m (lt x) (lt y) = snd C .snd .snd .snd m + subComplexFam≃Pushout m (lt x) (eq y) = snd C .snd .snd .snd m + subComplexFam≃Pushout m (lt x) (gt y) = ⊥.rec (¬squeeze {m} {n} (x , y)) + subComplexFam≃Pushout m (eq x) (lt y) = + ⊥.rec (¬m<ᵗm {n} (subst (_<ᵗ n) x (<ᵗ-trans {m} {suc m} {n} (<ᵗsucm {m}) y))) + subComplexFam≃Pushout m (eq x) (eq y) = + ⊥.rec (¬m<ᵗm {m} (subst (m <ᵗ_) (y ∙ sym x) (<ᵗsucm {m}))) + subComplexFam≃Pushout m (eq x) (gt y) = compEquiv (pathToEquiv (λ i → fst C (x (~ i)))) (isoToEquiv (PushoutEmptyFam (λ x → ¬Fin0 (fst x)) ¬Fin0)) - ... | gt x | lt x₁ = - ⊥.rec (¬squeeze (x₁ , <ᵗ-trans x (<ᵗ-trans <ᵗsucm <ᵗsucm))) - ... | gt x | eq x₁ = - ⊥.rec (¬m<ᵗm (subst (n <ᵗ_) x₁ (<ᵗ-trans x <ᵗsucm))) - ... | gt x | gt x₁ = isoToEquiv (PushoutEmptyFam (λ x → ¬Fin0 (fst x)) ¬Fin0) + subComplexFam≃Pushout m (gt x) (lt y) = + ⊥.rec (¬squeeze {suc m} {n} (y + , <ᵗ-trans {n} {m} {suc (suc m)} x + (<ᵗ-trans {m} {suc m} {suc (suc m)} (<ᵗsucm {m}) (<ᵗsucm {suc m})))) + subComplexFam≃Pushout m (gt x) (eq y) = + ⊥.rec (¬m<ᵗm {n} (subst (n <ᵗ_) y (<ᵗ-trans {n} {m} {suc m} x (<ᵗsucm {m})))) + subComplexFam≃Pushout m (gt x) (gt y) = + isoToEquiv (PushoutEmptyFam (λ x → ¬Fin0 (fst x)) ¬Fin0) + + subComplexCW↪Gen : (m : ℕ) (p : Trichotomyᵗ m n) (q : Trichotomyᵗ (suc m) n) + → subComplexFam m p → subComplexFam (suc m) q + subComplexCW↪Gen m p q x = invEq (subComplexFam≃Pushout m p q) (inl x) + + subComplexFinEquiv' : (m : ℕ) (ineq : n <ᵗ suc m) (p : _) + → isEquiv {B = Pushout (subComplexα m p) fst} inl + subComplexFinEquiv' m ineq (lt x) = ⊥.rec (¬squeeze {m} {n} (x , ineq)) + subComplexFinEquiv' m ineq (eq x) = + isoToIsEquiv (PushoutEmptyFam (¬Fin0 ∘ fst) ¬Fin0) + subComplexFinEquiv' m ineq (gt x) = + isoToIsEquiv (PushoutEmptyFam (¬Fin0 ∘ fst) ¬Fin0) + +CW↑Gen : ∀ {ℓ} (C : CWskel ℓ) (m n : ℕ) + → Trichotomyᵗ n (suc m) → m <ᵗ n → fst C m → fst C n +CW↑Gen C zero (suc n) s p q = ⊥.rec (C .snd .snd .snd .fst q) +CW↑Gen C (suc m) (suc n) (lt y) p x = ⊥.rec (¬squeeze {m} {n} (p , y)) +CW↑Gen C (suc m) (suc n) (eq y) p x = + CW↪ C n (subst (fst C) (sym (cong predℕ y)) x) +CW↑Gen C (suc m) (suc n) (gt y) p x = + CW↪ C n (CW↑Gen C (suc m) n (n ≟ᵗ suc (suc m)) y x) + +CW↑Gen≡ : ∀ {ℓ} (C : CWskel ℓ) (m n : ℕ) + (p : Trichotomyᵗ n (suc m)) (q : m <ᵗ n) (x : fst C m) + → Path (realise C) (incl x) (incl {n = n} (CW↑Gen C m n p q x)) +CW↑Gen≡ C zero (suc n) p q x = ⊥.rec (C .snd .snd .snd .fst x) +CW↑Gen≡ C (suc m) (suc n) (lt x₁) q x = ⊥.rec (¬squeeze {m} {n} (q , x₁)) +CW↑Gen≡ C (suc m) (suc n) (eq x₁) q x = help _ (cong predℕ (sym x₁)) + where + help : (n : ℕ) (p : suc m ≡ n) + → Path (realise C) (incl x) (incl (CW↪ C n (subst (fst C) p x))) + help = J> push x + ∙ λ i → incl {n = suc (suc m)} (CW↪ C (suc m) (transportRefl x (~ i))) +CW↑Gen≡ C (suc m) (suc n) (gt y) q x = + CW↑Gen≡ C (suc m) n (n ≟ᵗ suc (suc m)) y x + ∙ push (CW↑Gen C (suc m) n (n ≟ᵗ suc (suc m)) y x) + +module subComplexMapGen {ℓ : Level} (C : CWskel ℓ) where + subComplex→map' : (m n : ℕ) (p : Trichotomyᵗ n m) + → SubComplexGen.subComplexFam C m n p → fst C n + subComplex→map' m n (lt x) = idfun _ + subComplex→map' m n (eq x) = idfun _ + subComplex→map' m n (gt x) = CW↑Gen C m n (n ≟ᵗ suc m) x + + subComplex←map' : (m n : ℕ) (ineq : n <ᵗ suc m) (p : Trichotomyᵗ n m) + → fst C n → SubComplexGen.subComplexFam C m n p + subComplex←map' m n ineq (lt x) = idfun _ + subComplex←map' m n ineq (eq x) = idfun _ + subComplex←map' m n ineq (gt x) = ⊥.rec (¬squeeze {m} {n} (x , ineq)) + + private + retr-sect : (m n : ℕ) (ineq : n <ᵗ suc m) (p : Trichotomyᵗ n m) + → retract (subComplex→map' m n p) (subComplex←map' m n ineq p) + × section (subComplex→map' m n p) (subComplex←map' m n ineq p) + retr-sect m n ineq (lt x) = (λ x → refl) , λ x → refl + retr-sect m n ineq (eq x) = (λ x → refl) , (λ x → refl) + retr-sect m n ineq (gt x) = ⊥.rec (¬squeeze {m} {n} (x , ineq)) + + subComplexIso : (m n : ℕ) (ineq : n <ᵗ suc m) (p : Trichotomyᵗ n m) + → Iso (fst C n) (SubComplexGen.subComplexFam C m n p) + Iso.fun (subComplexIso m n ineq p) = subComplex←map' m n ineq p + Iso.inv (subComplexIso m n ineq p) = subComplex→map' m n p + Iso.rightInv (subComplexIso m n ineq p) = retr-sect m n ineq p .fst + Iso.leftInv (subComplexIso m n ineq p) = retr-sect m n ineq p .snd + +subComplex→map : ∀ {ℓ} (C : CWskel ℓ) (m n : ℕ) + → SubComplexGen.subComplexFam C m n (n ≟ᵗ m) → fst C n +subComplex→map C m n = subComplexMapGen.subComplex→map' C m n (n ≟ᵗ m) + + +subComplex→Equiv : ∀ {ℓ} (C : CWskel ℓ) (n m : ℕ) (p : m <ᵗ suc n) + → SubComplexGen.subComplexFam C n m (m ≟ᵗ n) ≃ fst C m +subComplex→Equiv C n m p = + isoToEquiv (invIso (subComplexMapGen.subComplexIso C n m p (m ≟ᵗ n))) + +complex≃subcomplex' : (C : CWskel ℓ) (n m : ℕ) (ineq : m <ᵗ suc n) (p : _) + → fst C m ≃ SubComplexGen.subComplexFam C n m p +complex≃subcomplex' C n m ineq p = + isoToEquiv (subComplexMapGen.subComplexIso C n m ineq p) + +module _ (C : CWskel ℓ) where + module _ (n : ℕ) where + module G = SubComplexGen C n + subComplexFam : ℕ → Type ℓ + subComplexFam m = G.subComplexFam m (m ≟ᵗ n) + + subComplexCard : (m : ℕ) → ℕ + subComplexCard m = G.subComplexCard m (m ≟ᵗ n) + + -- Attaching maps + subComplexα : (m : ℕ) + → Fin (subComplexCard m) × S⁻ m → subComplexFam m + subComplexα m = G.subComplexα m (m ≟ᵗ n) + + subComplex₀-empty : ¬ subComplexFam zero + subComplex₀-empty = G.subComplex₀-empty (0 ≟ᵗ n) + + -- Resulting complex has CW structure + subComplexFam≃Pushout : (m : ℕ) + → subComplexFam (suc m) ≃ Pushout (subComplexα m) fst + subComplexFam≃Pushout m = + G.subComplexFam≃Pushout m (m ≟ᵗ n) (suc m ≟ᵗ n) -- packaging it all together subComplexYieldsCWskel : (n : ℕ) → yieldsCWskel (subComplexFam n) - fst (subComplexYieldsCWskel n) = subComplexCard n - fst (snd (subComplexYieldsCWskel n)) = subComplexα n + fst (subComplexYieldsCWskel n) m = subComplexCard n m + fst (snd (subComplexYieldsCWskel n)) m = subComplexα n m fst (snd (snd (subComplexYieldsCWskel n))) = subComplex₀-empty n snd (snd (snd (subComplexYieldsCWskel n))) m = subComplexFam≃Pushout n m -- main definition subComplex : (n : ℕ) → CWskel ℓ - fst (subComplex n) = subComplexFam n + fst (subComplex n) m = subComplexFam n m snd (subComplex n) = subComplexYieldsCWskel n - -- Let us also show that this defines a _finite_ CW complex subComplexFin : (m : ℕ) (n : Fin (suc m)) - → isEquiv (CW↪ (subComplexFam (fst n) , subComplexYieldsCWskel (fst n)) m) - subComplexFin m (n , r) with (m ≟ᵗ n) | (suc m ≟ᵗ n) - ... | lt x | p = ⊥.rec (¬squeeze (x , r)) - ... | eq x | lt x₁ = - ⊥.rec (¬m<ᵗm (subst (_<ᵗ n) x (<ᵗ-trans <ᵗsucm x₁))) - ... | eq x | eq x₁ = ⊥.rec (¬m<ᵗm (subst (m <ᵗ_) (x₁ ∙ sym x) <ᵗsucm)) - ... | eq x | gt x₁ = - subst isEquiv (funExt λ x → cong (help .fst) - (retEq (isoToEquiv (PushoutEmptyFam {A = Fin 0 × fst C m} - (λ x₃ → ¬Fin0 (fst x₃)) ¬Fin0 {f = snd} {g = fst})) x)) - (help .snd) - where - help : fst C m ≃ fst C n - help = invEquiv (pathToEquiv (λ i → fst C (x (~ i)))) - ... | gt x | lt x₁ = - ⊥.rec (¬squeeze (x₁ , <ᵗ-trans x (<ᵗ-trans <ᵗsucm <ᵗsucm))) - ... | gt x | eq x₁ = - ⊥.rec (¬m<ᵗm (subst (n <ᵗ_) x₁ r)) - ... | gt x | gt x₁ = - subst isEquiv (funExt (retEq (isoToEquiv (PushoutEmptyFam {A = Fin 0 × fst C n} - (λ x₃ → ¬Fin0 (fst x₃)) ¬Fin0 {f = snd} {g = fst})))) - (idEquiv _ .snd) + → isEquiv (CW↪ (subComplex (fst n)) m) + subComplexFin m n = + compEquiv + (_ , SubComplexGen.subComplexFinEquiv' C (fst n) m (snd n) (m ≟ᵗ fst n)) + (invEquiv (subComplexFam≃Pushout (fst n) m)) .snd subComplexYieldsFinCWskel : (n : ℕ) → yieldsFinCWskel n (subComplexFam n) fst (subComplexYieldsFinCWskel n) = subComplexYieldsCWskel n - snd (subComplexYieldsFinCWskel n) k = subComplexFin (k + n) (n , <ᵗ-+) + snd (subComplexYieldsFinCWskel n) k = subComplexFin (k + n) (n , <ᵗ-+ {n} {k}) finSubComplex : (n : ℕ) → finCWskel ℓ n - fst (finSubComplex n) = subComplexFam n + fst (finSubComplex n) m = subComplexFam n m snd (finSubComplex n) = subComplexYieldsFinCWskel n - complex≃subcomplex : (n : ℕ) (m : Fin (suc n)) - → fst C (fst m) ≃ subComplex n .fst (fst m) - complex≃subcomplex n (m , p) with (m ≟ᵗ n) - ... | lt x = idEquiv _ - ... | eq x = idEquiv _ - ... | gt x = ⊥.rec (¬squeeze (x , p)) - -- C⁽ⁿ⁾ₙ ≃ Cₙ realiseSubComplex : (n : ℕ) (C : CWskel ℓ) → Iso (fst C n) (realise (subComplex C n)) realiseSubComplex n C = - compIso (equivToIso (complex≃subcomplex C n flast)) + compIso (equivToIso (complex≃subcomplex' C n n (<ᵗsucm {n}) (n ≟ᵗ n))) (realiseFin n (finSubComplex C n)) +subCWExplicit : ∀ {ℓ} (n : ℕ) → CWexplicit ℓ → CWexplicit ℓ +fst (subCWExplicit n (X , Xsk , e)) = Xsk .fst n +fst (snd (subCWExplicit n (X , Xsk , e))) = subComplex Xsk n +snd (snd (subCWExplicit n (X , Xsk , e))) = + isoToEquiv (realiseSubComplex n Xsk) + +subCW : ∀ {ℓ} (n : ℕ) → CWexplicit ℓ → CW ℓ +subCW n X = CWexplicit→CW (subCWExplicit n X) + -- Strictifying finCWskel niceFinCWskel : ∀ {ℓ} (n : ℕ) → finCWskel ℓ n → finCWskel ℓ n fst (niceFinCWskel n (A , AC , fin)) = finSubComplex (A , AC) n .fst snd (niceFinCWskel n (A , AC , fin)) = finSubComplex (A , AC) n .snd -makeNiceFinCWskel : ∀ {ℓ} {A : Type ℓ} → isFinCW A → isFinCW A +makeNiceFinCWskel : ∀ {ℓ} {A : Type ℓ} → hasFinCWskel A → hasFinCWskel A makeNiceFinCWskel {A = A} (dim , cwsk , e) = better where improved = finSubComplex (cwsk .fst , cwsk .snd .fst) dim - better : isFinCW A + better : hasFinCWskel A fst better = dim fst (snd better) = improved snd (snd better) = @@ -190,62 +270,256 @@ makeNiceFinCWElim : ∀ {ℓ ℓ'} {A : finCW ℓ → Type ℓ'} → (C : _) → A C makeNiceFinCWElim {A = A} ind C = subst A (makeNiceFinCW≡ C) (ind C) -makeNiceFinCWElim' : ∀ {ℓ ℓ'} {C : Type ℓ} {A : ∥ isFinCW C ∥₁ → Type ℓ'} +makeNiceFinCWElim' : ∀ {ℓ ℓ'} {C : Type ℓ} {A : isFinCW C → Type ℓ'} → ((x : _) → isProp (A x)) - → ((cw : isFinCW C) → A (makeNiceFinCW (C , ∣ cw ∣₁) .snd)) + → ((cw : hasFinCWskel C) → A (makeNiceFinCW (C , ∣ cw ∣₁) .snd)) → (cw : _) → A cw makeNiceFinCWElim' {A = A} pr ind = PT.elim pr λ cw → subst A (squash₁ _ _) (ind cw) --- Goal: Show that a cell complex C and its subcomplex C⁽ⁿ⁾ share --- homology in low enough dimensions -module _ (C : CWskel ℓ) where - ChainSubComplex : (n : ℕ) → snd C .fst n ≡ snd (subComplex C (suc n)) .fst n - ChainSubComplex n with (n ≟ᵗ suc n) - ... | lt x = refl - ... | eq x = ⊥.rec (¬m<ᵗm (subst (n <ᵗ_) (sym x) <ᵗsucm)) - ... | gt x = ⊥.rec (¬-suc-n<ᵗn x) - - ≤ChainSubComplex : (n : ℕ) (m : Fin n) - → snd C .fst (fst m) ≡ snd (subComplex C n) .fst (fst m) - ≤ChainSubComplex n (m , p) with (m ≟ᵗ n) - ... | lt x = refl - ... | eq x = ⊥.rec (¬m<ᵗm (subst (_<ᵗ n) x p)) - ... | gt x = ⊥.rec (¬m<ᵗm (<ᵗ-trans x p)) - -subChainIso : (C : CWskel ℓ) (n : ℕ) (m : Fin n) - → AbGroupIso (CW-AugChainComplex C .ChainComplex.chain (fst m)) - (CW-AugChainComplex (subComplex C n) .ChainComplex.chain (fst m)) -subChainIso C n (zero , p) = idGroupIso -subChainIso C n (suc m , p) with (m ≟ᵗ n) -... | lt x = idGroupIso -... | eq x = ⊥.rec (¬-suc-n<ᵗn ((subst (suc m <ᵗ_) (sym x) p))) -... | gt x = ⊥.rec (¬-suc-n<ᵗn ((<ᵗ-trans p x))) - --- main result -subComplexHomology : (C : CWskel ℓ) (n m : ℕ) (p : suc (suc m) <ᵗ n) - → GroupIso (H̃ˢᵏᵉˡ C m) (H̃ˢᵏᵉˡ (subComplex C n) m) -subComplexHomology C n m p = - homologyIso m _ _ - (subChainIso C n (suc (suc m) , p)) - (subChainIso C n (suc m , <ᵗ-trans <ᵗsucm p)) - (subChainIso C n (m , <ᵗ-trans <ᵗsucm (<ᵗ-trans <ᵗsucm p))) - (lem₁ m) - (lem₁ (suc m)) +-- Pointed subcomplexes +CWskel∙Gen : ∀ {ℓ} (X : CWskel ℓ) → fst X 1 → (n m : ℕ) (p : _) + → G.subComplexFam X (suc m) (suc n) p +CWskel∙Gen X x₀ n m (lt x) = CWskel∙ X x₀ n +CWskel∙Gen X x₀ n m (eq x) = CWskel∙ X x₀ n +CWskel∙Gen X x₀ n m (gt x) = CWskel∙ X x₀ m + +CWskel∙Gen≡CWskel∙ : ∀ {ℓ} (X : CWskel ℓ) (x : fst X 1) → (n m : ℕ) + → CWskel∙Gen X x n (suc m) (suc n ≟ᵗ suc (suc m)) + ≡ CWskel∙ (subComplex X (suc (suc m))) x n +CWskel∙Gen≡CWskel∙ X x zero m = refl +CWskel∙Gen≡CWskel∙ X x (suc n) m = + lem (suc (suc n) ≟ᵗ suc (suc m)) (suc n ≟ᵗ suc (suc m)) + ∙ cong (CW↪ (subComplex X (suc (suc m))) (suc n)) + (CWskel∙Gen≡CWskel∙ X x n m) + where + lem : (p : _) (q : _) → CWskel∙Gen X x (suc n) (suc m) p + ≡ invEq (G.subComplexFam≃Pushout X (suc (suc m)) (suc n) q p) + (inl (CWskel∙Gen X x n (suc m) q)) + lem (lt x) (lt y) = refl + lem (lt x) (eq y) = + ⊥.rec (¬m<ᵗm {suc (suc m)} + (subst (_<ᵗ suc (suc m)) y + (<ᵗ-trans {n} {suc n} {suc m} (<ᵗsucm {n}) x))) + lem (lt x) (gt y) = + ⊥.rec (¬squeeze {n} {m} (x + , <ᵗ-trans {suc m} {n} {suc (suc n)} y + (<ᵗ-trans {n} {suc n} {suc (suc n)} + (<ᵗsucm {n}) (<ᵗsucm {suc n})))) + lem (eq x) (lt y) = refl + lem (eq x) (eq y) = + ⊥.rec (¬m<ᵗm {suc (suc n)} + (subst (_<ᵗ suc (suc n)) (y ∙ sym x) (<ᵗsucm {n}))) + lem (eq x) (gt y) = + ⊥.rec (¬-suc-n<ᵗn {suc (suc n)} (subst (_<ᵗ suc n) (sym x) y)) + lem (gt x) (lt y) = + ⊥.rec (¬squeeze {n} {suc m} (y , x)) + lem (gt y) (eq z) = + ((λ j → transp (λ i → fst X (suc (predℕ (z (~ j ∨ i))))) (~ j) + (CWskel∙ X x (predℕ (z (~ j)))))) + ∙ cong (λ p → subst (fst X) p (CWskel∙ X x n)) (isSetℕ _ _ _ z) + ∙ sym (transportRefl _) + lem (gt x) (gt y) = refl + +CWskel∙GenSubComplex : ∀ {ℓ} (X : CWskel ℓ) (x₀ : fst X 1) {n m : ℕ} + (t : m <ᵗ suc (suc n)) (p : _) + → CWskel∙Gen X x₀ m (suc n) p + ≡ subComplexMapGen.subComplex←map' X (suc (suc n)) (suc m) t p + (CWskel∙ X x₀ m) +CWskel∙GenSubComplex X x₌ t (lt x) = refl +CWskel∙GenSubComplex X x₌ t (eq x) = refl +CWskel∙GenSubComplex X x₌ {n} {m} t (gt x) = + ⊥.rec (¬squeeze {suc n} {m} (x , t)) + +CWskel∙SubComplex : ∀ {ℓ} (X : CWskel ℓ) (x₀ : fst X 1) {n m : ℕ} + (t : m <ᵗ suc (suc n)) + → CWskel∙ (subComplex X (suc (suc n))) x₀ m + ≡ fst (complex≃subcomplex' X (suc (suc n)) (suc m) t + (suc m ≟ᵗ suc (suc n))) (CWskel∙ X x₀ m) +CWskel∙SubComplex X x₀ {n} {m} t = + sym (CWskel∙Gen≡CWskel∙ X x₀ m n) + ∙ CWskel∙GenSubComplex X x₀ t (suc m ≟ᵗ suc (suc n)) + +-- Homology of subcomplex +CW↪CommSubComplex : ∀ {ℓ} (C : CWskel ℓ) (m k : ℕ) → + CW↪ C m ∘ subComplex→map C k m + ≡ subComplex→map C k (suc m) ∘ CW↪ (subComplex C k) m +CW↪CommSubComplex C m k with (m ≟ᵗ k) | (suc m ≟ᵗ k) +... | lt x | lt y = refl +... | lt x | eq y = refl +... | lt x | gt y = ⊥.rec (¬squeeze {m} {k} (x , y)) +... | eq x | lt y = + ⊥.rec (¬m<ᵗm {k} (subst (_<ᵗ k) x (<ᵗ-trans {m} {suc m} {k} (<ᵗsucm {m}) y))) +... | eq x | eq y = + ⊥.rec (¬m<ᵗm {m} (subst (_<ᵗ_ m) (y ∙ (λ i → x (~ i))) (<ᵗsucm {m}))) +... | eq x | gt y = funExt λ s → help _ _ x y s (suc m ≟ᵗ suc k) + where + help : (m : ℕ) (k : ℕ) (x : m ≡ k) (y : k <ᵗ suc m) (s : fst C m) (p : _) + → CW↪ C m s ≡ CW↑Gen C k (suc m) p y (transport refl (subst (fst C) x s)) + help zero = λ k x y s → ⊥.rec (CW₀-empty C s) + help (suc m) = J> λ y s + → λ { (lt x) → ⊥.rec (¬squeeze {m} {suc m} (y , x)) + ; (eq x) → cong (CW↪ C (suc m)) (sym (transportRefl s) + ∙ λ i → subst (fst C) (isSetℕ _ _ refl (cong predℕ (sym x)) i) + (transportRefl (transportRefl s (~ i)) (~ i))) + ; (gt x) → ⊥.rec (¬m<ᵗm {m} x) } +... | gt x | lt y = + ⊥.rec (¬squeeze {suc m} {k} (y + , <ᵗ-trans {k} {m} {suc (suc m)} x + (<ᵗ-trans {m} {suc m} {suc (suc m)} (<ᵗsucm {m}) (<ᵗsucm {suc m})))) +... | gt x | eq y = + ⊥.rec (¬m<ᵗm {k} (subst (k <ᵗ_) y (<ᵗ-trans {k} {m} {suc m} x (<ᵗsucm {m})))) +... | gt x | gt y = funExt λ a → help _ _ x y (suc m ≟ᵗ suc k) a + where + help : (m k : ℕ) (x : k <ᵗ m) (y : k <ᵗ suc m) (p : _) (a : _) + → CW↪ C m (CW↑Gen C k m (m ≟ᵗ suc k) x a) ≡ CW↑Gen C k (suc m) p y a + help (suc m) zero x y p a = ⊥.rec (C .snd .snd .snd .fst a) + help (suc m) (suc k) x y (lt z) a = ⊥.rec (¬squeeze {k} {suc m} (y , z)) + help (suc m) (suc k) x y (eq z) a = + ⊥.rec (¬m<ᵗm {m} (subst (_<ᵗ m) (sym (cong (predℕ ∘ predℕ) z)) x)) + help (suc m) (suc k) x y (gt z) a = + cong (CW↪ C (suc m)) + λ i → CW↑Gen C (suc k) (suc m) + (Trichotomyᵗ-suc (m ≟ᵗ suc k)) (isProp<ᵗ {k} {m} x z i) a + +CW↪SubComplexCharac : ∀ {ℓ} (C : CWskel ℓ) (m k : ℕ) (q : m <ᵗ k) → + CW↪ (subComplex C k) m + ≡ invEq (subComplex→Equiv C k (suc m) q) + ∘ CW↪ C m ∘ subComplex→map C k m +CW↪SubComplexCharac C m k q = funExt λ x + → sym (retEq (subComplex→Equiv C k (suc m) q) _) + ∙ cong (invEq (subComplex→Equiv C k (suc m) q)) + (funExt⁻ (sym (CW↪CommSubComplex C m k)) x) + +CW↑GenComm : ∀ {ℓ} (C : CWskel ℓ) (m n k : ℕ) + (p : Trichotomyᵗ n (suc m)) (t : m <ᵗ n) + → CW↑Gen C m n p t ∘ subComplex→map C k m + ≡ subComplex→map C k n ∘ CW↑Gen (subComplex C k) m n p t +CW↑GenComm C zero (suc n) k p t = + funExt λ x → (⊥.rec (G.subComplex₀-empty C k (0 ≟ᵗ k) x)) +CW↑GenComm C (suc m) (suc n) k p t = + funExt (λ qa → (help n m k p _ t refl qa + ∙ cong ((subComplex→map C k (suc n) ∘ + CW↑Gen (subComplex C k) (suc m) (suc n) p t)) (transportRefl qa))) where - lem₁ : (m : ℕ) {q : _} {r : _} - → Iso.fun (subChainIso C n (m , q) .fst) ∘ CW-AugChainComplex C .ChainComplex.bdry m .fst - ≡ CW-AugChainComplex (subComplex C n) .ChainComplex.bdry m .fst - ∘ Iso.fun (subChainIso C n (suc m , r) .fst) - lem₁ zero with (0 ≟ᵗ n) - ... | lt x = refl - ... | eq x = ⊥.rec (subst (suc (suc m) <ᵗ_) (sym x) p) - ... | gt x = ⊥.rec x - lem₁ (suc m) {q} {r} with (m ≟ᵗ n) | (suc m ≟ᵗ n) | (suc (suc m) ≟ᵗ n) - ... | lt x | lt x₁ | lt x₂ = refl - ... | lt x | lt x₁ | eq x₂ = refl - ... | lt x | lt x₁ | gt x₂ = ⊥.rec (¬squeeze (x₁ , x₂)) - ... | lt x | eq x₁ | t = ⊥.rec ((¬m<ᵗm (subst (_<ᵗ n) x₁ q))) - ... | lt x | gt x₁ | t = ⊥.rec ((¬m<ᵗm (<ᵗ-trans x₁ q))) - ... | eq x | s | t = ⊥.rec (¬-suc-n<ᵗn (subst (suc m <ᵗ_) (sym x) q)) - ... | gt x | s | t = ⊥.rec ((¬-suc-n<ᵗn (<ᵗ-trans q x))) + help : (n m k : ℕ) (p : _) (s : _) (t : _) (pp : s ≡ (suc m ≟ᵗ k)) + → (x : G.subComplexFam C k (suc m) s) → + CW↑Gen C (suc m) (suc n) p t + (subComplexMapGen.subComplex→map' C k (suc m) s x) + ≡ + subComplexMapGen.subComplex→map' C k (suc n) (suc n ≟ᵗ k) + (CW↑Gen (subComplex C k) (suc m) (suc n) p t + (subst (G.subComplexFam C k (suc m)) pp x)) + help (suc n) m k (lt y) s t pp x = ⊥.rec (¬squeeze {m} {suc n} (t , y)) + help (suc n) m k (eq y) s t pp x = cong (CW↪ C (suc n)) + (cong (λ p → subst (fst C) p + (subComplexMapGen.subComplex→map' C k (suc m) s x)) (isSetℕ _ _ _ _) + ∙ lem m (cong predℕ (cong predℕ y)) s (sym pp) x + ∙ cong (subComplex→map C k (suc n)) + (cong (λ p → subst (subComplexFam C k) p + (subst (G.subComplexFam C k (suc m)) pp x)) + (isSetℕ _ _ (cong suc (sym (cong (predℕ ∘ predℕ) y))) + (cong predℕ (sym y))))) + ∙ funExt⁻ (CW↪CommSubComplex C (suc n) k) + ((subst (λ m₁ → subComplexFam C k m₁) (cong predℕ (sym y)) + (subst (G.subComplexFam C k (suc m)) pp x))) + where + lem : (m : ℕ) (y : n ≡ m) (s : _) (t : (suc m ≟ᵗ k) ≡ s) (x : _) + → subst (fst C) (cong suc (sym y)) + (subComplexMapGen.subComplex→map' C k (suc m) s x) + ≡ subComplex→map C k (suc n) + (subst (subComplexFam C k) (cong suc (sym y)) + (subst (G.subComplexFam C k (suc m)) (sym t) x)) + lem = J> (J> λ x → + transportRefl _ + ∙ sym (cong (subComplex→map C k (suc n)) + (transportRefl _ ∙ transportRefl x))) + help (suc n) m k (gt y) s t pp x = + cong (CW↪ C (suc n)) (help n m k (suc n ≟ᵗ suc (suc m)) s y pp x) + ∙ funExt⁻ (CW↪CommSubComplex C (suc n) k) + (CW↑Gen (subComplex C k) (suc m) (suc n) (suc n ≟ᵗ suc (suc m)) y + (subst (G.subComplexFam C k (suc m)) pp x)) + +subComplex→comm : ∀ {ℓ} (C : CWskel ℓ) (m n : ℕ) + (p : _) (q : _) (x : G.subComplexFam C n m p) + → CW↪ C m (subComplexMapGen.subComplex→map' C n m p x) + ≡ subComplexMapGen.subComplex→map' C n (suc m) q + (SubComplexGen.subComplexCW↪Gen C n m p q x) +subComplex→comm C m zero (eq y) s x = ⊥.rec (CW₀-empty C (subst (fst C) y x)) +subComplex→comm C m zero (gt y) q x = ⊥.rec (CW₀-empty C x) +subComplex→comm C m (suc n) (lt y) (lt z) x = refl +subComplex→comm C m (suc n) (lt y) (eq z) x = refl +subComplex→comm C m (suc n) (lt y) (gt z) x = + ⊥.rec (¬squeeze {m} {suc n} (y , z)) +subComplex→comm C m (suc n) (eq y) (lt z) x = + ⊥.rec (¬m<ᵗm {suc n} (transport (λ i → y i <ᵗ suc n) + (<ᵗ-trans {m} {suc m} {suc n} (<ᵗsucm {m}) z))) +subComplex→comm C m (suc n) (eq y) (eq z) x = + ⊥.rec ( falseDichotomies.eq-eq (sym y , sym z)) +subComplex→comm C m (suc n) (eq y) (gt z) x with (m ≟ᵗ suc n) +... | lt x₃ = ⊥.rec (¬squeeze {n} {m} (z , x₃)) +... | eq x₃ = cong (CW↪ C m) (sym (cong (subst (fst C) (sym x₃)) + (transportRefl _ + ∙ cong (λ p → subst (fst C) p x) (isSetℕ _ _ y x₃)) + ∙ subst⁻Subst (fst C) x₃ x)) +... | gt x₃ = ⊥.rec (¬m<ᵗm {suc n} (subst (suc n <ᵗ_) y x₃)) +subComplex→comm C m (suc n) (gt y) (lt z) x = + ⊥.rec (¬squeeze {m} {n} + (z , <ᵗ-trans {suc n} {m} {suc (suc m)} y + (<ᵗ-trans {m} {suc m} {suc (suc m)} (<ᵗsucm {m}) (<ᵗsucm {suc m})))) +subComplex→comm C m (suc n) (gt y) (eq z) x = (⊥.rec + (¬m<ᵗm {suc n} (transport (λ i → suc n <ᵗ z i) + (<ᵗ-trans {suc n} {m} {suc m} y ( <ᵗsucm {m}))))) +subComplex→comm C (suc m) (suc n) (gt y) (gt z) x with m ≟ᵗ n +... | lt x₃ = ⊥.rec (¬squeeze {n} {suc m} (z , x₃)) +... | eq x₃ = ⊥.rec (¬m<ᵗm {n} (subst (n <ᵗ_) x₃ y)) +... | gt x₃ = cong (CW↪ C (suc m)) + λ j → CW↑Gen C (suc n) (suc m) + (Trichotomyᵗ-suc (m ≟ᵗ suc n)) (isProp<ᵗ {n} {m} y x₃ j) x + +subComplex→Full : ∀ {ℓ} (C : CWskel ℓ) (m : ℕ) → cellMap (subComplex C m) C +SequenceMap.map (subComplex→Full C n) = subComplex→map C n +SequenceMap.comm (subComplex→Full C n) m = + subComplex→comm C m n (m ≟ᵗ n) (suc m ≟ᵗ n) + +subComplex→ : ∀ {ℓ} (C : CWskel ℓ) (m n : ℕ) + → finCellMap n (subComplex C m) C +FinSequenceMap.fmap (subComplex→ C m n) = subComplex→map C m ∘ fst +FinSequenceMap.fcomm (subComplex→ C m n) t = + subComplex→comm C (fst t) m (fst t ≟ᵗ m) (suc (fst t) ≟ᵗ m) + +subComplexFam↓ : ∀ {ℓ} (C : CWskel ℓ) (m : ℕ) (p : _) (q : _) + → G.subComplexFam C m (suc m) p → G.subComplexFam C m m q +subComplexFam↓ C m (lt x) q = ⊥.rec (¬-suc-n<ᵗn {m} x) +subComplexFam↓ C m (eq x) q = ⊥.rec (falseDichotomies.eq-eq(refl , sym x)) +subComplexFam↓ C m (gt x) (lt y) = ⊥.rec (¬m<ᵗm {m} y) +subComplexFam↓ C m (gt x) (eq y) = idfun _ +subComplexFam↓ C m (gt x) (gt y) = ⊥.rec (¬m<ᵗm {m} y) + +CW↪subComplexFam↓ : ∀ {ℓ} (C : CWskel ℓ) (m : ℕ) (p : _) (q : _) (x : _) + → SubComplexGen.subComplexCW↪Gen C m m p q (subComplexFam↓ C m q p x) ≡ x +CW↪subComplexFam↓ C m p (lt y) x = ⊥.rec (¬-suc-n<ᵗn {m} y) +CW↪subComplexFam↓ C m p (eq y) x = ⊥.rec (falseDichotomies.eq-eq(refl , sym y)) +CW↪subComplexFam↓ C m (lt z) (gt y) x = ⊥.rec (¬m<ᵗm {m} z) +CW↪subComplexFam↓ C m (eq z) (gt y) x = + transportRefl _ ∙ cong (λ p → subst (fst C) p x) (isSetℕ _ _ z refl) + ∙ transportRefl x +CW↪subComplexFam↓ C m (gt z) (gt y) x = ⊥.rec (¬m<ᵗm {m} z) + +subComplex→map'Charac : ∀ {ℓ} (C : CWskel ℓ) (m : ℕ) (p : _) (q : _) + → subComplexMapGen.subComplex→map' C m (suc m) p + ≡ CW↪ C m ∘ subComplexMapGen.subComplex→map' C m m q + ∘ subComplexFam↓ C m p q +subComplex→map'Charac C m p (lt x) = ⊥.rec (¬m<ᵗm {m} x) +subComplex→map'Charac C m (lt y) (eq x) = ⊥.rec (¬-suc-n<ᵗn {m} y) +subComplex→map'Charac C m (eq y) (eq x) = + ⊥.rec (falseDichotomies.eq-eq (refl , sym y)) +subComplex→map'Charac C zero (gt y) (eq x) = + funExt (λ q → ⊥.rec (C .snd .snd .snd .fst q)) +subComplex→map'Charac C (suc m) (gt y) (eq x) with (m ≟ᵗ m) +... | lt z = ⊥.rec (¬m<ᵗm {m} z) +... | eq z = funExt λ x + → cong (CW↪ C (suc m)) (cong (λ p → subst (fst C) p x) (isSetℕ _ _ _ _) + ∙ transportRefl x) +... | gt z = ⊥.rec (¬m<ᵗm {m} z) +subComplex→map'Charac C m p (gt x) = ⊥.rec (¬m<ᵗm {m} x) diff --git a/Cubical/Data/Fin/Inductive/Properties.agda b/Cubical/Data/Fin/Inductive/Properties.agda index ab5b9a5337..badbaf1ba1 100644 --- a/Cubical/Data/Fin/Inductive/Properties.agda +++ b/Cubical/Data/Fin/Inductive/Properties.agda @@ -217,6 +217,18 @@ Iso-Fin×Fin-Fin· {n = suc n} {m = m} = (⊎Iso idIso Iso-Fin×Fin-Fin·)) (Iso-Fin⊎Fin-Fin+ {n = m} {n · m}) +Iso-FinSuc→-Fin→× : ∀ {ℓ} (n : ℕ) {A : Fin (suc n) → Type ℓ} + → Iso ((x : Fin (suc n)) → A x) + (((x : _) → A (fsuc x)) × A fzero) +fst (Iso.fun (Iso-FinSuc→-Fin→× n) f) x = f (fsuc x) +snd (Iso.fun (Iso-FinSuc→-Fin→× n) f) = f fzero +Iso.inv (Iso-FinSuc→-Fin→× n) (f , s) (zero , w) = s +Iso.inv (Iso-FinSuc→-Fin→× (suc n)) (f , s) (suc t , w) = f (t , w) +fst (Iso.rightInv (Iso-FinSuc→-Fin→× (suc n)) (f , s) i) (w , t) = f (w , t) +snd (Iso.rightInv (Iso-FinSuc→-Fin→× n) (f , s) i) = s +Iso.leftInv (Iso-FinSuc→-Fin→× n) f i (zero , tt) = f fzero +Iso.leftInv (Iso-FinSuc→-Fin→× (suc n)) f i (suc s , t) = f (suc s , t) + Iso-Fin×Bool-Fin : {n : ℕ} → Iso (Fin n × Bool) (Fin (2 · n)) Iso-Fin×Bool-Fin = compIso (compIso Σ-swap-Iso diff --git a/Cubical/Data/Nat/Order/Inductive.agda b/Cubical/Data/Nat/Order/Inductive.agda index cd6db7afa1..1b5b808893 100644 --- a/Cubical/Data/Nat/Order/Inductive.agda +++ b/Cubical/Data/Nat/Order/Inductive.agda @@ -78,3 +78,34 @@ isProp<ᵗ {n = suc n} {suc m} = isProp<ᵗ {n = n} {m = m} <→<ᵗ {n = suc n} {m = zero} x = snotz (sym (+-suc (fst x) (suc n)) ∙ snd x) <→<ᵗ {n = suc n} {m = suc m} p = <→<ᵗ {n = n} {m = m} (pred-≤-pred p) + +module _ {n m : ℕ} where + isPropTrichotomyᵗ : isProp (Trichotomyᵗ n m) + isPropTrichotomyᵗ (lt x) (lt y) i = lt (isProp<ᵗ {n = n} {m} x y i) + isPropTrichotomyᵗ (lt x) (eq y) = ⊥.rec (¬m<ᵗm {m} (subst (_<ᵗ m) y x)) + isPropTrichotomyᵗ (lt x) (gt y) = ⊥.rec (¬m<ᵗm {m} (<ᵗ-trans {m} {n} {m} y x)) + isPropTrichotomyᵗ (eq x) (lt y) = + ⊥.rec (¬m<ᵗm {m} (subst (_<ᵗ m) x y)) + isPropTrichotomyᵗ (eq x) (eq y) i = eq (isSetℕ n m x y i) + isPropTrichotomyᵗ (eq x) (gt y) = ⊥.rec (¬m<ᵗm {n} (subst (_<ᵗ n) (sym x) y)) + isPropTrichotomyᵗ (gt x) (lt y) = ⊥.rec (¬m<ᵗm {n} (<ᵗ-trans {n} {m} {n} y x)) + isPropTrichotomyᵗ (gt x) (eq y) = ⊥.rec (¬m<ᵗm {n} (subst (_<ᵗ n) (sym y) x)) + isPropTrichotomyᵗ (gt x) (gt y) i = gt (isProp<ᵗ {n = m} {n} x y i) + +module falseDichotomies where + lt-eq : {n m : ℕ} → ¬ m <ᵗ n × (m ≡ suc n) + lt-eq {n = n} (p , q) = ¬-suc-n<ᵗn {n = n} (subst (_<ᵗ n) q p) + + lt-gt : {n m : ℕ} → ¬ m <ᵗ n × (suc n <ᵗ m) + lt-gt {n = n} {m} (p , q) = + ¬-suc-n<ᵗn {n = n} (<ᵗ-trans {n = suc n} {m} {n} q p) + + eq-eq : {n m : ℕ} → ¬ (m ≡ n) × (m ≡ suc n) + eq-eq {n = n} (p , q) = + ¬m<ᵗm {n} (subst (_<ᵗ suc n) (sym p ∙ q) (<ᵗsucm {n})) + + eq-gt : {n m : ℕ} → ¬ (m ≡ n) × (suc n <ᵗ m) + eq-gt (p , q) = lt-eq (q , cong suc (sym p)) + + gt-lt : {n m : ℕ} → ¬ (n <ᵗ m) × (m <ᵗ suc n) + gt-lt {n = n} {m = m} = ¬squeeze {n = n} {m = m} diff --git a/Cubical/Data/Nat/Properties.agda b/Cubical/Data/Nat/Properties.agda index 19a995fda8..2590a016bc 100644 --- a/Cubical/Data/Nat/Properties.agda +++ b/Cubical/Data/Nat/Properties.agda @@ -46,6 +46,10 @@ znots eq = subst (caseNat ℕ ⊥) eq 0 snotz : ¬ (suc n ≡ 0) snotz eq = subst (caseNat ⊥ ℕ) eq 0 +sucn≠n : {n : ℕ} → ¬ (suc n ≡ n) +sucn≠n {n = zero} = snotz +sucn≠n {n = suc n} p = sucn≠n {n = n} (cong predℕ p) + injSuc : suc m ≡ suc n → m ≡ n injSuc p = cong predℕ p diff --git a/Cubical/Data/Sum/Properties.agda b/Cubical/Data/Sum/Properties.agda index 35e0e132c7..8ce0928e77 100644 --- a/Cubical/Data/Sum/Properties.agda +++ b/Cubical/Data/Sum/Properties.agda @@ -356,3 +356,15 @@ Iso⊎→Iso {A = A} {C = C} {B = B} {D = D} f e p = Iso' Iso.inv Iso' = e⁻-pres-inr Iso.rightInv Iso' x = lem2 x (_ , refl) (_ , refl) Iso.leftInv Iso' x = lem1 x (_ , refl) (_ , refl) + +Lift⊎Iso : ∀ (ℓ : Level) + → Iso (Lift {j = ℓ} A ⊎ Lift {j = ℓ} B) + (Lift {j = ℓ} (A ⊎ B)) +fun (Lift⊎Iso ℓD) (inl x) = liftFun inl x +fun (Lift⊎Iso ℓD) (inr x) = liftFun inr x +inv (Lift⊎Iso ℓD) (lift (inl x)) = inl (lift x) +inv (Lift⊎Iso ℓD) (lift (inr x)) = inr (lift x) +rightInv (Lift⊎Iso ℓD) (lift (inl x)) = refl +rightInv (Lift⊎Iso ℓD) (lift (inr x)) = refl +leftInv (Lift⊎Iso ℓD) (inl x) = refl +leftInv (Lift⊎Iso ℓD) (inr x) = refl diff --git a/Cubical/Foundations/Pointed/Base.agda b/Cubical/Foundations/Pointed/Base.agda index 90b28d2830..91a7f020af 100644 --- a/Cubical/Foundations/Pointed/Base.agda +++ b/Cubical/Foundations/Pointed/Base.agda @@ -22,6 +22,10 @@ pt = str Pointed₀ = Pointed ℓ-zero +Lift∙ : ∀ {i j} → (A : Pointed i) → Pointed (ℓ-max i j) +fst (Lift∙ {j = j} A) = Lift {j = j} (typ A) +snd (Lift∙ A) = lift (pt A) + {- Pointed functions -} _→∙_ : (A : Pointed ℓ) (B : Pointed ℓ') → Type (ℓ-max ℓ ℓ') (A , a) →∙ (B , b) = Σ[ f ∈ (A → B) ] f a ≡ b diff --git a/Cubical/HITs/AbPath.agda b/Cubical/HITs/AbPath.agda new file mode 100644 index 0000000000..19efa8d403 --- /dev/null +++ b/Cubical/HITs/AbPath.agda @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.AbPath where + +open import Cubical.HITs.AbPath.Base public +open import Cubical.HITs.AbPath.Properties public diff --git a/Cubical/HITs/AbPath/Base.agda b/Cubical/HITs/AbPath/Base.agda new file mode 100644 index 0000000000..71ab681452 --- /dev/null +++ b/Cubical/HITs/AbPath/Base.agda @@ -0,0 +1,16 @@ +module Cubical.HITs.AbPath.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed + +-- 'Abelianised' path types (useful for encode-decode computations of π₁ᵃᵇ) +data _≡ᵃᵇ_ {ℓ} {A : Type ℓ} (x y : A) : Type ℓ + where + paths : x ≡ y → x ≡ᵃᵇ y + com : (p q r : x ≡ y) → paths (p ∙ sym q ∙ r) ≡ paths (r ∙ sym q ∙ p) + +Pathᵃᵇ : ∀ {ℓ} (A : Type ℓ) (x y : A) → Type ℓ +Pathᵃᵇ A = _≡ᵃᵇ_ + +Ωᵃᵇ : ∀ {ℓ} (A : Pointed ℓ) → Type ℓ +Ωᵃᵇ (A , a) = a ≡ᵃᵇ a diff --git a/Cubical/HITs/AbPath/Properties.agda b/Cubical/HITs/AbPath/Properties.agda new file mode 100644 index 0000000000..e02bbf80db --- /dev/null +++ b/Cubical/HITs/AbPath/Properties.agda @@ -0,0 +1,198 @@ +module Cubical.HITs.AbPath.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws as GLaws +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Pointed + +open import Cubical.HITs.AbPath.Base + +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.SetTruncation as ST + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.Abelianization.Base +open import Cubical.Algebra.Group.Abelianization.Properties as Abi +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Monoid.Base +open import Cubical.Algebra.Semigroup.Base +open import Cubical.Algebra.Group.Free + +open import Cubical.Homotopy.Group.Base + +open IsAbGroup +open IsGroup +open IsMonoid +open IsSemigroup +open AbGroupStr + +elimProp≡ᵃᵇ : ∀ {ℓ ℓ'} {A : Type ℓ} {x y : A} {B : x ≡ᵃᵇ y → Type ℓ'} + → ((s : _) → isProp (B s)) + → ((p : x ≡ y) → B (paths p)) + → (x : _) → B x +elimProp≡ᵃᵇ pr path* (paths x) = path* x +elimProp≡ᵃᵇ {B = B} pr path* (com p q r i) = help i + where + help : PathP (λ i → B (com p q r i)) + (path* (p ∙ sym q ∙ r)) + (path* (r ∙ sym q ∙ p)) + help = isProp→PathP (λ _ → pr _) _ _ + +private + pathsᵃᵇLemmaL : ∀ {ℓ} {A : Type ℓ} {x y : A} (z : _) + (p : x ≡ z) (q : x ≡ y)(r : x ≡ y) (s : x ≡ y) + → Path (z ≡ᵃᵇ y) (paths (sym p ∙ q ∙ sym r ∙ s)) + (paths (sym p ∙ s ∙ sym r ∙ q)) + pathsᵃᵇLemmaL = + J> λ q r s + → cong paths (sym (lUnit _)) ∙ com q r s ∙ cong paths (lUnit _) + + pathsᵃᵇLemmaR : ∀ {ℓ} {A : Type ℓ} {x y : A} (z : _) + (p : y ≡ z) (q : x ≡ y)(r : x ≡ y) (s : x ≡ y) + → Path (x ≡ᵃᵇ z) (paths ((q ∙ sym r ∙ s) ∙ p)) + (paths ((s ∙ sym r ∙ q) ∙ p)) + pathsᵃᵇLemmaR = + J> λ q r s + → cong paths (sym (rUnit _)) ∙ com q r s ∙ cong paths (rUnit _) + +-- 'left action of Ω (A , x) on x ≡ᵃᵇ y' +act≡ᵃᵇ : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ x) → x ≡ᵃᵇ y → x ≡ᵃᵇ y +act≡ᵃᵇ p (paths x₁) = paths (p ∙ x₁) +act≡ᵃᵇ {y = y} p (com q r s i) = pathsᵃᵇLemmaL _ (sym p) q r s i + +actL·πᵃᵇ : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) → y ≡ᵃᵇ z → x ≡ᵃᵇ z +actL·πᵃᵇ p (paths q) = paths (p ∙ q) +actL·πᵃᵇ p (com q r s i) = pathsᵃᵇLemmaL _ (sym p) q r s i + +·πᵃᵇ : ∀ {ℓ} {A : Type ℓ} {x y z : A} + → ∥ x ≡ᵃᵇ y ∥₂ → ∥ y ≡ᵃᵇ z ∥₂ → ∥ x ≡ᵃᵇ z ∥₂ +·πᵃᵇ {x = x} {y} {z} = ST.rec2 squash₂ preMult + where + preMult : x ≡ᵃᵇ y → y ≡ᵃᵇ z → ∥ x ≡ᵃᵇ z ∥₂ + preMult (paths p) q = ∣ actL·πᵃᵇ p q ∣₂ + preMult (com p q r i) s = lem s i + where + lem : (s : y ≡ᵃᵇ z) → ∣ actL·πᵃᵇ (p ∙ sym q ∙ r) s ∣₂ + ≡ ∣ actL·πᵃᵇ (r ∙ sym q ∙ p) s ∣₂ + lem = elimProp≡ᵃᵇ (λ _ → squash₂ _ _) + λ s i → ∣ pathsᵃᵇLemmaR _ s p q r i ∣₂ + +symᵃᵇ : ∀ {ℓ} {A : Type ℓ} {x y : A} → x ≡ᵃᵇ y → y ≡ᵃᵇ x +symᵃᵇ (paths x) = paths (sym x) +symᵃᵇ {A = A} {x} {y} (com p q r i) = lem i + where + lem : Path (y ≡ᵃᵇ x) (paths (sym (p ∙ sym q ∙ r))) + (paths (sym (r ∙ sym q ∙ p))) + lem = cong paths (symDistr _ _ ∙ cong₂ _∙_ (symDistr _ _) refl + ∙ sym (GLaws.assoc _ _ _)) + ∙∙ com _ _ _ + ∙∙ sym (cong paths (symDistr _ _ ∙ cong₂ _∙_ (symDistr _ _) refl + ∙ sym (GLaws.assoc _ _ _))) + +-πᵃᵇ : ∀ {ℓ} {A : Type ℓ} {x y : A} → ∥ x ≡ᵃᵇ y ∥₂ → ∥ y ≡ᵃᵇ x ∥₂ +-πᵃᵇ = ST.map symᵃᵇ + +reflᵃᵇ : ∀ {ℓ} {A : Type ℓ} {x : A} → x ≡ᵃᵇ x +reflᵃᵇ = paths refl + +-- Group laws +·πᵃᵇrUnit : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : ∥ x ≡ᵃᵇ y ∥₂) + → ·πᵃᵇ p ∣ reflᵃᵇ ∣₂ ≡ p +·πᵃᵇrUnit = ST.elim (λ _ → isSetPathImplicit) (elimProp≡ᵃᵇ (λ _ → squash₂ _ _) + λ p → cong ∣_∣₂ (cong paths (sym (rUnit _)))) + +·πᵃᵇlUnit : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : ∥ x ≡ᵃᵇ y ∥₂) + → ·πᵃᵇ ∣ reflᵃᵇ ∣₂ p ≡ p +·πᵃᵇlUnit = ST.elim (λ _ → isSetPathImplicit) (elimProp≡ᵃᵇ (λ _ → squash₂ _ _) + λ p → cong ∣_∣₂ (cong paths (sym (lUnit _)))) + +·πᵃᵇrCancel : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : ∥ x ≡ᵃᵇ y ∥₂) + → ·πᵃᵇ p (-πᵃᵇ p) ≡ ∣ reflᵃᵇ ∣₂ +·πᵃᵇrCancel = ST.elim (λ _ → isSetPathImplicit) (elimProp≡ᵃᵇ (λ _ → squash₂ _ _) + λ p → cong ∣_∣₂ (cong paths (rCancel _))) + +·πᵃᵇlCancel : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : ∥ x ≡ᵃᵇ y ∥₂) + → ·πᵃᵇ (-πᵃᵇ p) p ≡ ∣ reflᵃᵇ ∣₂ +·πᵃᵇlCancel = ST.elim (λ _ → isSetPathImplicit) (elimProp≡ᵃᵇ (λ _ → squash₂ _ _) + λ p → cong ∣_∣₂ (cong paths (lCancel _))) + +·πᵃᵇassoc : ∀ {ℓ} {A : Type ℓ} {x : A} (p q r : ∥ x ≡ᵃᵇ x ∥₂) + → ·πᵃᵇ p (·πᵃᵇ q r) ≡ ·πᵃᵇ (·πᵃᵇ p q) r +·πᵃᵇassoc = ST.elim3 (λ _ _ _ → isSetPathImplicit) + (elimProp≡ᵃᵇ (λ _ → isPropΠ2 λ _ _ → squash₂ _ _) + λ p → elimProp≡ᵃᵇ (λ _ → isPropΠ λ _ → squash₂ _ _) + λ q → elimProp≡ᵃᵇ (λ _ → squash₂ _ _) + λ r → cong ∣_∣₂ (cong paths (GLaws.assoc p q r))) + +·πᵃᵇcomm : ∀ {ℓ} {A : Type ℓ} {x : A} (p q : ∥ x ≡ᵃᵇ x ∥₂) → ·πᵃᵇ p q ≡ ·πᵃᵇ q p +·πᵃᵇcomm = ST.elim2 (λ _ _ → isSetPathImplicit) + (elimProp≡ᵃᵇ (λ _ → isPropΠ (λ _ → squash₂ _ _)) + λ p → elimProp≡ᵃᵇ (λ _ → squash₂ _ _) λ q + → cong ∣_∣₂ (cong paths (cong₂ _∙_ refl (lUnit q)) + ∙ com p refl q + ∙ cong paths (cong₂ _∙_ refl (sym (lUnit p))))) + +π₁ᵃᵇAbGroup : ∀ {ℓ} (A : Pointed ℓ) → AbGroup ℓ +fst (π₁ᵃᵇAbGroup (A , a)) = ∥ a ≡ᵃᵇ a ∥₂ +0g (snd (π₁ᵃᵇAbGroup A)) = ∣ reflᵃᵇ ∣₂ +_+_ (snd (π₁ᵃᵇAbGroup A)) = ·πᵃᵇ +- snd (π₁ᵃᵇAbGroup A) = -πᵃᵇ +is-set (isSemigroup (isMonoid (isGroup (isAbGroup (snd (π₁ᵃᵇAbGroup A)))))) = squash₂ +IsSemigroup.·Assoc (isSemigroup (isMonoid (isGroup (isAbGroup (snd (π₁ᵃᵇAbGroup A)))))) = + ·πᵃᵇassoc +IsMonoid.·IdR (isMonoid (isGroup (isAbGroup (snd (π₁ᵃᵇAbGroup A))))) = ·πᵃᵇrUnit +IsMonoid.·IdL (isMonoid (isGroup (isAbGroup (snd (π₁ᵃᵇAbGroup A))))) = ·πᵃᵇlUnit +·InvR (isGroup (isAbGroup (snd (π₁ᵃᵇAbGroup A)))) = ·πᵃᵇrCancel +·InvL (isGroup (isAbGroup (snd (π₁ᵃᵇAbGroup A)))) = ·πᵃᵇlCancel +IsAbGroup.+Comm (isAbGroup (snd (π₁ᵃᵇAbGroup A))) = ·πᵃᵇcomm + +-πᵃᵇinvDistr : ∀ {ℓ} {A : Pointed ℓ} (p q : ∥ Ωᵃᵇ A ∥₂) + → -πᵃᵇ {x = pt A} (·πᵃᵇ p q) ≡ ·πᵃᵇ (-πᵃᵇ p) (-πᵃᵇ q) +-πᵃᵇinvDistr {A = A} p q = + GroupTheory.invDistr (AbGroup→Group (π₁ᵃᵇAbGroup A)) p q + ∙ ·πᵃᵇcomm (-πᵃᵇ q) (-πᵃᵇ p) + +Ωᵃᵇ→Abelianizeπ₁ : ∀ {ℓ} {A : Pointed ℓ} → Ωᵃᵇ A → Abelianization (πGr 0 A) +Ωᵃᵇ→Abelianizeπ₁ (paths x) = η ∣ x ∣₂ +Ωᵃᵇ→Abelianizeπ₁ {A = A} (com p q r i) = help i + where + open AbelianizationGroupStructure (πGr 0 A) + help : Path (Abelianization (πGr 0 A)) + (η (·π 0 ∣ p ∣₂ (·π 0 ∣ sym q ∣₂ ∣ r ∣₂))) + (η (·π 0 ∣ r ∣₂ (·π 0 ∣ sym q ∣₂ ∣ p ∣₂))) + help = commAb (η ∣ p ∣₂) (η (·π 0 ∣ sym q ∣₂ ∣ r ∣₂)) + ∙ cong (_·Ab η ∣ p ∣₂) (commAb (η ∣ sym q ∣₂) (η ∣ r ∣₂)) + ∙ sym (assocAb (η ∣ r ∣₂) (η ∣ sym q ∣₂) (η ∣ p ∣₂)) + +π₁ᵃᵇ→Abelianizeπ₁ : ∀ {ℓ} {A : Pointed ℓ} + → ∥ Ωᵃᵇ A ∥₂ → Abelianization (πGr 0 A) +π₁ᵃᵇ→Abelianizeπ₁ = ST.rec isset Ωᵃᵇ→Abelianizeπ₁ + +π₁→π₁ᵃᵇHom : ∀ {ℓ} {A : Pointed ℓ} + → GroupHom (πGr 0 A) (AbGroup→Group (π₁ᵃᵇAbGroup A)) +fst π₁→π₁ᵃᵇHom = ST.map paths +snd π₁→π₁ᵃᵇHom = + makeIsGroupHom (ST.elim2 (λ _ _ → isSetPathImplicit) λ p q → refl) + +-- Below definition is verbose to speed up type checking +Abelianizeπ₁→π₁ᵃᵇ : ∀ {ℓ} {A : Pointed ℓ} + → AbGroupHom (AbelianizationAbGroup (πGr 0 A)) (π₁ᵃᵇAbGroup A) +fst (Abelianizeπ₁→π₁ᵃᵇ {A = A}) = + fst (fromAbelianization (π₁ᵃᵇAbGroup A) π₁→π₁ᵃᵇHom) +snd (Abelianizeπ₁→π₁ᵃᵇ {A = A}) = + makeIsGroupHom (elimProp2 _ (λ _ _ → squash₂ _ _ ) + (ST.elim2 (λ _ _ → isSetPathImplicit) + λ f g → refl)) + +Abelianizeπ₁≅π₁ᵃᵇ : ∀ {ℓ} (A : Pointed ℓ) + → AbGroupIso (AbelianizationAbGroup (πGr 0 A)) (π₁ᵃᵇAbGroup A) +Iso.fun (fst (Abelianizeπ₁≅π₁ᵃᵇ A)) = fst Abelianizeπ₁→π₁ᵃᵇ +Iso.inv (fst (Abelianizeπ₁≅π₁ᵃᵇ A)) = π₁ᵃᵇ→Abelianizeπ₁ +Iso.rightInv (fst (Abelianizeπ₁≅π₁ᵃᵇ A)) = + ST.elim (λ _ → isSetPathImplicit) (elimProp≡ᵃᵇ (λ _ → squash₂ _ _) λ p → refl) +Iso.leftInv (fst (Abelianizeπ₁≅π₁ᵃᵇ A)) = Abi.elimProp _ (λ _ → isset _ _) + (ST.elim (λ _ → isOfHLevelPath 2 isset _ _) λ p → refl) +snd (Abelianizeπ₁≅π₁ᵃᵇ A) = Abelianizeπ₁→π₁ᵃᵇ {A = A} .snd diff --git a/Cubical/HITs/Bouquet.agda b/Cubical/HITs/Bouquet.agda index b11d077e3e..e844aa730e 100644 --- a/Cubical/HITs/Bouquet.agda +++ b/Cubical/HITs/Bouquet.agda @@ -2,4 +2,5 @@ module Cubical.HITs.Bouquet where open import Cubical.HITs.Bouquet.Base public +open import Cubical.HITs.Bouquet.Properties public open import Cubical.HITs.Bouquet.FundamentalGroupProof public diff --git a/Cubical/HITs/Bouquet/FundamentalGroupProof.agda b/Cubical/HITs/Bouquet/FundamentalGroupProof.agda index 7995223ac4..be99b6721d 100644 --- a/Cubical/HITs/Bouquet/FundamentalGroupProof.agda +++ b/Cubical/HITs/Bouquet/FundamentalGroupProof.agda @@ -21,16 +21,27 @@ open import Cubical.Foundations.Univalence open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.GroupoidLaws renaming (assoc to pathAssoc) -open import Cubical.HITs.SetTruncation hiding (rec2) +open import Cubical.Foundations.HLevels +open import Cubical.Functions.Morphism + +open import Cubical.Data.Nat hiding (_·_ ; elim) +open import Cubical.Data.Fin.Inductive + +open import Cubical.HITs.Bouquet.Discrete +open import Cubical.HITs.FreeGroup as FG hiding (winding ; elimProp) +open import Cubical.HITs.FreeGroupoid.Base as FGrp +open import Cubical.HITs.SetTruncation as ST hiding (rec2) open import Cubical.HITs.PropositionalTruncation hiding (map ; elim) renaming (rec to propRec) +open import Cubical.HITs.Bouquet.Base +open import Cubical.HITs.Bouquet.Properties +open import Cubical.HITs.FreeGroupoid +open import Cubical.HITs.SphereBouquet.Base + open import Cubical.Algebra.Group + open import Cubical.Homotopy.Group.Base open import Cubical.Homotopy.Loopspace -open import Cubical.HITs.Bouquet.Base -open import Cubical.HITs.FreeGroup.Base -open import Cubical.HITs.FreeGroupoid - private variable ℓ : Level @@ -293,3 +304,70 @@ TruncatedFamilies≡ x = ua (TruncatedFamiliesEquiv x) ∥ FreeGroupoid A ∥₂ ≡⟨ sym freeGroupTruncIdempotent ⟩ FreeGroup A ∎ + +-- ΩBouquet(A) ≅ FinGroup A for A = Fin k +-- Todo: generalise to other discrete sets +Iso-ΩFinBouquet-FreeGroup : {n : ℕ} + → Iso (fst (Ω (Bouquet∙ (Fin n)))) (FreeGroup (Fin n)) +Iso-ΩFinBouquet-FreeGroup {n = n} = + compIso + (compIso (invIso (setTruncIdempotentIso (isOfHLevelPath' 2 + (isGroupoidBouquet (DiscreteFin {n} )) _ _))) + (equivToIso (TruncatedFamiliesEquiv base))) + (equivToIso (invEquiv freeGroupTruncIdempotent≃)) + +invIso-ΩFinBouquet-FreeGroupPresStr : {n : ℕ} (x y : FreeGroup (Fin n)) + → Iso.inv (Iso-ΩFinBouquet-FreeGroup {n}) (FG._·_ x y) + ≡ Iso.inv (Iso-ΩFinBouquet-FreeGroup {n}) x + ∙ Iso.inv (Iso-ΩFinBouquet-FreeGroup {n}) y +invIso-ΩFinBouquet-FreeGroupPresStr {n = n} x y = + cong (F ∘ G) (lem1 x y) ∙ lem2 (H x) (H y) + where + F = Iso.fun (setTruncIdempotentIso + (isOfHLevelPath' 2 (isGroupoidBouquet DiscreteFin) _ _)) + G = invEq (TruncatedFamiliesEquiv base) + H = fst freeGroupTruncIdempotent≃ + + lem2 : (x y : _) → F (G (ST.rec2 squash₂ (λ x y → ∣ x FGrp.· y ∣₂) x y)) + ≡ F (G x) ∙ F (G y) + lem2 = + ST.elim2 (λ _ _ → isOfHLevelPath 2 + (isOfHLevelPath' 2 (isGroupoidBouquet (DiscreteFin {n})) _ _) _ _) + λ _ _ → refl + + lem1 : (x t : _) → H (x FG.· t) + ≡ ST.rec2 squash₂ (λ x y → ∣ x FGrp.· y ∣₂) (H x) (H t) + lem1 x t = + cong H (cong₂ FG._·_ (sym (retEq freeGroupTruncIdempotent≃ x)) + (sym (retEq freeGroupTruncIdempotent≃ t))) + ∙ cong H (sym (lem3 (H x) (H t))) + ∙ secEq (freeGroupTruncIdempotent≃ {A = Fin n}) _ + where + lem3 : (x y : _) + → invEq freeGroupTruncIdempotent≃ + (ST.rec2 squash₂ (λ x y → ∣ x FGrp.· y ∣₂) x y) + ≡ invEq freeGroupTruncIdempotent≃ x FG.· invEq freeGroupTruncIdempotent≃ y + lem3 = ST.elim2 (λ _ _ → isOfHLevelPath 2 trunc _ _) λ _ _ → refl + +invIso-ΩFinBouquet-FreeGroupPresInv : {n : ℕ} (x : FreeGroup (Fin n)) + → Iso.inv (Iso-ΩFinBouquet-FreeGroup {n = n}) (FG.inv x) + ≡ sym (Iso.inv (Iso-ΩFinBouquet-FreeGroup {n = n}) x) +invIso-ΩFinBouquet-FreeGroupPresInv {n = n} = + morphLemmas.distrMinus FG._·_ _∙_ + (Iso.inv (Iso-ΩFinBouquet-FreeGroup {n = n})) + invIso-ΩFinBouquet-FreeGroupPresStr ε refl inv sym + (sym ∘ lUnit) (sym ∘ rUnit) FG.invl rCancel pathAssoc refl + +CharacFinBouquetFunIso : {m k : ℕ} + → Iso (Bouquet∙ (Fin m) →∙ Bouquet∙ (Fin k)) + (Fin m → FreeGroup (Fin k)) +CharacFinBouquetFunIso {m = m} {k} = + compIso (CharacBouquet∙Iso (Bouquet∙ (Fin k))) + (codomainIso (Iso-ΩFinBouquet-FreeGroup {n = k})) + +Iso-Bouquet→∙-SphereBouquet₁→∙ : {m k : ℕ} + → Iso (Bouquet∙ (Fin m) →∙ Bouquet∙ (Fin k)) + (SphereBouquet∙ (suc zero) (Fin m) →∙ SphereBouquet∙ (suc zero) (Fin k)) +Iso-Bouquet→∙-SphereBouquet₁→∙ = + compIso (pre∘∙equiv (invEquiv∙ Bouquet≃∙SphereBouquet)) + (post∘∙equiv (invEquiv∙ Bouquet≃∙SphereBouquet)) diff --git a/Cubical/HITs/Bouquet/Properties.agda b/Cubical/HITs/Bouquet/Properties.agda new file mode 100644 index 0000000000..10eba449e2 --- /dev/null +++ b/Cubical/HITs/Bouquet/Properties.agda @@ -0,0 +1,82 @@ +{- + +This file contains: + +- Definition of the Bouquet of circles of a type aka wedge of A circles + +-} +module Cubical.HITs.Bouquet.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Unit + +open import Cubical.HITs.Bouquet.Base +open import Cubical.HITs.SphereBouquet +open import Cubical.HITs.Pushout +open import Cubical.HITs.S1 + +open import Cubical.Homotopy.Loopspace + +private + variable + ℓ : Level + +elimPropBouquet : ∀ {ℓ ℓ'} {A : Type ℓ} + {B : Bouquet A → Type ℓ'} + (pr : (x : _) → isProp (B x)) + (b : B base) + → (x : Bouquet A) → B x +elimPropBouquet pr b base = b +elimPropBouquet {B = B} pr b (loop x i) = + isProp→PathP {B = λ i → B (loop x i)} (λ _ → pr _) b b i + +module _ {ℓ} {A : Type ℓ} where + SphereBouquet→Bouquet : SphereBouquet 1 A → Bouquet A + SphereBouquet→Bouquet (inl x) = base + SphereBouquet→Bouquet (inr (s , base)) = base + SphereBouquet→Bouquet (inr (s , loop i)) = loop s i + SphereBouquet→Bouquet (push a i) = base + + Bouquet→SphereBouquet : Bouquet A → SphereBouquet 1 A + Bouquet→SphereBouquet base = inl tt + Bouquet→SphereBouquet (loop x i) = + (push x ∙∙ (λ i → inr (x , loop i)) ∙∙ sym (push x)) i + + Iso-SphereBouquet-Bouquet : Iso (SphereBouquet 1 A) (Bouquet A) + Iso.fun Iso-SphereBouquet-Bouquet = SphereBouquet→Bouquet + Iso.inv Iso-SphereBouquet-Bouquet = Bouquet→SphereBouquet + Iso.rightInv Iso-SphereBouquet-Bouquet base = refl + Iso.rightInv Iso-SphereBouquet-Bouquet (loop x i) j = + SphereBouquet→Bouquet + (doubleCompPath-filler (push x) (λ i → inr (x , loop i)) (sym (push x)) (~ j) i) + Iso.leftInv Iso-SphereBouquet-Bouquet (inl x) = refl + Iso.leftInv Iso-SphereBouquet-Bouquet (inr (s , base)) = push s + Iso.leftInv Iso-SphereBouquet-Bouquet (inr (s , loop i)) j = + doubleCompPath-filler (push s) (λ i → inr (s , loop i)) (sym (push s)) (~ j) i + Iso.leftInv Iso-SphereBouquet-Bouquet (push a i) j = push a (i ∧ j) + + Bouquet≃∙SphereBouquet : SphereBouquet∙ 1 A ≃∙ Bouquet A , base + fst Bouquet≃∙SphereBouquet = isoToEquiv (Iso-SphereBouquet-Bouquet) + snd Bouquet≃∙SphereBouquet = refl + +module _ {ℓ ℓ'} {A : Type ℓ} (B : Pointed ℓ') where + BouquetFun∙→Ω : (Bouquet∙ A →∙ B) → (A → Ω B .fst) + BouquetFun∙→Ω f x = sym (snd f) ∙∙ cong (fst f) (loop x) ∙∙ snd f + + Ω→BouquetFun∙ : (A → Ω B .fst) → (Bouquet∙ A →∙ B) + fst (Ω→BouquetFun∙ f) base = pt B + fst (Ω→BouquetFun∙ f) (loop x i) = f x i + snd (Ω→BouquetFun∙ f) = refl + + CharacBouquet∙Iso : Iso (Bouquet∙ A →∙ B) (A → Ω B .fst) + Iso.fun CharacBouquet∙Iso = BouquetFun∙→Ω + Iso.inv CharacBouquet∙Iso = Ω→BouquetFun∙ + Iso.rightInv CharacBouquet∙Iso h i x j = + compPath-filler (h x) refl (~ i) j + fst (Iso.leftInv CharacBouquet∙Iso h i) base = snd h (~ i) + fst (Iso.leftInv CharacBouquet∙Iso h i) (loop x j) = + doubleCompPath-filler (sym (snd h)) (cong (fst h) (loop x)) (snd h) (~ i) j + snd (Iso.leftInv CharacBouquet∙Iso h i) j = snd h (~ i ∨ j) diff --git a/Cubical/HITs/Pushout/Base.agda b/Cubical/HITs/Pushout/Base.agda index 07afc42443..e75a0621f6 100644 --- a/Cubical/HITs/Pushout/Base.agda +++ b/Cubical/HITs/Pushout/Base.agda @@ -5,6 +5,7 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.Function open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Pointed open import Cubical.Data.Unit open import Cubical.Data.Sigma @@ -36,6 +37,10 @@ Pushout→ f₁ f₂ g₁ g₂ h₀ h₁ h₂ e₁ e₂ (push a i) = cofib : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Type _ cofib f = Pushout (λ _ → tt) f +cofib∙ : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Pointed _ +fst (cofib∙ f) = cofib f +snd (cofib∙ f) = inl tt + cfcod : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) → B → cofib f cfcod f = inr @@ -140,3 +145,38 @@ module _ {ℓ₁ ℓ₂ ℓ₃ : Level} {A : Type ℓ₁} {B : Type ℓ₂} {C : leftInv IsoPushoutPushoutGen (inl x) = refl leftInv IsoPushoutPushoutGen (inr x) = refl leftInv IsoPushoutPushoutGen (push a i) j = rUnit (push a) (~ j) i + + +-- Pushout along a composition +module _ {ℓ ℓ' ℓ'' ℓ'''} + {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} + (f1 : A → B) (f2 : B → C) {g : A → D} where + PushoutComp→IteratedPushout : + Pushout (f2 ∘ f1) g → Pushout {C = Pushout f1 g} f2 inl + PushoutComp→IteratedPushout (inl x) = inl x + PushoutComp→IteratedPushout (inr x) = inr (inr x) + PushoutComp→IteratedPushout (push a i) = (push (f1 a) ∙ λ i → inr (push a i)) i + + IteratedPushout→PushoutComp : + Pushout {C = Pushout f1 g} f2 inl → Pushout (f2 ∘ f1) g + IteratedPushout→PushoutComp (inl x) = inl x + IteratedPushout→PushoutComp (inr (inl x)) = inl (f2 x) + IteratedPushout→PushoutComp (inr (inr x)) = inr x + IteratedPushout→PushoutComp (inr (push a i)) = push a i + IteratedPushout→PushoutComp (push a i) = inl (f2 a) + + Iso-PushoutComp-IteratedPushout : + Iso (Pushout (f2 ∘ f1) g) (Pushout {C = Pushout f1 g} f2 inl) + Iso.fun Iso-PushoutComp-IteratedPushout = PushoutComp→IteratedPushout + Iso.inv Iso-PushoutComp-IteratedPushout = IteratedPushout→PushoutComp + Iso.rightInv Iso-PushoutComp-IteratedPushout (inl x) = refl + Iso.rightInv Iso-PushoutComp-IteratedPushout (inr (inl x)) = push x + Iso.rightInv Iso-PushoutComp-IteratedPushout (inr (inr x)) = refl + Iso.rightInv Iso-PushoutComp-IteratedPushout (inr (push a i)) j = + compPath-filler' (push (f1 a)) (λ i₁ → inr (push a i₁)) (~ j) i + Iso.rightInv Iso-PushoutComp-IteratedPushout (push a i) j = push a (i ∧ j) + Iso.leftInv Iso-PushoutComp-IteratedPushout (inl x) = refl + Iso.leftInv Iso-PushoutComp-IteratedPushout (inr x) = refl + Iso.leftInv Iso-PushoutComp-IteratedPushout (push a i) j = + (cong-∙ IteratedPushout→PushoutComp (push (f1 a)) (λ i → inr (push a i)) + ∙ sym (lUnit _)) j i diff --git a/Cubical/HITs/Pushout/Properties.agda b/Cubical/HITs/Pushout/Properties.agda index 8a05e0fc8f..62f4d99d0d 100644 --- a/Cubical/HITs/Pushout/Properties.agda +++ b/Cubical/HITs/Pushout/Properties.agda @@ -37,8 +37,10 @@ open import Cubical.Data.Sigma open import Cubical.Data.Unit open import Cubical.Data.Empty as ⊥ open import Cubical.Data.List +open import Cubical.Data.Sum open import Cubical.HITs.Pushout.Base +open import Cubical.HITs.Pushout.Flattening open import Cubical.HITs.Susp.Base private @@ -624,6 +626,34 @@ module _ {ℓA₁ ℓB₁ ℓC₁ ℓA₂ ℓB₂ ℓC₂} pushoutEquiv : Pushout f₁ g₁ ≃ Pushout f₂ g₂ pushoutEquiv = isoToEquiv pushoutIso +-- Pushouts commute with Σ +module _ {ℓ ℓ' ℓ'' ℓ'''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + (f : A → B) (g : A → C) (X : Pushout f g → Type ℓ''') where + open FlatteningLemma f g (X ∘ inl) (X ∘ inr) (λ a → substEquiv X (push a)) + + PushoutΣL : Σ[ a ∈ A ] X (inl (f a)) → Σ[ b ∈ B ] X (inl b) + PushoutΣL (a , x) = f a , x + + PushoutΣR : Σ[ a ∈ A ] X (inl (f a)) → Σ[ c ∈ C ] X (inr c) + PushoutΣR (a , x) = g a , subst X (push a) x + + PushoutΣ : Type _ + PushoutΣ = Pushout PushoutΣL PushoutΣR + + repairLeft : (a : Pushout f g) → X a ≃ E a + repairLeft (inl x) = idEquiv _ + repairLeft (inr x) = idEquiv _ + repairLeft (push a i) = help i + where + help : PathP (λ i → X (push a i) ≃ E (push a i)) (idEquiv _) (idEquiv _) + help = ΣPathPProp (λ _ → isPropIsEquiv _) + (toPathP (funExt λ x → + transportRefl _ ∙ substSubst⁻ X (push a) x)) + + ΣPushout≃PushoutΣ : Σ (Pushout f g) X ≃ PushoutΣ + ΣPushout≃PushoutΣ = + compEquiv (Σ-cong-equiv-snd repairLeft) flatten + module _ {C : Type ℓ} {B : Type ℓ'} where PushoutAlongEquiv→ : {A : Type ℓ} (e : A ≃ C) (f : A → B) → Pushout (fst e) f → B @@ -694,6 +724,17 @@ rightInv (PushoutEmptyFam {A = A} {B = B} ¬A ¬C {f = f} {g = g}) (push a i) j (¬A a) j i leftInv (PushoutEmptyFam {A = A} {B = B} ¬A ¬C) x = refl +PushoutEmptyDomainIso : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + → Iso (Pushout {A = ⊥} {B = A} {C = B} (λ()) (λ())) (A ⊎ B) +Iso.fun PushoutEmptyDomainIso (inl x) = inl x +Iso.fun PushoutEmptyDomainIso (inr x) = inr x +Iso.inv PushoutEmptyDomainIso (inl x) = inl x +Iso.inv PushoutEmptyDomainIso (inr x) = inr x +Iso.rightInv PushoutEmptyDomainIso (inl x) = refl +Iso.rightInv PushoutEmptyDomainIso (inr x) = refl +Iso.leftInv PushoutEmptyDomainIso (inl x) = refl +Iso.leftInv PushoutEmptyDomainIso (inr x) = refl + PushoutCompEquivIso : ∀ {ℓA ℓA' ℓB ℓB' ℓC} {A : Type ℓA} {A' : Type ℓA'} {B : Type ℓB} {B' : Type ℓB'} {C : Type ℓC} @@ -1142,3 +1183,20 @@ module PushoutPasteLeft {ℓ₀ ℓ₂ ℓ₄ ℓP ℓA ℓB : Level} isPushoutTotSquare→isPushoutRightSquare e = rotatePushoutSquare (_ , help) .snd where help = M.isPushoutTotSquare→isPushoutBottomSquare (rotatePushoutSquare (_ , e) .snd) + +LiftPushoutIso : (ℓP : Level) {f : A → B} {g : A → C} + → Iso (Pushout (liftFun {ℓ'' = ℓP} {ℓ''' = ℓP} f) + (liftFun {ℓ'' = ℓP} {ℓ''' = ℓP} g)) + (Lift {j = ℓP} (Pushout f g)) +fun (LiftPushoutIso ℓP) (inl (lift x)) = lift (inl x) +fun (LiftPushoutIso ℓP) (inr (lift x)) = lift (inr x) +fun (LiftPushoutIso ℓP) (push (lift a) i) = lift (push a i) +inv (LiftPushoutIso ℓP) (lift (inl x)) = inl (lift x) +inv (LiftPushoutIso ℓP) (lift (inr x)) = inr (lift x) +inv (LiftPushoutIso ℓP) (lift (push a i)) = push (lift a) i +rightInv (LiftPushoutIso ℓP) (lift (inl x)) = refl +rightInv (LiftPushoutIso ℓP) (lift (inr x)) = refl +rightInv (LiftPushoutIso ℓP) (lift (push a i)) = refl +leftInv (LiftPushoutIso ℓP) (inl (lift x)) = refl +leftInv (LiftPushoutIso ℓP) (inr (lift x)) = refl +leftInv (LiftPushoutIso ℓP) (push (lift a) i) = refl diff --git a/Cubical/HITs/SequentialColimit/Base.agda b/Cubical/HITs/SequentialColimit/Base.agda index 0b9a0355d7..303b3edd7c 100644 --- a/Cubical/HITs/SequentialColimit/Base.agda +++ b/Cubical/HITs/SequentialColimit/Base.agda @@ -34,3 +34,8 @@ realiseSequenceMap : {C : Sequence ℓ} {D : Sequence ℓ'} realiseSequenceMap (sequencemap map comm) (incl x) = incl (map _ x) realiseSequenceMap (sequencemap map comm) (push {n = n} x i) = (push (map n x) ∙ λ i → incl (comm n x i)) i + +LiftSequence : ∀ {ℓA} (ℓ↑ : Level) (S : Sequence ℓA) → + Sequence (ℓ-max ℓA ℓ↑) +obj (LiftSequence ℓ↑ S) n = Lift {j = ℓ↑} (obj S n) +map (LiftSequence ℓ↑ S) = liftFun (map S) diff --git a/Cubical/HITs/SequentialColimit/Properties.agda b/Cubical/HITs/SequentialColimit/Properties.agda index 2a52c0040d..dcaaeae5a9 100644 --- a/Cubical/HITs/SequentialColimit/Properties.agda +++ b/Cubical/HITs/SequentialColimit/Properties.agda @@ -24,6 +24,7 @@ open import Cubical.Data.Fin.Inductive open import Cubical.Data.Sigma open import Cubical.HITs.SequentialColimit.Base +open import Cubical.HITs.Pushout open import Cubical.Homotopy.Connected private @@ -612,3 +613,200 @@ SeqColimIso S n = compIso (Iso-SeqColim→SeqColimShift S n) (sequenceEquiv→ColimIso (SequenceIso→SequenceEquiv (ShiftSequenceIso n))) + +module _ {ℓ ℓ' ℓ'' : Level} {A : Sequence ℓ} {B : Sequence ℓ'} {C : Sequence ℓ''} + (F : SequenceMap A B) (G : SequenceMap A C) where + + PushoutSequenceFam : ℕ → Type _ + PushoutSequenceFam n = Pushout (SequenceMap.map F n) (SequenceMap.map G n) + + PushoutSequenceMap : (n : ℕ) → PushoutSequenceFam n → PushoutSequenceFam (suc n) + PushoutSequenceMap n (inl x) = inl (map B x) + PushoutSequenceMap n (inr x) = inr (map C x) + PushoutSequenceMap n (push a i) = + ((λ i → inl (SequenceMap.comm F n a i)) + ∙∙ push (map A a) + ∙∙ (λ i → inr (SequenceMap.comm G n a (~ i)))) i + + PushoutSequence : Sequence _ + obj PushoutSequence = PushoutSequenceFam + map PushoutSequence = PushoutSequenceMap _ + + PushoutColim : Type _ + PushoutColim = Pushout (realiseSequenceMap F) (realiseSequenceMap G) + + PushoutSequenceFam→PushoutColim : (n : ℕ) → PushoutSequenceFam n → PushoutColim + PushoutSequenceFam→PushoutColim n (inl x) = inl (incl {n = n} x) + PushoutSequenceFam→PushoutColim n (inr x) = inr (incl {n = n} x) + PushoutSequenceFam→PushoutColim n (push a i) = push (incl {n = n} a) i + + PushoutSequenceFam→PushoutColimCommFill : (n : ℕ) (a : _) (i j k : I) → PushoutColim + PushoutSequenceFam→PushoutColimCommFill n a i j k = + hfill (λ k → + λ { (i = i0) → inl (compPath-filler' + (push {n = n} (SequenceMap.map F n a)) + (λ i₁ → incl {n = suc n} (SequenceMap.comm F n a i₁)) + (~ j) (~ k)) + ; (i = i1) → inr (compPath-filler' + (push {n = n} (SequenceMap.map G n a)) + (λ i₁ → incl {n = suc n} (SequenceMap.comm G n a i₁)) + (~ j) (~ k)) + ; (j = i0) → push (push {n = n} a (~ k)) i + ; (j = i1) → PushoutSequenceFam→PushoutColim (suc n) + (doubleCompPath-filler + (λ i → inl (SequenceMap.comm F n a i)) + (push (map A a)) + (λ i → inr (SequenceMap.comm G n a (~ i))) k i)}) + (inS (push (incl {n = suc n} (map A a)) i)) k + + PushoutSequenceFam→PushoutColimComm : (n : ℕ) (x : PushoutSequenceFam n) + → PushoutSequenceFam→PushoutColim n x + ≡ PushoutSequenceFam→PushoutColim (suc n) + (PushoutSequenceMap n x) + PushoutSequenceFam→PushoutColimComm n (inl x) i = inl (push {n = n} x i) + PushoutSequenceFam→PushoutColimComm n (inr x) i = inr (push {n = n} x i) + PushoutSequenceFam→PushoutColimComm n (push a i) j = + PushoutSequenceFam→PushoutColimCommFill n a i j i1 + + ColimPushout→PushoutColim : SeqColim PushoutSequence → PushoutColim + ColimPushout→PushoutColim (incl {n = n} x) = + PushoutSequenceFam→PushoutColim n x + ColimPushout→PushoutColim (push {n = n} x i) = + PushoutSequenceFam→PushoutColimComm n x i + + PushoutColim→ColimPushoutL : SeqColim B → SeqColim PushoutSequence + PushoutColim→ColimPushoutL (incl {n = n} x) = incl {n = n} (inl x) + PushoutColim→ColimPushoutL (push {n = n} x i) = push {n = n} (inl x) i + + PushoutColim→ColimPushoutR : SeqColim C → SeqColim PushoutSequence + PushoutColim→ColimPushoutR (incl {n = n} x) = incl {n = n} (inr x) + PushoutColim→ColimPushoutR (push {n = n} x i) = push {n = n} (inr x) i + + PushoutColim→ColimPushoutLRFill : (n : ℕ) (x : _) (i j k : I) + → SeqColim PushoutSequence + PushoutColim→ColimPushoutLRFill n x i j k = + hfill (λ k → + λ {(i = i0) → push {n = n} (push x j) (~ k) + ; (i = i1) → p i0 j + ; (j = i0) → PushoutColim→ColimPushoutL + (compPath-filler' + (push {n = n} (SequenceMap.map F n x)) + (λ i → incl {n = suc n} (SequenceMap.comm F n x i)) k i) + ; (j = i1) → PushoutColim→ColimPushoutR + (compPath-filler' + (push {n = n} (SequenceMap.map G n x)) + (λ i → incl {n = suc n} (SequenceMap.comm G n x i)) k i)}) + (inS (p (~ i) j)) k + where + p : (i j : I) → SeqColim PushoutSequence + p i j = incl {n = suc n} (doubleCompPath-filler + (λ i₁ → inl (SequenceMap.comm F n x i₁)) + (push (map A x)) + (λ i₁ → inr (SequenceMap.comm G n x (~ i₁))) i j) + + PushoutColim→ColimPushoutLR : (a : SeqColim A) + → PushoutColim→ColimPushoutL (realiseSequenceMap F a) + ≡ PushoutColim→ColimPushoutR (realiseSequenceMap G a) + PushoutColim→ColimPushoutLR (incl {n = n} x) i = incl {n = n} (push x i) + PushoutColim→ColimPushoutLR (push {n = n} x i) j = + PushoutColim→ColimPushoutLRFill n x i j i1 + + PushoutColim→ColimPushout : PushoutColim → SeqColim PushoutSequence + PushoutColim→ColimPushout (inl x) = PushoutColim→ColimPushoutL x + PushoutColim→ColimPushout (inr x) = PushoutColim→ColimPushoutR x + PushoutColim→ColimPushout (push a i) = PushoutColim→ColimPushoutLR a i + + PushoutColim→ColimPushout→PushoutColim-Incl : (n : ℕ) (x : PushoutSequenceFam n) + → PushoutColim→ColimPushout (ColimPushout→PushoutColim (incl {n = n} x)) + ≡ incl {n = n} x + PushoutColim→ColimPushout→PushoutColim-Incl n (inl x) = refl + PushoutColim→ColimPushout→PushoutColim-Incl n (inr x) = refl + PushoutColim→ColimPushout→PushoutColim-Incl n (push a i) = refl + + PushoutColim→ColimPushout→PushoutColim-Push : (n : ℕ) (x : PushoutSequenceFam n) + → Square {A = SeqColim PushoutSequence} + (cong PushoutColim→ColimPushout + (PushoutSequenceFam→PushoutColimComm n x)) + (push {n = n} x) + (PushoutColim→ColimPushout→PushoutColim-Incl n x) + (PushoutColim→ColimPushout→PushoutColim-Incl (suc n) + (PushoutSequenceMap n x)) + PushoutColim→ColimPushout→PushoutColim-Push n (inl x) = refl + PushoutColim→ColimPushout→PushoutColim-Push n (inr x) = refl + PushoutColim→ColimPushout→PushoutColim-Push n (push a i) j k = + hcomp (λ r → + λ {(i = i0) → PushoutColim→ColimPushoutL + (compPath-filler' + (push {n = n} (SequenceMap.map F n a)) + (λ i → incl {n = suc n} (SequenceMap.comm F n a i)) + (~ k) (~ r)) + ; (i = i1) → PushoutColim→ColimPushoutR + (compPath-filler' + (push {n = n} (SequenceMap.map G n a)) + (λ i → incl {n = suc n} (SequenceMap.comm G n a i)) + (~ k) (~ r)) + ; (j = i0) → PushoutColim→ColimPushout + (PushoutSequenceFam→PushoutColimCommFill n a i k r) + ; (j = i1) → PushoutColim→ColimPushoutLRFill n a (~ r) i (~ k) + ; (k = i0) → PushoutColim→ColimPushoutLRFill n a (~ r) i i1 + ; (k = i1) → PushoutColim→ColimPushout→PushoutColim-Incl (suc n) + ((doubleCompPath-filler + (λ i₁ → inl (SequenceMap.comm F n a i₁)) (push (map A a)) + (λ i₁ → inr (SequenceMap.comm G n a (~ i₁))) r i)) j}) + (incl {n = suc n} (push (map A a) i)) + + PushoutColim→ColimPushout→PushoutColim : (x : SeqColim PushoutSequence) + → PushoutColim→ColimPushout (ColimPushout→PushoutColim x) ≡ x + PushoutColim→ColimPushout→PushoutColim (incl x) = + PushoutColim→ColimPushout→PushoutColim-Incl _ x + PushoutColim→ColimPushout→PushoutColim (push x i) j = + PushoutColim→ColimPushout→PushoutColim-Push _ x j i + + ColimPushout→PushoutColim→ColimPushout-inl : (x : _) + → ColimPushout→PushoutColim (PushoutColim→ColimPushoutL x) ≡ inl x + ColimPushout→PushoutColim→ColimPushout-inl (incl x) = refl + ColimPushout→PushoutColim→ColimPushout-inl (push x i) = refl + + ColimPushout→PushoutColim→ColimPushout-inr : (x : _) + → ColimPushout→PushoutColim (PushoutColim→ColimPushoutR x) ≡ inr x + ColimPushout→PushoutColim→ColimPushout-inr (incl x) = refl + ColimPushout→PushoutColim→ColimPushout-inr (push x i) = refl + + Iso-PushoutColim-ColimPushout : Iso PushoutColim (SeqColim PushoutSequence) + Iso.fun Iso-PushoutColim-ColimPushout = PushoutColim→ColimPushout + Iso.inv Iso-PushoutColim-ColimPushout = ColimPushout→PushoutColim + Iso.rightInv Iso-PushoutColim-ColimPushout x = + PushoutColim→ColimPushout→PushoutColim x + Iso.leftInv Iso-PushoutColim-ColimPushout (inl x) = + ColimPushout→PushoutColim→ColimPushout-inl x + Iso.leftInv Iso-PushoutColim-ColimPushout (inr x) = + ColimPushout→PushoutColim→ColimPushout-inr x + Iso.leftInv Iso-PushoutColim-ColimPushout (push (incl {n = n} x) i) j = + push (incl {n = n} x) i + Iso.leftInv Iso-PushoutColim-ColimPushout (push (push {n = n} x k) i) j = + hcomp (λ r → + λ {(i = i0) → ColimPushout→PushoutColim→ColimPushout-inl + (compPath-filler' + (push {n = n} (SequenceMap.map F n x)) + (λ i → incl {n = suc n} (SequenceMap.comm F n x i)) r k) j + ; (i = i1) → ColimPushout→PushoutColim→ColimPushout-inr + (compPath-filler' + (push {n = n} (SequenceMap.map G n x)) + (λ i → incl {n = suc n} (SequenceMap.comm G n x i)) r k) j + ; (j = i0) → ColimPushout→PushoutColim + (PushoutColim→ColimPushoutLRFill n x k i r) + ; (j = i1) → PushoutSequenceFam→PushoutColimCommFill n x i (~ r) (~ k) + ; (k = i0) → PushoutSequenceFam→PushoutColimCommFill n x i (~ r) i1 + ; (k = i1) → push (incl {n = suc n} (map A x)) i}) + (PushoutSequenceFam→PushoutColimCommFill n x i i1 (~ k)) + +SeqColimLift : (S : Sequence ℓ') + → Iso (SeqColim (LiftSequence ℓ S)) (SeqColim S) +Iso.fun (SeqColimLift S) (incl (lift x)) = incl x +Iso.fun (SeqColimLift S) (push (lift x) i) = push x i +Iso.inv (SeqColimLift S) (incl x) = incl (lift x) +Iso.inv (SeqColimLift S) (push x i) = push (lift x) i +Iso.rightInv (SeqColimLift S) (incl x) = refl +Iso.rightInv (SeqColimLift S) (push x i) = refl +Iso.leftInv (SeqColimLift S) (incl x) = refl +Iso.leftInv (SeqColimLift S) (push x i) = refl diff --git a/Cubical/HITs/Sn/Degree.agda b/Cubical/HITs/Sn/Degree.agda index 027dc6f4ed..fe87723591 100644 --- a/Cubical/HITs/Sn/Degree.agda +++ b/Cubical/HITs/Sn/Degree.agda @@ -8,11 +8,13 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv open import Cubical.Foundations.HLevels open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws open import Cubical.Data.Nat open import Cubical.Data.Bool -open import Cubical.Data.Int renaming (_·_ to _·ℤ_) +open import Cubical.Data.Int renaming (_·_ to _·ℤ_ ; _+_ to _+ℤ_) +open import Cubical.HITs.S1 open import Cubical.HITs.Sn open import Cubical.HITs.Susp open import Cubical.HITs.Truncation as TR @@ -24,7 +26,10 @@ open import Cubical.Algebra.Group.Morphisms open import Cubical.Homotopy.Group.Base open import Cubical.Homotopy.Group.PinSn +open import Cubical.Homotopy.Loopspace +open import Cubical.ZCohomology.Base +open import Cubical.ZCohomology.GroupStructure open import Cubical.ZCohomology.Groups.Sn -- prelims for definition in 0 case @@ -160,3 +165,54 @@ degreeIdfun (suc n) = cong (degree (suc n)) (sym (cong fst suspFunS∙Id)) ∙∙ (sym (degreeSusp n (idfun _))) ∙∙ degreeIdfun n + +degreeHom : {n : ℕ} (f g : S₊∙ (suc n) →∙ S₊∙ (suc n)) + → degree (suc n) (∙Π f g .fst) + ≡ degree (suc n) (fst f) +ℤ degree (suc n) (fst g) +degreeHom {n = n} f g = + cong (Iso.fun (Hⁿ-Sⁿ≅ℤ n .fst)) (cong ∣_∣₂ (funExt + λ x → cong ∣_∣ₕ (cong₂ (λ f g → ∙Π f g .fst x) + (sym (Iso.rightInv (sphereFunIso n) f)) + (sym (Iso.rightInv (sphereFunIso n) g))) + ∙∙ help n _ _ x + ∙∙ cong₂ (λ x y → ∣ x ∣ₕ +[ suc n ]ₖ ∣ y ∣ₕ) + (funExt⁻ (cong fst (Iso.rightInv (sphereFunIso n) f)) x) + (funExt⁻ (cong fst (Iso.rightInv (sphereFunIso n) g)) x))) + ∙ IsGroupHom.pres· (Hⁿ-Sⁿ≅ℤ n .snd) _ _ + where + help : (n : ℕ) (f g : S₊∙ n →∙ Ω (S₊∙ (suc n))) (x : S₊ (suc n)) + → Path (coHomK (suc n)) + ∣ ∙Π (Iso.fun (sphereFunIso n) f) + (Iso.fun (sphereFunIso n) g) .fst x ∣ₕ + (∣ fst (Iso.fun (sphereFunIso n) f) x ∣ + +[ suc n ]ₖ ∣ fst (Iso.fun (sphereFunIso n) g) x ∣) + help zero f g base = refl + help zero f g (loop i) j = lem j i + where + lem : cong ∣_∣ₕ ((fst f false ∙ refl) ∙ fst g false ∙ refl) + ≡ cong₂ (λ x y → ∣ x ∣ₕ +[ suc zero ]ₖ ∣ y ∣ₕ) + (fst f false) (fst g false) + lem = cong-∙ ∣_∣ₕ _ _ + ∙ cong₂ _∙_ ((λ i j → ∣ rUnit (fst f false) (~ i) j ∣ₕ) + ∙ (λ j → cong (λ x → rUnitₖ 1 ∣ x ∣ₕ (~ j)) (fst f false))) + ((λ i j → ∣ rUnit (fst g false) (~ i) j ∣ₕ) + ∙ (λ j → cong (λ x → lUnitₖ 1 ∣ x ∣ₕ (~ j)) (fst g false))) + ∙ sym (cong₂Funct (λ x y → ∣ x ∣ₕ +[ suc zero ]ₖ ∣ y ∣ₕ) + (fst f false) (fst g false)) + help (suc n) f g north = refl + help (suc n) f g south = refl + help (suc n) f g (merid a i) j = lem j i + where + lem : cong ∣_∣ₕ ((cong (fst (toΩ→fromSusp f)) (σS a) ∙ refl) + ∙ (cong (fst (toΩ→fromSusp g)) (σS a) ∙ refl)) + ≡ cong₂ (λ x y → ∣ x ∣ₕ +[ suc (suc n) ]ₖ ∣ y ∣ₕ) + (fst f a) (fst g a) + lem = cong-∙ ∣_∣ₕ _ _ + ∙ cong₂ _∙_ (cong (cong ∣_∣ₕ) + (funExt⁻ (cong fst (Iso.leftInv ΩSuspAdjointIso f)) a) + ∙ λ j i → rUnitₖ (suc (suc n)) ∣ fst f a i ∣ₕ (~ j)) + (cong (cong ∣_∣ₕ) + (funExt⁻ (cong fst (Iso.leftInv ΩSuspAdjointIso g)) a) + ∙ (λ j i → lUnitₖ (suc (suc n)) ∣ fst g a i ∣ₕ (~ j))) + ∙ sym (cong₂Funct (λ x y → ∣ x ∣ₕ +[ suc (suc n) ]ₖ ∣ y ∣ₕ) + (fst f a) (fst g a)) diff --git a/Cubical/HITs/Sn/Properties.agda b/Cubical/HITs/Sn/Properties.agda index 823cce02b0..002013f137 100644 --- a/Cubical/HITs/Sn/Properties.agda +++ b/Cubical/HITs/Sn/Properties.agda @@ -46,6 +46,13 @@ IsoSucSphereSusp : (n : ℕ) → Iso (S₊ (suc n)) (Susp (S₊ n)) IsoSucSphereSusp zero = S¹IsoSuspBool IsoSucSphereSusp (suc n) = idIso +IsoSphereSusp : (n : ℕ) → Iso (S₊ n) (Susp (S⁻ n)) +IsoSphereSusp zero = BoolIsoSusp⊥ +IsoSphereSusp (suc n) = IsoSucSphereSusp n + +EquivSphereSusp : (n : ℕ) → Susp (S⁻ n) ≃ S₊ n +EquivSphereSusp n = isoToEquiv (invIso (IsoSphereSusp n)) + IsoSucSphereSusp∙ : (n : ℕ) → Iso.inv (IsoSucSphereSusp n) north ≡ ptSn (suc n) IsoSucSphereSusp∙ zero = refl @@ -327,14 +334,15 @@ sphereConnected : (n : HLevel) → isConnected (suc n) (S₊ n) sphereConnected n = ∣ ptSn n ∣ , elim (λ _ → isOfHLevelPath (suc n) (isOfHLevelTrunc (suc n)) _ _) (λ a → sym (spoke ∣_∣ (ptSn n)) ∙ spoke ∣_∣ a) -sphereToTrunc : (n : ℕ) {A : S₊ n → Type ℓ} - → ((x : S₊ n) → hLevelTrunc (suc n) (A x)) +sphereToTrunc : (n : ℕ) {A : S⁻ n → Type ℓ} + → ((x : S⁻ n) → hLevelTrunc n (A x)) → ∥ ((x : _) → A x) ∥₁ -sphereToTrunc zero {A = A} indr = +sphereToTrunc zero _ = ∣ (λ{()}) ∣₁ +sphereToTrunc (suc zero) {A = A} indr = rec squash₁ (λ p → rec squash₁ (λ q → ∣ (λ { false → q ; true → p}) ∣₁) (indr false)) (indr true) -sphereToTrunc (suc zero) {A = A} indr = +sphereToTrunc (suc (suc zero)) {A = A} indr = lem (indr base) (cong indr loop) where lem : (x : hLevelTrunc 2 (A base)) @@ -345,8 +353,8 @@ sphereToTrunc (suc zero) {A = A} indr = ; (loop i) → toPathP {A = λ i → A (loop i)} q i}) ∣₁) (PathIdTruncIso 1 .Iso.fun (fromPathP p)) -sphereToTrunc (suc (suc n)) {A = A} indr = - lem (sphereToTrunc (suc n)) (indr north) (indr south) +sphereToTrunc (suc (suc (suc n))) {A = A} indr = + lem (sphereToTrunc (suc (suc n))) (indr north) (indr south) λ a → cong indr (merid a) where lem : ({A : S₊ (suc n) → Type _} → diff --git a/Cubical/HITs/SphereBouquet/Base.agda b/Cubical/HITs/SphereBouquet/Base.agda index b60c1dc818..66fe3c62d2 100644 --- a/Cubical/HITs/SphereBouquet/Base.agda +++ b/Cubical/HITs/SphereBouquet/Base.agda @@ -5,6 +5,7 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Pointed open import Cubical.Data.Nat +open import Cubical.Data.Fin.Inductive open import Cubical.HITs.Wedge open import Cubical.HITs.Sn @@ -14,3 +15,11 @@ SphereBouquet n A = ⋁gen A λ a → S₊∙ n SphereBouquet∙ : ∀ {ℓ} (n : ℕ) (A : Type ℓ) → Pointed ℓ SphereBouquet∙ n A = ⋁gen∙ A λ a → S₊∙ n + +FinSphereBouquetMap : (c1 c2 : ℕ) (n : ℕ) → Type +FinSphereBouquetMap c1 c2 n = + SphereBouquet (suc n) (Fin c1) → SphereBouquet (suc n) (Fin c2) + +FinSphereBouquetMap∙ : (c1 c2 : ℕ) (n : ℕ) → Type +FinSphereBouquetMap∙ c1 c2 n = + SphereBouquet∙ (suc n) (Fin c1) →∙ SphereBouquet∙ (suc n) (Fin c2) diff --git a/Cubical/HITs/SphereBouquet/Degree.agda b/Cubical/HITs/SphereBouquet/Degree.agda index 6902c7fd6d..1657b4cde9 100644 --- a/Cubical/HITs/SphereBouquet/Degree.agda +++ b/Cubical/HITs/SphereBouquet/Degree.agda @@ -42,6 +42,8 @@ open import Cubical.ZCohomology.GroupStructure open import Cubical.ZCohomology.Groups.Sn open import Cubical.ZCohomology.Properties +open import Cubical.Homotopy.Group.Base + private variable ℓ ℓ' : Level @@ -489,3 +491,51 @@ bouquetDegreeComp∙ {n = zero} {m} {k = k} {l = l} f g = ∙ cong fst (cong₂ compGroupHom (sym (bouquetDegreeSusp (fst g))) (sym (bouquetDegreeSusp (fst f))))) bouquetDegreeComp∙ {n = suc n} = bouquetDegreeComp∙Suc + +bouquetDegree∙Π : (n m k : ℕ) + (f g : SphereBouquet∙ (suc n) (Fin m) →∙ SphereBouquet∙ (suc n) (Fin k)) + → bouquetDegree (SphereBouquet∙Π f g .fst) + ≡ addGroupHom ℤ[Fin m ] ℤ[Fin k ] (bouquetDegree (fst f)) (bouquetDegree (fst g)) +bouquetDegree∙Π n m k f g = + agreeOnℤFinGenerator→≡ + λ s → funExt λ y → (sym (isGeneratorℤFinGenerator' _ s) + ∙ cong (degree (suc n)) (funExt (main n s y f g)) + ∙ degreeHom {n = n} + ((λ x₁ → pickPetal y (fst f (inr (s , x₁)))) + , cong (pickPetal y) (cong (fst f) (sym (push s)) ∙ snd f)) + ((λ x₁ → pickPetal y (fst g (inr (s , x₁)))) + , cong (pickPetal y) (cong (fst g) (sym (push s)) ∙ snd g)) + ∙ isGeneratorℤFinGenerator' _ s + ∙ cong sumFinℤ (funExt λ x → + ·DistR+ (ℤFinGenerator s x) + (degree (suc n) (λ x₁ → pickPetal y (fst f (inr (x , x₁))))) + (degree (suc n) (λ x₁ → pickPetal y (fst g (inr (x , x₁)))))) + ∙ sumFinℤHom _ _) -- + where + main : (n : ℕ) (s : Fin m) (y : _) + (f g : SphereBouquet∙ (suc n) (Fin m) + →∙ SphereBouquet∙ (suc n) (Fin k)) (x : S₊ (suc n)) + → pickPetal y (SphereBouquet∙Π f g .fst (inr (s , x))) + ≡ (∙Π ((λ x₁ → pickPetal y (fst f (inr (s , x₁)))) , + (λ i → pickPetal y (((λ i₁ → fst f (push s (~ i₁))) ∙ snd f) i))) + ((λ x₁ → pickPetal y (fst g (inr (s , x₁)))) , + (λ i → pickPetal y (((λ i₁ → fst g (push s (~ i₁))) ∙ snd g) i))) + .fst x) + main zero s y f g base = refl + main zero s y f g (loop i) = refl + main (suc n) s y f g north = refl + main (suc n) s y f g south = refl + main (suc n) s y f g (merid a i) j = lem j i + where + lem : cong (pickPetal y) (cong (λ x → SphereBouquet∙Π f g .fst (inr (s , x))) + (merid a)) + ≡ (sym (cong (pickPetal y) ((λ i₁ → fst f (push s (~ i₁))) ∙ snd f)) + ∙∙ cong (pickPetal y) (cong (λ x → fst f (inr (s , x))) (σS a)) + ∙∙ cong (pickPetal y) ((λ i₁ → fst f (push s (~ i₁))) ∙ snd f)) + ∙ (sym (cong (pickPetal y) ((λ i₁ → fst g (push s (~ i₁))) ∙ snd g)) + ∙∙ cong (pickPetal y) (cong (λ x → fst g (inr (s , x))) (σS a)) + ∙∙ cong (pickPetal y) ((λ i₁ → fst g (push s (~ i₁))) ∙ snd g)) + lem = (cong-∙ (pickPetal y) _ _ + ∙ cong₂ _∙_ + (cong-∙∙ (pickPetal y) (sym ((λ i₁ → fst f (push s (~ i₁))) ∙ snd f)) _ _) + (cong-∙∙ (pickPetal y) (sym ((λ i₁ → fst g (push s (~ i₁))) ∙ snd g)) _ _)) diff --git a/Cubical/HITs/SphereBouquet/Properties.agda b/Cubical/HITs/SphereBouquet/Properties.agda index 5c38145268..ec2cd8e756 100644 --- a/Cubical/HITs/SphereBouquet/Properties.agda +++ b/Cubical/HITs/SphereBouquet/Properties.agda @@ -1,5 +1,3 @@ -{-# OPTIONS --cubical #-} - module Cubical.HITs.SphereBouquet.Properties where open import Cubical.Foundations.Prelude @@ -28,6 +26,7 @@ open import Cubical.HITs.Wedge open import Cubical.HITs.SphereBouquet.Base open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Group.Base private variable @@ -389,3 +388,61 @@ module _ {Cₙ Cₙ₊₁ : Type ℓ} (n mₙ : ℕ) Bouquet≃-gen : cofib (invEq e ∘ inl) ≃ SphereBouquet n (Fin mₙ) Bouquet≃-gen = isoToEquiv BouquetIso-gen + +-- A 'normal form' for functions of type ⋁Sⁿ → ⋁Sⁿ +normalFormCofibFun : ∀ {n m k : ℕ} + (α : SphereBouquet∙ (suc n) (Fin m) →∙ SphereBouquet∙ (suc n) (Fin k)) + (f : S₊∙ (suc n) →∙ (cofib (fst α) , inl tt)) + → ∃[ f' ∈ S₊∙ (suc n) →∙ SphereBouquet∙ (suc n) (Fin k) ] + (((inr , (λ i → inr (α .snd (~ i))) ∙ sym (push (inl tt))) ∘∙ f') ≡ f) +normalFormCofibFun {n = n} {m} {k} α f = + PT.rec squash₁ + (λ g → TR.rec (isProp→isOfHLevelSuc n squash₁) + (λ gid → ∣ ((λ x → g x .fst) , (cong fst gid)) + , ΣPathP ((λ i x → g x .snd i) + , (lem _ _ _ _ _ _ _ _ _ _ (cong snd gid))) ∣₁) + (isConnectedPath (suc n) + (help (fst f (ptSn (suc n)))) (g (ptSn (suc n))) + ((inl tt) , (((λ i → inr (α .snd (~ i))) + ∙ sym (push (inl tt))) ∙ sym (snd f))) .fst)) + makefun + where + lem : ∀ {ℓ} {A : Type ℓ} (x y : A) (inrgid : x ≡ y) + (z : _) (inrα : y ≡ z) (w : _) (pushtt : z ≡ w) + (t : _) (snf : w ≡ t) (s : x ≡ t) + → Square s ((inrα ∙ pushtt) ∙ snf) inrgid refl + → Square (inrgid ∙ inrα ∙ pushtt) (sym snf) s refl + lem x = J> (J> (J> (J> λ s sq → (sym (lUnit _) ∙ sym (rUnit _)) + ◁ λ i j → (sq ∙ sym (rUnit _) ∙ sym (rUnit _)) j i))) + cool : (x : S₊ (suc n)) → Type + cool x = + Σ[ x' ∈ SphereBouquet (suc n) (Fin k) ] + Σ[ y ∈ ((ptSn (suc n) ≡ x) → inl tt ≡ x') ] + Σ[ feq ∈ inr x' ≡ fst f x ] + ((p : ptSn (suc n) ≡ x) + → Square ((λ i → inr (snd α (~ i))) ∙ sym (push (inl tt))) (snd f) + ((λ i → inr (y p i)) ∙∙ feq ∙∙ cong (fst f) (sym p)) refl) + + inr' : SphereBouquet (suc n) (Fin k) → cofib (fst α) + inr' = inr + + help : isConnectedFun (suc (suc n)) inr' + help = inrConnected _ _ _ + (isConnected→isConnectedFun _ isConnectedSphereBouquet') + + makefun : ∥ ((x : _) + → Σ[ x' ∈ SphereBouquet (suc n) (Fin k) ] inr x' ≡ fst f x) ∥₁ + makefun = sphereToTrunc _ λ x → help (fst f x) .fst + +-- H-space structure +SphereBouquet∙Π : ∀ {ℓ ℓ'} {n : ℕ} {A : Type ℓ} {B : Pointed ℓ'} + → (f g : SphereBouquet∙ (suc n) A →∙ B) + → (SphereBouquet∙ (suc n) A →∙ B) +fst (SphereBouquet∙Π {B = B} f g) (inl x) = pt B +fst (SphereBouquet∙Π {B = B} f g) (inr (a , s)) = + ∙Π ((λ x → fst f (inr (a , x))) , cong (fst f) (sym (push a)) ∙ snd f) + ((λ x → fst g (inr (a , x))) , cong (fst g) (sym (push a)) ∙ snd g) .fst s +fst (SphereBouquet∙Π {B = B} f g) (push a i) = + ∙Π ((λ x → fst f (inr (a , x))) , cong (fst f) (sym (push a)) ∙ snd f) + ((λ x → fst g (inr (a , x))) , cong (fst g) (sym (push a)) ∙ snd g) .snd (~ i) +snd (SphereBouquet∙Π f g) = refl diff --git a/Cubical/HITs/Susp/Base.agda b/Cubical/HITs/Susp/Base.agda index de686c20bb..492847d629 100644 --- a/Cubical/HITs/Susp/Base.agda +++ b/Cubical/HITs/Susp/Base.agda @@ -9,6 +9,7 @@ open import Cubical.Foundations.Pointed open import Cubical.Data.Unit open import Cubical.Data.Bool open import Cubical.Data.Empty +open import Cubical.Data.Nat open import Cubical.HITs.S1 open import Cubical.HITs.S2.Base @@ -49,6 +50,23 @@ suspFun↑ b f north = b suspFun↑ b f south = b suspFun↑ b f (merid a i) = f a i +-- Iterated suspensions +Susp^ : ℕ → Type ℓ → Type ℓ +Susp^ 0 X = X +Susp^ (suc n) X = Susp^ n (Susp X) + +Susp^' : ℕ → Type ℓ → Type ℓ +Susp^' zero x = x +Susp^' (suc n) x = Susp (Susp^' n x) + +Susp^'≡Susp^ : (n : ℕ) {A : Type ℓ} → Susp^' n A ≡ Susp^ n A +Susp^'≡Susp^ zero {A = A} = refl +Susp^'≡Susp^ (suc n) {A = A} = lem n ∙ Susp^'≡Susp^ n {A = Susp A} + where + lem : (n : ℕ) → Susp (Susp^' n A) ≡ Susp^' n (Susp A) + lem zero = refl + lem (suc n) = cong Susp (lem n) + UnitIsoSuspUnit : Iso Unit (Susp Unit) fun UnitIsoSuspUnit _ = north inv UnitIsoSuspUnit _ = tt diff --git a/Cubical/HITs/Wedge/Base.agda b/Cubical/HITs/Wedge/Base.agda index 712fe6b718..9ede908253 100644 --- a/Cubical/HITs/Wedge/Base.agda +++ b/Cubical/HITs/Wedge/Base.agda @@ -25,6 +25,18 @@ A ⋁∙ᵣ B = (A ⋁ B) , (inr (snd B)) ⋁gen∙ : ∀ {ℓ ℓ'} (A : Type ℓ) (B : A → Pointed ℓ') → Pointed (ℓ-max ℓ ℓ') ⋁gen∙ A B = ⋁gen A B , inl tt +-- Projections +module _ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} where + proj⋁ₗ : A ⋁ B → fst A + proj⋁ₗ (inl x) = x + proj⋁ₗ (inr x) = pt A + proj⋁ₗ (push a i) = pt A + + proj⋁ᵣ : A ⋁ B → fst B + proj⋁ᵣ (inl x) = pt B + proj⋁ᵣ (inr x) = x + proj⋁ᵣ (push a i) = pt B + -- Wedge sums of functions _∨→_ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} → (f : A →∙ C) (g : B →∙ C) diff --git a/Cubical/HITs/Wedge/Properties.agda b/Cubical/HITs/Wedge/Properties.agda index 1672d8ecef..f0e00e45be 100644 --- a/Cubical/HITs/Wedge/Properties.agda +++ b/Cubical/HITs/Wedge/Properties.agda @@ -10,6 +10,8 @@ open import Cubical.Foundations.Function open import Cubical.Data.Sigma open import Cubical.Data.Unit open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Nat +open import Cubical.Data.Fin.Inductive.Base open import Cubical.HITs.Wedge.Base open import Cubical.HITs.Susp @@ -17,7 +19,6 @@ open import Cubical.HITs.Pushout open import Cubical.Homotopy.Loopspace - private variable ℓ ℓ' ℓ'' : Level @@ -494,6 +495,50 @@ module _ {A : Type ℓ} {B : Type ℓ'} compPath-filler' (push tt) (λ i → inr (push a i)) (~ j) i Iso.leftInv ⋁gen⊎Iso (push a i) j = push a (i ∧ j) +-- Chacaterisation of cofibres of first projections +-- cofib (Σ[ x ∈ A ] (B x) --fst→ A) ≃ ⋁[ x ∈ X ] (Susp (B x)) +module _ {ℓ ℓ'} {A : Type ℓ} (B : A → Pointed ℓ') + where + cofibFst : Type _ + cofibFst = cofib {A = Σ A (fst ∘ B)} {B = A} fst + + cofibFst→⋁ : cofibFst → ⋁gen A λ a → Susp∙ (fst (B a)) + cofibFst→⋁ (inl x) = inl x + cofibFst→⋁ (inr a) = inr (a , north) + cofibFst→⋁ (push (a , b) i) = (push a ∙ λ i → inr (a , toSusp (B a) b i)) i + + ⋁→cofibFst : ⋁gen A (λ a → Susp∙ (fst (B a))) → cofibFst + ⋁→cofibFst (inl x) = inl x + ⋁→cofibFst (inr (x , north)) = inl tt + ⋁→cofibFst (inr (x , south)) = inr x + ⋁→cofibFst (inr (x , merid a i)) = push (x , a) i + ⋁→cofibFst (push a i) = inl tt + + Iso-cofibFst-⋁ : Iso cofibFst (⋁gen A (λ a → Susp∙ (fst (B a)))) + Iso.fun Iso-cofibFst-⋁ = cofibFst→⋁ + Iso.inv Iso-cofibFst-⋁ = ⋁→cofibFst + Iso.rightInv Iso-cofibFst-⋁ (inl x) = refl + Iso.rightInv Iso-cofibFst-⋁ (inr (x , north)) = push x + Iso.rightInv Iso-cofibFst-⋁ (inr (x , south)) i = inr (x , merid (pt (B x)) i) + Iso.rightInv Iso-cofibFst-⋁ (inr (x , merid a i)) j = + hcomp (λ k → λ {(i = i0) → push x (j ∨ ~ k) + ; (i = i1) → inr (x , merid (pt (B x)) j) + ; (j = i0) → compPath-filler' (push x) + (λ i₁ → inr (x , toSusp (B x) a i₁)) k i + ; (j = i1) → inr (x , merid a i)}) + (inr (x , compPath-filler (merid a) (sym (merid (pt (B x)))) (~ j) i)) + Iso.rightInv Iso-cofibFst-⋁ (push a i) j = push a (i ∧ j) + Iso.leftInv Iso-cofibFst-⋁ (inl x) = refl + Iso.leftInv Iso-cofibFst-⋁ (inr x) = push (x , snd (B x)) + Iso.leftInv Iso-cofibFst-⋁ (push (a , b) i) j = help j i + where + help : Square (cong ⋁→cofibFst ((push a ∙ λ i → inr (a , toSusp (B a) b i)))) + (push (a , b)) refl (push (a , (snd (B a)))) + help = (cong-∙ ⋁→cofibFst (push a) (λ i → inr (a , toSusp (B a) b i)) + ∙ sym (lUnit _) + ∙ cong-∙ (⋁→cofibFst ∘ inr ∘ (a ,_)) (merid b) (sym (merid (snd (B a))))) + ◁ λ i j → compPath-filler (push (a , b)) (sym (push (a , pt (B a)))) (~ i) j + -- f : ⋁ₐ Bₐ → C has cofibre the pushout of cofib (f ∘ inr) ← Σₐ → A module _ {ℓA ℓB ℓC : Level} {A : Type ℓA} {B : A → Pointed ℓB} (C : Pointed ℓC) (f : (⋁gen A B , inl tt) →∙ C) where @@ -647,3 +692,76 @@ module _ (X∙ @ (X , x₀) : Pointed ℓ) (Y∙ @ (Y , y₀) : Pointed ℓ') wh Pushout⋁≃Unit : Pushout (⋁proj₁ X∙ Y∙) (⋁proj₂ X∙ Y∙) ≃ Unit Pushout⋁≃Unit = _ , botPushout + +-- Pushout along projections is contractible +module _ (A : Pointed ℓ) (B : Pointed ℓ') where + private + push-inl : (a : typ A) + → Path (Pushout (proj⋁ₗ {A = A}) proj⋁ᵣ) (inl a) (inr (pt B)) + push-inl x = push (inl x) + + push-inr : (b : typ B) + → Path (Pushout (proj⋁ₗ {B = B}) proj⋁ᵣ) (inl (pt A)) (inr b) + push-inr x = push (inr x) + + push-push : push-inl (pt A) ≡ push-inr (pt B) + push-push i = push (push tt i) + + F : ∀ {ℓ} {A : Type ℓ} {x : A} (y : A) + (push-inl push-inr : x ≡ y) (q : push-inl ≡ push-inr) + → (z : A) (m : x ≡ z) + → Square (push-inl ∙ sym push-inr) m refl m + F y push-inl push-inr q z m = (cong₂ _∙_ q refl ∙ rCancel push-inr) + ◁ λ i j → m (i ∧ j) + + F' : ∀ {ℓ} {A : Type ℓ} {x : A} + → F x refl refl refl _ refl ≡ sym (rUnit refl) + F' = sym (compPath≡compPath' _ _) ∙ sym (rUnit _) ∙ sym (lUnit _) + + H : ∀ {ℓ} {A : Type ℓ} {x : A} (y : A) + (push-inl push-inr : x ≡ y) (q : push-inl ≡ push-inr) + → Cube (λ k j → compPath-filler push-inr (sym push-inl) (~ j) k) + (λ k j → F _ push-inr push-inl (sym q) _ push-inr j k) + (λ i j → x) q + (λ i j → (push-inr ∙ sym push-inl) j) λ i j → push-inr j + H {x = x} = J> (J> (λ k i j → F' {x = x} (~ k) j i)) + + isContrPushout-proj⋁ : isContr (Pushout proj⋁ₗ proj⋁ᵣ) + fst isContrPushout-proj⋁ = inl (pt A) + snd isContrPushout-proj⋁ (inl x) = push-inr (pt B) ∙ sym (push-inl x) + snd isContrPushout-proj⋁ (inr x) = push (inr x) + snd isContrPushout-proj⋁ (push (inl x) i) = + compPath-filler (push-inr (pt B)) (sym (push-inl x)) (~ i) + snd isContrPushout-proj⋁ (push (inr x) i) = + F _ (push-inr (pt B)) (push-inl (pt A)) (sym push-push) _ (push (inr x)) i + snd isContrPushout-proj⋁ (push (push a i) j) k = + H _ (push-inl (pt A)) (push-inr (pt B)) push-push i k j + +open Iso +Iso-⋁genFinSuc-⋁genFin⋁ : ∀ {ℓ} {n : ℕ} {A : Fin (suc n) → Pointed ℓ} + → Iso (⋁gen (Fin (suc n)) A) (⋁gen∙ (Fin n) (A ∘ fsuc) ⋁ A fzero) +fun Iso-⋁genFinSuc-⋁genFin⋁ (inl x) = inl (inl tt) +fun Iso-⋁genFinSuc-⋁genFin⋁ (inr ((zero , w) , t)) = inr t +fun (Iso-⋁genFinSuc-⋁genFin⋁ {n = suc n}) (inr ((suc f , w) , t)) = + inl (inr ((f , w) , t)) +fun Iso-⋁genFinSuc-⋁genFin⋁ (push (zero , w) i) = push tt i +fun (Iso-⋁genFinSuc-⋁genFin⋁ {n = suc n}) (push (suc x , w) i) = + inl (push (x , w) i) +inv Iso-⋁genFinSuc-⋁genFin⋁ (inl (inl x)) = inl tt +inv Iso-⋁genFinSuc-⋁genFin⋁ (inl (inr x)) = inr ((fsuc (fst x)) , (snd x)) +inv Iso-⋁genFinSuc-⋁genFin⋁ (inl (push a i)) = push (fsuc a) i +inv Iso-⋁genFinSuc-⋁genFin⋁ (inr x) = inr (fzero , x) +inv Iso-⋁genFinSuc-⋁genFin⋁ (push a i) = push fzero i +rightInv Iso-⋁genFinSuc-⋁genFin⋁ (inl (inl tt)) i = inl (inl tt) +rightInv (Iso-⋁genFinSuc-⋁genFin⋁ {n = suc n}) (inl (inr ((zero , w) , t))) = refl +rightInv (Iso-⋁genFinSuc-⋁genFin⋁ {n = suc n}) (inl (inr ((suc a , w) , t))) = refl +rightInv (Iso-⋁genFinSuc-⋁genFin⋁ {n = suc n}) (inl (push (zero , w) i)) = refl +rightInv (Iso-⋁genFinSuc-⋁genFin⋁ {n = suc n}) (inl (push (suc a , w) i)) = refl +rightInv Iso-⋁genFinSuc-⋁genFin⋁ (inr x) i = inr x +rightInv Iso-⋁genFinSuc-⋁genFin⋁ (push a i) j = push a i +leftInv Iso-⋁genFinSuc-⋁genFin⋁ (inl x) i = inl tt +leftInv Iso-⋁genFinSuc-⋁genFin⋁ (inr ((zero , tt) , t)) i = inr ((0 , tt) , t) +leftInv (Iso-⋁genFinSuc-⋁genFin⋁ {n = suc n}) (inr ((suc x , w) , t)) i = + inr ((suc x , w) , t) +leftInv Iso-⋁genFinSuc-⋁genFin⋁ (push (zero , w) i) j = push (0 , w) i +leftInv (Iso-⋁genFinSuc-⋁genFin⋁ {n = suc n}) (push (suc x , w) i) = refl diff --git a/Cubical/Homotopy/BlakersMassey.agda b/Cubical/Homotopy/BlakersMassey.agda index 4634ebc4b2..2dd48e2433 100644 --- a/Cubical/Homotopy/BlakersMassey.agda +++ b/Cubical/Homotopy/BlakersMassey.agda @@ -20,6 +20,7 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Transport open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Pointed open import Cubical.Foundations.Equiv open import Cubical.Foundations.Equiv.Properties @@ -35,6 +36,7 @@ open import Cubical.Data.Unit open import Cubical.HITs.Truncation renaming (hLevelTrunc to Trunc) open import Cubical.HITs.Pushout hiding (PushoutGenFib) +open import Cubical.HITs.Wedge open import Cubical.Homotopy.Connected open import Cubical.Homotopy.WedgeConnectivity @@ -807,3 +809,68 @@ module BlakersMassey□ {ℓ ℓ' ℓ'' : Level} (isEquiv→isConnected _ (isoToIsEquiv TotalPathGen×Iso) (n + m)) isConnectedTotalFun) (isEquiv→isConnected _ (isoToIsEquiv (invIso Totalfib×Iso)) (n + m))) + +-- Consequence of Blakers-Massey: connectedness of ⋁↪ (wedge inclusion) +isConnected⋁↪ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} {n m : ℕ} + → isConnected (suc (suc n)) (fst A) + → isConnected (suc (suc m)) (fst B) + → isConnectedFun (suc m + suc n) (⋁↪ {A = A} {B}) +isConnected⋁↪ {A = A} {B} {n} {m} cA cB = + subst (isConnectedFun (suc m + suc n)) (sym main) + (isConnectedComp _ _ _ + (isEquiv→isConnected _ (isoToIsEquiv lem) _) + isConnected-toPullback) + where + isConnectedFoldL : (cB : isConnected (suc (suc m)) (fst B)) + → isConnectedFun (suc (suc m)) proj⋁ₗ + isConnectedFoldL cB = subst (isConnectedFun (suc (suc m))) + (funExt (λ { (inl x) → refl + ; (inr x) → refl + ; (push a i) → refl})) + conf + where + f' : A ⋁ B → fst A + f' = Iso.fun cofibInr-⋁ ∘ inr + + conf : isConnectedFun (suc (suc m)) f' + conf = isConnectedComp (Iso.fun cofibInr-⋁) inr + (suc (suc m)) + (isEquiv→isConnected _ (isoToIsEquiv cofibInr-⋁) (suc (suc m))) + (inrConnected (suc (suc m)) _ _ + (isConnected→isConnectedFun (suc (suc m)) cB)) + + isConnectedFoldR : (cA : isConnected (suc (suc n)) (fst A)) + → isConnectedFun (suc (suc n)) proj⋁ᵣ + isConnectedFoldR cA = + subst (isConnectedFun (suc (suc n))) + (funExt (λ { (inl x) → refl ; (inr x) → refl ; (push a i) → refl})) + conf + where + f' : A ⋁ B → fst B + f' = (Iso.fun cofibInr-⋁ ∘ inr) ∘ fst (symPushout _ _) + + conf : isConnectedFun (suc (suc n)) f' + conf = + isConnectedComp (Iso.fun cofibInr-⋁ ∘ inr) (fst (symPushout _ _)) + (suc (suc n)) + (isConnectedComp _ inr (suc (suc n)) + (isEquiv→isConnected _ (isoToIsEquiv cofibInr-⋁) (suc (suc n))) + (inrConnected (suc (suc n)) _ _ + (isConnected→isConnectedFun (suc (suc n)) cA))) + (isEquiv→isConnected _ (snd (symPushout _ _)) (suc (suc n))) + + open BlakersMassey□ {A = A ⋁ B} {B = fst A} {C = fst B} + proj⋁ₗ proj⋁ᵣ (suc m) (suc n) + (isConnectedFoldL cB) (isConnectedFoldR cA) + + lem : Iso (Σ (fst A × fst B) PushoutPath×) (fst A × fst B) + lem = compIso + (Σ-cong-iso-snd (λ _ → + equivToIso (isContr→≃Unit + (isOfHLevelPath 0 (isContrPushout-proj⋁ A B) _ _)))) + rUnit×Iso + + main : ⋁↪ ≡ Iso.fun lem ∘ toPullback + main = funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a i) → refl} diff --git a/Cubical/Homotopy/Connected.agda b/Cubical/Homotopy/Connected.agda index 42d474b7ca..998269ff32 100644 --- a/Cubical/Homotopy/Connected.agda +++ b/Cubical/Homotopy/Connected.agda @@ -941,3 +941,14 @@ module _ {ℓ ℓ' ℓ'' : Level} {A : Type ℓ} {B : A → Type ℓ'} {C : A FunConnected→TotalFunConnected n con r = isConnectedRetractFromIso n (Iso-fibTotalFun-fibFun r) (con (fst r) (snd r)) + + +isConnectedCofib : ∀ {ℓ} {A B : Type ℓ} (n : ℕ) {f : A → B} + → isConnectedFun (suc n) f → isConnected (suc (suc n)) (cofib f) +isConnectedCofib n {f = f} cf = + isConnectedPoint2 (suc n) (inl tt) (inlConnected (suc n) (λ _ → tt) f cf) + +connectedFunPresConnected : ∀ {ℓ} {A B : Type ℓ} (n : ℕ) {f : A → B} + → isConnected n B → isConnectedFun n f → isConnected n A +connectedFunPresConnected n {f = f} conB conf = + isOfHLevelRetractFromIso 0 (connectedTruncIso n f conf) conB diff --git a/Cubical/Homotopy/Group/Base.agda b/Cubical/Homotopy/Group/Base.agda index 879da17fff..ab2cd2092c 100644 --- a/Cubical/Homotopy/Group/Base.agda +++ b/Cubical/Homotopy/Group/Base.agda @@ -642,6 +642,40 @@ snd (π'Gr≅πGr n A) = makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) λ p q i → ∣ IsoSphereMapΩ-pres∙Π n p q i ∣₂) +-- Proof that π'Gr preserves universe lifts +π'GrLiftIso : ∀ {ℓ} (ℓ' : Level) {A : Pointed ℓ} (n : ℕ) + → GroupIso (π'Gr n (Lift∙ {j = ℓ'} A)) (π'Gr n A) +fun (fst (π'GrLiftIso ℓ' n)) = + sMap λ f → (λ x → lower (fst f x)) + , (cong lower (snd f)) +inv (fst (π'GrLiftIso ℓ' n)) = + sMap λ f → (λ x → lift (fst f x)) + , (cong lift (snd f)) +rightInv (fst (π'GrLiftIso ℓ' n)) = + sElim (λ _ → isSetPathImplicit) λ f → refl +leftInv (fst (π'GrLiftIso ℓ' n)) = + sElim (λ _ → isSetPathImplicit) λ f → refl +snd (π'GrLiftIso ℓ' zero) = + makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) + λ f g → cong ∣_∣₂ (ΣPathP ((funExt + λ { base → refl + ; (loop i) j → (cong-∙ lower (Ω→ f .fst loop) (Ω→ g .fst loop) + ∙ cong₂ _∙_ + (cong-∙∙ lower (sym (snd f)) (cong (fst f) loop) (snd f)) + (cong-∙∙ lower (sym (snd g)) (cong (fst g) loop) (snd g))) j i}) + , refl))) +snd (π'GrLiftIso ℓ' {A = A} (suc n)) = + makeIsGroupHom (sElim2 (λ _ _ → isSetPathImplicit) + λ f g → cong ∣_∣₂ (ΣPathP ((funExt ( + λ { north → refl + ; south → refl + ; (merid a i) j + → (cong-∙ lower (Ω→ f .fst (σS a)) (Ω→ g .fst (σS a)) + ∙ cong₂ _∙_ + (cong-∙∙ lower (sym (snd f)) (cong (fst f) (σS a)) (snd f)) + (cong-∙∙ lower (sym (snd g)) (cong (fst g) (σS a)) (snd g))) j i})) + , refl))) + {- Proof of πₙ(ΩA) = πₙ₊₁(A) -} Iso-πΩ-π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) → Iso (π n (Ω A)) (π (suc n) A) @@ -958,11 +992,11 @@ v f∘_ v → π' (suc n) A → π' (suc n) B π'∘∙fun n f = sMap (f ∘∙_) -GroupHomπ≅π'PathP : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') (n : ℕ) - → GroupHom (πGr n A) (πGr n B) ≡ GroupHom (π'Gr n A) (π'Gr n B) -GroupHomπ≅π'PathP A B n i = +GroupHomπ≅π'PathP : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') (n m : ℕ) + → GroupHom (πGr n A) (πGr m B) ≡ GroupHom (π'Gr n A) (π'Gr m B) +GroupHomπ≅π'PathP A B n m i = GroupHom (fst (GroupPath _ _) (GroupIso→GroupEquiv (π'Gr≅πGr n A)) (~ i)) - (fst (GroupPath _ _) (GroupIso→GroupEquiv (π'Gr≅πGr n B)) (~ i)) + (fst (GroupPath _ _) (GroupIso→GroupEquiv (π'Gr≅πGr m B)) (~ i)) πFun : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) → π (suc n) A → π (suc n) B @@ -979,7 +1013,7 @@ snd (πHom n f) = π'∘∙Hom' : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) → GroupHom (π'Gr n A) (π'Gr n B) π'∘∙Hom' {A = A} {B = B} n f = - transport (λ i → GroupHomπ≅π'PathP A B n i) + transport (λ i → GroupHomπ≅π'PathP A B n n i) (πHom n f) π'∘∙Hom'≡π'∘∙fun : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} @@ -1007,9 +1041,17 @@ snd (π'∘∙Hom {A = A} {B = B} n f) = isHom∘∙ (π'Gr n B .snd)) (π'∘∙Hom' n f .snd) +GroupHomπ≅π'PathP-hom : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + (n : ℕ) (f : A →∙ B) + → PathP (λ i → GroupHomπ≅π'PathP A B n n i) (πHom n f) (π'∘∙Hom n f) +GroupHomπ≅π'PathP-hom {A = A} {B = B} n f = + (λ j → transp (λ i → GroupHomπ≅π'PathP A B n n (i ∧ j)) (~ j) + (πHom n f)) + ▷ Σ≡Prop (λ _ → isPropIsGroupHom _ _) (π'∘∙Hom'≡π'∘∙fun n f) + -- post composition with an equivalence induces an -- isomorphism of homotopy groups -π'eqFun : ∀ {ℓ} {A : Pointed ℓ} {B : Pointed ℓ} (n : ℕ) +π'eqFun : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) → A ≃∙ B → (π' (suc n) A) → π' (suc n) B π'eqFun n p = π'∘∙fun n (≃∙map p) @@ -1027,39 +1069,70 @@ invEquiv∙idEquiv∙≡idEquiv : ∀ {ℓ} {A : Pointed ℓ} invEquiv∙idEquiv∙≡idEquiv = ΣPathP ((Σ≡Prop (λ _ → isPropIsEquiv _) refl) , (sym (lUnit refl))) π'eqFunIsEquiv : - ∀ {ℓ} {A : Pointed ℓ} {B : Pointed ℓ} (n : ℕ) + ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) → (e : A ≃∙ B) → isEquiv (π'eqFun n e) -π'eqFunIsEquiv {B = B} n = - Equiv∙J (λ A e → isEquiv (π'eqFun n e)) - (subst isEquiv (sym (π'eqFun-idEquiv n)) +π'eqFunIsEquiv {ℓ = ℓ} {ℓ'} {A} {B} n e = + subst isEquiv + (funExt (sElim (λ _ → isSetPathImplicit) + (λ f → cong ∣_∣₂ + (ΣPathP (refl + , (cong-∙ lower (cong (lift ∘ (fst (fst e))) (snd f)) _)))))) + (πA≃πB .snd) + where + e' : Lift∙ {j = ℓ'} A ≃∙ Lift∙ {j = ℓ} B + fst e' = + compEquiv (invEquiv LiftEquiv) (compEquiv (fst e) LiftEquiv) + snd e' = cong lift (snd e) + + main : ∀ {ℓ} {A B : Pointed ℓ} (n : ℕ) + → (e : A ≃∙ B) + → isEquiv (π'eqFun n e) + main {B = B} n = + Equiv∙J (λ A e → isEquiv (π'eqFun n e)) + (subst isEquiv (sym (π'eqFun-idEquiv n)) (idIsEquiv (π' (suc n) B))) -π'eqFunIsHom : ∀ {ℓ} {A B : Pointed ℓ}(n : ℕ) + πA≃πB : π' (suc n) A ≃ π' (suc n) B + πA≃πB = + compEquiv (invEquiv (isoToEquiv (fst (π'GrLiftIso _ n)))) + (compEquiv (_ , main n e') + (isoToEquiv (fst (π'GrLiftIso _ n)))) + +π'eqFunIsHom : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) → (e : A ≃∙ B) → IsGroupHom (π'Gr n A .snd) (π'eqFun n e) (π'Gr n B .snd) -π'eqFunIsHom {B = B} n = - Equiv∙J (λ A e → IsGroupHom (π'Gr n A .snd) (π'eqFun n e) (π'Gr n B .snd)) +π'eqFunIsHom {ℓ = ℓ} {ℓ'} {A} {B} n e = + subst (λ ϕ → IsGroupHom (π'Gr n A .snd) + ϕ (π'Gr n B .snd)) + (funExt (sElim (λ _ → isSetPathImplicit) + (λ f → cong ∣_∣₂ (ΣPathP + (refl + , cong-∙ lower (cong (lift ∘ (fst (fst e))) (snd f)) _))))) + (compGroupHom + (GroupIso→GroupHom (invGroupIso (π'GrLiftIso _ n))) + (compGroupHom (_ , main n e') + (GroupIso→GroupHom (π'GrLiftIso _ n))) .snd) + where + e' : Lift∙ {j = ℓ'} A ≃∙ Lift∙ {j = ℓ} B + fst e' = + compEquiv (invEquiv LiftEquiv) (compEquiv (fst e) LiftEquiv) + snd e' = cong lift (snd e) + + main : ∀ {ℓ} {A B : Pointed ℓ} (n : ℕ) + → (e : A ≃∙ B) + → IsGroupHom (π'Gr n A .snd) (π'eqFun n e) + (π'Gr n B .snd) + main {B = B} n = + Equiv∙J (λ A e → IsGroupHom (π'Gr n A .snd) (π'eqFun n e) (π'Gr n B .snd)) (subst (λ x → IsGroupHom (π'Gr n B .snd) x (π'Gr n B .snd)) (sym (π'eqFun-idEquiv n)) (makeIsGroupHom λ _ _ → refl)) -π'GrIso : ∀ {ℓ} {A : Pointed ℓ} {B : Pointed ℓ} (n : ℕ) - → A ≃∙ B - → GroupIso (π'Gr n A) (π'Gr n B) -fun (fst (π'GrIso n e)) = π'eqFun n e -inv (fst (π'GrIso n e)) = π'eqFun n (invEquiv∙ e) -rightInv (fst (π'GrIso {B = B} n e)) = - Equiv∙J (λ A e → (f : _) → π'eqFun n e (π'eqFun n (invEquiv∙ e) f) ≡ f) - (λ f → (λ i → π'eqFun-idEquiv n i (π'eqFun n (invEquiv∙idEquiv∙≡idEquiv i) f)) - ∙ funExt⁻ (π'eqFun-idEquiv n) f) - e -leftInv (fst (π'GrIso n e)) = - Equiv∙J (λ A e → (f : _) → π'eqFun n (invEquiv∙ e) (π'eqFun n e f) ≡ f) - (λ f → (λ i → π'eqFun n (invEquiv∙idEquiv∙≡idEquiv i) (π'eqFun-idEquiv n i f)) - ∙ funExt⁻ (π'eqFun-idEquiv n) f) - e +π'GrIso : {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) + (e : A ≃∙ B) → GroupIso (π'Gr n A) (π'Gr n B) +fst (π'GrIso n e) = setTruncIso (pre∘∙equiv e) snd (π'GrIso n e) = π'eqFunIsHom n e π'Iso : ∀ {ℓ} {A : Pointed ℓ} {B : Pointed ℓ} (n : ℕ) @@ -1089,3 +1162,49 @@ is-set (isSemigroup (isMonoid (isGroup (snd (hGroupoidπ₁ A a))))) = snd A a a ·IdL (isMonoid (isGroup (snd (hGroupoidπ₁ A a)))) = sym ∘ lUnit ·InvR (isGroup (snd (hGroupoidπ₁ A a))) = rCancel ·InvL (isGroup (snd (hGroupoidπ₁ A a))) = lCancel + +-- Adjunction +sphereFunIso : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) + → Iso (S₊∙ n →∙ (Path (fst A) (pt A) (pt A) , refl)) (S₊∙ (suc n) →∙ A) +sphereFunIso zero = compIso IsoBool→∙ (invIso (IsoSphereMapΩ 1)) +sphereFunIso (suc n) = ΩSuspAdjointIso + +-- +∙Π∘∙ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + (n : ℕ) (f g : S₊∙ (suc n) →∙ A) (h : A →∙ B) + → h ∘∙ ∙Π f g ≡ ∙Π (h ∘∙ f) (h ∘∙ g) +∙Π∘∙ {A = A} n f g h = + cong (h ∘∙_) (cong₂ ∙Π (sym (Iso.rightInv (sphereFunIso n) f)) + (sym (Iso.rightInv (sphereFunIso n) g))) + ∙∙ lem2 n (Iso.inv (sphereFunIso n) f) (Iso.inv (sphereFunIso n) g) + ∙∙ cong₂ (λ f g → ∙Π (h ∘∙ f) (h ∘∙ g)) + (Iso.rightInv (sphereFunIso n) f) + (Iso.rightInv (sphereFunIso n) g) + where + lem : ∀ {ℓ} {A : Type ℓ} {x y : A} (p : x ≡ y) → Square p refl (refl ∙ p) refl + lem p = lUnit p ◁ λ i j → (refl ∙ p) (i ∨ j) + + mainEq : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) (a : A) (b : B) + (fp : f a ≡ b) (l1 l2 : a ≡ a) + → Square (cong f ((l1 ∙ refl) ∙ (l2 ∙ refl))) + ((sym (refl ∙ fp) ∙∙ cong f l1 ∙∙ (refl ∙ fp)) + ∙ (sym (refl ∙ fp) ∙∙ cong f l2 ∙∙ (refl ∙ fp))) + fp fp + mainEq f a = J> λ l1 l2 → cong-∙ f _ _ + ∙ cong₂ _∙_ (cong-∙ f l1 refl ∙ cong₃ _∙∙_∙∙_ (rUnit refl) refl (rUnit refl)) + (cong-∙ f l2 refl ∙ cong₃ _∙∙_∙∙_ (rUnit refl) refl (rUnit refl)) + + lem2 : (n : ℕ) (f g : S₊∙ n →∙ Ω A) + → (h ∘∙ ∙Π (Iso.fun (sphereFunIso n) f) (Iso.fun (sphereFunIso n) g)) + ≡ ∙Π (h ∘∙ Iso.fun (sphereFunIso n) f) (h ∘∙ Iso.fun (sphereFunIso n) g) + fst (lem2 zero f g i) base = snd h i + fst (lem2 zero f g i) (loop i₁) = + mainEq (fst h) _ _ (snd h) (fst f false) (fst g false) i i₁ + fst (lem2 (suc n) f g i) north = snd h i + fst (lem2 (suc n) f g i) south = snd h i + fst (lem2 (suc n) f g i) (merid a i₁) = + mainEq (fst h) _ _ (snd h) + (cong (Iso.fun (sphereFunIso (suc n)) f .fst) (σS a)) + (cong (Iso.fun (sphereFunIso (suc n)) g .fst) (σS a)) i i₁ + snd (lem2 zero f g i) j = lem (snd h) j i + snd (lem2 (suc n) f g i) j = lem (snd h) j i diff --git a/Cubical/Homotopy/Group/LES.agda b/Cubical/Homotopy/Group/LES.agda index 929e2cef7b..09465ae854 100644 --- a/Cubical/Homotopy/Group/LES.agda +++ b/Cubical/Homotopy/Group/LES.agda @@ -644,14 +644,14 @@ private of homotopy groups defined using (Sⁿ →∙ A) -} π∘∙A→B-PathP : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) - → PathP (λ i → GroupHomπ≅π'PathP A B n i) + → PathP (λ i → GroupHomπ≅π'PathP A B n n i) (πLES.A→B f n) (π'∘∙Hom n f) π∘∙A→B-PathP n f = toPathP (Σ≡Prop (λ _ → isPropIsGroupHom _ _) (π'∘∙Hom'≡π'∘∙fun n f)) π∘∙fib→A-PathP : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (n : ℕ) (f : A →∙ B) - → PathP (λ i → GroupHomπ≅π'PathP (ΩLES.fibf f) A n i) + → PathP (λ i → GroupHomπ≅π'PathP (ΩLES.fibf f) A n n i) (πLES.fib→A f n) (π'∘∙Hom n (fst , refl)) π∘∙fib→A-PathP {A = A} {B = B} n f = @@ -666,3 +666,73 @@ of homotopy groups defined using (Sⁿ →∙ A) -} where lem : πLES.fib→A f n .fst ≡ sMap (Ω^→ (suc n) (fst , refl) .fst) lem = cong sMap (cong fst (Ω^fibf→A≡ (suc n) f)) + + + +module π'LES {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) where + module M = πLES f + fib : Pointed _ + fib = (fiber (fst f) (pt B)) , (pt A , snd f) + + fib→A : (n : ℕ) → GroupHom (π'Gr n fib) (π'Gr n A) + fib→A n = π'∘∙Hom n (fst , refl) + + A→B : (n : ℕ) → GroupHom (π'Gr n A) (π'Gr n B) + A→B n = π'∘∙Hom n f + + -- todo: improve + B→fib : (n : ℕ) → GroupHom (π'Gr (suc n) B) (π'Gr n fib) + B→fib n = transport (GroupHomπ≅π'PathP B fib (suc n) n) (M.B→fib n) + + private + P : (n : ℕ) → PathP (λ i → GroupHomπ≅π'PathP B fib (suc n) n i) + (M.B→fib n) (B→fib n) + P n = toPathP refl + + Ker-A→B⊂Im-fib→A : (n : ℕ) (x : π' (suc n) A) + → isInKer (A→B n) x + → isInIm (fib→A n) x + Ker-A→B⊂Im-fib→A n = + transport (λ i → (x : _) → isInKer (π∘∙A→B-PathP n f i) x + → isInIm (π∘∙fib→A-PathP n f i) x) + (M.Ker-A→B⊂Im-fib→A n) + + Im-fib→A⊂Ker-A→B : (n : ℕ) (x : π' (suc n) A) + → isInIm (fib→A n) x + → isInKer (A→B n) x + Im-fib→A⊂Ker-A→B n = + transport (λ i → (x : _) → isInIm (π∘∙fib→A-PathP n f i) x + → isInKer (π∘∙A→B-PathP n f i) x) + (M.Im-fib→A⊂Ker-A→B n) + + Ker-fib→A⊂Im-B→fib : (n : ℕ) (x : π' (suc n) fib) + → isInKer (fib→A n) x + → isInIm (B→fib n) x + Ker-fib→A⊂Im-B→fib n = + transport (λ i → (x : _) → isInKer (π∘∙fib→A-PathP n f i) x + → isInIm (P n i) x) + (M.Ker-fib→A⊂Im-B→fib n) + + Im-B→fib⊂Ker-fib→A : (n : ℕ) (x : π' (suc n) fib) + → isInIm (B→fib n) x + → isInKer (fib→A n) x + Im-B→fib⊂Ker-fib→A n = + transport (λ i → (x : _) → isInIm (P n i) x + → isInKer (π∘∙fib→A-PathP n f i) x) + (M.Im-B→fib⊂Ker-fib→A n) + + Im-A→B⊂Ker-B→fib : (n : ℕ) (x : π' (suc (suc n)) B) + → isInIm (A→B (suc n)) x + → isInKer (B→fib n) x + Im-A→B⊂Ker-B→fib n = + transport (λ i → (x : _) → isInIm (π∘∙A→B-PathP (suc n) f i) x + → isInKer (P n i) x) + (M.Im-A→B⊂Ker-B→fib n) + + Ker-B→fib⊂Im-A→B : (n : ℕ) (x : π' (suc (suc n)) B) + → isInKer (B→fib n) x + → isInIm (A→B (suc n)) x + Ker-B→fib⊂Im-A→B n = + transport (λ i → (x : _) → isInKer (P n i) x + → isInIm (π∘∙A→B-PathP (suc n) f i) x) + (M.Ker-B→fib⊂Im-A→B n) diff --git a/Cubical/Homotopy/Group/Pi4S3/BrunerieNumber.agda b/Cubical/Homotopy/Group/Pi4S3/BrunerieNumber.agda index d3add9b0b9..4d3b34727d 100644 --- a/Cubical/Homotopy/Group/Pi4S3/BrunerieNumber.agda +++ b/Cubical/Homotopy/Group/Pi4S3/BrunerieNumber.agda @@ -289,7 +289,7 @@ isSurjective-π₃S³→π₃TotalPushoutPath× = (((isConnectedΩ^→ 3 3 (S³→TotalPushoutPath× , refl) isConnected-toPullback) p) .fst) - transportLem : PathP (λ i → GroupHomπ≅π'PathP (S₊∙ 3) TotalPushoutPath×∙ 2 i) + transportLem : PathP (λ i → GroupHomπ≅π'PathP (S₊∙ 3) TotalPushoutPath×∙ 2 2 i) π₃S³→π₃TotalPushoutPath× π₃S³→π₃P transportLem = toPathP (Σ≡Prop (λ _ → isPropIsGroupHom _ _) diff --git a/Cubical/Homotopy/Group/PiAbCofibFinSphereBouquetMap.agda b/Cubical/Homotopy/Group/PiAbCofibFinSphereBouquetMap.agda new file mode 100644 index 0000000000..c28db47930 --- /dev/null +++ b/Cubical/Homotopy/Group/PiAbCofibFinSphereBouquetMap.agda @@ -0,0 +1,954 @@ +{-# OPTIONS --lossy-unification #-} +{- +This file computes πₙᵃᵇ(cofib α) and α : ⋁Sⁿ →∙ ⋁Sⁿ +-} +module Cubical.Homotopy.Group.PiAbCofibFinSphereBouquetMap where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.GroupoidLaws as GLaws +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Int renaming (_·_ to _·ℤ_) +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_ ; _·_ to _·ℕ_) +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Fin.Inductive +open import Cubical.Data.Nat.Order.Inductive +open import Cubical.Data.Sigma + +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.Sn.Degree +open import Cubical.HITs.SphereBouquet +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.SphereBouquet.Degree +open import Cubical.HITs.Pushout as PO +open import Cubical.HITs.Bouquet as Bouq +open import Cubical.HITs.FreeGroup as FG +open import Cubical.HITs.SetQuotients.Base renaming (_/_ to _/s_) +open import Cubical.HITs.SetQuotients.Properties as SQ +open import Cubical.HITs.AbPath + +open import Cubical.Relation.Nullary hiding (⟪_⟫) + +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup as FAB +open import Cubical.Algebra.Group.QuotientGroup +open import Cubical.Algebra.Group.Abelianization.Base +open import Cubical.Algebra.Group.Abelianization.Properties as Abi +open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.IsomorphismTheorems + +open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.Group.PiSphereBouquet +open import Cubical.Homotopy.Group.PiCofibFinSphereBouquetMap + +open Iso renaming (inv to inv') + +-- Part 1: Computation of πᵃᵇ₁(cofib α) for α +-- α : ⋁S¹ →∙ ⋁S¹ +-- Preliminary definitions: +module _ {m k : ℕ} (f : Fin m → FreeGroup (Fin k)) where + gens→finSphereBouquetMap : SphereBouquet 1 (Fin m) → SphereBouquet 1 (Fin k) + gens→finSphereBouquetMap = fun Iso-Bouquet→∙-SphereBouquet₁→∙ + (inv' CharacFinBouquetFunIso f) .fst + + AbFreeGroup≅ℤ[] : (m : _) + → AbGroupIso (AbelianizationAbGroup (freeGroupGroup (Fin m))) + (ℤ[Fin m ]) + AbFreeGroup≅ℤ[] m = compGroupIso GroupIso-AbelienizeFreeGroup→FreeAbGroup + (invGroupIso (ℤFin≅FreeAbGroup m)) + + AbFreeGroup→ℤ[] : (m : _) → + AbGroupHom (AbelianizationAbGroup (freeGroupGroup (Fin m))) + (ℤ[Fin m ]) + AbFreeGroup→ℤ[] m = GroupIso→GroupHom (AbFreeGroup≅ℤ[] m) + + + bouquetDegree-AbFreeGroup→ℤ[] : (x : _) + → fst (bouquetDegree gens→finSphereBouquetMap) (AbFreeGroup→ℤ[] m .fst x) + ≡ AbFreeGroup→ℤ[] k .fst (AbelianizationFun (FG.rec f) .fst x) + bouquetDegree-AbFreeGroup→ℤ[] = + Abi.elimProp _ (λ _ → isSetΠ (λ _ → isSetℤ) _ _) + (funExt⁻ (cong fst (help main))) + where + help = freeGroupHom≡ + {f = compGroupHom + (compGroupHom (AbelianizationGroupStructure.ηAsGroupHom _) + (AbFreeGroup→ℤ[] m)) + (bouquetDegree gens→finSphereBouquetMap)} + {g = compGroupHom + (compGroupHom (AbelianizationGroupStructure.ηAsGroupHom _) + (AbelianizationFun (FG.rec f))) (AbFreeGroup→ℤ[] k)} + + main : (a : _) → bouquetDegree gens→finSphereBouquetMap .fst + (AbFreeGroup→ℤ[] m .fst (η (η a))) + ≡ AbFreeGroup→ℤ[] k .fst + (fst (AbelianizationFun (FG.rec f)) (η (η a))) + main a = funExt λ s + → sumFin-choose _ _ (λ _ → refl) +Comm _ _ a + (characdiag s) + λ x p → cong₂ _·ℤ_ (charac¬ x p) refl + where + charac¬ : (x' : Fin m) → ¬ x' ≡ a + → fst (AbFreeGroup→ℤ[] m) (η (η a)) x' ≡ pos 0 + charac¬ x' p with (fst a ≟ᵗ fst x') + ... | lt x = refl + ... | eq x = ⊥.rec (p (Σ≡Prop (λ _ → isProp<ᵗ) (sym x))) + ... | gt x = refl + + lem : ℤFinGenerator a a ≡ 1 + lem with (fst a ≟ᵗ fst a) + ... | lt x = ⊥.rec (¬m<ᵗm x) + ... | eq x = refl + ... | gt x = ⊥.rec (¬m<ᵗm x) + + aux : (x : FreeGroup (Fin k)) (y : S¹) + → fst (SphereBouquet∙ 1 (Fin k)) + aux b base = inl tt + aux b (loop i) = + Bouquet→SphereBouquet (inv' Iso-ΩFinBouquet-FreeGroup b i) + auxId : (x : _) → gens→finSphereBouquetMap (inr (a , x)) ≡ aux (f a) x + auxId base = refl + auxId (loop i) = refl + + characdiagMain : (w : _) + → (λ s → degree (suc zero) (λ x → pickPetal s (aux w x))) + ≡ fst (AbFreeGroup→ℤ[] k) (η w) + characdiagMain = + funExt⁻ (cong fst (freeGroupHom≡ {Group = AbGroup→Group ℤ[Fin k ]} + {f = _ , makeIsGroupHom λ f g + → funExt λ t → cong (degree 1) + (funExt (λ { base → refl + ; (loop i) j → lem2 t f g j i})) + ∙ (degreeHom {n = 0} + ((λ x → pickPetal t (aux f x)) , refl) + ((λ x → pickPetal t (aux g x)) , refl))} + {g = _ , compGroupHom (AbelianizationGroupStructure.ηAsGroupHom _) + (AbFreeGroup→ℤ[] k) .snd} + λ s → funExt λ w → final s w ∙ ℤFinGeneratorComm w s)) + where + final : (s w : Fin k) → degree 1 (λ x → pickPetal w (aux (η s) x)) + ≡ ℤFinGenerator w s + final s w with (fst w ≟ᵗ fst s) + ... | lt x = refl + ... | eq x = refl + ... | gt x = refl + + lem2 : (t : _) (f g : FreeGroup (Fin k)) + → cong (pickPetal t ∘ aux (f FG.· g)) loop + ≡ (cong (pickPetal t ∘ aux f) loop ∙ refl) + ∙ cong (pickPetal t ∘ aux g) loop ∙ refl + lem2 t f g = + cong (cong (pickPetal t ∘ Bouquet→SphereBouquet)) + (invIso-ΩFinBouquet-FreeGroupPresStr f g) + ∙ cong-∙ (pickPetal t ∘ Bouquet→SphereBouquet) + (inv' Iso-ΩFinBouquet-FreeGroup f) + (inv' Iso-ΩFinBouquet-FreeGroup g) + ∙ cong₂ _∙_ (rUnit _) (rUnit _) + + characdiag : (s : _) → + ℤFinGenerator a a + ·ℤ degree 1 (λ x → pickPetal s (gens→finSphereBouquetMap (inr (a , x)))) + ≡ fst (AbFreeGroup→ℤ[] k) + (fst (AbelianizationFun (FG.rec f)) (η (η a))) s + characdiag s = cong₂ _·ℤ_ lem refl + ∙ cong (degree (suc zero)) + (funExt λ x → cong (pickPetal s) (auxId x)) + ∙ funExt⁻ (characdiagMain (f a)) s + +-- Part 2: Computation of πᵃᵇ₁(cofib α) for α : ⋁S¹ →∙ ⋁S¹ +-- with α 'strictified' (induced by a set of generators in +-- FreeGroup (Fin k)) +module _ {m k : ℕ} (α' : Fin m → FreeGroup (Fin k)) where + private + α : Bouquet∙ (Fin m) →∙ Bouquet∙ (Fin k) + α = inv' CharacFinBouquetFunIso α' + + αSphereBouquet : SphereBouquet∙ (suc zero) (Fin m) + →∙ SphereBouquet∙ (suc zero) (Fin k) + αSphereBouquet = fun Iso-Bouquet→∙-SphereBouquet₁→∙ α + + _·f_ : ∀ {ℓ} {A : Type ℓ} → FreeGroup A → FreeGroup A → FreeGroup A + _·f_ = FG._·_ + + Bouquet→CofibFinBouquetMap : Bouquet (Fin k) → cofib (fst α) + Bouquet→CofibFinBouquetMap = inr + + Freeᵃᵇ/ImFinBouquetMap : AbGroup ℓ-zero + Freeᵃᵇ/ImFinBouquetMap = + AbelianizationAbGroup (freeGroupGroup (Fin k)) + /Im AbelianizationFun (FG.rec α') + + FinBouquetCode : Bouquet (Fin k) → Type + FinBouquetCode base = Freeᵃᵇ/ImFinBouquetMap .fst + FinBouquetCode (loop x i) = + ua (isoToEquiv (·GroupAutomorphismR (AbGroup→Group (Freeᵃᵇ/ImFinBouquetMap)) + [ η (η x) ])) i + + substFinBouquetCode : (p : _) (x : _) + → subst FinBouquetCode (inv' Iso-ΩFinBouquet-FreeGroup p) [ η x ] + ≡ [ η (x FG.· p) ] + substFinBouquetCode = FG.elimProp (λ _ → isPropΠ (λ _ → squash/ _ _)) + (λ a x i → [ η (transportRefl x i FG.· η (transportRefl a i)) ]) + (λ g1 g2 p q x + → cong (λ P → subst FinBouquetCode P [ η x ]) + (invIso-ΩFinBouquet-FreeGroupPresStr g1 g2) + ∙ substComposite FinBouquetCode + (inv' Iso-ΩFinBouquet-FreeGroup g1) + (inv' Iso-ΩFinBouquet-FreeGroup g2) [ η x ] + ∙ cong (subst FinBouquetCode (inv' Iso-ΩFinBouquet-FreeGroup g2)) + (p x) + ∙ q (x FG.· g1) + ∙ λ i → [ η (FG.assoc x g1 g2 (~ i)) ]) + (λ x i → [ η ((transportRefl x ∙ FG.idr x) i) ]) + λ g p x + → cong (λ P → subst FinBouquetCode P [ η x ]) + (invIso-ΩFinBouquet-FreeGroupPresInv g) + ∙ cong (subst FinBouquetCode (sym (inv' Iso-ΩFinBouquet-FreeGroup g))) + (λ i → [ η ((FG.idr x + ∙ cong₂ FG._·_ refl (sym (FG.invl g)) + ∙ (FG.assoc x (inv g) g)) i) ]) + ∙ sym (cong (subst FinBouquetCode (sym (inv' Iso-ΩFinBouquet-FreeGroup g))) + (p (x FG.· inv g))) + ∙ subst⁻Subst FinBouquetCode (inv' Iso-ΩFinBouquet-FreeGroup g ) + [ η (x FG.· inv g) ] + + CofibFinBoquetFunCode : cofib (fst α) → Type + CofibFinBoquetFunCode (inl x) = Freeᵃᵇ/ImFinBouquetMap .fst + CofibFinBoquetFunCode (inr x) = FinBouquetCode x + CofibFinBoquetFunCode (push base i) = Freeᵃᵇ/ImFinBouquetMap .fst + CofibFinBoquetFunCode (push (loop x j) i) = lem i j + where + open AbelianizationGroupStructure (freeGroupGroup (Fin k)) + lem : refl ≡ cong (FinBouquetCode) (inv' Iso-ΩFinBouquet-FreeGroup (α' x)) + lem = sym uaIdEquiv + ∙ cong ua (Σ≡Prop (λ _ → isPropIsEquiv _) + (funExt (SQ.elimProp (λ _ → squash/ _ _) + (Abi.elimProp _ (λ _ → squash/ _ _) + λ g → sym (substFinBouquetCode (α' x) g + ∙ eq/ _ _ ∣ (η (η x)) + , (sym (ridAb (η (α' x))) + ∙ commAb (η (α' x)) 1Ab) + ∙ cong₂ _·Ab_ (sym (rinvAb (η g))) refl + ∙ sym (assocAb (η g) (η (inv g)) (η (α' x))) + ∙ cong₂ _·Ab_ refl (commAb (η (inv g)) (η (α' x))) + ∙ assocAb (η g) (η (α' x)) (η (inv g)) ∣₁))))) + ∙ retEq univalence _ + + isSetCofibFinBoquetFunCode : (x : _) → isSet (CofibFinBoquetFunCode x) + isSetCofibFinBoquetFunCode = + PO.elimProp _ (λ _ → isPropIsSet) + (λ _ → AbGroupStr.is-set (Freeᵃᵇ/ImFinBouquetMap .snd)) + (elimPropBouquet + (λ _ → isPropIsSet) + (AbGroupStr.is-set (Freeᵃᵇ/ImFinBouquetMap .snd))) + + FG→π₁cofibFinBouquetMap : + GroupHom (freeGroupGroup (Fin k)) + (πGr 0 (cofib (fst α) , inr base)) + fst FG→π₁cofibFinBouquetMap b = + ∣ (λ i → inr (inv' Iso-ΩFinBouquet-FreeGroup b i)) ∣₂ + snd FG→π₁cofibFinBouquetMap = makeIsGroupHom λ a b + → cong ∣_∣₂ ((λ i j → inr (invIso-ΩFinBouquet-FreeGroupPresStr a b i j)) + ∙ cong-∙ inr (inv' Iso-ΩFinBouquet-FreeGroup a) + (inv' Iso-ΩFinBouquet-FreeGroup b)) + + private + loopCofibFinBouquetMap : (x : Fin m) + → Path (Path (cofib (fst α)) (inr base) (inr base)) + (λ i → inr (inv' Iso-ΩFinBouquet-FreeGroup (α' x) i)) + refl + loopCofibFinBouquetMap x i j = + hcomp (λ k → λ {(i = i0) → push (loop x j) k + ; (i = i1) → push base k + ; (j = i0) → push base k + ; (j = i1) → push base k}) + (inl tt) + + AbFG→π₁ᵃᵇCofibFinBouquetMap : Abelianization (freeGroupGroup (Fin k)) + → ∥ Pathᵃᵇ (cofib (fst α)) (inr base) (inr base) ∥₂ + AbFG→π₁ᵃᵇCofibFinBouquetMap = fst Abelianizeπ₁→π₁ᵃᵇ + ∘ AbelianizationFun FG→π₁cofibFinBouquetMap .fst + + Hom-AbFG-π₁ᵃᵇCofibFinBouquetMap : + AbGroupHom (AbelianizationAbGroup (freeGroupGroup (Fin k))) + (π₁ᵃᵇAbGroup (cofib (fst α) , inr base)) + Hom-AbFG-π₁ᵃᵇCofibFinBouquetMap = + compGroupHom (AbelianizationFun FG→π₁cofibFinBouquetMap) Abelianizeπ₁→π₁ᵃᵇ + + AbFG→π₁ᵃᵇCofibFinBouquetMap≡' : (x : Abelianization (freeGroupGroup (Fin m))) + → (a : FreeGroup (Fin k)) + → ·πᵃᵇ (AbFG→π₁ᵃᵇCofibFinBouquetMap ((AbelianizationFun (FG.rec α')) .fst x)) + (AbFG→π₁ᵃᵇCofibFinBouquetMap (η a)) + ≡ AbFG→π₁ᵃᵇCofibFinBouquetMap (η a) + AbFG→π₁ᵃᵇCofibFinBouquetMap≡' = + Abi.elimProp _ (λ _ → isPropΠ λ _ → squash₂ _ _) + (FG.elimProp (λ _ → isPropΠ λ _ → squash₂ _ _) + (λ c a → (λ i → ·πᵃᵇ ∣ paths (loopCofibFinBouquetMap c i) ∣₂ + (AbFG→π₁ᵃᵇCofibFinBouquetMap (η a))) + ∙ cong ∣_∣₂ (cong paths (sym (lUnit _)))) + (λ g1 g2 p q a + → cong₂ ·πᵃᵇ (cong AbFG→π₁ᵃᵇCofibFinBouquetMap + (IsGroupHom.pres· + (snd (AbelianizationFun (FG.rec α'))) (η g1) (η g2)) + ∙ IsGroupHom.pres· + (snd (compGroupHom (AbelianizationFun FG→π₁cofibFinBouquetMap) + (Abelianizeπ₁→π₁ᵃᵇ))) + (fst (AbelianizationFun (FG.rec α')) (η g1)) + (fst (AbelianizationFun (FG.rec α')) (η g2))) + (λ _ → (AbFG→π₁ᵃᵇCofibFinBouquetMap (η a))) + ∙ sym (·πᵃᵇassoc (AbFG→π₁ᵃᵇCofibFinBouquetMap + (AbelianizationFun (FG.rec α') .fst (η g1))) + (AbFG→π₁ᵃᵇCofibFinBouquetMap + (AbelianizationFun (FG.rec α') .fst (η g2))) + (AbFG→π₁ᵃᵇCofibFinBouquetMap (η a))) + ∙ cong (·πᵃᵇ (AbFG→π₁ᵃᵇCofibFinBouquetMap + (AbelianizationFun (FG.rec α') .fst (η g1)))) (q a) + ∙ p a) + (λ a → ·πᵃᵇlUnit (AbFG→π₁ᵃᵇCofibFinBouquetMap (η a))) + λ g p a → cong₂ ·πᵃᵇ + (sym (sym (IsGroupHom.presinv (snd (compGroupHom + (AbelianizationFun FG→π₁cofibFinBouquetMap) + (Abelianizeπ₁→π₁ᵃᵇ))) + (AbelianizationFun (FG.rec α') .fst (η g))) + ∙ cong AbFG→π₁ᵃᵇCofibFinBouquetMap + (IsGroupHom.presinv + (snd (AbelianizationFun (FG.rec α'))) (η g)))) + (sym (sym (IsGroupHom.presinv (snd (compGroupHom + (AbelianizationFun FG→π₁cofibFinBouquetMap) + (Abelianizeπ₁→π₁ᵃᵇ))) (η (inv a))) + ∙ cong (AbFG→π₁ᵃᵇCofibFinBouquetMap ∘ η) + (GroupTheory.invInv (freeGroupGroup (Fin k)) a))) + ∙ sym (-πᵃᵇinvDistr (AbFG→π₁ᵃᵇCofibFinBouquetMap + ((AbelianizationFun (FG.rec α')) .fst (η g))) + (AbFG→π₁ᵃᵇCofibFinBouquetMap (η (inv a)))) + ∙ cong -πᵃᵇ (p (inv a)) + ∙ sym (IsGroupHom.presinv (snd (compGroupHom + (AbelianizationFun FG→π₁cofibFinBouquetMap) + (Abelianizeπ₁→π₁ᵃᵇ))) (η (inv a))) + ∙ cong (AbFG→π₁ᵃᵇCofibFinBouquetMap ∘ η) + (GroupTheory.invInv (freeGroupGroup (Fin k)) a)) + + AbFG→π₁ᵃᵇCofibFinBouquetMap≡ : (x : Abelianization (freeGroupGroup (Fin m))) + (a b : FreeGroup (Fin k)) + (q : ·πᵃᵇ (AbFG→π₁ᵃᵇCofibFinBouquetMap (AbelianizationFun (FG.rec α') .fst x)) + (AbFG→π₁ᵃᵇCofibFinBouquetMap (η b)) + ≡ AbFG→π₁ᵃᵇCofibFinBouquetMap (η a)) + → Path (∥ Pathᵃᵇ (cofib (fst α)) (inr base) (inr base) ∥₂) + (AbFG→π₁ᵃᵇCofibFinBouquetMap (η a)) + (AbFG→π₁ᵃᵇCofibFinBouquetMap (η b)) + AbFG→π₁ᵃᵇCofibFinBouquetMap≡ x a b q = + sym q ∙ AbFG→π₁ᵃᵇCofibFinBouquetMap≡' x b + + decodeCofibαinl : Abelianization (freeGroupGroup (Fin k)) + → ∥ inr base ≡ᵃᵇ inl tt ∥₂ + decodeCofibαinl = + Abi.rec _ squash₂ + (λ s → ·πᵃᵇ (∣ paths (λ i → inr (inv' Iso-ΩFinBouquet-FreeGroup s i)) ∣₂) + ∣ paths (sym (push base)) ∣₂) + λ a b c → cong₂ ·πᵃᵇ (cong ∣_∣₂ (lem a b c) + ∙ cong (·πᵃᵇ (∣ paths (t a) ∣₂)) + (·πᵃᵇcomm ∣ paths (t b) ∣₂ ∣ paths (t c) ∣₂) + ∙ sym (cong ∣_∣₂ (lem a c b))) + (λ _ → ∣ paths (sym (push base)) ∣₂) + where + t : (x : _) → Path (cofib (fst α)) _ _ + t x i = inr (inv' Iso-ΩFinBouquet-FreeGroup x i) + + lem : (x y z : _) + → Path (Pathᵃᵇ (cofib (fst α)) _ _) + (paths (t (x ·f (y ·f z)))) + (paths (t x ∙ t y ∙ t z)) + lem x y z = + cong paths (((λ j i → inr (invIso-ΩFinBouquet-FreeGroupPresStr x (y ·f z) j i)) + ∙ cong-∙ inr (inv' Iso-ΩFinBouquet-FreeGroup x) + (inv' Iso-ΩFinBouquet-FreeGroup (y ·f z))) + ∙ cong₂ _∙_ refl + ((λ j i → inr (invIso-ΩFinBouquet-FreeGroupPresStr y z j i)) + ∙ cong-∙ inr (inv' Iso-ΩFinBouquet-FreeGroup y) + (inv' Iso-ΩFinBouquet-FreeGroup z))) + + Hom-Freeᵃᵇ/ImFinBouquetMap-π₁ᵃᵇCofibFinBouquetMap : + AbGroupHom Freeᵃᵇ/ImFinBouquetMap + (π₁ᵃᵇAbGroup (cofib (fst α) , inr base)) + fst Hom-Freeᵃᵇ/ImFinBouquetMap-π₁ᵃᵇCofibFinBouquetMap = + SQ.rec squash₂ AbFG→π₁ᵃᵇCofibFinBouquetMap main + where + main : (a b : fst (AbelianizationAbGroup (freeGroupGroup (Fin k)))) + (q : ∃[ x ∈ fst (AbelianizationAbGroup (freeGroupGroup (Fin m))) ] + _ ≡ _) + → AbFG→π₁ᵃᵇCofibFinBouquetMap a ≡ AbFG→π₁ᵃᵇCofibFinBouquetMap b + main = Abi.elimProp2 _ (λ _ _ → isPropΠ (λ _ → squash₂ _ _)) + λ a b → PT.elim (λ _ → squash₂ _ _) + λ {(x , q) + → AbFG→π₁ᵃᵇCofibFinBouquetMap≡ x a b + (cong (λ s → ·πᵃᵇ s (AbFG→π₁ᵃᵇCofibFinBouquetMap (η b))) + (cong AbFG→π₁ᵃᵇCofibFinBouquetMap q + ∙ IsGroupHom.pres· (snd Hom-AbFG-π₁ᵃᵇCofibFinBouquetMap) + (η a) (η (inv b))) + ∙ (sym (·πᵃᵇassoc (AbFG→π₁ᵃᵇCofibFinBouquetMap (η a)) + (AbFG→π₁ᵃᵇCofibFinBouquetMap (η (inv b))) + (AbFG→π₁ᵃᵇCofibFinBouquetMap (η b)))) + ∙ cong (·πᵃᵇ (AbFG→π₁ᵃᵇCofibFinBouquetMap (η a))) + (cong₂ ·πᵃᵇ (IsGroupHom.presinv + (snd Hom-AbFG-π₁ᵃᵇCofibFinBouquetMap) (η b)) refl + ∙ ·πᵃᵇlCancel ((AbFG→π₁ᵃᵇCofibFinBouquetMap (η b)))) + ∙ ·πᵃᵇrUnit (AbFG→π₁ᵃᵇCofibFinBouquetMap (η a)))} + snd Hom-Freeᵃᵇ/ImFinBouquetMap-π₁ᵃᵇCofibFinBouquetMap = + makeIsGroupHom (SQ.elimProp2 + (λ _ _ → squash₂ _ _) + (IsGroupHom.pres· (snd Hom-AbFG-π₁ᵃᵇCofibFinBouquetMap))) + + decodeFinBouquetCode : (x : _) → FinBouquetCode x + → ∥ Pathᵃᵇ (cofib (fst α)) (inr base) (inr x) ∥₂ + decodeFinBouquetCode base = fst Hom-Freeᵃᵇ/ImFinBouquetMap-π₁ᵃᵇCofibFinBouquetMap + decodeFinBouquetCode (loop x i) = help i + where + lem : (p : _) + → transport (λ i → ∥ Pathᵃᵇ (cofib (fst α)) (inr base) (inr (loop x i)) ∥₂) + (·πᵃᵇ p (-πᵃᵇ (AbFG→π₁ᵃᵇCofibFinBouquetMap (η (η x))))) ≡ p + lem = ST.elim (λ _ → isSetPathImplicit) + (elimProp≡ᵃᵇ (λ _ → squash₂ _ _) λ p + → (λ j → transp (λ i → ∥ Pathᵃᵇ (cofib (fst α)) + (inr base) (inr (loop x (i ∨ j))) ∥₂) j + ∣ paths (p ∙ (λ i → inr (loop x (~ i ∨ j)))) ∣₂) + ∙ cong ∣_∣₂ (cong paths (sym (rUnit p)))) + + help : PathP (λ i → ua (isoToEquiv (·GroupAutomorphismR + (AbGroup→Group Freeᵃᵇ/ImFinBouquetMap) + [ η (η x) ])) i + → ∥ Pathᵃᵇ (cofib (fst α)) (inr base) (inr (loop x i)) ∥₂) + (fst Hom-Freeᵃᵇ/ImFinBouquetMap-π₁ᵃᵇCofibFinBouquetMap) + (fst Hom-Freeᵃᵇ/ImFinBouquetMap-π₁ᵃᵇCofibFinBouquetMap) + help = toPathP (funExt (SQ.elimProp (λ _ → squash₂ _ _) + λ s → cong (transport (λ i → ∥ Pathᵃᵇ (cofib (fst α)) + (inr base) (inr (loop x i)) ∥₂)) + ((cong (fst Hom-Freeᵃᵇ/ImFinBouquetMap-π₁ᵃᵇCofibFinBouquetMap) + (cong (inv' (·GroupAutomorphismR + (AbGroup→Group Freeᵃᵇ/ImFinBouquetMap) + [ η (η x) ])) + (transportRefl [ s ])) + ∙ IsGroupHom.pres· (snd Hom-Freeᵃᵇ/ImFinBouquetMap-π₁ᵃᵇCofibFinBouquetMap) + [ s ] [ η (inv (η x)) ] + ∙ cong (·πᵃᵇ (AbFG→π₁ᵃᵇCofibFinBouquetMap s)) + (IsGroupHom.presinv + (snd Hom-Freeᵃᵇ/ImFinBouquetMap-π₁ᵃᵇCofibFinBouquetMap) + [ η (η x) ]))) + ∙ lem (AbFG→π₁ᵃᵇCofibFinBouquetMap s))) + + decodeCofibFinBoquetFun : (x : _) → CofibFinBoquetFunCode x + → ∥ inr base ≡ᵃᵇ x ∥₂ + decodeCofibFinBoquetFun (inl x) p = + ·πᵃᵇ (decodeFinBouquetCode base p) ∣ paths (sym (push base)) ∣₂ + decodeCofibFinBoquetFun (inr x) = decodeFinBouquetCode x + decodeCofibFinBoquetFun (push a i) = help a i + where + help : (a : Bouquet (Fin m)) + → PathP (λ i → CofibFinBoquetFunCode (push a i) + → ∥ Pathᵃᵇ (cofib (fst α)) (inr base) (push a i) ∥₂) + (λ p → ·πᵃᵇ (decodeFinBouquetCode base p) + ∣ paths (sym (push base)) ∣₂) + (decodeFinBouquetCode (inv' CharacFinBouquetFunIso α' .fst a)) + help = + elimPropBouquet (λ _ → isOfHLevelPathP' 1 (isSetΠ (λ _ → squash₂)) _ _) + (funExt (SQ.elimProp (λ _ → isOfHLevelPathP' 1 squash₂ _ _) + (Abi.elimProp _ (λ _ → isOfHLevelPathP' 1 squash₂ _ _) + λ g → λ i → ∣ paths (compPath-filler + (λ i₂ → inr (inv' Iso-ΩFinBouquet-FreeGroup g i₂)) + (sym (push base)) (~ i)) ∣₂))) + + FinBouquetCode+ : (x : _) → Freeᵃᵇ/ImFinBouquetMap .fst + → FinBouquetCode x → FinBouquetCode x + FinBouquetCode+ base p q = AbGroupStr._+_ (snd Freeᵃᵇ/ImFinBouquetMap) p q + FinBouquetCode+ (loop x i) p = commPathP i + where + typecheck : ∀ {ℓ} (A B : Type ℓ) (p : A ≡ B) + (f : A → A) (g : B → B) + → ((x : _) → transport p (f (transport (sym p) x)) ≡ g x) + → PathP (λ i → p i → p i) f g + typecheck = λ A → J> λ f g p + → funExt λ x → sym (transportRefl _ ∙ cong f (transportRefl x)) ∙ p x + + typecheck' : ∀ {ℓ} {A B : Type ℓ} (p : A ≃ B) + {f : A → A} {g : B → B} → ((x : _) → fst p (f (invEq p x)) ≡ g x) + → PathP (λ i → ua p i → ua p i) f g + typecheck' p {f = f} {g} h = + typecheck _ _ (ua p) f g λ b + → transportRefl _ + ∙ cong (fst p) (cong f (cong (invEq p) (transportRefl b))) ∙ h b + + + commPathP : PathP (λ i → ua (isoToEquiv (·GroupAutomorphismR + (AbGroup→Group Freeᵃᵇ/ImFinBouquetMap) + [ η (η x) ])) i + → ua (isoToEquiv (·GroupAutomorphismR + (AbGroup→Group Freeᵃᵇ/ImFinBouquetMap) + [ η (η x) ])) i) + (AbGroupStr._+_ (snd Freeᵃᵇ/ImFinBouquetMap) p) + (AbGroupStr._+_ (snd Freeᵃᵇ/ImFinBouquetMap) p) + commPathP = + typecheck' (isoToEquiv (·GroupAutomorphismR + (AbGroup→Group Freeᵃᵇ/ImFinBouquetMap) + [ η (η x) ])) + (SQ.elimProp (λ _ → squash/ _ _) + (Abi.elimProp _ (λ _ → squash/ _ _) λ g + → sym (AbGroupStr.+Assoc (snd Freeᵃᵇ/ImFinBouquetMap) p + (invEq (isoToEquiv (·GroupAutomorphismR (AbGroup→Group Freeᵃᵇ/ImFinBouquetMap) + [ η (η x) ])) [ η g ]) + [ η (η x) ]) + ∙ cong (snd Freeᵃᵇ/ImFinBouquetMap AbGroupStr.+ p) + (sym (AbGroupStr.+Assoc (snd Freeᵃᵇ/ImFinBouquetMap) + [ η g ] [ η (inv (η x)) ] [ η (η x) ]) + ∙ cong (snd Freeᵃᵇ/ImFinBouquetMap AbGroupStr.+ [ η g ]) + (AbGroupStr.+InvL (snd Freeᵃᵇ/ImFinBouquetMap) [ η (η x) ]) + ∙ AbGroupStr.+IdR (snd Freeᵃᵇ/ImFinBouquetMap) [ η g ]))) + + CofibFinBoquetFunCode+ : (x : _) → Freeᵃᵇ/ImFinBouquetMap .fst + → CofibFinBoquetFunCode x → CofibFinBoquetFunCode x + CofibFinBoquetFunCode+ (inl x) p q = + AbGroupStr._+_ (snd Freeᵃᵇ/ImFinBouquetMap) p q + CofibFinBoquetFunCode+ (inr x) = FinBouquetCode+ x + CofibFinBoquetFunCode+ (push x i) p = help x i + where + help : (x : Bouquet (Fin m)) + → PathP (λ i → CofibFinBoquetFunCode (push x i) + → CofibFinBoquetFunCode (push x i)) + (snd Freeᵃᵇ/ImFinBouquetMap AbGroupStr.+ p) + (FinBouquetCode+ (α .fst x) p) + help = elimPropBouquet (λ _ → isOfHLevelPathP' 1 + (isSetΠ (λ _ → isSetCofibFinBoquetFunCode _)) _ _) + refl + + encodeCofibFinBoquetFun' : (x : _) + → Pathᵃᵇ (cofib (fst α)) (inr base) x → CofibFinBoquetFunCode x + encodeCofibFinBoquetFun' x (paths q) = subst CofibFinBoquetFunCode q [ η ε ] + encodeCofibFinBoquetFun' x (com p q r i) = + encodeCofibFinBoquetFun'-comm _ q p r i + where + Code' : (x : _) (p : Path (cofib (fst α)) (inr base) (inr x)) → hProp ℓ-zero + fst (Code' base p) = + ∃[ r ∈ Path (Bouquet (Fin k)) base base ] p ≡ λ i → inr (r i) + snd (Code' base p) = squash₁ + fst (Code' (loop x i) p) = + ∃[ r ∈ Path (Bouquet (Fin k)) base (loop x i) ] p ≡ λ i → inr (r i) + snd (Code' (loop x i) p) = squash₁ + + Code : (x : cofib (fst α)) (p : inr base ≡ x) → Type + Code (inl x) p = ∃[ r ∈ Path (Bouquet (Fin k)) base base ] + (p ≡ (λ i → inr (r i)) ∙ sym (push base)) + Code (inr x) p = fst (Code' x p) + Code (push a i) p = help a i p .fst + where + help : (a : Bouquet (Fin m)) + → PathP (λ i → Path (cofib (fst α)) (inr base) (push a i) → + hProp ℓ-zero) + (λ a → (∃[ r ∈ Path (Bouquet (Fin k)) base base ] + a ≡ (λ i → inr (r i)) ∙ sym (push base)) , squash₁) + (Code' (fst α a)) + help = + elimPropBouquet (λ _ → isOfHLevelPathP' 1 + (isSetΠ (λ _ → isSetHProp)) _ _) + λ i p → (∃[ r ∈ Path (Bouquet (Fin k)) base base ] + p ≡ compPath-filler (λ i → inr (r i)) (sym (push base)) (~ i)) + , squash₁ + + aux : (x : cofib (fst α)) (p : inr base ≡ x) → Code x p + aux = J> ∣ refl , refl ∣₁ + + encodeCofibFinBoquetFun'-comm : (x : _) (q p r : inr base ≡ x) + → subst CofibFinBoquetFunCode (p ∙ sym q ∙ r) [ η ε ] + ≡ subst CofibFinBoquetFunCode (r ∙ sym q ∙ p) [ η ε ] + encodeCofibFinBoquetFun'-comm = J> λ p q + → cong (λ p → subst CofibFinBoquetFunCode p [ η ε ]) + (cong (p ∙_) (sym (lUnit q))) + ∙ substComposite CofibFinBoquetFunCode p q [ η ε ] + ∙ PT.rec2 (squash/ _ _) + (λ {(x' , xe) (y' , ye) + → lem (fun Iso-ΩFinBouquet-FreeGroup x') + (fun Iso-ΩFinBouquet-FreeGroup y') + p + (cong (cong Bouquet→CofibFinBouquetMap) + (leftInv Iso-ΩFinBouquet-FreeGroup x') ∙ sym xe) + q + (cong (cong Bouquet→CofibFinBouquetMap) + (leftInv Iso-ΩFinBouquet-FreeGroup y') ∙ sym ye)}) + (aux _ p) + (aux _ q) + ∙ sym (substComposite CofibFinBoquetFunCode q p [ η ε ]) + ∙ cong (λ p → subst CofibFinBoquetFunCode p [ η ε ]) + (cong (q ∙_) (lUnit p)) + where + lem : (x y : FreeGroup (Fin k)) + (p : Path (cofib (fst α)) (inr base) (inr base)) + (s : (λ i → inr (inv' Iso-ΩFinBouquet-FreeGroup x i)) ≡ p) + (q : Path (cofib (fst α)) (inr base) (inr base)) + (s' : (λ i → inr (inv' Iso-ΩFinBouquet-FreeGroup y i)) ≡ q) + → subst CofibFinBoquetFunCode q (subst CofibFinBoquetFunCode p [ η ε ]) + ≡ subst CofibFinBoquetFunCode p (subst CofibFinBoquetFunCode q [ η ε ]) + lem x y = + J> (J> cong (subst CofibFinBoquetFunCode y') + (substFinBouquetCode x ε + ∙ AbGroupStr.+IdL (snd Freeᵃᵇ/ImFinBouquetMap) [ η x ]) + ∙ substFinBouquetCode y x + ∙ cong [_] (AbelianizationGroupStructure.commAb + (freeGroupGroup (Fin k)) (η x) (η y)) + ∙ sym (substFinBouquetCode x y) + ∙ cong (subst CofibFinBoquetFunCode x') + (sym (substFinBouquetCode y ε + ∙ AbGroupStr.+IdL (snd Freeᵃᵇ/ImFinBouquetMap) [ η y ])) ) + where + x' y' : Path (cofib (fst α)) (inr base) (inr base) + x' = (λ i → inr (inv' Iso-ΩFinBouquet-FreeGroup x i)) + y' = (λ i → inr (inv' Iso-ΩFinBouquet-FreeGroup y i)) + + encodeCofibFinBoquetFun : (x : _) + → ∥ Pathᵃᵇ (cofib (fst α)) (inr base) x ∥₂ → CofibFinBoquetFunCode x + encodeCofibFinBoquetFun x = + ST.rec (isSetCofibFinBoquetFunCode _) (encodeCofibFinBoquetFun' x) + + decodeEncodeCofibFinBouquetMap : (x : _) + (p : ∥ Pathᵃᵇ (cofib (fst α)) (inr base) x ∥₂) + → decodeCofibFinBoquetFun x (encodeCofibFinBoquetFun x p) ≡ p + decodeEncodeCofibFinBouquetMap x = + ST.elim (λ _ → isSetPathImplicit) (elimProp≡ᵃᵇ (λ _ → squash₂ _ _) + (J (λ x p → decodeCofibFinBoquetFun x + (encodeCofibFinBoquetFun x ∣ paths p ∣₂) ≡ ∣ paths p ∣₂) + refl)) + + encodeDecodeCofibFinBouquetMap : (p : _) + → encodeCofibFinBoquetFun (inr base) (decodeCofibFinBoquetFun (inr base) p) + ≡ p + encodeDecodeCofibFinBouquetMap = SQ.elimProp (λ _ → squash/ _ _) + (Abi.elimProp _ (λ _ → squash/ _ _) + λ w → substFinBouquetCode w ε ∙ λ i → [ η (FG.idl w (~ i)) ]) + + --- Main results --- + Freeᵃᵇ/ImFinBouquetMap≅π₁ᵃᵇCofibFinBouquetMap : + AbGroupIso Freeᵃᵇ/ImFinBouquetMap + (π₁ᵃᵇAbGroup (cofib (fst α) , inr base)) + fun (fst Freeᵃᵇ/ImFinBouquetMap≅π₁ᵃᵇCofibFinBouquetMap) = + Hom-Freeᵃᵇ/ImFinBouquetMap-π₁ᵃᵇCofibFinBouquetMap .fst + inv' (fst Freeᵃᵇ/ImFinBouquetMap≅π₁ᵃᵇCofibFinBouquetMap) = + encodeCofibFinBoquetFun (inr base) + rightInv (fst Freeᵃᵇ/ImFinBouquetMap≅π₁ᵃᵇCofibFinBouquetMap) = + decodeEncodeCofibFinBouquetMap (inr base) + leftInv (fst Freeᵃᵇ/ImFinBouquetMap≅π₁ᵃᵇCofibFinBouquetMap) = + encodeDecodeCofibFinBouquetMap + snd Freeᵃᵇ/ImFinBouquetMap≅π₁ᵃᵇCofibFinBouquetMap = + Hom-Freeᵃᵇ/ImFinBouquetMap-π₁ᵃᵇCofibFinBouquetMap .snd + + Freeᵃᵇ/ImFinBouquetMap≅π₁ᵃᵇCofibFinBouquetMap' : + AbGroupIso Freeᵃᵇ/ImFinBouquetMap + (AbelianizationAbGroup (πGr 0 (cofib (fst α) , inr base))) + Freeᵃᵇ/ImFinBouquetMap≅π₁ᵃᵇCofibFinBouquetMap' = + compGroupIso Freeᵃᵇ/ImFinBouquetMap≅π₁ᵃᵇCofibFinBouquetMap + (invGroupIso (Abelianizeπ₁≅π₁ᵃᵇ (_ , inr base))) + + Iso-CofibFinBouquetMap-CofibFinSphereBouquetMap : + Iso (cofib (fst α)) (cofib (fst αSphereBouquet)) + Iso-CofibFinBouquetMap-CofibFinSphereBouquetMap = + pushoutIso _ _ _ _ + (invEquiv (Bouquet≃∙SphereBouquet .fst)) (idEquiv _) + (invEquiv (Bouquet≃∙SphereBouquet .fst)) + (λ i x → tt) + (funExt λ x → cong (invEq (isoToEquiv Iso-SphereBouquet-Bouquet)) + (cong (fst α) (sym (rightInv Iso-SphereBouquet-Bouquet x)))) + + -- π₁ᵃᵇ(cofib α) ≅ Freeᵃᵇ/Im(deg(α)) + Freeᵃᵇ/ImFinBouquetMap≅π'₁ᵃᵇCofibFinSphereBouquetMap : + AbGroupIso Freeᵃᵇ/ImFinBouquetMap + (AbelianizationAbGroup + (π'Gr 0 (cofib (fst αSphereBouquet) , inl tt))) + Freeᵃᵇ/ImFinBouquetMap≅π'₁ᵃᵇCofibFinSphereBouquetMap = + compGroupIso Freeᵃᵇ/ImFinBouquetMap≅π₁ᵃᵇCofibFinBouquetMap' + (GroupEquiv→GroupIso (AbelianizationEquiv + (compGroupEquiv + (GroupIso→GroupEquiv + (invGroupIso (π'Gr≅πGr 0 (cofib (fst α) , inr base)))) + (GroupIso→GroupEquiv + (π'GrIso 0 (isoToEquiv + (Iso-CofibFinBouquetMap-CofibFinSphereBouquetMap) + , sym (push (inl tt)))))))) + + -- Freeᵃᵇ/Im(deg(α)) ≅ ℤ[]/Im(deg(α)) + Freeᵃᵇ/ImFinBouquetMap≅ℤ[]/ImBouquetDegree : + AbGroupIso Freeᵃᵇ/ImFinBouquetMap + (ℤ[Fin k ] /Im bouquetDegree (fst αSphereBouquet)) + Freeᵃᵇ/ImFinBouquetMap≅ℤ[]/ImBouquetDegree = + Hom/ImIso _ _ (AbFreeGroup≅ℤ[] α' m) (AbFreeGroup≅ℤ[] α' k) + (bouquetDegree-AbFreeGroup→ℤ[] α') + + -- Locked versions for faster type checking + Freeᵃᵇ/ImFinBouquetMap≅π'₁ᵃᵇCofibFinSphereBouquetMap-Lock : lockUnit {ℓ-zero} + → AbGroupIso (AbelianizationAbGroup + (π'Gr 0 (cofib (fst αSphereBouquet) , inl tt))) + Freeᵃᵇ/ImFinBouquetMap + Freeᵃᵇ/ImFinBouquetMap≅π'₁ᵃᵇCofibFinSphereBouquetMap-Lock unlock = + invGroupIso Freeᵃᵇ/ImFinBouquetMap≅π'₁ᵃᵇCofibFinSphereBouquetMap + + Freeᵃᵇ/ImFinBouquetMap≅ℤ[]/ImBouquetDegree-Lock : lockUnit {ℓ-zero} + → AbGroupIso Freeᵃᵇ/ImFinBouquetMap + (ℤ[Fin k ] /Im (bouquetDegree (fst αSphereBouquet))) + Freeᵃᵇ/ImFinBouquetMap≅ℤ[]/ImBouquetDegree-Lock unlock = + Freeᵃᵇ/ImFinBouquetMap≅ℤ[]/ImBouquetDegree + + π'BoquetFunCofib≅Freeᵃᵇ/ImFinBouquetMap>1-Lock : lockUnit {ℓ-zero} + → AbGroupIso (AbelianizationAbGroup (π'Gr 0 (cofib∙ (fst αSphereBouquet)))) + (ℤ[Fin k ] /Im (bouquetDegree (fst αSphereBouquet))) + π'BoquetFunCofib≅Freeᵃᵇ/ImFinBouquetMap>1-Lock t = + compGroupIso (Freeᵃᵇ/ImFinBouquetMap≅π'₁ᵃᵇCofibFinSphereBouquetMap-Lock t) + (Freeᵃᵇ/ImFinBouquetMap≅ℤ[]/ImBouquetDegree-Lock t) + + π'BoquetFunCofib≅Freeᵃᵇ/ImFinBouquetMap>1 : + AbGroupIso (AbelianizationAbGroup + (π'Gr 0 (cofib (fst αSphereBouquet) , inl tt))) + (ℤ[Fin k ] /Im (bouquetDegree (fst αSphereBouquet))) + π'BoquetFunCofib≅Freeᵃᵇ/ImFinBouquetMap>1 = + π'BoquetFunCofib≅Freeᵃᵇ/ImFinBouquetMap>1-Lock unlock + +-- Part 3: Same thing for arbitrary α : ⋁Sⁿ →∙ ⋁Sⁿ +-- (arbtirary n and α) +module _ {n m k : ℕ} + (α : SphereBouquet∙ (suc n) (Fin m) →∙ SphereBouquet∙ (suc n) (Fin k)) where + + preπ'FinSphereBouquetMapGenerator : (k : Fin k) → S₊∙ (suc n) →∙ cofib∙ (fst α) + preπ'FinSphereBouquetMapGenerator k = + (λ x → inr (inr (k , x))) , + (λ i → inr (push k (~ i))) ∙ (λ i → inr (snd α (~ i))) ∙ sym (push (inl tt)) + + π'FinSphereBouquetMapGenerator : (k : Fin k) → π'Gr n (cofib∙ (fst α)) .fst + π'FinSphereBouquetMapGenerator k = + ∣ preπ'FinSphereBouquetMapGenerator k ∣₂ + + πᵃᵇFinSphereBouquetMapGenerator : (k : Fin k) + → Abelianization (π'Gr n (cofib∙ (fst α))) + πᵃᵇFinSphereBouquetMapGenerator k = η (π'FinSphereBouquetMapGenerator k) + +private + π'ᵃᵇCofibFinSphereBouquetMap≅ℤ[]/BouquetDegreeMain : {n m k : ℕ} + (α : SphereBouquet∙ (suc n) (Fin m) →∙ SphereBouquet∙ (suc n) (Fin k)) + → Σ[ ϕ ∈ AbGroupIso + (AbelianizationAbGroup (π'Gr n (cofib∙ (fst α)))) + (ℤ[Fin k ] /Im (bouquetDegree (fst α))) ] + ((w : Fin k) → fun (fst ϕ) (πᵃᵇFinSphereBouquetMapGenerator α w) + ≡ [ ℤFinGenerator w ]) + π'ᵃᵇCofibFinSphereBouquetMap≅ℤ[]/BouquetDegreeMain {n = zero} {m} {k} α = + lem (inv' (compIso (invIso CharacFinBouquetFunIso) + Iso-Bouquet→∙-SphereBouquet₁→∙) α) α + (rightInv (compIso (invIso CharacFinBouquetFunIso) + Iso-Bouquet→∙-SphereBouquet₁→∙) α) + where + Goal : (α : _) (s : (w : _) → _) → Type + Goal α s = + Σ[ ϕ ∈ AbGroupIso + (AbelianizationAbGroup (π'Gr zero (cofib∙ (fst α)))) + (ℤ[Fin k ] /Im (bouquetDegree (fst α))) ] + ((w : Fin k) → fun (fst ϕ) (s w) ≡ [ ℤFinGenerator w ]) + module _ (α' : Fin m → FreeGroup (Fin k)) where + gens→finSphereBouquetMap∙ : SphereBouquet∙ 1 (Fin m) + →∙ SphereBouquet∙ 1 (Fin k) + gens→finSphereBouquetMap∙ = fun Iso-Bouquet→∙-SphereBouquet₁→∙ + (inv' CharacFinBouquetFunIso α') + + gens→finSphereBouquetMap∙snd : snd gens→finSphereBouquetMap∙ ≡ refl + gens→finSphereBouquetMap∙snd = cong₃ _∙∙_∙∙_ (λ _ → refl) + (cong (cong (fst (fun (pre∘∙equiv (invEquiv∙ Bouquet≃∙SphereBouquet)) + (inv' CharacFinBouquetFunIso α')))) + (cong₂ _∙_ (cong sym (cong₂ _∙_ refl (∙∙lCancel _) ∙ sym (rUnit refl))) + (cong₃ _∙∙_∙∙_ (sym (rUnit _) + ∙ cong (cong (inv' (invIso (equivToIso (fst Bouquet≃∙SphereBouquet))))) + lem) refl (sym (rUnit _)) + ∙ sym (rUnit refl)) + ∙ sym (rUnit refl) )) + (cong₃ _∙∙_∙∙_ refl refl + (sym (lUnit _) ∙ ∙∙lCancel _)) + ∙ cong₃ _∙∙_∙∙_ refl refl (sym (rUnit refl)) + ∙ sym (rUnit refl) + where + lem : rightInv (invIso (equivToIso (fst Bouquet≃∙SphereBouquet))) + ((fst (isoToEquiv (invIso (equivToIso (fst Bouquet≃∙SphereBouquet)))) + (pt (Bouquet∙ (Fin m))))) ≡ refl + lem = ∙∙lCancel _ + + module _ (lock : lockUnit {ℓ-zero}) where + lockIso : Iso _ _ + lockIso = + fst (invGroupIso ((π'BoquetFunCofib≅Freeᵃᵇ/ImFinBouquetMap>1-Lock α') + lock)) + + η' : _ → Abelianization (π'Gr 0 (cofib∙ (fst gens→finSphereBouquetMap∙))) + η' = η + + presGen : (lock : _) (w : Fin k) (t : _) + (p : t ≡ πᵃᵇFinSphereBouquetMapGenerator gens→finSphereBouquetMap∙ w) + → inv' (lockIso lock) t ≡ [ ℤFinGenerator w ] + presGen l w t p = + (cong (fun (fst (Freeᵃᵇ/ImFinBouquetMap≅ℤ[]/ImBouquetDegree-Lock α' l))) + (lem3 l)) + ∙ lem l + where + lem : (l : _) + → fun (fst (Freeᵃᵇ/ImFinBouquetMap≅ℤ[]/ImBouquetDegree-Lock α' l)) + [ η (η w) ] ≡ [ ℤFinGenerator w ] + lem unlock = refl + + lem2 : inv' (Freeᵃᵇ/ImFinBouquetMap≅π'₁ᵃᵇCofibFinSphereBouquetMap-Lock + α' unlock .fst) [ η (η w) ] + ≡ πᵃᵇFinSphereBouquetMapGenerator gens→finSphereBouquetMap∙ w + lem2 = cong η' (cong ∣_∣₂ + (ΣPathP (funExt + (λ { base i → inr (push w i) + ; (loop i) j → inr (doubleCompPath-filler + (push w) + (λ j → inr (w , loop j)) + (sym (push w)) (~ j) i)}) + , ((sym (lUnit (sym (push (inl tt))))) + ◁ compPath-filler' (λ i → inr (push w (~ i))) (sym (push (inl tt)))) + ▷ (cong₂ _∙_ refl (lUnit (sym (push (inl tt)))))) + ∙ λ i → preπ'FinSphereBouquetMapGenerator + (gens→finSphereBouquetMap α' + , gens→finSphereBouquetMap∙snd (~ i)) w)) + + lem3 : (l : _) + → fun (Freeᵃᵇ/ImFinBouquetMap≅π'₁ᵃᵇCofibFinSphereBouquetMap-Lock α' l + .fst) t ≡ [ η (η w) ] + lem3 unlock = + cong (fun (Freeᵃᵇ/ImFinBouquetMap≅π'₁ᵃᵇCofibFinSphereBouquetMap-Lock + α' unlock .fst)) p + ∙ cong (fun (Freeᵃᵇ/ImFinBouquetMap≅π'₁ᵃᵇCofibFinSphereBouquetMap-Lock + α' unlock .fst)) (sym lem2) + ∙ rightInv (Freeᵃᵇ/ImFinBouquetMap≅π'₁ᵃᵇCofibFinSphereBouquetMap-Lock + α' unlock .fst) [ η (η w) ] + + presGen' : (lock : _) (w : Fin k) + → inv' (lockIso lock) + (πᵃᵇFinSphereBouquetMapGenerator gens→finSphereBouquetMap∙ w) + ≡ [ ℤFinGenerator w ] + presGen' l w = presGen l w _ refl + + presGen⁻ : (lock : _)(w : Fin k) + → fun (lockIso lock) [ ℤFinGenerator w ] + ≡ (πᵃᵇFinSphereBouquetMapGenerator gens→finSphereBouquetMap∙ w) + presGen⁻ lock w = + cong (fun (lockIso lock)) (sym (presGen lock w _ refl)) + ∙ rightInv (lockIso lock) + (πᵃᵇFinSphereBouquetMapGenerator gens→finSphereBouquetMap∙ w) + + abstract + lockIso' : (lock : lockUnit {ℓ-zero}) + → AbGroupIso (ℤ[Fin k ] /Im (bouquetDegree (gens→finSphereBouquetMap α'))) + (AbelianizationAbGroup + (π'Gr zero (cofib∙ (gens→finSphereBouquetMap α')))) + fst (lockIso' l) = lockIso l + snd (lockIso' l) = + isGroupHomInv (GroupIso→GroupEquiv + (π'BoquetFunCofib≅Freeᵃᵇ/ImFinBouquetMap>1-Lock α' l)) + + G : lockUnit → (a : (w : _) → _) (t : (w : _) + → a w ≡ (πᵃᵇFinSphereBouquetMapGenerator gens→finSphereBouquetMap∙ w)) + → Goal gens→finSphereBouquetMap∙ a + fst (G l a t) = π'BoquetFunCofib≅Freeᵃᵇ/ImFinBouquetMap>1-Lock α' l + snd (G l a t) w = cong (inv' (lockIso l)) (t w) ∙ presGen l w _ refl + + + lem : (w : Fin m → FreeGroup (Fin k)) + (α : SphereBouquet∙ (suc zero) (Fin m) + →∙ SphereBouquet∙ (suc zero) (Fin k)) + (s : fun Iso-Bouquet→∙-SphereBouquet₁→∙ + (inv' CharacFinBouquetFunIso w) ≡ α) + → Goal α (πᵃᵇFinSphereBouquetMapGenerator α) + lem w = J> G w unlock _ (λ _ → refl) + π'ᵃᵇCofibFinSphereBouquetMap≅ℤ[]/BouquetDegreeMain {n = suc n} {m} {k} α = + →Goal unlock _ (λ _ → refl) + where + open πCofibFinSphereBouquetMap _ _ _ α + + lockIso : (lock : lockUnit {ℓ-zero}) + → GroupIso (π'Gr (suc n) (cofib∙ (fst α))) + (AbGroup→Group (ℤ[Fin k ] /Im (bouquetDegree (fst α)))) + lockIso unlock = π'CofibFinSphereBouquetMap≅ℤ[]/BouquetDegree α + + lockInvIso : (lock : lockUnit {ℓ-zero}) → _ + lockInvIso lock = invGroupIso (lockIso lock) + + lockInvHom = GroupIso→GroupHom (lockInvIso unlock) + + lockInvIso≡ : (l : _) (f : _) + → inv' (fst (lockInvIso l)) f + ≡ fun (fst Iso3) (fun (fst Iso2) (fun (fst Iso1) f)) + lockInvIso≡ unlock f = refl + + lockInvIso≡' : (f : _) (r : _) (q : _) + → fun (fst (surjImIso (π'∘∙Hom (suc n) inr') isSurjective-π'∘∙Hominr)) + (∣ r ∣₂ , ∣ q ∣₁) ≡ ∣ f ∣₂ + → fun (fst Iso3) (fun (fst Iso2) (fun (fst Iso1) ∣ f ∣₂)) + ≡ [ fst (GroupIso→GroupHom (πₙ⋁Sⁿ≅ℤ[] n k)) (fst q) ] + lockInvIso≡' f r q t = + cong (fun (fst (πCofibFinSphereBouquetMap.Iso3 n k _ α))) + (cong (fun (fst (πCofibFinSphereBouquetMap.Iso2 n k _ α))) + (cong (fun (fst (isoThm1 _))) + (sym (cong (inv' (fst (surjImIso (π'∘∙Hom (suc n) inr') + isSurjective-π'∘∙Hominr))) t) + ∙ leftInv (fst (surjImIso (π'∘∙Hom (suc n) inr') + isSurjective-π'∘∙Hominr)) + (∣ r ∣₂ , ∣ q ∣₁)))) + + lockInvHomGen : (l : _) (w : _) + → inv' (fst (lockInvIso l)) (π'FinSphereBouquetMapGenerator α w) + ≡ [ ℤFinGenerator w ] + lockInvHomGen l w = + lockInvIso≡ l (∣ preπ'FinSphereBouquetMapGenerator α w ∣₂) + ∙ lockInvIso≡' _ (preπ'FinSphereBouquetMapGenerator α w) + (∣ (λ x → inr (w , x)) , sym (push w) ∣₂ , refl) + refl + ∙ cong [_] (πₙ⋁Sⁿ≅ℤ[]Generator _ _ _) + + lockInvHomGen' : (l : lockUnit {ℓ-zero}) (w : _) (s : _) + (q : π'FinSphereBouquetMapGenerator α w ≡ s) + → fun (fst (lockIso l)) s ≡ [ ℤFinGenerator w ] + lockInvHomGen' l w = J> lockInvHomGen l w + + Goal : (s : (w : _) → _) → Type + Goal s = + Σ[ ϕ ∈ (AbGroupIso (AbelianizationAbGroup (π'Gr (suc n) (cofib∙ (fst α)))) + (ℤ[Fin k ] /Im (bouquetDegree (fst α)))) ] + ((w : Fin k) → fun (fst ϕ) (s w) ≡ [ ℤFinGenerator w ]) + + →Goal : lockUnit {ℓ-zero} → (s : (w : _) → _) + (p : (w : _) → s w ≡ πᵃᵇFinSphereBouquetMapGenerator α w) → Goal s + fst (→Goal l s p) = + compGroupIso (invGroupIso (AbelianizationIdempotent + (Group→AbGroup _ (π'-comm _)))) + (lockIso l) + snd (→Goal l s p) w = + cong (fun (fst (lockIso l))) + (cong (inv' (fst (AbelianizationIdempotent + (Group→AbGroup _ (π'-comm _))))) (p w)) + ∙ lockInvHomGen' l w _ refl + +-- Main results of file -- +module _ {n m k : ℕ} + (α : SphereBouquet∙ (suc n) (Fin m) →∙ SphereBouquet∙ (suc n) (Fin k)) where + + -- πₙᵃᵇ(cofib α) ≅ ℤ[]/Im α + π'ᵃᵇCofibFinSphereBouquetMap≅ℤ[]/BouquetDegree : + AbGroupIso (AbelianizationAbGroup (π'Gr n (cofib∙ (fst α)))) + (ℤ[Fin k ] /Im (bouquetDegree (fst α))) + π'ᵃᵇCofibFinSphereBouquetMap≅ℤ[]/BouquetDegree = + fst (π'ᵃᵇCofibFinSphereBouquetMap≅ℤ[]/BouquetDegreeMain α) + + -- Characterisation of generators + π'ᵃᵇCofibFinSphereBouquetMap≅ℤ[]/BouquetDegreePresGens : (w : Fin k) → + fun (fst π'ᵃᵇCofibFinSphereBouquetMap≅ℤ[]/BouquetDegree) + (πᵃᵇFinSphereBouquetMapGenerator α w) ≡ [ ℤFinGenerator w ] + π'ᵃᵇCofibFinSphereBouquetMap≅ℤ[]/BouquetDegreePresGens = + snd (π'ᵃᵇCofibFinSphereBouquetMap≅ℤ[]/BouquetDegreeMain α) diff --git a/Cubical/Homotopy/Group/PiCofibFinSphereBouquetMap.agda b/Cubical/Homotopy/Group/PiCofibFinSphereBouquetMap.agda new file mode 100644 index 0000000000..328b96ef01 --- /dev/null +++ b/Cubical/Homotopy/Group/PiCofibFinSphereBouquetMap.agda @@ -0,0 +1,224 @@ +{-# OPTIONS --lossy-unification #-} +{- +This file computes πₙ(cofib α) for n ≥ 2 and α : ⋁Sⁿ →∙ ⋁Sⁿ +-} +module Cubical.Homotopy.Group.PiCofibFinSphereBouquetMap where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Pointed + +open import Cubical.Data.Nat renaming (_+_ to _+ℕ_ ; _·_ to _·ℕ_) +open import Cubical.Data.Sigma +open import Cubical.Data.Fin.Inductive +open import Cubical.Data.Nat.Order.Inductive +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Int +open import Cubical.Data.Unit + +open import Cubical.HITs.Sn +open import Cubical.HITs.Sn.Degree +open import Cubical.HITs.Susp +open import Cubical.HITs.SphereBouquet +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.SphereBouquet.Degree +open import Cubical.HITs.Truncation as TR +open import Cubical.HITs.Pushout + +open import Cubical.Algebra.Group +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.Group.Subgroup +open import Cubical.Algebra.Group.IsomorphismTheorems +open import Cubical.Algebra.Group.QuotientGroup +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup as FAB +open import Cubical.Algebra.AbGroup.FinitePresentation + +open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.Group.Properties +open import Cubical.Homotopy.Group.PiSphereBouquet +open import Cubical.Homotopy.BlakersMassey +open import Cubical.Homotopy.Group.LES +open import Cubical.Homotopy.Connected + +open import Cubical.Relation.Nullary + +open FinitePresentation + +module _ {n m k : ℕ} + (α : SphereBouquet∙ (suc (suc n)) (Fin m) + →∙ SphereBouquet∙ (suc (suc n)) (Fin k)) where + + BouquetDegreeSubGroup : Subgroup (AbGroup→Group ℤ[Fin k ]) + BouquetDegreeSubGroup = imSubgroup (bouquetDegree (fst α)) + + BouquetDegreeNormalSubGroup : NormalSubgroup (AbGroup→Group ℤ[Fin k ]) + fst BouquetDegreeNormalSubGroup = BouquetDegreeSubGroup + snd BouquetDegreeNormalSubGroup = + isNormalIm _ λ f g i x → +Comm (f x) (g x) i + + ℤ[]/BouquetDegree : Group ℓ-zero + ℤ[]/BouquetDegree = AbGroup→Group ℤ[Fin k ] / BouquetDegreeNormalSubGroup + + +module πCofibFinSphereBouquetMap (n k m : ℕ) + (α : SphereBouquet∙ (suc (suc n)) (Fin m) + →∙ SphereBouquet∙ (suc (suc n)) (Fin k)) where + + inr' : SphereBouquet∙ (suc (suc n)) (Fin k) →∙ (cofib (fst α) , inl tt) + fst inr' = inr + snd inr' = (λ i → inr (α .snd (~ i))) ∙ sym (push (inl tt)) + + conα : isConnectedFun (suc (suc n)) (fst α) + conα b = + isOfHLevelRetractFromIso 0 + (compIso (truncOfΣIso (suc (suc n))) + (mapCompIso (compIso (Σ-cong-iso-snd + (λ _ → equivToIso (isContr→≃Unit + (isConnectedPath (suc (suc n)) + (isConnectedSphereBouquet' {n = suc n}) _ _)))) rUnit×Iso))) + (isConnectedSubtr (suc (suc n)) 1 isConnectedSphereBouquet') + + coninr : isConnectedFun (suc (suc (suc n))) (fst inr') + coninr = inrConnected (suc (suc (suc n))) _ _ + (isConnected→isConnectedFun (suc (suc (suc n))) + isConnectedSphereBouquet') + + open BlakersMassey□ (λ _ → tt) (fst α) (suc (suc n)) (suc n) + (isConnected→isConnectedFun _ (isConnectedSphereBouquet' {n = suc n})) + conα + + α∘inr : SphereBouquet∙ (suc (suc n)) (Fin m) + →∙ (fiber (fst inr') (inl tt) , (inl tt) , inr' .snd) + fst α∘inr x = (fst α x) , sym (push x) + snd α∘inr = ΣPathP ((snd α) + , (compPath-filler' (λ i → inr (α .snd (~ i))) (sym (push (inl tt))))) + + open π'LES inr' + isSurjective-π'∘∙Hom : isSurjective (π'∘∙Hom (suc n) α∘inr) + isSurjective-π'∘∙Hom = + connectedFun→π'Surj (suc n) _ + λ b → isConnectedSubtr' n (suc (suc (suc n))) + (subst (λ n → isConnected (suc (suc n)) (fiber (fst α∘inr) b)) + (+-suc n n) (lem b)) + where + is1 : Iso (Σ (Unit × fst (SphereBouquet∙ (suc (suc n)) (Fin k))) + PushoutPath×) + (fiber (fst inr') (inl tt)) + Iso.fun is1 ((tt , s) , p) = s , (sym p) + Iso.inv is1 (s , p) = (tt , s) , sym p + Iso.rightInv is1 (s , p) = refl + Iso.leftInv is1 ((tt , s) , p) = refl + + lem : isConnectedFun (suc (suc n +ℕ suc n)) (α∘inr .fst) + lem = isConnectedComp _ _ _ (isEquiv→isConnected _ (isoToIsEquiv is1) _) + isConnected-toPullback + + isSurjective-π'∘∙Hominr : isSurjective (π'∘∙Hom (suc n) inr') + isSurjective-π'∘∙Hominr = connectedFun→π'Surj (suc n) _ coninr + + Imα⊂Kerinr : (x : _) → isInIm (π'∘∙Hom (suc n) α) x + → isInKer (π'∘∙Hom (suc n) inr') x + Imα⊂Kerinr x p = Im-fib→A⊂Ker-A→B (suc n) x + (PT.rec squash₁ + (uncurry (ST.elim (λ _ → isSetΠ λ _ → isProp→isSet squash₁) + (λ a → J (λ x _ → isInIm (fib→A (suc n)) x) + ∣ (π'∘∙Hom (suc n) α∘inr .fst ∣ a ∣₂) + , (cong ∣_∣₂ (ΣPathP (refl , (sym (rUnit _) + ∙ cong-∙ fst (ΣPathP ((cong (fst α) (snd a)) + , λ i j → push (snd a i) (~ j))) _)))) ∣₁))) p) + + Kerinr⊂Imα : (x : _) → isInKer (π'∘∙Hom (suc n) inr') x + → isInIm (π'∘∙Hom (suc n) α) x + Kerinr⊂Imα x p = + PT.rec squash₁ + (uncurry ( λ f → J (λ x _ → isInIm (π'∘∙Hom (suc n) α) x) + (PT.rec squash₁ (uncurry + (ST.elim (λ _ → isSetΠ λ _ → isProp→isSet squash₁) + (λ g s → ∣ ∣ g ∣₂ , cong ∣_∣₂ + (ΣPathP (refl + , sym (cong-∙ fst (ΣPathP ((cong (fst α) (snd g)) + , (λ i j → push (snd g i) (~ j)))) _) ∙ rUnit _)) + ∙ cong (fib→A (suc n) .fst) s ∣₁))) (isSurjective-π'∘∙Hom f)))) + (Ker-A→B⊂Im-fib→A (suc n) x p) + + -- πₙCofib≅πₙ⋁Sⁿ/ker(πₙinr) + Iso1 : GroupIso (π'Gr (suc n) (cofib (fst α) , inl tt)) + (π'Gr (suc n) (SphereBouquet∙ (suc (suc n)) (Fin k)) + / kerNormalSubgroup (π'∘∙Hom (suc n) inr')) + Iso1 = + compGroupIso (invGroupIso (surjImIso (π'∘∙Hom (suc n) inr') + isSurjective-π'∘∙Hominr)) + (isoThm1 _) + + -- πₙ⋁Sⁿ/ker(πₙinr)≅πₙ⋁Sⁿ/im(πₙα) + Iso2 : GroupIso (π'Gr (suc n) (SphereBouquet∙ (suc (suc n)) (Fin k)) + / kerNormalSubgroup (π'∘∙Hom (suc n) inr')) + (π'Gr (suc n) (SphereBouquet∙ (suc (suc n)) (Fin k)) + / imNormalSubgroup (π'∘∙Hom (suc n) α) (π'-comm n)) + Iso2 = Hom/Iso idGroupIso (λ a b → Kerinr⊂Imα _) λ a b → Imα⊂Kerinr _ + + -- πₙ⋁Sⁿ/im(πₙα)≅ℤ[]/BouquetDegree-α + Iso3 : GroupIso ((π'Gr (suc n) (SphereBouquet∙ (suc (suc n)) (Fin k)) + / imNormalSubgroup (π'∘∙Hom (suc n) α) (π'-comm n))) + (ℤ[]/BouquetDegree α) + Iso3 = (Hom/ImIso _ _ ( (πₙ⋁Sⁿ≅ℤ[] n m)) ( (πₙ⋁Sⁿ≅ℤ[] n k)) + (funExt⁻ (cong fst (πₙ⋁SⁿHomElim ϕ ψ + λ s → funExt (λ x + → sumFinℤId m (λ r → sym (degreeComp' (suc (suc n)) _ _)) + ∙ sumFin-choose _+_ 0 (λ _ → refl) +Comm _ _ s + (cong (degree (suc (suc n))) + (funExt (λ w → cong (pickPetal x ∘ fst α ∘ inr) + (ΣPathP (refl , lem1 s w))))) + (λ w p → cong (degree (suc (suc n))) + (funExt (λ r → cong (pickPetal x ∘ fst α) (lem2 s x w r p) + ∙ (cong (pickPetal x) (snd α)))) + ∙ degreeConst (suc (suc n))) + ∙ cong (degree (suc (suc n))) refl))))) + where + lem1 : (s : Fin m) (w : S₊ (suc (suc n))) → pickPetal s (inr (s , w)) ≡ w + lem1 s w with (fst s ≟ᵗ fst s) + ... | lt x = ⊥.rec (¬m<ᵗm x) + ... | eq x = refl + ... | gt x = ⊥.rec (¬m<ᵗm x) + + ϕ = compGroupHom (GroupIso→GroupHom (πₙ⋁Sⁿ≅ℤ[] n m)) (bouquetDegree (fst α)) + ψ = compGroupHom (π'∘∙Hom (suc n) α) ((GroupIso→GroupHom (πₙ⋁Sⁿ≅ℤ[] n k))) + + lem2 : (s : Fin m) (x : Fin k) (w : Fin m) (r : Susp (S₊ (suc n))) + (p : ¬ w ≡ s) + → Path (SphereBouquet (suc (suc n)) (Fin m)) + (inr (w , ⋁Sphere→ΠSphere (inr (s , r)) w)) (inl tt) + lem2 s x w r p with (fst w ≟ᵗ fst s) + ... | lt x₁ = sym (push w) + ... | eq x₁ = ⊥.rec (p (Σ≡Prop (λ _ → isProp<ᵗ) x₁)) + ... | gt x₁ = sym (push w) + + +-- Main results +π'CofibFinSphereBouquetMap≅ℤ[]/BouquetDegree : {n m k : ℕ} + (α : SphereBouquet∙ (suc (suc n)) (Fin m) + →∙ SphereBouquet∙ (suc (suc n)) (Fin k)) + → GroupIso (π'Gr (suc n) (cofib (fst α) , inl tt)) + (ℤ[]/BouquetDegree α) +π'CofibFinSphereBouquetMap≅ℤ[]/BouquetDegree {n = n} {m} {k} α = + compGroupIso (compGroupIso (πCofibFinSphereBouquetMap.Iso1 n k m α) + (πCofibFinSphereBouquetMap.Iso2 n k m α)) + (πCofibFinSphereBouquetMap.Iso3 n k m α) + +hasFPπ'CofibFinSphereBouquetMap : {n m k : ℕ} + (α : FinSphereBouquetMap∙ m k (suc n)) + → FinitePresentation (Group→AbGroup (π'Gr (suc n) (cofib (fst α) , inl tt)) + (π'-comm n)) +nGens (hasFPπ'CofibFinSphereBouquetMap {n = n} {m = m} {k = k} α) = k +nRels (hasFPπ'CofibFinSphereBouquetMap {n = n} {m = m} {k = k} α) = m +rels (hasFPπ'CofibFinSphereBouquetMap {n = n} {m = m} {k = k} α) = + bouquetDegree (fst α) +fpiso (hasFPπ'CofibFinSphereBouquetMap {n = n} {m = m} {k = k} α) = + π'CofibFinSphereBouquetMap≅ℤ[]/BouquetDegree α diff --git a/Cubical/Homotopy/Group/PiSphereBouquet.agda b/Cubical/Homotopy/Group/PiSphereBouquet.agda new file mode 100644 index 0000000000..dcc22e8064 --- /dev/null +++ b/Cubical/Homotopy/Group/PiSphereBouquet.agda @@ -0,0 +1,261 @@ +{-# OPTIONS --lossy-unification #-} +{- +This file computes πₙ(⋁ₐSⁿ) for wedge sums over finite sets and n ≥ 2. +-} +module Cubical.Homotopy.Group.PiSphereBouquet where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.HLevels + +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.Group.MorphismProperties +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.AbGroup.Instances.FreeAbGroup as FAB +open import Cubical.Algebra.Group.Instances.Pi + +open import Cubical.Axiom.Choice + +open import Cubical.Data.Nat +open import Cubical.Data.Fin.Inductive +open import Cubical.Data.Nat.Order.Inductive +open import Cubical.Data.Sigma +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Unit + +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.Sn.Degree +open import Cubical.HITs.Susp +open import Cubical.HITs.Pushout +open import Cubical.HITs.SphereBouquet +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.Truncation as TR +open import Cubical.HITs.Wedge +open import Cubical.HITs.SphereBouquet.Degree + +open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.Group.Properties +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.BlakersMassey +open import Cubical.Homotopy.Group.PinSn + +open import Cubical.ZCohomology.Groups.Sn + +-- ⋁Sⁿ → ΠSⁿ +⋁Sphere→ΠSphere : {n k : ℕ} → SphereBouquet n (Fin k) → (Fin k → S₊ n) +⋁Sphere→ΠSphere a b = pickPetal b a + +isConnected⋁Sphere→ΠSphere : {n k : ℕ} + → isConnectedFun (suc n + suc n) (⋁Sphere→ΠSphere {n = suc n} {suc k}) +isConnected⋁Sphere→ΠSphere {n = n} {k = zero} = + subst (isConnectedFun (suc n + suc n)) + (funExt (λ x → funExt (lem x))) + (isEquiv→isConnected _ (isoToIsEquiv main) _) + where + cntr = isContr→≃Unit (inhProp→isContr fzero isPropFin1) + + main : Iso (SphereBouquet (suc n) (Fin 1)) (Fin 1 → S₊ (suc n)) + main = compIso (compIso (compIso + ((PushoutAlongEquiv cntr _)) + (compIso (Σ-cong-iso (equivToIso cntr) λ _ → idIso) lUnit×Iso)) + (invIso (ΠUnitIso ( λ _ → S₊ (suc n))))) + (domIso (equivToIso (invEquiv cntr))) + + lem : (x : _) (y : _) → Iso.fun main x y ≡ ⋁Sphere→ΠSphere x y + lem (inl x) y = refl + lem (inr ((zero , tt) , x)) (zero , tt) = refl + lem (push (zero , tt) i) (zero , tt) = refl +isConnected⋁Sphere→ΠSphere {n = n} {k = suc k} = + subst (isConnectedFun (suc n + suc n)) + (λ i x y → g≡ (isConnected⋁Sphere→ΠSphere {k = k}) x y i) + (isConnectedg (isConnected⋁Sphere→ΠSphere {k = k})) + where + conLem : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {A' : Type ℓ'} {B : Type ℓ''} + (f : A → A') {n} + → isConnectedFun n f → isConnectedFun n (map-× f (idfun B)) + conLem f cf (a , b) = isConnectedRetractFromIso _ + (compIso (Σ-cong-iso-snd (λ _ → invIso ΣPathIsoPathΣ)) + (compIso Σ-assoc-Iso + (Σ-cong-iso-snd + (λ x → compIso (invIso Σ-assoc-Iso) + (compIso (Σ-cong-iso-fst Σ-swap-Iso) + (compIso Σ-assoc-Iso + (compIso (Σ-cong-iso-snd + (λ s → + compIso (Σ-cong-iso-snd + λ b → iso sym sym (λ _ → refl) (λ _ → refl)) + (equivToIso (isContr→≃Unit + (isContrSingl _))))) + rUnit×Iso))))))) (cf a) + + module _ (ind : isConnectedFun (suc n + suc n) + (⋁Sphere→ΠSphere {n = suc n} {suc k})) where + g : SphereBouquet (suc n) (Fin (suc (suc k))) + → Fin (suc (suc k)) → S₊ (suc n) + g = Iso.inv (Iso-FinSuc→-Fin→× (suc k)) + ∘ map-× (⋁Sphere→ΠSphere {n = suc n} {suc k}) (idfun _) + ∘ ⋁↪ + ∘ Iso.fun Iso-⋁genFinSuc-⋁genFin⋁ + + g≡inl : (y : _) → g (inl tt) y ≡ ⋁Sphere→ΠSphere (inl tt) y + g≡inl (zero , y) = refl + g≡inl (suc s , y) = refl + + g≡inr : (x : _) (y : _) → g (inr x) y ≡ ⋁Sphere→ΠSphere (inr x) y + g≡inr ((zero , t) , q) (zero , p) = refl + g≡inr ((zero , t) , q) (suc x , p) = refl + g≡inr ((suc s , t) , q) (zero , p) = refl + g≡inr ((suc s , t) , q) (suc x , p) with (x ≟ᵗ s) + ... | lt x₁ = refl + ... | eq x₁ = refl + ... | gt x₁ = refl + + g≡inlr : (x : _) (y : _) + → Square (λ i → g (push x i) y) (λ i → ⋁Sphere→ΠSphere (push x i) y) + (g≡inl y) (g≡inr (x , ptSn (suc n)) y) + g≡inlr (zero , t) (zero , p) = refl + g≡inlr (suc s , t) (zero , p) = refl + g≡inlr (zero , t) (suc x , p) = refl + g≡inlr (suc s , t) (suc x , p) with (x ≟ᵗ s) + ... | lt x₁ = refl + ... | eq x₁ = refl + ... | gt x₁ = refl + + g≡ : (x : _) (y : _) → g x y ≡ ⋁Sphere→ΠSphere x y + g≡ (inl x) = g≡inl + g≡ (inr x) = g≡inr x + g≡ (push x i) y j = g≡inlr x y j i + + isConnectedg : isConnectedFun (suc n + suc n) g + isConnectedg = + isConnectedComp (Iso.inv (Iso-FinSuc→-Fin→× (suc k))) _ _ + (isEquiv→isConnected _ + (isoToIsEquiv (invIso (Iso-FinSuc→-Fin→× (suc k)))) (suc n + suc n)) + (isConnectedComp + (map-× (⋁Sphere→ΠSphere {n = suc n} {suc k}) (idfun _)) + _ (suc n + suc n) + (conLem _ ind) + (isConnectedComp ⋁↪ _ _ + (isConnected⋁↪ + isConnectedSphereBouquet' (sphereConnected (suc n))) + (isEquiv→isConnected _ (isoToIsEquiv Iso-⋁genFinSuc-⋁genFin⋁) + (suc n + suc n)))) + +-- πₙ(ΠₐSⁿ) ≅ℤ[a] +πₙΠSⁿ≅ℤ[] : (n k : ℕ) + → GroupIso (π'Gr n ((Fin k → S₊ (suc n)) , (λ _ → ptSn (suc n)))) + (AbGroup→Group ℤ[Fin k ]) +πₙΠSⁿ≅ℤ[] n k = compGroupIso is3 is4 + where + is1 : (n : ℕ) + → Iso (S₊∙ (suc n) →∙ ((Fin k → S₊ (suc n)) , (λ _ → ptSn (suc n)))) + (Fin k → S₊∙ (suc n) →∙ S₊∙ (suc n)) + fst (Iso.fun (is1 n) (f , s) x) y = f y x + snd (Iso.fun (is1 n) (f , s) x) = funExt⁻ s x + fst (Iso.inv (is1 n) f) y x = f x .fst y + snd (Iso.inv (is1 n) f) i x = f x .snd i + Iso.rightInv (is1 n) f = refl + Iso.leftInv (is1 n) f = refl + + is2 : (n : ℕ) + → Iso (π'Gr n ((Fin k → S₊ (suc n)) , (λ _ → ptSn (suc n))) .fst) + (Fin k → π'Gr n (S₊∙ (suc n)) .fst) + is2 n = compIso (setTruncIso (is1 n)) + (compIso + setTruncTrunc2Iso + (compIso (equivToIso (_ , InductiveFinSatAC 2 k _)) + (codomainIso (invIso setTruncTrunc2Iso)))) + + is3 : GroupIso (π'Gr n ((Fin k → S₊ (suc n)) , (λ _ → ptSn (suc n)))) + (ΠGroup (λ (x : Fin k) → π'Gr n (S₊∙ (suc n)))) + fst is3 = is2 n + snd is3 = makeIsGroupHom (ST.elim2 + (λ _ _ → isOfHLevelPath 2 (isSetΠ (λ _ → squash₂)) _ _) + λ f g → funExt λ x → cong ∣_∣₂ (lem n f g x)) + where + lem : (n : ℕ) + → (f g : S₊∙ (suc n) →∙ ((Fin k → S₊ (suc n)) , (λ _ → ptSn (suc n)))) + → (x : _) → Iso.fun (is1 n) (∙Π f g) x + ≡ ∙Π (Iso.fun (is1 n) f x) (Iso.fun (is1 n) g x) + lem zero f g x = + ΣPathP ((funExt (λ { base → refl ; (loop i) → refl})) , refl) + lem (suc n) f g x = + ΣPathP ((funExt (λ { north → refl ; south → refl ; (merid a i) → refl})) + , refl) + + is4 : GroupIso (ΠGroup (λ (x : Fin k) → π'Gr n (S₊∙ (suc n)))) + (AbGroup→Group ℤ[Fin k ]) + is4 = ΠGroupIso λ _ + → compGroupIso (GroupEquiv→GroupIso (πₙSⁿ≅HⁿSⁿ n)) + (Hⁿ-Sⁿ≅ℤ n) + +πₙ⋁Sⁿ≅ℤ[] : (n k : ℕ) + → GroupIso (π'Gr (suc n) (SphereBouquet∙ (suc (suc n)) (Fin k))) + (AbGroup→Group ℤ[Fin k ]) +πₙ⋁Sⁿ≅ℤ[] n k = + compGroupIso + (GroupEquiv→GroupIso (connectedFun→π'Equiv (suc n) + (⋁Sphere→ΠSphere , refl) (con k))) + (πₙΠSⁿ≅ℤ[] (suc n) k) + where + con : (k : _) + → isConnectedFun (suc (suc (suc (suc n)))) (⋁Sphere→ΠSphere {k = k}) + con zero b = ∣ (inl tt) , (funExt λ ()) ∣ + , TR.elim + (λ _ → isOfHLevelPath _ (isOfHLevelTrunc (suc (suc (suc (suc n))))) _ _) + λ {((inl tt) , q) → cong ∣_∣ₕ (ΣPathP (refl , cong funExt (funExt λ())))} + con (suc k) b = isConnectedSubtr (suc (suc (suc (suc n)))) n + (subst (λ p → isConnected p (fiber ⋁Sphere→ΠSphere b)) + (cong suc (sym (+-suc _ _)) ∙ sym (+-suc _ _)) + (isConnected⋁Sphere→ΠSphere b)) + +-- Generator of πₙ⋁Sⁿ +genπₙ⋁Sⁿ : {n k : ℕ} (x : Fin k) + → π'Gr (suc n) (SphereBouquet∙ (suc (suc n)) (Fin k)) .fst +genπₙ⋁Sⁿ x = ∣ (λ s → inr (x , s)) , (sym (push x)) ∣₂ + +πₙ⋁Sⁿ≅ℤ[]Generator : (n k : ℕ) (x : Fin k) + → Iso.fun (fst (πₙ⋁Sⁿ≅ℤ[] n k)) (genπₙ⋁Sⁿ x) + ≡ ℤFinGenerator x +πₙ⋁Sⁿ≅ℤ[]Generator n k x = funExt pickPetalId + where + pickPetalId : (w : _) + → degree (suc (suc n)) (λ z → ⋁Sphere→ΠSphere (inr (x , z)) w) + ≡ ℤFinGenerator x w + pickPetalId w with (fst x ≟ᵗ fst w) | (fst w ≟ᵗ fst x) + ... | lt x | lt x₁ = degreeConst (suc (suc n)) + ... | lt p | eq q = ⊥.rec (¬m<ᵗm (subst (_<ᵗ fst w) (sym q) p)) + ... | lt x | gt x₁ = degreeConst (suc (suc n)) + ... | eq p | lt q = ⊥.rec (⊥.rec (¬m<ᵗm (subst (fst w <ᵗ_) p q))) + ... | eq x | eq x₁ = degreeIdfun (suc (suc n)) + ... | eq p | gt q = ⊥.rec (¬m<ᵗm (subst (fst x <ᵗ_) (sym p) q)) + ... | gt x | lt x₁ = degreeConst (suc (suc n)) + ... | gt p | eq q = ⊥.rec (¬m<ᵗm (subst (_<ᵗ fst x) q p)) + ... | gt x | gt x₁ = degreeConst (suc (suc n)) + +-- Elimination principles for homomorphisms out of πₙ⋁Sⁿ +πₙ⋁SⁿHomElim : {n k k' : ℕ} + (ϕ ψ : GroupHom (π'Gr (suc n) (SphereBouquet∙ (suc (suc n)) (Fin k))) + (AbGroup→Group ℤ[Fin k' ])) + → ((x : Fin k) → fst ϕ (genπₙ⋁Sⁿ x) ≡ fst ψ (genπₙ⋁Sⁿ x)) + → ϕ ≡ ψ +πₙ⋁SⁿHomElim {n = n} {k} {k'} ϕ ψ ind = + Σ≡Prop (λ _ → isPropIsGroupHom _ _) (funExt λ x + → cong (fst ϕ) (sym (Iso.leftInv (fst (πₙ⋁Sⁿ≅ℤ[] n k)) x)) + ∙ funExt⁻ (cong fst lem) (Iso.fun (fst (πₙ⋁Sⁿ≅ℤ[] n k)) x) + ∙ cong (fst ψ) (Iso.leftInv (fst (πₙ⋁Sⁿ≅ℤ[] n k)) x)) + where + lem : compGroupHom (GroupIso→GroupHom (invGroupIso (πₙ⋁Sⁿ≅ℤ[] n k))) ϕ + ≡ compGroupHom (GroupIso→GroupHom (invGroupIso (πₙ⋁Sⁿ≅ℤ[] n k))) ψ + lem = agreeOnℤFinGenerator→≡ + λ x → cong (fst ϕ) (cong (Iso.inv (fst (πₙ⋁Sⁿ≅ℤ[] n k))) + (sym (πₙ⋁Sⁿ≅ℤ[]Generator n k x)) + ∙ Iso.leftInv (fst (πₙ⋁Sⁿ≅ℤ[] n k)) (genπₙ⋁Sⁿ x)) + ∙ ind x + ∙ cong (fst ψ) (sym (Iso.leftInv (fst (πₙ⋁Sⁿ≅ℤ[] n k)) (genπₙ⋁Sⁿ x)) + ∙ sym (cong (Iso.inv (fst (πₙ⋁Sⁿ≅ℤ[] n k))) + (sym (πₙ⋁Sⁿ≅ℤ[]Generator n k x)))) diff --git a/Cubical/Homotopy/Group/Properties.agda b/Cubical/Homotopy/Group/Properties.agda new file mode 100644 index 0000000000..56aaaa4062 --- /dev/null +++ b/Cubical/Homotopy/Group/Properties.agda @@ -0,0 +1,82 @@ +{-# OPTIONS --lossy-unification #-} +module Cubical.Homotopy.Group.Properties where + +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Group.Base + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Equiv + +open import Cubical.HITs.PropositionalTruncation +open import Cubical.HITs.SetTruncation as ST +open import Cubical.HITs.Truncation as TR + +open import Cubical.Data.Nat + +open import Cubical.Algebra.Group.Morphisms +open import Cubical.Algebra.AbGroup +open import Cubical.Algebra.Group.Abelianization.Properties + +connectedFun→πEquiv : ∀ {ℓ} {A B : Pointed ℓ} (n : ℕ) (f : A →∙ B) + → isConnectedFun (3 + n) (fst f) + → GroupEquiv (πGr n A) (πGr n B) +connectedFun→πEquiv {ℓ = ℓ} {A = A} {B = B} n f conf = + (πHom n f .fst + , subst isEquiv (funExt (ST.elim (λ _ → isSetPathImplicit) (λ _ → refl))) + (πA≃πB .snd)) + , πHom n f .snd + where + lem : (n : ℕ) → suc (suc n) ∸ n ≡ 2 + lem zero = refl + lem (suc n) = lem n + + connΩ^→f : isConnectedFun 2 (fst (Ω^→ (suc n) f)) + connΩ^→f = subst (λ k → isConnectedFun k (fst (Ω^→ (suc n) f))) + (lem n) + (isConnectedΩ^→ (suc (suc n)) (suc n) f conf) + + πA≃πB : π (suc n) A ≃ π (suc n) B + πA≃πB = compEquiv setTrunc≃Trunc2 + (compEquiv (connectedTruncEquiv 2 + (fst (Ω^→ (suc n) f)) connΩ^→f) + (invEquiv setTrunc≃Trunc2)) + +connectedFun→π'Equiv : ∀ {ℓ} {A B : Pointed ℓ} (n : ℕ) (f : A →∙ B) + → isConnectedFun (3 + n) (fst f) + → GroupEquiv (π'Gr n A) (π'Gr n B) +fst (fst (connectedFun→π'Equiv {ℓ = ℓ} {A = A} {B = B} n f conf)) = π'∘∙Hom n f .fst +snd (fst (connectedFun→π'Equiv {ℓ = ℓ} {A = A} {B = B} n f conf)) = + transport (λ i → isEquiv (GroupHomπ≅π'PathP-hom n f i .fst)) + (connectedFun→πEquiv n f conf .fst .snd) +snd (connectedFun→π'Equiv {ℓ = ℓ} {A = A} {B = B} n f conf) = π'∘∙Hom n f .snd + +connected→Abπ'Equiv : ∀ {ℓ} {A B : Pointed ℓ} (n : ℕ) (f : A →∙ B) + → isConnectedFun (3 + n) (fst f) + → AbGroupEquiv (AbelianizationAbGroup (π'Gr n A)) (AbelianizationAbGroup (π'Gr n B)) +connected→Abπ'Equiv n f isc = AbelianizationEquiv (connectedFun→π'Equiv n f isc) + +connectedFun→πSurj : ∀ {ℓ} {A B : Pointed ℓ} (n : ℕ) (f : A →∙ B) + → isConnectedFun (2 + n) (fst f) + → isSurjective {G = πGr n A} {H = πGr n B} (πHom n f) +connectedFun→πSurj {ℓ = ℓ} {A = A} {B = B} n f conf = + ST.elim (λ _ → isProp→isSet squash₁) + λ p → TR.rec squash₁ (λ {(q , r) → ∣ ∣ q ∣₂ , (cong ∣_∣₂ r) ∣₁}) + (connΩ^→f p .fst) + where + lem : (n : ℕ) → suc n ∸ n ≡ 1 + lem zero = refl + lem (suc n) = lem n + + connΩ^→f : isConnectedFun 1 (fst (Ω^→ (suc n) f)) + connΩ^→f = subst (λ k → isConnectedFun k (fst (Ω^→ (suc n) f))) + (lem n) + (isConnectedΩ^→ (suc n) (suc n) f conf) + +connectedFun→π'Surj : ∀ {ℓ} {A B : Pointed ℓ} (n : ℕ) (f : A →∙ B) + → isConnectedFun (2 + n) (fst f) + → isSurjective (π'∘∙Hom n f) +connectedFun→π'Surj n f conf = + transport (λ i → isSurjective (GroupHomπ≅π'PathP-hom n f i)) + (connectedFun→πSurj n f conf) diff --git a/Cubical/Homotopy/HopfInvariant/Brunerie.agda b/Cubical/Homotopy/HopfInvariant/Brunerie.agda index 5ed4653336..145356f14f 100644 --- a/Cubical/Homotopy/HopfInvariant/Brunerie.agda +++ b/Cubical/Homotopy/HopfInvariant/Brunerie.agda @@ -64,7 +64,7 @@ open PlusBis -- Some abstract versions of imported lemmas/definitions from -- ZCohomology.Groups.SphereProduct for faster type checking. -abstract +opaque H²-genₗabs : coHom 2 (S₊ 2 × S₊ 2) H²-genₗabs = H²-S²×S²-genₗ