diff --git a/Cubical/Algebra/AbGroup/Properties.agda b/Cubical/Algebra/AbGroup/Properties.agda index 6feb7ca2c4..8cad131409 100644 --- a/Cubical/Algebra/AbGroup/Properties.agda +++ b/Cubical/Algebra/AbGroup/Properties.agda @@ -37,6 +37,17 @@ module AbGroupTheory (A : AbGroup ℓ) where implicitInverse : ∀ {a b} → a + b ≡ 0g → b ≡ - a implicitInverse b+a≡0 = invUniqueR b+a≡0 + invDistrAb : ∀ a b → - (a + b) ≡ ((- a) + (- b)) + invDistrAb a b = invDistr a b ∙ +Comm (- b) (- a) + + inv^Distr : ∀ a b → (n : ℕ) → iter n -_ (a + b) ≡ (iter n -_ a + iter n -_ b) + inv^Distr a b zero = refl + inv^Distr a b (suc zero) = invDistrAb a b + inv^Distr a b (suc (suc n)) = + invInv (iter n -_ (a + b)) + ∙ inv^Distr a b n + ∙ cong₂ _+_ (sym (invInv (iter n -_ a))) (sym (invInv (iter n -_ b))) + addGroupHom : (A : AbGroup ℓ) (B : AbGroup ℓ') (ϕ ψ : AbGroupHom A B) → AbGroupHom A B fst (addGroupHom A B ϕ ψ) x = AbGroupStr._+_ (snd B) (ϕ .fst x) (ψ .fst x) snd (addGroupHom A B ϕ ψ) = makeIsGroupHom diff --git a/Cubical/Algebra/Group/Properties.agda b/Cubical/Algebra/Group/Properties.agda index c2cb2a51f8..130078a2ea 100644 --- a/Cubical/Algebra/Group/Properties.agda +++ b/Cubical/Algebra/Group/Properties.agda @@ -7,6 +7,9 @@ open import Cubical.Foundations.Structure open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.GroupoidLaws hiding (assoc) open import Cubical.Data.Sigma +open import Cubical.Data.Nat hiding (_·_) +open import Cubical.Data.Sum +open import Cubical.Data.Nat.IsEven open import Cubical.Algebra.Semigroup open import Cubical.Algebra.Monoid.Base @@ -89,6 +92,27 @@ module GroupTheory (G : Group ℓ) where a · (1g · inv a) ≡⟨ congR _·_ (·IdL (inv a)) ∙ ·InvR a ⟩ 1g ∎ + -iter-odd+even : (n m : ℕ) (g : fst G) + → isEvenT n ≡ isOddT m + → GroupStr._·_ (snd G) + (iter n (GroupStr.inv (snd G)) g) + (iter m (GroupStr.inv (snd G)) g) + ≡ GroupStr.1g (snd G) + -iter-odd+even n m g p with (evenOrOdd n) + ... | inl q = cong₂ (GroupStr._·_ (snd G)) + (iterEvenInv (GroupStr.inv (snd G)) + invInv n q g) + (iterOddInv (GroupStr.inv (snd G)) + invInv m (transport p q) g) + ∙ GroupStr.·InvR (snd G) g + ... | inr q = cong₂ (GroupStr._·_ (snd G)) + (iterOddInv (GroupStr.inv (snd G)) + invInv n q g) + (iterEvenInv (GroupStr.inv (snd G)) + invInv m + (transport (even≡odd→odd≡even n m p) q) g) + ∙ GroupStr.·InvL (snd G) g + congIdLeft≡congIdRight : (_·G_ : G → G → G) (-G_ : G → G) (0G : G) (rUnitG : (x : G) → x ·G 0G ≡ x) diff --git a/Cubical/Data/Nat/Base.agda b/Cubical/Data/Nat/Base.agda index 12b21c4c86..b786f08a62 100644 --- a/Cubical/Data/Nat/Base.agda +++ b/Cubical/Data/Nat/Base.agda @@ -34,6 +34,11 @@ iter : ∀ {ℓ} {A : Type ℓ} → ℕ → (A → A) → A → A iter zero f z = z iter (suc n) f z = f (iter n f z) +iter+ : ∀ {ℓ} {A : Type ℓ} (n m : ℕ) (f : A → A) (z : A) + → iter (n + m) f z ≡ iter n f (iter m f z) +iter+ zero m f z _ = iter m f z +iter+ (suc n) m f z i = f (iter+ n m f z i) + elim : ∀ {ℓ} {A : ℕ → Type ℓ} → A zero → ((n : ℕ) → A n → A (suc n)) @@ -56,10 +61,9 @@ isOdd zero = false isOdd (suc n) = isEven n --Typed version -private - toType : Bool → Type - toType false = ⊥ - toType true = Unit +toType : Bool → Type +toType false = ⊥ +toType true = Unit isEvenT : ℕ → Type isEvenT n = toType (isEven n) @@ -67,6 +71,9 @@ isEvenT n = toType (isEven n) isOddT : ℕ → Type isOddT n = isEvenT (suc n) +evenOrOddT : (n : ℕ) → Type +evenOrOddT n = isEvenT n ⊎ isOddT n + isZero : ℕ → Bool isZero zero = true isZero (suc n) = false diff --git a/Cubical/Data/Nat/IsEven.agda b/Cubical/Data/Nat/IsEven.agda index 9a8fc9a6fc..d3c4f0d195 100644 --- a/Cubical/Data/Nat/IsEven.agda +++ b/Cubical/Data/Nat/IsEven.agda @@ -1,11 +1,15 @@ module Cubical.Data.Nat.IsEven where open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Path +open import Cubical.Foundations.GroupoidLaws open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Bool open import Cubical.Data.Nat open import Cubical.Data.Sum +open import Cubical.Data.Nat.Mod -- negation result ¬IsEvenFalse : (n : ℕ) → (isEven n ≡ false) → isOdd n ≡ true @@ -70,3 +74,168 @@ evenOrOdd-Odd : (n : ℕ) → (isEven n ≡ false) → Σ[ x ∈ isOddT n ] even evenOrOdd-Odd zero p = ⊥.rec (true≢false p) evenOrOdd-Odd (suc zero) p = tt , refl evenOrOdd-Odd (suc (suc n)) p = evenOrOdd-Odd n p + +isEvenT→isEvenTrue : (n : ℕ) → isEvenT n → (isEven n ≡ true) +isEvenT→isEvenTrue zero x = refl +isEvenT→isEvenTrue (suc (suc n)) x = isEvenT→isEvenTrue n x + +isEvenT→isOddFalse : (n : ℕ) → isEvenT n → (isOdd n ≡ false) +isEvenT→isOddFalse zero x = refl +isEvenT→isOddFalse (suc (suc n)) x = isEvenT→isOddFalse n x + +isOddT→isEvenFalse : (n : ℕ) → isOddT n → (isEven n ≡ false) +isOddT→isEvenFalse (suc n) x = isEvenT→isOddFalse n x + +isOddT→isOddTrue : (n : ℕ) → isOddT n → (isOdd n ≡ true) +isOddT→isOddTrue (suc n) x = isEvenT→isEvenTrue n x + +------------ + +even+even≡even : (n m : ℕ) → isEvenT n → isEvenT m → isEvenT (n + m) +even+even≡even zero m p q = q +even+even≡even (suc (suc n)) m p q = even+even≡even n m p q + +even+odd≡odd : (n m : ℕ) → isEvenT n → isOddT m → isOddT (n + m) +even+odd≡odd zero m p q = q +even+odd≡odd (suc (suc n)) m p q = even+odd≡odd n m p q + +odd+even≡odd : (n m : ℕ) → isOddT n → isEvenT m → isOddT (n + m) +odd+even≡odd n m p q = subst isOddT (+-comm m n) (even+odd≡odd m n q p) + +odd+odd≡even : (n m : ℕ) → isOddT n → isOddT m → isEvenT (n + m) +odd+odd≡even (suc n) (suc m) p q = + subst isEvenT (cong suc (sym (+-suc n m))) + (even+even≡even n m p q) + +isEven·2 : (n : ℕ) → isEvenT (n + n) +isEven·2 zero = tt +isEven·2 (suc n) = + subst isEvenT (cong suc (+-comm (suc n) n)) (isEven·2 n) + +even·x≡even : (n m : ℕ) → isEvenT n → isEvenT (n · m) +even·x≡even zero m p = tt +even·x≡even (suc (suc n)) m p = + subst isEvenT (sym (+-assoc m m (n · m))) + (even+even≡even (m + m) (n · m) (isEven·2 m) (even·x≡even n m p)) + +x·even≡even : (n m : ℕ) → isEvenT m → isEvenT (n · m) +x·even≡even n m p = subst isEvenT (·-comm m n) (even·x≡even m n p) + +odd·odd≡odd : (n m : ℕ) → isOddT n → isOddT m → isOddT (n · m) +odd·odd≡odd (suc n) (suc m) p q = + subst isOddT t (even+even≡even m _ q + (even+even≡even n _ p (even·x≡even n m p))) + where + t : suc (m + (n + n · m)) ≡ suc (m + n · suc m) + t = cong suc (cong (m +_) (cong (n +_) (·-comm n m) ∙ ·-comm (suc m) n)) + +even≡odd→odd≡even : (n m : ℕ) → isEvenT n ≡ isOddT m → isOddT n ≡ isEvenT m +even≡odd→odd≡even zero zero p = sym p +even≡odd→odd≡even zero (suc zero) p = refl +even≡odd→odd≡even zero (suc (suc m)) p = even≡odd→odd≡even zero m p +even≡odd→odd≡even (suc zero) zero p = refl +even≡odd→odd≡even (suc zero) (suc zero) p = sym p +even≡odd→odd≡even (suc zero) (suc (suc m)) p = even≡odd→odd≡even (suc zero) m p +even≡odd→odd≡even (suc (suc n)) m p = even≡odd→odd≡even n m p + +-- Relation to mod 2 +isEvenT↔≡0 : (n : ℕ) → isEvenT n ↔ ((n mod 2) ≡ 0) +isEvenT↔≡0 zero .fst _ = refl +isEvenT↔≡0 zero .snd _ = tt +isEvenT↔≡0 (suc zero) .fst () +isEvenT↔≡0 (suc zero) .snd = snotz +isEvenT↔≡0 (suc (suc n)) = + comp↔ (isEvenT↔≡0 n) + (subst (λ x → (modInd 1 n ≡ 0) ↔ (x ≡ 0)) + (sym (mod+mod≡mod 2 2 n + ∙ cong (_mod 2) main ∙ mod-idempotent n)) id↔) + where + main : ((2 mod 2) + (n mod 2)) ≡ n mod 2 + main = cong (_+ (n mod 2)) (zero-charac 2) + +isOddT↔≡1 : (n : ℕ) → isOddT n ↔ ((n mod 2) ≡ 1) +isOddT↔≡1 zero .fst () +isOddT↔≡1 zero .snd x = snotz (sym x) +isOddT↔≡1 (suc zero) .fst _ = refl +isOddT↔≡1 (suc zero) .snd _ = tt +isOddT↔≡1 (suc (suc n)) = + comp↔ (isOddT↔≡1 n) + (subst (λ x → (modInd 1 n ≡ 1) ↔ (x ≡ 1)) + (sym (mod-idempotent n) ∙ sym (mod+mod≡mod 2 2 n)) id↔) + +-- characterisation of even/odd iterations of involutions +module _ {ℓ} {A : Type ℓ} (invA : A → A) + (invol : (a : A) → invA (invA a) ≡ a) where + iterEvenInv : (k : ℕ) → isEvenT k → (a : A) → iter k invA a ≡ a + iterEvenInv zero evk a = refl + iterEvenInv (suc (suc k)) evk a = + invol (iter k invA a) ∙ iterEvenInv k evk a + + iterOddInv : (k : ℕ) → isOddT k → (a : A) → iter k invA a ≡ invA a + iterOddInv (suc k) p a = cong invA (iterEvenInv k p a) + + iter+iter : (k : ℕ) (a : A) → iter k invA (iter k invA a) ≡ a + iter+iter k a = sym (iter+ k k invA a) ∙ iterEvenInv (k + k) (isEven·2 k) a + +-- pointed versions +module _ {ℓ} {A : Pointed ℓ} (invA : A →∙ A) + (invol : invA ∘∙ invA ≡ idfun∙ A) where + private -- technical lemmas + sqEven' : (k : ℕ) (p : isEvenT k) + → iterEvenInv (fst invA) (funExt⁻ (cong fst invol)) k p (pt A) + ≡ iter∙ k invA .snd + sqEven' zero p = refl + sqEven' (suc (suc k)) p = + (rUnit _ + ∙ (λ i → ((λ j → fst (invol (~ i ∧ j)) (iter k (fst invA) (pt A))) + ∙ (λ j → fst (invol (~ i)) (iterEvenInv (fst invA) + (funExt⁻ (λ i₁ → fst (invol i₁))) k p (pt A) j))) + ∙ λ j → fst (invol (~ i ∨ j)) (pt A) ) + ∙ cong₂ _∙_ (sym (lUnit _)) + (rUnit (funExt⁻ (cong fst invol) (pt A)) + ∙ cong₂ _∙_ refl (rUnit refl) + ∙ PathP→compPathL (symP (cong snd invol))) + ∙ assoc (cong (fst invA) (cong (fst invA) (iterEvenInv (fst invA) + (funExt⁻ (cong fst invol)) k p (pt A)))) + (cong (fst invA) (snd invA)) + (snd invA)) + ∙ cong₂ _∙_ (sym (cong-∙ (fst invA) + (cong (fst invA) (iterEvenInv (fst invA) + (funExt⁻ (cong fst invol)) k p (pt A))) + (snd invA)) + ∙ cong (congS (fst invA)) + (cong₂ _∙_ (cong (cong (fst invA)) + (sqEven' k p)) refl)) refl + + sqOdd' : (k : ℕ) (p : isOddT k) + → iterOddInv (fst invA) (funExt⁻ (cong fst invol)) k p (pt A) ∙ snd invA + ≡ iter∙ k invA .snd + sqOdd' (suc zero) p = refl + sqOdd' (suc (suc (suc k))) p = + cong₂ _∙_ (cong (congS (fst invA)) + (sqEven' (suc (suc k)) p)) + refl + + sqEven : (k : ℕ) (p : isEvenT k) + → Square (iterEvenInv (fst invA) (funExt⁻ (cong fst invol)) k p (pt A)) + refl (iter∙ k invA .snd) refl + sqEven k p = sqEven' k p ◁ λ i j → iter∙ k invA .snd (i ∨ j) + + + sqOdd : (k : ℕ) (p : isOddT k) + → Square (iterOddInv (fst invA) (funExt⁻ (cong fst invol)) k p (pt A)) + refl (iter∙ k invA .snd) (snd invA) + sqOdd k p = flipSquare (sym (sqOdd' k p) + ◁ symP (compPath-filler' (iterOddInv (fst invA) + (funExt⁻ (cong (λ r → fst r) invol)) k p (pt A)) + (snd invA))) + + iter∙EvenInv : (k : ℕ) → isEvenT k → iter∙ k invA ≡ idfun∙ A + iter∙EvenInv k p i .fst x = + iterEvenInv (fst invA) (funExt⁻ (cong fst invol)) k p x i + iter∙EvenInv k p i .snd j = sqEven k p j i + + iter∙OddInv : (k : ℕ) → isOddT k → iter∙ k invA ≡ invA + iter∙OddInv k p i .fst x = + iterOddInv (fst invA) (funExt⁻ (cong fst invol)) k p x i + iter∙OddInv k p i .snd j = sqOdd k p j i diff --git a/Cubical/Data/Vec/Base.agda b/Cubical/Data/Vec/Base.agda index be5b35d011..dc358c6b64 100644 --- a/Cubical/Data/Vec/Base.agda +++ b/Cubical/Data/Vec/Base.agda @@ -35,6 +35,12 @@ map : ∀ {A : Type ℓ} {B : Type ℓ'} {n} → (A → B) → Vec A n → Vec B map f [] = [] map f (x ∷ xs) = f x ∷ map f xs +compMap : {C : Type ℓ''} {n : ℕ} + → (f : B → C) (g : A → B) (v : Vec A n) + → map (λ x → f (g x)) v ≡ map f (map g v) +compMap f g [] = refl +compMap f g (x ∷ v) i = f (g x) ∷ compMap f g v i + replicate : ∀ {n} {A : Type ℓ} → A → Vec A n replicate {n = zero} x = [] replicate {n = suc n} x = x ∷ replicate x @@ -63,4 +69,3 @@ concat (xs ∷ xss) = xs ++ concat xss lookup : ∀ {n} {A : Type ℓ} → Fin n → Vec A n → A lookup zero (x ∷ xs) = x lookup (suc i) (x ∷ xs) = lookup i xs - diff --git a/Cubical/Foundations/Equiv.agda b/Cubical/Foundations/Equiv.agda index 0c78a79647..bcb85ed46c 100644 --- a/Cubical/Foundations/Equiv.agda +++ b/Cubical/Foundations/Equiv.agda @@ -19,6 +19,7 @@ module Cubical.Foundations.Equiv where open import Cubical.Foundations.Function open import Cubical.Foundations.Prelude open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Equiv.Base public open import Cubical.Data.Sigma.Base @@ -323,3 +324,29 @@ isEquiv≃isEquiv' f = isoToEquiv (isEquiv-isEquiv'-Iso f) -- The fact that funExt is an equivalence can be found in Cubical.Functions.FunExtEquiv +-- characterisation of retract proof produced by isoToEquiv +retEqIsoToEquiv : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} + (is : Iso A B) (x : _) + → retEq (isoToEquiv is) x + ≡ ((sym (leftInv is (inv is (fun is x))) + ∙ cong (inv is) ((rightInv is (fun is x))))) + ∙ leftInv is x +retEqIsoToEquiv is x i j = + hcomp (λ k → λ {(i = i1) + → compPath-filler (sym (leftInv is (inv is (fun is x))) + ∙ cong (inv is) ((rightInv is (fun is x)))) + (leftInv is x) k j + ; (j = i0) → (cong (inv is) (sym (rightInv is (fun is x))) + ∙ leftInv is (inv is (fun is x))) (i ∨ k) + ; (j = i1) → lUnit (leftInv is x) (~ i) k}) + (lemma j i) + where + p = sym (symDistr (sym (leftInv is (inv is (fun is x)))) + (cong (inv is) (rightInv is (fun is x)))) + + lemma : Square (cong (inv is) (sym (rightInv is (fun is x))) + ∙ leftInv is (inv is (fun is x))) + refl refl + (sym (leftInv is (inv is (fun is x))) + ∙ cong (inv is) ((rightInv is (fun is x)))) + lemma = p ◁ λ i j → p i1 (~ i ∧ j) diff --git a/Cubical/Foundations/Pointed/Base.agda b/Cubical/Foundations/Pointed/Base.agda index 91a7f020af..62718042fe 100644 --- a/Cubical/Foundations/Pointed/Base.agda +++ b/Cubical/Foundations/Pointed/Base.agda @@ -10,6 +10,8 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Univalence open import Cubical.Foundations.Function +open import Cubical.Data.Nat + private variable ℓ ℓ' : Level @@ -39,6 +41,16 @@ idfun∙ : (A : Pointed ℓ) → A →∙ A idfun∙ A .fst x = x idfun∙ A .snd = refl +iter∙ : {A : Pointed ℓ} (k : ℕ) → A →∙ A → A →∙ A +iter∙ k f .fst = iter k (fst f) +iter∙ zero f .snd = refl +iter∙ (suc k) f .snd = cong (fst f) (iter∙ k f .snd) ∙ snd f + +subst∙ : ∀ {ℓ ℓA} {X : Type ℓ} (A : X → Pointed ℓA) + → {x y : X} (p : x ≡ y) → A x →∙ A y +subst∙ A p .fst = subst (fst ∘ A) p +subst∙ A p .snd i = transp (λ j → fst (A (p (i ∨ j)))) i (pt (A (p i))) + infix 3 _≃∙_ {-Pointed equivalences -} _≃∙_ : (A : Pointed ℓ) (B : Pointed ℓ') → Type (ℓ-max ℓ ℓ') diff --git a/Cubical/Foundations/Pointed/Properties.agda b/Cubical/Foundations/Pointed/Properties.agda index 79fffdf880..a849d2820e 100644 --- a/Cubical/Foundations/Pointed/Properties.agda +++ b/Cubical/Foundations/Pointed/Properties.agda @@ -262,3 +262,20 @@ constantPointed≡ {A = A} {B = B , b} f Hf i = ; (i = i1) → snd f k ; (j = i1) → snd f k }) (Hf ((~ i) ∧ (~ j)) (pt A)) + + +-- Properties of subst∙ +subst∙refl : ∀ {ℓ ℓA} {X : Type ℓ} (A : X → Pointed ℓA) + → {x : X} → subst∙ A (refl {x = x}) ≡ idfun∙ (A x) +subst∙refl A {x} = ΣPathP ((funExt transportRefl) + , (λ j i → transp (λ t → fst (A x)) (j ∨ i) (pt (A x)))) + +subst∙Id : ∀ {ℓ ℓA ℓB} {X : Type ℓ} (A : X → Pointed ℓA) {B : Pointed ℓB} + → {x y : X} (p : x ≡ y) (f : A x →∙ B) + → f ∘∙ subst∙ A (sym p) ≡ transport (λ i → A (p i) →∙ B) f +subst∙Id A {B = B} {x = x} = + J (λ y p → (f : A x →∙ B) + → f ∘∙ subst∙ A (sym p) + ≡ transport (λ i → A (p i) →∙ B) f) + λ f → (cong₂ _∘∙_ refl (subst∙refl A) ∙ ∘∙-idˡ _) + ∙ sym (transportRefl f) diff --git a/Cubical/Foundations/Prelude.agda b/Cubical/Foundations/Prelude.agda index a9454e5875..92545c2ed8 100644 --- a/Cubical/Foundations/Prelude.agda +++ b/Cubical/Foundations/Prelude.agda @@ -19,6 +19,8 @@ This file proves a variety of basic results about paths: - Export universe lifting +- Biimplication + -} module Cubical.Foundations.Prelude where @@ -637,3 +639,16 @@ liftExt x i = lift (x i) liftFun : ∀ {ℓ ℓ' ℓ'' ℓ'''} {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Lift {j = ℓ''} A → Lift {j = ℓ'''} B liftFun f (lift a) = lift (f a) + +-- Biimplication +_↔_ : (A : Type ℓ) (B : Type ℓ') → Type (ℓ-max ℓ ℓ') +A ↔ B = Σ (A → B) λ _ → (B → A) + +id↔ : A ↔ A +id↔ .fst x = x +id↔ .snd x = x + +comp↔ : {B : Type ℓ'} {C : Type ℓ''} + → A ↔ B → B ↔ C → A ↔ C +comp↔ (f , g) (h , r) .fst x = h (f x) +comp↔ (f , g) (h , r) .snd x = g (r x) diff --git a/Cubical/HITs/Join/CoHSpace.agda b/Cubical/HITs/Join/CoHSpace.agda new file mode 100644 index 0000000000..087c56486d --- /dev/null +++ b/Cubical/HITs/Join/CoHSpace.agda @@ -0,0 +1,350 @@ +{- +This file contains a definition of the co-H-space structure on +joins and a proof that it is equivalent to that on suspensions +-} + +{-# OPTIONS --safe --lossy-unification #-} + +module Cubical.HITs.Join.CoHSpace where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Function +open import Cubical.Foundations.Path +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Path + +open import Cubical.Data.Sigma renaming (fst to proj₁; snd to proj₂) +open import Cubical.Data.Unit +open import Cubical.Data.Nat +open import Cubical.Data.Sum + +open import Cubical.HITs.Join +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp renaming (toSusp to σ) + +open import Cubical.HITs.SmashProduct +open import Cubical.HITs.Pushout +open import Cubical.Foundations.Pointed.Homogeneous + +open import Cubical.Homotopy.Loopspace + +private + variable + ℓ ℓ' ℓ'' : Level + +open Iso + +-- Standard loop in Ω (join A B) +ℓ* : (A : Pointed ℓ) (B : Pointed ℓ') + → fst A → fst B → Ω (join∙ A B) .fst +ℓ* A B a b = push (pt A) (pt B) + ∙ (push a (pt B) ⁻¹ ∙∙ push a b ∙∙ (push (pt A) b ⁻¹)) + +ℓ*IdL : (A : Pointed ℓ) (B : Pointed ℓ') + → (b : fst B) → ℓ* A B (pt A) b ≡ refl +ℓ*IdL A B b = + cong₂ _∙_ refl + (doubleCompPath≡compPath _ _ _ + ∙ cong (push (pt A) (pt B) ⁻¹ ∙_) (rCancel _) + ∙ sym (rUnit _)) + ∙ rCancel _ + +ℓ*IdR : (A : Pointed ℓ) (B : Pointed ℓ') + → (a : typ A) → ℓ* A B a (pt B) ≡ refl +ℓ*IdR A B a = + cong₂ _∙_ refl + (doubleCompPath≡compPath _ _ _ + ∙ assoc _ _ _ + ∙ cong (_∙ push (pt A) (pt B) ⁻¹) (rCancel _) + ∙ sym (lUnit _)) + ∙ rCancel _ + +-- Addition of functions join A B → C +_+*_ : {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + (f g : join∙ A B →∙ C) → join∙ A B →∙ C +fst (_+*_ {C = C} f g) (inl x) = pt C +fst (_+*_ {C = C} f g) (inr x) = pt C +fst (_+*_ {A = A} {B = B} f g) (push a b i) = + (Ω→ f .fst (ℓ* A B a b) ∙ Ω→ g .fst (ℓ* A B a b)) i +snd (f +* g) = refl + +-- Inversion +-* : {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + (f : join∙ A B →∙ C) → join∙ A B →∙ C +fst (-* {C = C} f) (inl x) = pt C +fst (-* {C = C} f) (inr x) = pt C +fst (-* {A = A} {B} f) (push a b i) = Ω→ f .fst (ℓ* A B a b) (~ i) +snd (-* f) = refl + +-- Iterated inversion +-*^ : {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + (n : ℕ) (f : join∙ A B →∙ C) + → join∙ A B →∙ C +-*^ n = iter n -* + +-- Neutral element +0* : {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + → join∙ A B →∙ C +0* = const∙ _ _ + +-- Mapping space join∙ A B →∙ C is equivalent to Susp (A ⋀ B) →∙ C +fromSusp≅fromJoin : {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + → Iso (Susp∙ (A ⋀ B) →∙ C) (join∙ A B →∙ C) +fromSusp≅fromJoin {A = A} {B = B} = + post∘∙equiv (isoToEquiv SmashJoinIso , sym (push (pt A) (pt B))) + +-- Goal now: show that this is an equivalence of H-spaces + +-- Technical lemma +Ω→ℓ⋀ : {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + (f : join∙ A B →∙ C) + (x : fst A) (y : fst B) + → Ω→ f .fst (ℓ* A B x y) + ≡ Ω→ (inv fromSusp≅fromJoin f) .fst (σ (A ⋀∙ B) (inr (x , y))) +Ω→ℓ⋀ f x y = + (cong₃ _∙∙_∙∙_ (symDistr _ _) + (cong-∙ (fst (inv fromSusp≅fromJoin f)) _ _ ∙ cong₂ _∙_ refl refl) + refl + ∙ doubleCompPath≡compPath _ _ _ + ∙ sym (assoc _ _ _) + ∙ cong (snd f ⁻¹ ∙_) (assoc _ _ _ ∙ assoc _ _ _) + ∙ sym (doubleCompPath≡compPath _ _ _) + ∙ cong₃ _∙∙_∙∙_ + refl + (sym (assoc _ _ _) + ∙ cong₂ _∙_ refl + ((sym (assoc _ _ _) ∙ cong₂ _∙_ refl (rCancel _)) ∙ sym (rUnit _)) + ∙ sym (cong-∙ (fst f) _ _)) + refl) ⁻¹ + +-- We'll show it for the inverse function +-- Inverse function preseves neutral elements +fromSusp≅fromJoin⁻Pres0* : (A : Pointed ℓ) (B : Pointed ℓ') {C : Pointed ℓ''} + → inv fromSusp≅fromJoin (const∙ (join∙ A B) C) ≡ const∙ _ _ +fromSusp≅fromJoin⁻Pres0* A B = ΣPathP (refl , (sym (rUnit refl))) + +-- Inverse function preseves +* +fromSusp≅fromJoin⁻Pres+* : {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + → (f g : join∙ A B →∙ C) + → Iso.inv fromSusp≅fromJoin (f +* g) + ≡ ·Susp (A ⋀∙ B) (Iso.inv fromSusp≅fromJoin f) + (Iso.inv fromSusp≅fromJoin g) +fromSusp≅fromJoin⁻Pres+* {A = A} {B = B} f g = + ΣPathP ((funExt λ { north → refl + ; south → refl + ; (merid a i) j → main j a i}) + , (sym (rUnit _) + ∙ cong sym (cong₂ _∙_ (cong (Ω→ f .fst) (ℓ*IdR A B (pt A)) ∙ Ω→ f .snd) + (cong (Ω→ g .fst) (ℓ*IdR A B (pt A)) ∙ Ω→ g .snd)) + ∙ sym (rUnit _))) + where + main : cong (fst (f +* g) ∘ (SuspSmash→Join {A = A} {B = B})) ∘ merid + ≡ λ a → Ω→ (Iso.inv fromSusp≅fromJoin f) .fst (σ (A ⋀∙ B) a) + ∙ Ω→ (Iso.inv fromSusp≅fromJoin g) .fst (σ (A ⋀∙ B) a) + main = ⋀→Homogeneous≡ (isHomogeneousPath _ _) + λ x y → cong-∙∙ (fst (f +* g)) _ _ _ + ∙ cong₃ _∙∙_∙∙_ + (cong sym (cong₂ _∙_ (cong (Ω→ f .fst) (ℓ*IdR A B x) + ∙ Ω→ f .snd) + (cong (Ω→ g .fst) (ℓ*IdR A B x) + ∙ Ω→ g .snd) + ∙ sym (rUnit refl))) + refl + (cong sym ((cong₂ _∙_ (cong (Ω→ f .fst) (ℓ*IdL A B y) + ∙ Ω→ f .snd) + (cong (Ω→ g .fst) (ℓ*IdL A B y) + ∙ Ω→ g .snd) + ∙ sym (rUnit refl)))) + ∙ sym (rUnit _) + ∙ cong₂ _∙_ (Ω→ℓ⋀ f x y) (Ω→ℓ⋀ g x y) + +-- Inverse function preseves inversion +fromSusp≅fromJoin⁻Pres-* : + {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + → (f : join∙ A B →∙ C) + → inv fromSusp≅fromJoin (-* f) + ≡ -Susp (A ⋀∙ B) (inv fromSusp≅fromJoin f) +fromSusp≅fromJoin⁻Pres-* {A = A} {B = B} {C = C} f = + ΣPathP ((funExt (λ { north → refl + ; south → refl + ; (merid a i) j → lem j a i})) + , sym (rUnit _) + ∙ (cong (Ω→ f .fst) (ℓ*IdR A B (pt A)) + ∙ Ω→ f .snd)) + where + lem : cong (fst (-* f) ∘ SuspSmash→Join) ∘ merid + ≡ λ a → cong (-Susp (A ⋀∙ B) (inv fromSusp≅fromJoin f) .fst) (merid a) + lem = ⋀→Homogeneous≡ (isHomogeneousPath _ _) + λ x y → cong-∙∙ (fst (-* f)) _ _ _ + ∙∙ cong₃ _∙∙_∙∙_ ((cong (Ω→ f .fst) (ℓ*IdR A B x) ∙ Ω→ f .snd)) + refl + (cong (Ω→ f .fst) (ℓ*IdL A B y) ∙ Ω→ f .snd) + ∙ cong₂ _∙_ (doubleCompPath≡compPath _ _ _ + ∙ (cong (sym (snd f) ∙_) + (cong (_∙ snd f) + (cong sym (cong-∙ (fst f) _ _) + ∙ symDistr _ _ + ∙ (cong₂ _∙_ (lUnit _ + ∙ cong₂ _∙_ (sym (rCancel _)) refl + ∙ sym (assoc _ _ _)) refl + ∙ sym (assoc _ _ _)) + ∙ cong₂ _∙_ + refl + (cong₂ _∙_ (sym (symDistr _ _)) refl))) + ∙ (assoc _ _ _ + ∙ cong₂ _∙_ (assoc _ _ _ ∙ assoc _ _ _) refl) + ∙ sym (assoc _ _ _)) + ∙ sym (assoc _ _ _) + ∙ doubleCompPath≡compPath _ _ _ ⁻¹ -- + ∙ cong₃ _∙∙_∙∙_ + (sym (symDistr _ _)) + (cong sym (sym (cong-∙ (fst (inv fromSusp≅fromJoin f)) _ _))) + refl) + (sym ((cong (Ω→ (inv fromSusp≅fromJoin f) .fst) + (rCancel (merid (inl tt)))) + ∙ ∙∙lCancel _)) + ∙∙ sym (cong-∙ (-Susp (A ⋀∙ B) (inv fromSusp≅fromJoin f) .fst) _ _) + ∙∙ rUnit _ + ∙∙ Ω→-Susp _ (inv fromSusp≅fromJoin f) (inr (x , y)) -- + +-- Same lemmas again, this time the other direction +fromSusp≅fromJoinPres+* : + {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + → (f g : Susp∙ (A ⋀ B) →∙ C) + → Iso.fun fromSusp≅fromJoin (·Susp (A ⋀∙ B) f g) + ≡ (Iso.fun fromSusp≅fromJoin f) +* (Iso.fun fromSusp≅fromJoin g) +fromSusp≅fromJoinPres+* {A = A} {B = B} f g = + cong (fun fromSusp≅fromJoin) + (cong₂ (·Susp (A ⋀∙ B)) + (sym (leftInv fromSusp≅fromJoin f)) + (sym (leftInv fromSusp≅fromJoin g))) + ∙∙ sym (cong (fun fromSusp≅fromJoin) + (fromSusp≅fromJoin⁻Pres+* + (fun fromSusp≅fromJoin f) + (fun fromSusp≅fromJoin g))) + ∙∙ rightInv fromSusp≅fromJoin _ + +fromSusp≅fromJoinPres-* : + {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + → (f : Susp∙ (A ⋀ B) →∙ C) + → fun fromSusp≅fromJoin (-Susp (A ⋀∙ B) f) + ≡ -* (fun fromSusp≅fromJoin f) +fromSusp≅fromJoinPres-* {A = A} {B = B} f = + cong (fun fromSusp≅fromJoin ∘ -Susp (A ⋀∙ B)) + (sym (leftInv fromSusp≅fromJoin f)) + ∙ sym (cong (fun fromSusp≅fromJoin) + (fromSusp≅fromJoin⁻Pres-* (fun fromSusp≅fromJoin f))) + ∙ rightInv fromSusp≅fromJoin _ + +fromSusp≅fromJoinPres0* : (A : Pointed ℓ) (B : Pointed ℓ') {C : Pointed ℓ''} + → fun fromSusp≅fromJoin (const∙ (Susp∙ (A ⋀ B)) C) + ≡ const∙ _ _ +fromSusp≅fromJoinPres0* A B = ΣPathP (refl , (sym (rUnit refl))) + + +--- We use the above to transport the HSpace proofs from suspensions to joins +module _ {ℓ ℓ' ℓ'' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} where + private + P = isoToPath (fromSusp≅fromJoin {A = A} {B = B} {C = C}) + ·Susp-+*-PathP : PathP (λ i → P i → P i → P i) (·Susp (A ⋀∙ B)) _+*_ + ·Susp-+*-PathP = + toPathP (funExt λ f → funExt λ g + → transportRefl _ + ∙ (cong₂ _∘∙_ (cong₂ (·Susp (A ⋀∙ B)) + (cong₂ _∘∙_ (transportRefl f) refl) + (cong₂ _∘∙_ (transportRefl g) refl)) refl) + ∙ fromSusp≅fromJoinPres+* (inv fromSusp≅fromJoin f) (inv fromSusp≅fromJoin g) + ∙ cong₂ _+*_ (rightInv fromSusp≅fromJoin f) (rightInv fromSusp≅fromJoin g)) + + ·Susp--*-PathP : PathP (λ i → P i → P i) (-Susp (A ⋀∙ B)) -* + ·Susp--*-PathP = + toPathP (funExt λ x → transportRefl _ + ∙ cong₂ _∘∙_ (cong (-Susp (A ⋀∙ B)) (cong₂ _∘∙_ (transportRefl _) refl)) refl + ∙ fromSusp≅fromJoinPres-* (inv fromSusp≅fromJoin x) + ∙ cong -* (rightInv fromSusp≅fromJoin x)) + + ·Susp-0*-PathP : PathP (λ i → P i) (const∙ _ _) (const∙ _ _) + ·Susp-0*-PathP = symP (toPathP (cong₂ _∘∙_ (transportRefl _) refl + ∙ ΣPathP (refl , (sym (rUnit _))))) + + -- The laws + +*HSpace : Σ[ r ∈ ((f : join∙ A B →∙ C) → (f +* 0*) ≡ f) ] + Σ[ l ∈ ((f : join∙ A B →∙ C) → (0* +* f) ≡ f) ] + r 0* ≡ l 0* + +*HSpace = + transport (λ i + → Σ[ r ∈ ((f : P i) → ·Susp-+*-PathP i f (·Susp-0*-PathP i) ≡ f) ] + Σ[ l ∈ ((f : P i) → ·Susp-+*-PathP i (·Susp-0*-PathP i) f ≡ f) ] + r (·Susp-0*-PathP i) ≡ l (·Susp-0*-PathP i)) + ((·SuspIdR _) , (·SuspIdL _) , (sym (·SuspIdL≡·SuspIdR _))) + + +*IdR : (f : join∙ A B →∙ C) → (f +* 0*) ≡ f + +*IdR = +*HSpace .fst + + +*IdL : (f : join∙ A B →∙ C) → (0* +* f) ≡ f + +*IdL = +*HSpace .snd .fst + + +*InvR : (f : join∙ A B →∙ C) → (f +* (-* f)) ≡ 0* + +*InvR = transport (λ i → (f : P i) → ·Susp-+*-PathP i f (·Susp--*-PathP i f) + ≡ ·Susp-0*-PathP i) + (·SuspInvR _) + + +*InvL : (f : join∙ A B →∙ C) → ((-* f) +* f) ≡ 0* + +*InvL = transport (λ i → (f : P i) → ·Susp-+*-PathP i (·Susp--*-PathP i f) f + ≡ ·Susp-0*-PathP i) + (·SuspInvL _) + + +*Assoc : (f g h : join∙ A B →∙ C) → (f +* (g +* h)) ≡ ((f +* g) +* h) + +*Assoc = + transport (λ i → (f g h : P i) + → ·Susp-+*-PathP i f (·Susp-+*-PathP i g h) + ≡ ·Susp-+*-PathP i (·Susp-+*-PathP i f g) h) + (·SuspAssoc _) + + +*Comm : (Σ[ A' ∈ Pointed ℓ ] A ≃∙ Susp∙ (fst A')) + ⊎ (Σ[ B' ∈ Pointed ℓ' ] B ≃∙ Susp∙ (fst B')) + → (f g : join∙ A B →∙ C) → (f +* g) ≡ (g +* f) + +*Comm (inl (A' , e)) = + transport (λ i → (f g : P i) + → ·Susp-+*-PathP i f g + ≡ ·Susp-+*-PathP i g f) + (·SuspComm' _ (A' ⋀∙ B) + (compEquiv∙ (⋀≃ e (idEquiv∙ B) , refl) + (isoToEquiv SuspSmashCommIso , refl))) + +*Comm (inr (B' , e)) = + transport (λ i → (f g : P i) + → ·Susp-+*-PathP i f g + ≡ ·Susp-+*-PathP i g f) + (·SuspComm' _ (A ⋀∙ B') + (compEquiv∙ + (compEquiv∙ (compEquiv∙ (⋀≃ (idEquiv∙ A) e , refl) + (isoToEquiv ⋀CommIso , refl)) + (isoToEquiv SuspSmashCommIso , refl)) + (congSuspEquiv (isoToEquiv ⋀CommIso) , refl))) + + -*² : (f : join∙ A B →∙ C) → -* (-* f) ≡ f + -*² f = + sym (+*IdR (-* (-* f))) + ∙ cong (-* (-* f) +*_) (sym (+*InvL f)) + ∙ +*Assoc _ _ _ + ∙ cong (_+* f) ( +*InvL (-* f)) + ∙ +*IdL f + where + help : (-* (-* f)) +* (-* f) ≡ f +* (-* f) + help = +*InvL (-* f) ∙ sym (+*InvR f) + +join-commFun-ℓ* : (A : Pointed ℓ) (B : Pointed ℓ') (a : fst A) (b : fst B) + → cong join-commFun (ℓ* A B a b) + ≡ (sym (push (pt B) (pt A)) ∙∙ sym (ℓ* B A b a) ∙∙ push (pt B) (pt A)) +join-commFun-ℓ* A B a b = + cong-∙ join-commFun (push (pt A) (pt B)) _ + ∙ cong₂ _∙_ refl (cong-∙∙ join-commFun _ _ _) + ∙ sym (doubleCompPath≡compPath _ _ _ + ∙ cong₂ _∙_ refl + (cong₂ _∙_ (symDistr _ _) refl + ∙ sym (assoc _ _ _) ∙ cong₂ _∙_ refl (lCancel _) + ∙ sym (rUnit _))) diff --git a/Cubical/HITs/Join/Properties.agda b/Cubical/HITs/Join/Properties.agda index 2cd9d519ef..44b0b5c35e 100644 --- a/Cubical/HITs/Join/Properties.agda +++ b/Cubical/HITs/Join/Properties.agda @@ -455,6 +455,10 @@ join-commFun (inl x) = inr x join-commFun (inr x) = inl x join-commFun (push a b i) = push b a (~ i) +join-commFun∙ : ∀ {ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} → join∙ A B →∙ join∙ B A +proj₁ join-commFun∙ = join-commFun +snd (join-commFun∙ {A = A} {B = B}) = push (pt B) (pt A) ⁻¹ + join-commFun² : ∀ {ℓ'} {A : Type ℓ} {B : Type ℓ'} (x : join A B) → join-commFun (join-commFun x) ≡ x join-commFun² (inl x) = refl @@ -711,3 +715,26 @@ module _ {A : Pointed ℓ} {B : Pointed ℓ'} (f : A →∙ B) where inv GaneaIso = join→GaneaFib rightInv GaneaIso = join→GaneaFib→join leftInv GaneaIso = GaneaFib→join→GaneaFib + +-- Pinch map +joinPinch : ∀ {ℓ''} {A : Type ℓ} {B : Type ℓ'} (C : Pointed ℓ'') + → ((a : A) (b : B) → Ω C .fst) → join A B → fst C +joinPinch C p (inl x) = pt C +joinPinch C p (inr x) = pt C +joinPinch C p (push a b i) = p a b i + +joinPinch∙ : ∀ {ℓ''} (A : Pointed ℓ) (B : Pointed ℓ') (C : Pointed ℓ'') + → ((a : typ A) (b : typ B) → Ω C .fst) → join∙ A B →∙ C +proj₁ (joinPinch∙ A B C p) = joinPinch C p +snd (joinPinch∙ A B C p) = refl + +joinPinchComp : ∀ {ℓ ℓ' ℓ'' ℓA ℓB} {X : Pointed ℓ} + {A : Type ℓA} {B : Type ℓB} + {A' : Type ℓ'} {B' : Type ℓ''} + (g : A → A') (h : B → B') + → (f : A' → B' → Ω X .fst) (x : join A B) + → joinPinch X f (join→ g h x) + ≡ joinPinch X (λ a b → f (g a) (h b)) x +joinPinchComp {X = X} g h f (inl x) = refl +joinPinchComp {X = X} g h f (inr x) = refl +joinPinchComp {X = X} g h f (push a b i) = refl diff --git a/Cubical/HITs/SmashProduct/Base.agda b/Cubical/HITs/SmashProduct/Base.agda index 17bc33b0c4..0dc721cfce 100644 --- a/Cubical/HITs/SmashProduct/Base.agda +++ b/Cubical/HITs/SmashProduct/Base.agda @@ -194,7 +194,7 @@ prod→⋀^ zero A x = x prod→⋀^ (suc n) A x = inr ((prod→⋀^ n (predFinFamily∙ A) (fst x)) , (snd x)) -⋀→Homogeneous≡ : {A B : Pointed ℓ} {D : Type ℓ'} +⋀→Homogeneous≡ : ∀ {ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {D : Type ℓ''} → {f g : A ⋀ B → D} → (isHomogeneous (D , f (inl tt))) → ((x : _) (y : _) → f (inr (x , y)) ≡ g (inr (x , y))) @@ -210,6 +210,21 @@ prod→⋀^ (suc n) A x = f^≡g^ : f^ ≡ g^ f^≡g^ = ⋀→∙Homogeneous≡ hom p +⋀→HomogeneousPathP : ∀ {ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {D D' : Type ℓ''} (P : D ≡ D') + → {f : A ⋀ B → D} + → {g : A ⋀ B → D'} + → (isHomogeneous (D , f (inl tt))) + → ((x : _) (y : _) + → PathP (λ i → P i) + (f (inr (x , y))) (g (inr (x , y)))) + → (x : _) → PathP (λ i → P i) (f x) (g x) +⋀→HomogeneousPathP {A = A} {B} {D} P {f = f} {g} hom p x = + toPathP (funExt⁻ (⋀→Homogeneous≡ {f = transport P ∘ f} {g} + (subst isHomogeneous + (ΣPathP (P , toPathP refl)) + hom) + λ x y → fromPathP (p x y)) x) + ⋀^→Homogeneous≡ : (n : ℕ) (A : Fin (suc n) → Pointed ℓ) {B : Type ℓ'} → {f g : ⋀^ n A .fst → B} → isHomogeneous (B , f (⋀^ n A .snd)) @@ -282,8 +297,11 @@ snd (f ⋀∙→refl g) = refl ⋀≃ : ∀ {ℓ ℓ'} {A B : Pointed ℓ} {C D : Pointed ℓ'} → (f : A ≃∙ B) (g : C ≃∙ D) → (A ⋀ C) ≃ (B ⋀ D) -⋀≃ {ℓ = ℓ} {ℓ'} {B = B} {D = D} f g = _ , ⋀≃-isEq f g +⋀≃ {ℓ = ℓ} {ℓ'} {A = A} {B = B} {C = C} {D = D} f g = isoToEquiv mainIso where + back : (B ⋀ D) → (A ⋀ C) + back = ≃∙map (invEquiv∙ f) ⋀→ ≃∙map (invEquiv∙ g) + help : (x : _) → (idfun∙ B ⋀→ idfun∙ D) x ≡ x help (inl x) = refl help (inr x) = refl @@ -298,13 +316,39 @@ snd (f ⋀∙→refl g) = refl (push (push tt i) j) ⋀≃-isEq : {A : Pointed ℓ} {C : Pointed ℓ'} - (f : A ≃∙ B) (g : C ≃∙ D) → isEquiv (≃∙map f ⋀→ ≃∙map g) + (f : A ≃∙ B) (g : C ≃∙ D) + → retract (≃∙map f ⋀→ ≃∙map g) + (≃∙map (invEquiv∙ f) ⋀→ ≃∙map (invEquiv∙ g)) + × section (≃∙map f ⋀→ ≃∙map g) + (≃∙map (invEquiv∙ f) ⋀→ ≃∙map (invEquiv∙ g)) ⋀≃-isEq {C = C} = Equiv∙J (λ A f → (g : C ≃∙ D) - → isEquiv (≃∙map f ⋀→ ≃∙map g)) - (Equiv∙J (λ _ g → isEquiv (idfun∙ _ ⋀→ ≃∙map g)) - (subst isEquiv (sym (funExt help)) (idIsEquiv _))) - + → retract (≃∙map f ⋀→ ≃∙map g) + (≃∙map (invEquiv∙ f) ⋀→ ≃∙map (invEquiv∙ g)) + × section (≃∙map f ⋀→ ≃∙map g) + (≃∙map (invEquiv∙ f) ⋀→ ≃∙map (invEquiv∙ g))) + (Equiv∙J (λ C g → + retract (idfun∙ B ⋀→ ≃∙map g) + (≃∙map (invEquiv∙ (idEquiv∙ B)) ⋀→ ≃∙map (invEquiv∙ g)) + × section (idfun∙ B ⋀→ ≃∙map g) + (≃∙map (invEquiv∙ (idEquiv∙ B)) ⋀→ ≃∙map (invEquiv∙ g))) + (subst2 (λ f g → + retract (idfun∙ B ⋀→ idfun∙ D) + (f ⋀→ g) + × section (idfun∙ B ⋀→ idfun∙ D) + (f ⋀→ g)) + (sym (lem B)) (sym (lem D)) + (subst (λ f → retract f f × section f f) (sym (funExt help)) + ((λ _ → refl) , (λ _ → refl))))) + where + lem : ∀ {ℓ} (B : Pointed ℓ) → ≃∙map (invEquiv∙ (idEquiv∙ B)) ≡ idfun∙ B + lem B = ΣPathP (refl , sym (rUnit refl)) + + mainIso : Iso (A ⋀ C) (B ⋀ D) + mainIso .Iso.fun = ≃∙map f ⋀→ ≃∙map g + mainIso .Iso.inv = ≃∙map (invEquiv∙ f) ⋀→ ≃∙map (invEquiv∙ g) + mainIso .Iso.rightInv = ⋀≃-isEq f g .snd + mainIso .Iso.leftInv = ⋀≃-isEq f g .fst ⋀→Smash : A ⋀ B → Smash A B ⋀→Smash (inl x) = basel @@ -1052,6 +1096,29 @@ module _ {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} where Iso.rightInv SmashJoinIso = Join→SuspSmash→Join Iso.leftInv SmashJoinIso = SuspSmash→Join→SuspSmash +-- Pointed versions +Join→SuspSmash∙ : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') + → join∙ A B →∙ Susp∙ (A ⋀ B) +Join→SuspSmash∙ A B = Join→SuspSmash , refl + +SuspSmash→Join∙ : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') + → Susp∙ (A ⋀ B) →∙ join∙ A B +SuspSmash→Join∙ A B = SuspSmash→Join , push (pt A) (pt B) ⁻¹ + +permute⋀JoinIso : ∀ {ℓ ℓ' ℓ''} + (A : Pointed ℓ) (B : Pointed ℓ') (C : Pointed ℓ'') + → Iso (join (A ⋀ B) (typ C)) (join (typ A) (B ⋀ C)) +permute⋀JoinIso A B C = + compIso (invIso (SmashJoinIso {A = A ⋀∙ B} {C})) + (compIso (congSuspIso (invIso SmashAssocIso)) + (SmashJoinIso {A = A} {B ⋀∙ C})) + +permute⋀Join≃∙ : ∀ {ℓ ℓ' ℓ''} + (A : Pointed ℓ) (B : Pointed ℓ') (C : Pointed ℓ'') + → join∙ (A ⋀∙ B) C ≃∙ join∙ A (B ⋀∙ C) +fst (permute⋀Join≃∙ A B C) = isoToEquiv (permute⋀JoinIso A B C) +snd (permute⋀Join≃∙ A B C) = sym (push (pt A) (inl tt)) + -- Suspension commutes with smash products module _ {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} where @@ -1285,3 +1352,14 @@ module _ {ℓ ℓ' : Level} {A : Pointed ℓ} {B : Pointed ℓ'} where main y north = push (inr y) main y south = push (inr y) ∙ λ i → inr (merid (pt A) i , y) main y (merid a i) j = filler a y i j i1 + +-- Join→SuspSmash commutes with suspension +SuspFun-Join→SuspSmash≡ : ∀ {ℓA ℓA' ℓB ℓB'} + {A : Pointed ℓA} {A' : Pointed ℓA'} + {B : Pointed ℓB} {B' : Pointed ℓB'} + (f : A →∙ A') (g : B →∙ B') (x : _) + → Join→SuspSmash {A = A'} {B = B'} (join→ (fst f) (fst g) x) + ≡ suspFun (f ⋀→ g) (Join→SuspSmash {A = A} {B} x) +SuspFun-Join→SuspSmash≡ f g (inl x) = refl +SuspFun-Join→SuspSmash≡ f g (inr x) = refl +SuspFun-Join→SuspSmash≡ f g (push a b i) = refl diff --git a/Cubical/HITs/Sn/Multiplication.agda b/Cubical/HITs/Sn/Multiplication.agda index 4d9c9a6a99..cc497f2814 100644 --- a/Cubical/HITs/Sn/Multiplication.agda +++ b/Cubical/HITs/Sn/Multiplication.agda @@ -21,9 +21,11 @@ open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Pointed.Homogeneous open import Cubical.Data.Sum -open import Cubical.Data.Bool hiding (elim) +open import Cubical.Data.Bool hiding (elim ; _≤_) open import Cubical.Data.Nat hiding (elim) +open import Cubical.Data.Nat.Order open import Cubical.Data.Sigma +open import Cubical.Data.Empty as ⊥ open import Cubical.HITs.S1 renaming (_·_ to _*_) open import Cubical.HITs.S2 @@ -298,6 +300,17 @@ isEquiv-⋀S (suc n) m = SphereSmashIso : (n m : ℕ) → Iso (S₊∙ n ⋀ S₊∙ m) (S₊ (n + m)) SphereSmashIso n m = equivToIso (⋀S n m , isEquiv-⋀S n m) +SphereSmashIso∙ : (n m : ℕ) + → Iso.fun (SphereSmashIso n m) (inl tt) ≡ ptSn (n + m) +SphereSmashIso∙ zero m = refl +SphereSmashIso∙ (suc n) m = refl + +SphereSmashIso⁻∙ : (n m : ℕ) + → Iso.inv (SphereSmashIso n m) (ptSn (n + m)) ≡ inl tt +SphereSmashIso⁻∙ n m = + sym (cong (Iso.inv (SphereSmashIso n m)) (SphereSmashIso∙ n m)) + ∙ Iso.leftInv (SphereSmashIso n m) (inl tt) + -- Proof that the pinch map Sⁿ * Sᵐ → Sⁿ⁺ᵐ⁺¹ is an equivalence join→Sphere : (n m : ℕ) → join (S₊ n) (S₊ m) → S₊ (suc (n + m)) @@ -339,8 +352,7 @@ join→Sphere≡ zero (suc m) (inl x) = refl join→Sphere≡ (suc n) m (inl x) = refl join→Sphere≡ zero zero (inr x) = refl join→Sphere≡ zero (suc m) (inr x) = merid (ptSn (suc m)) -join→Sphere≡ (suc n) zero (inr x) = merid (ptSn (suc n + zero)) -join→Sphere≡ (suc n) (suc m) (inr x) = merid (ptSn (suc n + suc m)) +join→Sphere≡ (suc n) m (inr x) = merid (ptSn (suc n + m)) join→Sphere≡ zero zero (push false false i) j = loop i join→Sphere≡ zero zero (push false true i) j = base join→Sphere≡ zero zero (push true b i) j = base @@ -476,12 +488,6 @@ assoc⌣S {n = suc (suc n)} {m = m} {l} (merid a i) y z j = (sym (merid (ptSn (suc (n + suc m))))))) -- Goal: graded commutativity - --- To state it: we'll need iterated negations --S^ : {k : ℕ} (n : ℕ) → S₊ k → S₊ k --S^ zero x = x --S^ (suc n) x = invSphere (-S^ n x) - -- The folowing is an explicit definition of -S^ (n · m) which is -- often easier to reason about -S^-expl : {k : ℕ} (n m : ℕ) @@ -502,37 +508,6 @@ assoc⌣S {n = suc (suc n)} {m = m} {l} (merid a i) y z j = -S^-expl {k = suc (suc k)} n m (inr x) (inl x₁) (merid a i) = σS a i -S^-expl {k = suc (suc k)} n m (inr x) (inr x₁) (merid a i) = σS a (~ i) --- invSphere commutes with S^ -invSphere-S^ : {k : ℕ} (n : ℕ) (x : S₊ k) - → invSphere (-S^ n x) ≡ -S^ n (invSphere x) -invSphere-S^ zero x = refl -invSphere-S^ (suc n) x = cong invSphere (invSphere-S^ n x) - --S^² : {k : ℕ} (n : ℕ) (x : S₊ k) → -S^ n (-S^ n x) ≡ x --S^² zero x = refl --S^² (suc n) x = - cong invSphere (sym (invSphere-S^ n (-S^ n x))) - ∙ invSphere² _ (-S^ n (-S^ n x)) - ∙ -S^² n x - --S^Iso : {k : ℕ} (n : ℕ) → Iso (S₊ k) (S₊ k) -fun (-S^Iso n) = -S^ n -inv (-S^Iso n) = -S^ n -rightInv (-S^Iso n) = -S^² n -leftInv (-S^Iso n) = -S^² n - --S^-comp : {k : ℕ} (n m : ℕ) (x : S₊ k) - → -S^ n (-S^ m x) ≡ -S^ (n + m) x --S^-comp zero m x = refl --S^-comp (suc n) m x = cong invSphere (-S^-comp n m x) - --S^·2 : {k : ℕ} (n : ℕ) (x : S₊ k) → -S^ (n + n) x ≡ x --S^·2 zero x = refl --S^·2 (suc n) x = - cong invSphere (λ i → -S^ (+-comm n (suc n) i) x) - ∙ invSphere² _ (-S^ (n + n) x) - ∙ -S^·2 n x - -- technical transport lemma private -S^-transp : {k : ℕ} (m : ℕ) (p : k ≡ m) (n : ℕ) (x : S₊ k) @@ -543,8 +518,11 @@ private -- Iterated path inversion sym^ : ∀ {ℓ} {A : Type ℓ} {x : A} (n : ℕ) → x ≡ x → x ≡ x -sym^ zero p = p -sym^ (suc n) p = sym (sym^ n p) +sym^ n = iter n sym + +sym^-refl : ∀ {ℓ} {A : Type ℓ} {x : A} (n : ℕ) → sym^ {x = x} n refl ≡ refl +sym^-refl zero = refl +sym^-refl (suc n) = cong sym (sym^-refl n) -- Interaction between -S^ and sym^ σS-S^ : {k : ℕ} (n : ℕ) (x : S₊ (suc k)) @@ -1021,6 +999,24 @@ diag⌣ {n = suc n} (merid a i) j = help j i (+-comm (suc m) (suc n) ⁻¹))) ∙ subst⁻Subst S₊ (+-comm (suc m) (suc n) ⁻¹) (x ⌣S y)) +⌣Sinvₗ^ : {n m : ℕ} (k : ℕ) (x : S₊ (suc n)) (y : S₊ (suc m)) + → (-S^ k x) ⌣S y ≡ -S^ k (x ⌣S y) +⌣Sinvₗ^ {n = n} {m} zero x y = refl +⌣Sinvₗ^ {n = n} {m} (suc zero) x y = ⌣Sinvₗ x y +⌣Sinvₗ^ {n = n} {m} (suc (suc k)) x y = + cong (_⌣S y) (-S^² 1 (-S^ k x)) + ∙ ⌣Sinvₗ^ k x y + ∙ sym (-S^² 1 (-S^ k (x ⌣S y))) + +⌣Sinvᵣ^ : {n m : ℕ} (k : ℕ) (x : S₊ (suc n)) (y : S₊ (suc m)) + → x ⌣S (-S^ k y) ≡ -S^ k (x ⌣S y) +⌣Sinvᵣ^ {n = n} {m} zero x y = refl +⌣Sinvᵣ^ {n = n} {m} (suc zero) x y = ⌣Sinvᵣ x y +⌣Sinvᵣ^ {n = n} {m} (suc (suc k)) x y = + cong (x ⌣S_) (-S^² 1 (-S^ k y)) + ∙ ⌣Sinvᵣ^ k x y + ∙ sym (-S^² 1 (-S^ k (x ⌣S y))) + -- Interaction between ⌣S, SuspS¹→S² and SuspS¹→S² SuspS¹→S²-S¹×S¹→S² : (a b : S¹) → SuspS¹→S² (a ⌣S b) ≡ S¹×S¹→S² a b @@ -1035,3 +1031,146 @@ SuspS¹→S²-S¹×S¹→S² (loop i) b j = main b j i main b = cong-∙ SuspS¹→S² (merid b) (merid base ⁻¹) ∙ sym (rUnit _) ∙ lem b + +-- Sphere inversion in terms of swapping coordinates in joins of spheres +join-commFun-sphere→Join : (n m : ℕ) (x : _) + → PathP (λ i → S₊ (suc (+-comm n m i))) + (join→Sphere n m (join-commFun x)) + (-S^ (suc (m · n)) (join→Sphere m n x)) +join-commFun-sphere→Join n m (inl x) = + (λ i → ptSn (suc (+-comm n m i))) + ▷ sym (-S^pt (suc (m · n))) +join-commFun-sphere→Join n m (inr x) = + (λ i → ptSn (suc (+-comm n m i))) + ▷ sym (-S^pt (suc (m · n))) +join-commFun-sphere→Join zero zero (push a b i) j = lem j i + where + main : (a b : Bool) → sym (σS (b ⌣S a)) ≡ cong (-S^ 1) (σS (a ⌣S b)) + main false false = refl + main false true = refl + main true false = refl + main true true = refl + + lem : Square (sym (σS (b ⌣S a))) (cong (-S^ 1) (σS (a ⌣S b))) + (refl ∙ (refl ∙ refl)) (refl ∙ (refl ∙ refl)) + lem = flipSquare (sym (rUnit refl ∙ cong₂ _∙_ refl (rUnit refl)) + ◁ flipSquare (main a b) + ▷ (rUnit refl ∙ cong₂ _∙_ refl (rUnit refl))) + +join-commFun-sphere→Join zero (suc m) (push a b i) j = + comp (λ k → S₊ (suc (+-comm zero (suc m) (j ∧ k)))) + (λ k → + λ{(i = i0) → ((λ i → ptSn (suc (+-comm zero (suc m) (i ∧ k)))) + ▷ sym (-S^pt (suc (·-comm zero m k)))) j + ; (i = i1) → ((λ i → ptSn (suc (+-comm zero (suc m) (i ∧ k)))) + ▷ sym (-S^pt (suc (·-comm zero m k)))) j + ; (j = i0) → σSn (suc m) (b ⌣S a) (~ i) + ; (j = i1) → -S^ (suc (·-comm zero m k)) + (σS (toPathP {A = λ i → S₊ (+-comm zero (suc m) i)} + (sym (comm⌣S a b)) k) i)}) + (hcomp (λ k → + λ{(i = i0) → lUnit (λ r → -S^pt (suc zero) (~ r ∨ ~ k)) k j + ; (i = i1) → lUnit (λ r → -S^pt (suc zero) (~ r ∨ ~ k)) k j + ; (j = i0) → σS-S^ 1 (b ⌣S a) k i + ; (j = i1) → cong-S^σ m (suc zero) (b ⌣S a) k i}) + (σ (S₊∙ (suc m)) (-S^ 1 (b ⌣S a)) i)) + where + n = zero + lem : -S^ (m · n) (-S^ (n · m) (b ⌣S a)) ≡ b ⌣S a + lem = cong (-S^ (m · n)) (cong₂ -S^ (·-comm n m) refl) + ∙ -S^-comp (m · n) (m · n) (b ⌣S a) + ∙ -S^·2 (m · n) (b ⌣S a) +join-commFun-sphere→Join (suc n') m (push a b i) j = + comp (λ k → S₊ (suc (+-comm n m (j ∧ k)))) + (λ k → + λ{(i = i0) → ((λ i → ptSn (suc (+-comm n m (i ∧ k)))) + ▷ sym (-S^pt (suc (m · n)))) j + ; (i = i1) → ((λ i → ptSn (suc (+-comm n m (i ∧ k)))) + ▷ sym (-S^pt (suc (m · n)))) j + ; (j = i0) → σSn (n + m) (b ⌣S a) (~ i) + ; (j = i1) → -S^ (suc (m · n)) + (σS (toPathP {A = λ i → S₊ (+-comm n m i)} + (sym (comm⌣S a b)) k) i)}) + (hcomp (λ k → + λ{(i = i0) → lUnit (λ r → -S^pt (suc (m · n)) (~ r ∨ ~ k)) k j + ; (i = i1) → lUnit (λ r → -S^pt (suc (m · n)) (~ r ∨ ~ k)) k j + ; (j = i0) → σS-S^ 1 (b ⌣S a) k i + ; (j = i1) → cong-S^σ (n' + m) (suc (m · n)) + (-S^ (n · m) (b ⌣S a)) k i}) + (σ (S₊∙ (suc (n' + m))) (invSphere (lem (~ j))) i)) + where + n = suc n' + lem : -S^ (m · n) (-S^ (n · m) (b ⌣S a)) ≡ b ⌣S a + lem = cong (-S^ (m · n)) (cong₂ -S^ (·-comm n m) refl) + ∙ -S^-comp (m · n) (m · n) (b ⌣S a) + ∙ -S^·2 (m · n) (b ⌣S a) + +private + -S^σS-lem : (n m : ℕ) (a : S₊ n) (b : S₊ m) + → (1 ≤ (n + m)) + → PathP + (λ i₁ → -S^∙ {k = +-comm m n (~ i₁)} (suc (m · n)) .snd i₁ + ≡ -S^∙ (suc (m · n)) .snd i₁) + ((cong (-S^ (suc (m · n))) + (σS (subst S₊ (+-comm m n) (-S^ (m · n) (b ⌣S a)))))) + (σS (-S^ (suc (m · n)) (-S^ (m · n) (b ⌣S a)))) + -S^σS-lem zero zero a b ineq = + ⊥.rec (snotz (+-comm 1 (ineq .fst) ∙ snd ineq)) + -S^σS-lem zero (suc m) a b ineq i j = + cong-S^σ _ (suc (m · zero)) + (transp (λ j → S₊ (+-comm (suc m) zero (j ∧ ~ i))) + i (-S^ (suc m · zero) (b ⌣S a))) (~ i) j + -S^σS-lem (suc n) zero a b ineq i j = + cong-S^σ _ (suc zero) + (transp (λ j → S₊ (+-comm zero (suc n) (j ∧ ~ i))) + i (b ⌣S a)) (~ i) j + -S^σS-lem (suc n) (suc m) a b ineq i j = + cong-S^σ _ (suc (suc m · suc n)) + (transp (λ j → S₊ (+-comm (suc m) (suc n) (j ∧ ~ i))) + i (-S^ (suc m · suc n) (b ⌣S a))) (~ i) j + +join→Sphere∘join-commFunId : (n m : ℕ) (x : _) + → PathP (λ i → S₊ (suc (+-comm m n (~ i)))) + (-S^ (suc (m · n)) (join→Sphere n m x)) + (join→Sphere m n (join-commFun x)) +join→Sphere∘join-commFunId n m (inl x) i = -S^∙ (suc (m · n)) .snd i +join→Sphere∘join-commFunId n m (inr x) i = -S^∙ (suc (m · n)) .snd i +join→Sphere∘join-commFunId zero zero (push a b i) j = + (sym (rUnit refl) ◁ flipSquare (lem a b) ▷ rUnit refl) i j + where + lem : (a b : Bool) → cong invSphere (σS (a ⌣S b)) ≡ sym (σS (b ⌣S a)) + lem false false = refl + lem false true = refl + lem true false = refl + lem true true = refl +join→Sphere∘join-commFunId (suc n') zero (push a b i) j = lem j i + where + n = suc n' + m = zero + lem : SquareP (λ i j → S₊ (suc (+-comm m n (~ i)))) + (cong (-S^ (suc (m · n))) (σS (a ⌣S b))) + (sym (σS (b ⌣S a))) + (λ i → -S^∙ (suc (m · n)) .snd i) + λ i → -S^∙ (suc (m · n)) .snd i + lem = cong (congS (-S^ (suc (m · n))) ∘ σS) + (comm⌣S a b) + ◁ -S^σS-lem n zero a b (n' + zero , +-comm (n' + zero) 1) + ▷ (cong σS ((λ i → -S^ (suc (m · n)) (-S^ ((m · n)) (b ⌣S a))) + ∙ cong invSphere (-S^-comp (m · n) (m · n) (b ⌣S a) + ∙ -S^·2 (m · n) (b ⌣S a))) + ∙ σ-invSphere _ (b ⌣S a)) +join→Sphere∘join-commFunId n (suc m') (push a b i) j = lem j i + where + m = suc m' + lem : SquareP (λ i j → S₊ (suc (+-comm m n (~ i)))) + (cong (-S^ (suc (m · n))) (σS (a ⌣S b))) + (sym (σS (b ⌣S a))) + (λ i → -S^∙ (suc (m · n)) .snd i) + λ i → -S^∙ (suc (m · n)) .snd i + lem = cong (congS (-S^ (suc (m · n))) ∘ σS) + (comm⌣S a b) + ◁ -S^σS-lem n (suc m') a b (n + m' , +-comm (n + m') 1 ∙ sym (+-suc n m')) + ▷ (cong σS ((λ i → -S^ (suc (m · n)) (-S^ ((m · n)) (b ⌣S a))) + ∙ cong invSphere (-S^-comp (m · n) (m · n) (b ⌣S a) + ∙ -S^·2 (m · n) (b ⌣S a))) + ∙ σ-invSphere _ (b ⌣S a)) diff --git a/Cubical/HITs/Sn/Properties.agda b/Cubical/HITs/Sn/Properties.agda index 002013f137..1032ee76dc 100644 --- a/Cubical/HITs/Sn/Properties.agda +++ b/Cubical/HITs/Sn/Properties.agda @@ -22,6 +22,7 @@ open import Cubical.HITs.PropositionalTruncation as PT hiding (rec ; elim) open import Cubical.HITs.SmashProduct.Base open import Cubical.HITs.Pushout.Base open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Loopspace open import Cubical.HITs.Join renaming (joinS¹S¹→S³ to joinS¹S¹→S3) open import Cubical.Data.Bool hiding (elim) @@ -42,6 +43,7 @@ open Iso σS∙ {n = zero} = refl σS∙ {n = suc n} = rCancel (merid (ptSn (suc n))) +-- Iso between Sⁿ⁺¹ and Susp Sⁿ (and some properties/related constructions) IsoSucSphereSusp : (n : ℕ) → Iso (S₊ (suc n)) (Susp (S₊ n)) IsoSucSphereSusp zero = S¹IsoSuspBool IsoSucSphereSusp (suc n) = idIso @@ -63,6 +65,28 @@ IsoSucSphereSusp∙' : (n : ℕ) IsoSucSphereSusp∙' zero = refl IsoSucSphereSusp∙' (suc n) = refl +IsoSucSphereSusp≃∙ : (n : ℕ) → S₊∙ (suc n) ≃∙ (Susp∙ (S₊ n)) +IsoSucSphereSusp≃∙ n .fst = isoToEquiv (IsoSucSphereSusp n) +IsoSucSphereSusp≃∙ n .snd = IsoSucSphereSusp∙' n + +IsoSucSphereSuspInv≃∙ : (n : ℕ) → Susp∙ (S₊ n) ≃∙ S₊∙ (suc n) +IsoSucSphereSuspInv≃∙ n .fst = isoToEquiv (invIso (IsoSucSphereSusp n)) +IsoSucSphereSuspInv≃∙ n .snd = IsoSucSphereSusp∙ n + +IsoSucSphereSusp≃∙CompL : (n : ℕ) + → (≃∙map (IsoSucSphereSusp≃∙ n) + ∘∙ ≃∙map (IsoSucSphereSuspInv≃∙ n)) ≡ idfun∙ _ +IsoSucSphereSusp≃∙CompL n i .fst x = rightInv (IsoSucSphereSusp n) x i +IsoSucSphereSusp≃∙CompL zero i .snd j = rUnit (λ _ → north) (~ i) j +IsoSucSphereSusp≃∙CompL (suc n) i .snd j = rUnit (λ _ → north) (~ i) j + +IsoSucSphereSusp≃∙CompR : (n : ℕ) + → (≃∙map (IsoSucSphereSuspInv≃∙ n)) + ∘∙ ≃∙map (IsoSucSphereSusp≃∙ n) ≡ idfun∙ _ +IsoSucSphereSusp≃∙CompR n i .fst x = leftInv (IsoSucSphereSusp n) x i +IsoSucSphereSusp≃∙CompR zero i .snd j = rUnit (λ _ → base) (~ i) j +IsoSucSphereSusp≃∙CompR (suc n) i .snd j = rUnit (λ _ → north) (~ i) j + suspFunS∙ : {n : ℕ} → (S₊ n → S₊ n) → S₊∙ (suc n) →∙ S₊∙ (suc n) suspFunS∙ {n = zero} f = (λ x → Iso.inv S¹IsoSuspBool (suspFun f (Iso.fun S¹IsoSuspBool x))) , refl @@ -511,6 +535,15 @@ invSphere² (suc zero) base = refl invSphere² (suc zero) (loop i) = refl invSphere² (suc (suc n)) = invSusp² +-- pointed version of inversion +invSpherePt : {k : ℕ} → invSphere (ptSn (suc k)) ≡ ptSn (suc k) +invSpherePt {k = zero} = refl +invSpherePt {k = suc k} = sym (merid (ptSn (suc k))) + +invSphere∙ : {k : ℕ} → S₊∙ (suc k) →∙ S₊∙ (suc k) +invSphere∙ {k = k} .fst = invSphere +invSphere∙ {k = k} .snd = invSpherePt + -- Interaction between σ and invSphere σ-invSphere : (n : ℕ) (x : S₊ (suc n)) → σSn (suc n) (invSphere x) @@ -530,3 +563,189 @@ invSphere² (suc (suc n)) = invSusp² ∙∙ (λ i → (σSn 1 (loop i))) ∙∙ (rCancel (merid base))) j i) σ-invSphere (suc n) x = toSusp-invSusp (S₊∙ (suc n)) x + +-- cong version +cong-invsphere-σS : {k : ℕ} (x : S₊ (suc k)) + → Square (cong invSphere (σS x)) (σS (invSphere x)) + invSpherePt invSpherePt +cong-invsphere-σS {k = k} x = + (cong-∙ invSusp (merid x) (sym (merid (ptSn (suc k)))) + ∙ refl) + ◁ ((λ i → (λ j → merid (ptSn (suc k)) (~ i ∨ j)) + ∙∙ sym (merid x) + ∙∙ (λ j → merid (ptSn (suc k)) (~ i ∧ j))) + ▷ (sym (compPath≡compPath' + (merid (ptSn (suc k))) (sym (merid x))) + ∙ sym (symDistr (merid x) (sym (merid (ptSn (suc k))))) + ∙ sym (σ-invSphere _ x))) + +-- constrution of and lemmas about iterated version of invSphere +-S^ : {k : ℕ} (n : ℕ) → S₊ k → S₊ k +-S^ n = iter n invSphere + +-S^pt : {k : ℕ} (n : ℕ) → -S^ {k = suc k} n (ptSn (suc k)) ≡ ptSn (suc k) +-S^pt {k = k} n = iter∙ n (invSphere , invSpherePt) .snd + +-S^∙ : {k : ℕ} (n : ℕ) → S₊∙ (suc k) →∙ S₊∙ (suc k) +-S^∙ n .fst = -S^ n +-S^∙ n .snd = -S^pt n + +invSphere-S^ : {k : ℕ} (n : ℕ) (x : S₊ k) + → invSphere (-S^ n x) ≡ -S^ n (invSphere x) +invSphere-S^ zero x = refl +invSphere-S^ (suc n) x = cong invSphere (invSphere-S^ n x) + +-S^² : {k : ℕ} (n : ℕ) (x : S₊ k) → -S^ n (-S^ n x) ≡ x +-S^² zero x = refl +-S^² (suc n) x = + cong invSphere (sym (invSphere-S^ n (-S^ n x))) + ∙ invSphere² _ (-S^ n (-S^ n x)) + ∙ -S^² n x + +-S^Iso : {k : ℕ} (n : ℕ) → Iso (S₊ k) (S₊ k) +fun (-S^Iso n) = -S^ n +inv (-S^Iso n) = -S^ n +rightInv (-S^Iso n) = -S^² n +leftInv (-S^Iso n) = -S^² n + +-S^-comp : {k : ℕ} (n m : ℕ) (x : S₊ k) + → -S^ n (-S^ m x) ≡ -S^ (n + m) x +-S^-comp zero m x = refl +-S^-comp (suc n) m x = cong invSphere (-S^-comp n m x) + +-S^·2 : {k : ℕ} (n : ℕ) (x : S₊ k) → -S^ (n + n) x ≡ x +-S^·2 zero x = refl +-S^·2 (suc n) x = + cong invSphere (λ i → -S^ (+-comm n (suc n) i) x) + ∙ invSphere² _ (-S^ (n + n) x) + ∙ -S^·2 n x + +invSphere∙² : {k : ℕ} → invSphere∙ {k = k} ∘∙ invSphere∙ {k = k} + ≡ idfun∙ (S₊∙ (suc k)) +invSphere∙² {k = k} i .fst x = invSphere² _ x i +invSphere∙² {k = zero} i .snd j = rUnit (λ _ → ptSn 1) (~ i) j +invSphere∙² {k = suc k} i .snd j = rCancel (merid (ptSn (suc k))) i j + +-S^∙+1 : {k : ℕ} (n : ℕ) + → (-S^∙ {k = k} 1 ∘∙ -S^∙ {k = k} (suc n)) ≡ -S^∙ {k = k} n +-S^∙+1 {k = k} n i .fst x = invSphere² _ (-S^ n x) i +-S^∙+1 {k = k} n i .snd j = + hcomp (λ r → + λ{(i = i1) → invSphere² (suc k) (-S^pt n j) r + ; (j = i0) → invSphere² (suc k) (-S^ n (S₊∙ (suc k) .snd)) (i ∧ r) + ; (j = i1) → lem k i r }) + (-S^ 1 (compPath-filler (λ r → invSphere (-S^pt n r)) invSpherePt (~ i) j)) + where + lem : (k : ℕ) + → Square (refl ∙ invSpherePt) + (invSphere² (suc k) (ptSn (suc k))) + (cong invSphere (sym (invSpherePt {k = k}))) refl + lem zero = sym (lUnit refl) + lem (suc k) = sym (lUnit _) ◁ λ i j → merid (ptSn (suc k)) (~ i ∧ ~ j) + +private + -S^∙suspFun₁ : {k : ℕ} → -S^∙ {k = suc k} 1 ≡ suspFun∙ (-S^ {k = suc k} 1) + -S^∙suspFun₁ {k = k} = + ΣPathP ((funExt λ { north → sym (merid (ptSn (suc k))) + ; south → merid (ptSn (suc k)) + ; (merid a i) k → lem a k i}) + , (sym (lUnit _) + ◁ λ i j → merid (ptSn (suc k)) (~ j ∧ ~ i))) + where + lem : (a : S₊ (suc k)) + → Square (sym (merid a)) (merid (invSphere a)) + (sym (merid (ptSn (suc k)))) (merid (ptSn (suc k))) + lem a = + (λ i → compPath-filler + (sym (compPath-filler (merid a) + (sym (merid (ptSn (suc k)))) i)) + (merid (ptSn (suc k))) i) + ▷ (refl + ∙ cong (_∙ merid (ptSn (suc k))) (sym (σ-invSphere _ a)) + ∙ sym (assoc _ _ _) + ∙ cong₂ _∙_ refl (lCancel _) + ∙ sym (rUnit _)) + +-S^∙suspFun : {k : ℕ} (n : ℕ) + → -S^∙ {k = suc k} n ≡ suspFun∙ (-S^ {k = suc k} n) +-S^∙suspFun zero = ΣPathP (sym suspFunIdFun , refl) +-S^∙suspFun (suc n) = + ΣPathP (refl , cong₂ _∙_ refl (lUnit _)) + ∙ cong₂ _∘∙_ -S^∙suspFun₁ (-S^∙suspFun n) + ∙ sym (suspFun∙Comp (-S^ 1) (-S^ n)) + +cong-S^σ : (n k : ℕ) (a : S₊ (suc n)) + → Square (σSn (suc n) (-S^ k a)) + (cong (-S^ k) (σS a)) + (sym (-S^pt k)) (sym (-S^pt k)) +cong-S^σ n zero a = refl +cong-S^σ n (suc k) a i j = + hcomp (λ r → λ{(i = i0) → cong-invsphere-σS (-S^ k a) r j + ; (i = i1) → -S^ (suc k) (σS a j) + ; (j = i0) → compPath-filler (cong invSphere (-S^pt k)) + invSpherePt r (~ i) + ; (j = i1) → compPath-filler (cong invSphere (-S^pt k)) + invSpherePt r (~ i)}) + (invSphere (cong-S^σ n k a i j)) + +-- even and odd powers of -S^ +-S^-even : {k : ℕ} (n : ℕ) → isEvenT n → (x : S₊ (suc k)) → -S^ n x ≡ x +-S^-even zero p x = refl +-S^-even (suc (suc n)) p x = + cong (invSphere ∘ invSphere) (-S^-even n p x) + ∙ invSphere² _ x + +-S^∙-even : {k : ℕ} (n : ℕ) → isEvenT n → -S^∙ {k = k} n ≡ idfun∙ (S₊∙ (suc k)) +-S^∙-even {k = k} n p = ΣPathP ((funExt (-S^-even n p)) , lem k n p) + where + lem : (k n : ℕ) (p : _) + → PathP (λ i → -S^-even n p (ptSn (suc k)) i ≡ ptSn (suc k)) + (-S^pt n) (λ _ → ptSn (suc k)) + lem k zero p = refl + lem zero (suc (suc n)) p = + (sym (rUnit _) ∙ sym (rUnit _)) + ◁ (flipSquare (sym (rUnit _) + ◁ λ i j → invLooper (invLooper (lem zero n p j i)))) + lem (suc k) (suc (suc n)) p = + (cong₂ _∙_ (cong-∙ invSphere (cong invSphere (-S^pt n)) invSpherePt) refl + ∙ sym (assoc _ _ _) + ∙ cong₂ _∙_ refl (rCancel _) + ∙ sym (rUnit _)) + ◁ flipSquare (sym (rUnit _) + ◁ λ i j → invSphere (invSphere (lem (suc k) n p j i))) + +-S^∙-odd : {k : ℕ} (n : ℕ) → isOddT n + → -S^∙ {k = k} n ≡ (invSphere , invSpherePt) +-S^∙-odd {k = k} (suc n) o = + cong ((invSphere , invSpherePt) ∘∙_) (-S^∙-even {k = k} n o) ∙ ∘∙-idˡ _ + +-S^-odd : {k : ℕ} (n : ℕ) → isOddT n → (x : S₊ (suc k)) → -S^ n x ≡ invSphere x +-S^-odd (suc zero) p x = refl +-S^-odd (suc (suc (suc n))) p x = + cong (invSphere ∘ invSphere) (-S^-odd (suc n) p x) + ∙ invSphere² _ (invSphere x) + +-- Commutativity of ·→Ω with Sⁿ domain. +EH-ΠΩ : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} (f g : S₊∙ (suc n) →∙ Ω A) + → ·→Ω f g ≡ ·→Ω g f +EH-ΠΩ {A = A} {n = n} = + subst (λ T → (f g : T →∙ Ω A) → ·→Ω f g ≡ ·→Ω g f) + (ua∙ (isoToEquiv (invIso (IsoSucSphereSusp n))) (IsoSucSphereSusp∙ n)) + (Susp·→Ωcomm (S₊∙ n)) + +-- Interaction between Ω→ and σS +Ω→comp-σS : ∀ {ℓ} {X : Pointed ℓ} (n : ℕ) (f : S₊∙ (suc n) →∙ X) (a : S₊ n) + → Ω→ f .fst (σS a) + ≡ Ω→ (f ∘∙ ≃∙map (IsoSucSphereSuspInv≃∙ n)) .fst (toSusp (S₊∙ n) a) +Ω→comp-σS zero f true = + cong₃ _∙∙_∙∙_ (cong sym (lUnit (snd f))) + (sym (cong (congS (fst f ∘ fst (fst (IsoSucSphereSuspInv≃∙ zero)))) + (rCancel (merid true)))) + (lUnit (snd f)) +Ω→comp-σS zero f false = + cong₃ _∙∙_∙∙_ (cong sym (lUnit (snd f))) + (cong (congS (fst f)) + (rUnit _ + ∙ sym (cong-∙ SuspBool→S¹ (merid false) (sym (merid true))))) + (lUnit (snd f)) +Ω→comp-σS (suc n) f a = cong (λ f → Ω→ f .fst (σS a)) (sym (∘∙-idˡ f)) diff --git a/Cubical/HITs/Susp/Properties.agda b/Cubical/HITs/Susp/Properties.agda index c9f84f9c23..48852a4e0e 100644 --- a/Cubical/HITs/Susp/Properties.agda +++ b/Cubical/HITs/Susp/Properties.agda @@ -33,6 +33,10 @@ suspFunComp f g i north = north suspFunComp f g i south = south suspFunComp f g i (merid a i₁) = merid (f (g a)) i₁ +suspFun∙Comp : ∀ {ℓ ℓ' ℓ''} {A : Type ℓ} {B : Type ℓ'} {C : Type ℓ''} + (g : B → C) (f : A → B) → suspFun∙ (g ∘ f) ≡ (suspFun∙ g ∘∙ suspFun∙ f) +suspFun∙Comp f g = ΣPathP ((suspFunComp f g) , rUnit refl) + suspFunConst : (b : B) → suspFun (λ (_ : A) → b) ≡ λ _ → north suspFunConst b i north = north suspFunConst b i south = merid b (~ i) @@ -43,6 +47,10 @@ suspFunIdFun i north = north suspFunIdFun i south = south suspFunIdFun i (merid a j) = merid a j +suspFun∙IdFun : ∀ {ℓ} {A : Type ℓ} + → suspFun∙ (idfun A) ≡ idfun∙ _ +suspFun∙IdFun = ΣPathP (suspFunIdFun , refl) + Susp-iso-joinBool : ∀ {ℓ} {A : Type ℓ} → Iso (Susp A) (join A Bool) fun Susp-iso-joinBool north = inr true fun Susp-iso-joinBool south = inr false @@ -129,6 +137,11 @@ leftInv (congSuspIso is) (merid a i) j = merid (leftInv is a j) i congSuspEquiv : ∀ {ℓ} {A B : Type ℓ} → A ≃ B → Susp A ≃ Susp B congSuspEquiv {ℓ} {A} {B} h = isoToEquiv (congSuspIso (equivToIso h)) +suspFun∙Cancel : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : Iso A B) + → suspFun∙ (fun f) ∘∙ suspFun∙ (inv f) ≡ id∙ (Susp∙ B) +suspFun∙Cancel f = ΣPathP ((funExt (rightInv (congSuspIso f))) + , sym (rUnit refl)) + suspToPropElim : ∀ {ℓ ℓ'} {A : Type ℓ} {B : Susp A → Type ℓ'} (a : A) → ((x : Susp A) → isProp (B x)) → B north @@ -423,6 +436,7 @@ toSusp-invSusp A (merid a i) j = ▷ rUnit refl))) -- co-H-space structure + ·Susp : ∀ {ℓ'} (A : Pointed ℓ) {B : Pointed ℓ'} (f g : Susp∙ (typ A) →∙ B) → Susp∙ (typ A) →∙ B fst (·Susp A {B = B} f g) north = pt B @@ -430,3 +444,281 @@ fst (·Susp A {B = B} f g) south = pt B fst (·Susp A {B = B} f g) (merid a i) = (Ω→ f .fst (toSusp A a) ∙ Ω→ g .fst (toSusp A a)) i snd (·Susp A f g) = refl + +-Susp : ∀ {ℓ'} (A : Pointed ℓ) {B : Pointed ℓ'} + → Susp∙ (typ A) →∙ B → Susp∙ (typ A) →∙ B +fst (-Susp A {B = B} f) north = pt B +fst (-Susp A {B = B} f) south = pt B +fst (-Susp A {B = B} f) (merid a i) = + (Ω→ f .fst (toSusp A a) (~ i)) +snd (-Susp A f) = refl + +·Suspσ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (f g : Susp∙ (typ A) →∙ B) + → (a : typ A) → cong (·Susp A f g .fst) (toSusp A a) + ≡ Ω→ f .fst (toSusp A a) ∙ Ω→ g .fst (toSusp A a) +·Suspσ {A = A} f g a = cong-∙ (·Susp A f g .fst) (merid a) (merid (pt A) ⁻¹) + ∙ cong₂ _∙_ refl + (cong _⁻¹ (cong₂ _∙_ + (cong (Ω→ f .fst) (rCancel (merid (pt A))) ∙ Ω→ f .snd) + (cong (Ω→ g .fst) (rCancel (merid (pt A))) ∙ Ω→ g .snd) + ∙ sym (rUnit _))) + ∙ sym (rUnit _) + +·Suspσ' : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} (f g : Susp∙ (typ A) →∙ B) + → (a : typ A) → Ω→ (·Susp A f g) .fst (toSusp A a) + ≡ Ω→ f .fst (toSusp A a) ∙ Ω→ g .fst (toSusp A a) +·Suspσ' f g a = sym (rUnit _) ∙ ·Suspσ f g a + +-Susp-∘∙ : ∀ {ℓ' ℓ''} (A : Pointed ℓ) {B : Pointed ℓ'} {C : Pointed ℓ''} + (f : Susp∙ (typ A) →∙ B) (g : B →∙ C) + → -Susp A (g ∘∙ f) ≡ (g ∘∙ -Susp A f) +-Susp-∘∙ A f g = ΣPathP + ((funExt (λ { north → sym (snd g) + ; south → sym (snd g) + ; (merid a i) j + → hcomp (λ k → λ {(i = i0) → snd g (~ j ∧ k) + ; (i = i1) → snd g (~ j ∧ k) + ; (j = i0) → + (sym (compPath-filler (cong (fst g) (snd f)) + (snd g) k) + ∙∙ cong (fst g ∘ fst f) (toSusp A a) + ∙∙ compPath-filler (cong (fst g) (snd f)) + (snd g) k) (~ i) + ; (j = i1) → cong-∙∙ (fst g) (sym (snd f)) + (cong (fst f) (toSusp A a)) + (snd f) (~ k) (~ i)}) + (cong-∙∙ (fst g) (sym (snd f)) + (cong (fst f) (toSusp A a)) + (snd f) i1 (~ i))})) + , ((λ i j → snd g (~ i ∨ j)) ▷ lUnit _)) + +·Susp-suspFun : ∀ {ℓ ℓ'} (A A' : Pointed ℓ) {X : Pointed ℓ'} (e : A ≃∙ A') + → (f g : Susp∙ (typ A) →∙ X) + → ·Susp A f g + ≡ (·Susp A' (f ∘∙ suspFun∙ (invEq (fst e))) + (g ∘∙ suspFun∙ (invEq (fst e))) + ∘∙ suspFun∙ (fst (fst e))) +·Susp-suspFun A A' {X} = Equiv∙J (λ A e → (f g : Susp∙ (typ A) →∙ X) → + ·Susp A f g + ≡ (·Susp A' (f ∘∙ suspFun∙ (invEq (fst e))) + (g ∘∙ suspFun∙ (invEq (fst e))) + ∘∙ suspFun∙ (fst (fst e)))) + + λ f g → sym (cong₂ {x = ·Susp A' (f ∘∙ suspFun∙ (idfun (fst A'))) + (g ∘∙ suspFun∙ (idfun (fst A')))} + {y = ·Susp A' f g} _∘∙_ + (cong₂ (·Susp A') (cong (f ∘∙_) suspFun∙IdFun ∙ ∘∙-idˡ f) + (cong (g ∘∙_) suspFun∙IdFun ∙ ∘∙-idˡ g)) + {u = suspFun∙ (idfun (fst A'))} {v = idfun∙ _} + suspFun∙IdFun + ∙ ∘∙-idˡ (·Susp A' f g)) + +Ω→-Susp : ∀ {ℓ'} (A : Pointed ℓ) {B : Pointed ℓ'} + (f : Susp∙ (typ A) →∙ B) + (a : typ A) + → Ω→ (-Susp A f) .fst (toSusp A a) ≡ Ω→ f .fst (toSusp A a) ⁻¹ +Ω→-Susp A f a = sym (rUnit _) + ∙ cong-∙ (fst (-Susp A f)) (merid a) (merid (pt A) ⁻¹) + ∙ cong₂ _∙_ refl (cong (Ω→ f .fst) (rCancel (merid (pt A))) ∙ Ω→ f .snd) + ∙ sym (rUnit _) + +Ω∘∙σ-comp : ∀ {ℓA ℓB ℓC} {A : Pointed ℓA} {B : Pointed ℓB} {C : Pointed ℓC} + (f : Susp∙ (typ A) →∙ B) + (g : C →∙ A) + → (Ω→ f ∘∙ (((toSusp A) , (rCancel _)) ∘∙ g)) + ≡ (Ω→ (f ∘∙ suspFun∙ (fst g)) ∘∙ (toSusp C , rCancel _)) +Ω∘∙σ-comp {A = A} {B} {C} f g = + →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt (λ x → + cong (Ω→ f .fst) + (sym (cong-∙ (suspFun (fst g)) (merid x) (sym (merid (pt C))) + ∙ cong₂ _∙_ refl (cong (sym ∘ merid) (snd g)))))) + ∙ cong (_∘∙ (toSusp C , rCancel (merid (pt C)))) + (cong Ω→ (ΣPathP + {x = (fst f ∘ suspFun (fst g)) , snd f} + {y = f ∘∙ suspFun∙ (fst g)} + (refl , lUnit (snd f)))) + +·SuspInvR : ∀ {ℓ'} (A : Pointed ℓ) {B : Pointed ℓ'} + (f : Susp∙ (typ A) →∙ B) + → ·Susp A f (-Susp A f) ≡ const∙ _ _ +fst (·SuspInvR A {B = B} f i) north = pt B +fst (·SuspInvR A {B = B} f i) south = pt B +fst (·SuspInvR A {B = B} f i) (merid a j) = main i j + where + main : (Ω→ f .fst (toSusp A a) ∙ Ω→ (-Susp A f) .fst (toSusp A a)) ≡ refl + main = cong₂ _∙_ refl (Ω→-Susp A f a) ∙ rCancel _ +snd (·SuspInvR A {B = B} f i) = refl + +·SuspInvL : ∀ {ℓ'} (A : Pointed ℓ) {B : Pointed ℓ'} + (f : Susp∙ (typ A) →∙ B) + → ·Susp A (-Susp A f) f ≡ const∙ _ _ +fst (·SuspInvL A {B = B} f i) north = pt B +fst (·SuspInvL A {B = B} f i) south = pt B +fst (·SuspInvL A {B = B} f i) (merid a j) = main i j + where + main : Ω→ (-Susp A f) .fst (toSusp A a) ∙ Ω→ f .fst (toSusp A a) ≡ refl + main = cong₂ _∙_ (Ω→-Susp A f a) refl ∙ rCancel _ +snd (·SuspInvL A {B = B} f i) = refl + +·SuspAssoc : ∀ {ℓ'} (A : Pointed ℓ) {B : Pointed ℓ'} + (f g h : Susp∙ (typ A) →∙ B) + → ·Susp A f (·Susp A g h) ≡ ·Susp A (·Susp A f g) h +·SuspAssoc A {B = B} f g h = + ΣPathP ((funExt (λ { north → refl + ; south → refl + ; (merid a i) j → main a j i})) + , refl) + where + main : (a : typ A) → Ω→ f .fst (toSusp A a) + ∙ Ω→ (·Susp A g h) .fst (toSusp A a) + ≡ Ω→ (·Susp A f g) .fst (toSusp A a) + ∙ Ω→ h .fst (toSusp A a) + main a = cong₂ _∙_ refl (·Suspσ' g h a) + ∙ assoc _ _ _ + ∙ cong₂ _∙_ (sym (·Suspσ' f g a)) refl + +·SuspIdR : ∀ {ℓ'} (A : Pointed ℓ) {B : Pointed ℓ'} + (f : Susp∙ (typ A) →∙ B) + → ·Susp A f (const∙ _ _) ≡ f +·SuspIdR A f = + ΣPathP ((funExt + (λ { north → refl + ; south → refl + ; (merid a i) j → (cong₂ _∙_ refl (sym (rUnit refl)) + ∙ sym (rUnit (Ω→ f .fst (toSusp A a)))) j i})) + , refl) + ∙ Iso.rightInv (ΩSuspAdjointIso {A = A}) f + + +·SuspIdL : ∀ {ℓ'} (A : Pointed ℓ) {B : Pointed ℓ'} (f : Susp∙ (typ A) →∙ B) + → ·Susp A (const∙ _ _) f ≡ f +·SuspIdL A f = + ΣPathP ((funExt + (λ { north → refl + ; south → refl + ; (merid a i) j → (cong₂ _∙_ (sym (rUnit refl)) refl + ∙ sym (lUnit (Ω→ f .fst (toSusp A a)))) j i})) + , refl) + ∙ Iso.rightInv (ΩSuspAdjointIso {A = A}) f + +·SuspIdL≡·SuspIdR : ∀ {ℓ'} (A : Pointed ℓ) {B : Pointed ℓ'} + → ·SuspIdL A (const∙ _ B) ≡ ·SuspIdR A (const∙ _ B) +·SuspIdL≡·SuspIdR A {B = B} = + cong₂ _∙_ (cong ΣPathP (ΣPathP (cong funExt (funExt + (λ { north → refl + ; south → refl + ; (merid a i) k j + → lem (pt B) (refl ∙ refl) (rUnit refl) + (refl ∙ refl) (rUnit refl) k j i})) + , refl))) + refl + where + lem : ∀ {ℓ} {A : Type ℓ} (x : A) (p : x ≡ x) (q : refl ≡ p) + (rr : x ≡ x) (q : refl ≡ rr) + → cong₂ _∙_ (sym q) (refl {x = rr}) ∙ sym (lUnit rr) + ≡ cong₂ _∙_ (refl {x = rr}) (sym q) ∙ sym (rUnit rr) + lem x = J> (J> refl) + +-Susp² : ∀ {ℓ'} (A : Pointed ℓ) {B : Pointed ℓ'} (f : Susp∙ (typ A) →∙ B) + → -Susp A (-Susp A f) ≡ f +-Susp² A f = + sym (·SuspAssoc A _ _ _ + ∙ cong₂ (·Susp A) (·SuspInvR A f) refl + ∙ ·SuspIdL A _) + ∙ cong (·Susp A f) (·SuspInvR A (-Susp A f)) + ∙ ·SuspIdR A f + +open import Cubical.Foundations.Pointed.Homogeneous +·Susp≡·→Ω : ∀ {ℓ ℓ'} (A : Pointed ℓ) {B : Pointed ℓ'} + (f g : Susp∙ (typ A) →∙ Ω B) + → ·Susp A f g ≡ ·→Ω f g +·Susp≡·→Ω A {B = B} f g = + →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt λ { north → rUnit refl + ∙ cong₂ _∙_ (sym (snd f)) (sym (snd g)) + ; south → rUnit refl + ∙ cong₂ _∙_ (snd f ⁻¹ ∙ cong (fst f) (merid (pt A))) + (snd g ⁻¹ ∙ cong (fst g) (merid (pt A))) + ; (merid a i) j → + (cong₂ _∙_ (cong (sym (snd f) ∙∙_∙∙ snd f) (cong-∙ (fst f) _ _)) + (cong (sym (snd g) ∙∙_∙∙ snd g) (cong-∙ (fst g) _ _)) + ◁ lem B _ (sym (snd f)) _ + (cong (fst f) (merid a)) + (cong (fst f) (merid (pt A))) + _ (sym (snd g)) _ + (cong (fst g) (merid a)) + (cong (fst g) (merid (pt A)))) j i}) + where + open import Cubical.Foundations.Path + lem : ∀ {ℓ} (A : Pointed ℓ) (⋆f : Ω A .fst) (sndf : refl ≡ ⋆f) + (⋆fs : Ω A .fst) (mfa mfb : ⋆f ≡ ⋆fs) + (⋆g : Ω A .fst) (sndg : refl ≡ ⋆g) + (⋆gs : Ω A .fst) (mga mgb : ⋆g ≡ ⋆gs) + → Square ((sndf ∙∙ mfa ∙ mfb ⁻¹ ∙∙ sym sndf) + ∙ (sndg ∙∙ mga ∙ mgb ⁻¹ ∙∙ sym sndg)) + (cong₂ _∙_ mfa mga) + (rUnit refl ∙ cong₂ _∙_ sndf sndg) + (rUnit refl ∙ cong₂ _∙_ (sndf ∙ mfb) (sndg ∙ mgb)) + lem A = J> (J> λ mfb → J> (J> + λ mgb → cong₂ _∙_ (sym (rUnit _) ∙ sym (lUnit (mfb ⁻¹))) + (sym (rUnit _) ∙ sym (lUnit (mgb ⁻¹))) + ◁ flipSquare (sym (rUnit (rUnit refl)) + ◁ (flipSquare (sym (symDistr mgb mfb) + ◁ flipSquare (compPath-filler' (mgb ∙ mfb) (rUnit refl))) + ▷ (cong (_∙ rUnit refl) (EH 0 mgb mfb) + ∙ lUnit _ + ∙ (λ j → (λ i → rUnit refl (i ∧ j)) + ∙ (cong (λ x → rUnit x j) mfb + ∙ cong (λ x → lUnit x j) mgb) + ∙ λ i → rUnit refl (i ∨ j)) + ∙ cong (rUnit refl ∙_) + (sym (rUnit _) + ∙ sym (cong₂Funct _∙_ mfb mgb) + ∙ (λ j → cong₂ _∙_ (lUnit mfb j) (lUnit mgb j)))))))) + +Susp·→Ωcomm : ∀ {ℓ ℓ'} (A : Pointed ℓ) {B : Pointed ℓ'} + (f g : Susp∙ (typ A) →∙ Ω B) → ·→Ω f g ≡ ·→Ω g f +Susp·→Ωcomm A {B = B} f g = isoInvInjective ΩSuspAdjointIso _ _ + (cong fromSusp→toΩ (·Susp≡·→Ω _ f g ⁻¹) + ∙∙ help + ∙∙ cong fromSusp→toΩ (·Susp≡·→Ω _ g f)) + where + help : Iso.inv ΩSuspAdjointIso (·Susp A f g) + ≡ Iso.inv ΩSuspAdjointIso (·Susp A g f) + help = →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt λ x → sym (rUnit _) + ∙ ·Suspσ f g x + ∙ EH 0 _ _ + ∙ sym (·Suspσ g f x) + ∙ rUnit _) + +Susp·→Ωcomm' : ∀ {ℓ ℓ'} (A A' : Pointed ℓ) + → A ≃∙ Susp∙ (typ A') → {B : Pointed ℓ'} + (f g : A →∙ Ω B) → ·→Ω f g ≡ ·→Ω g f +Susp·→Ωcomm' A A' e {B = B} = + Equiv∙J (λ A _ → (f g : A →∙ Ω B) → ·→Ω f g ≡ ·→Ω g f) + (Susp·→Ωcomm A') e + +·SuspComm : ∀ {ℓ ℓ'} (A : Pointed ℓ) {B : Pointed ℓ'} + (f g : Susp∙ (Susp (typ A)) →∙ B) + → ·Susp (Susp∙ (typ A)) f g ≡ ·Susp (Susp∙ (typ A)) g f +·SuspComm A {B = B} f g = isoInvInjective ΩSuspAdjointIso _ _ cool + where + cool : Iso.inv ΩSuspAdjointIso (·Susp (Susp∙ (typ A)) f g) + ≡ Iso.inv ΩSuspAdjointIso (·Susp (Susp∙ (typ A)) g f) + cool = →∙Homogeneous≡ (isHomogeneousPath _ _) + (funExt λ x → sym (rUnit _) + ∙ ·Suspσ f g x + ∙ funExt⁻ (cong fst (Susp·→Ωcomm A + (Iso.inv ΩSuspAdjointIso f) (Iso.inv ΩSuspAdjointIso g))) x + ∙ sym (·Suspσ g f x) + ∙ rUnit _) + +·SuspComm' : ∀ {ℓ ℓ'} (A A' : Pointed ℓ) + → A ≃∙ Susp∙ (typ A') → {B : Pointed ℓ'} + (f g : Susp∙ (typ A) →∙ B) + → ·Susp A f g ≡ ·Susp A g f +·SuspComm' A A' e {B = B} = + Equiv∙J (λ A _ → (f g : Susp∙ (typ A) →∙ B) + → ·Susp A f g ≡ ·Susp A g f) (·SuspComm A') e diff --git a/Cubical/Homotopy/Group/Base.agda b/Cubical/Homotopy/Group/Base.agda index ab2cd2092c..76f9c17c49 100644 --- a/Cubical/Homotopy/Group/Base.agda +++ b/Cubical/Homotopy/Group/Base.agda @@ -132,6 +132,16 @@ fst (-Π {A = A} {n = suc (suc n)} f) (merid a j) = fst f ((merid a ∙ sym (merid (ptSn _))) (~ j)) snd (-Π {A = A} {n = suc (suc n)} f) = snd f +-- Action of ∙Π on σ +∙Πσ : ∀ {ℓ} {A : Pointed ℓ} {n : ℕ} (f g : S₊∙ (suc n) →∙ A) (t : S₊ n) + → Square (cong (fst (∙Π f g)) (σS t)) + (Ω→ f .fst (σS t) ∙ Ω→ g .fst (σS t)) + (snd (∙Π f g)) (snd (∙Π f g)) +∙Πσ {A = A} {zero} f g false = refl +∙Πσ {A = A} {zero} f g true = + rUnit refl + ∙ cong₂ _∙_ (sym (Ω→ f .snd)) (sym (Ω→ g .snd)) +∙Πσ {A = A} {suc n} f g t = ·Suspσ f g t -- to prove that this gives a group structure on π', we first -- prove that Ωⁿ A ≃ (Sⁿ →∙ A). @@ -600,6 +610,9 @@ snd (∙Π-lCancel {A = A} {n = suc n} f i) = refl -π' : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} → π' (suc n) A → π' (suc n) A -π' n = sMap -Π +-π^ : ∀ {ℓ} {k : ℕ} {A : Pointed ℓ} (n : ℕ) → π' (suc k) A → π' (suc k) A +-π^ {k = k} n = iter n (-π' k) + π'-rUnit : ∀ {ℓ} (n : ℕ) {A : Pointed ℓ} (x : π' (suc n) A) → (·π' n x (1π' (suc n))) ≡ x π'-rUnit n = sElim (λ _ → isSetPathImplicit) λ p i → ∣ ∙Π-rUnit p i ∣₂ diff --git a/Cubical/Homotopy/Group/Join.agda b/Cubical/Homotopy/Group/Join.agda index d2e6a04758..d2ce55654b 100644 --- a/Cubical/Homotopy/Group/Join.agda +++ b/Cubical/Homotopy/Group/Join.agda @@ -14,9 +14,14 @@ open import Cubical.Foundations.Pointed open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Path +open import Cubical.Data.Sum +open import Cubical.Data.Empty as ⊥ open import Cubical.Data.Sigma open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order open import Cubical.Data.Bool open import Cubical.HITs.SetTruncation as ST @@ -24,6 +29,7 @@ open import Cubical.HITs.Sn open import Cubical.HITs.Susp renaming (toSusp to σ) open import Cubical.HITs.S1 open import Cubical.HITs.Join +open import Cubical.HITs.Join.CoHSpace open import Cubical.HITs.Sn.Multiplication open import Cubical.Algebra.Group @@ -34,32 +40,9 @@ open import Cubical.Algebra.Group.GroupPath open Iso open GroupStr --- Standard loop in Ω (join A B) -ℓ* : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') - → fst A → fst B → Ω (join∙ A B) .fst -ℓ* A B a b = push (pt A) (pt B) - ∙ (push a (pt B) ⁻¹ ∙∙ push a b ∙∙ (push (pt A) b ⁻¹)) - ℓS : {n m : ℕ} → S₊ n → S₊ m → Ω (join∙ (S₊∙ n) (S₊∙ m)) .fst ℓS {n = n} {m} = ℓ* (S₊∙ n) (S₊∙ m) --- Addition of functions join A B → C -_+*_ : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} - (f g : join∙ A B →∙ C) → join∙ A B →∙ C -fst (_+*_ {C = C} f g) (inl x) = pt C -fst (_+*_ {C = C} f g) (inr x) = pt C -fst (_+*_ {A = A} {B = B} f g) (push a b i) = - (Ω→ f .fst (ℓ* A B a b) ∙ Ω→ g .fst (ℓ* A B a b)) i -snd (f +* g) = refl - --- Inversion --* : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} - (f : join∙ A B →∙ C) → join∙ A B →∙ C -fst (-* {C = C} f) (inl x) = pt C -fst (-* {C = C} f) (inr x) = pt C -fst (-* {A = A} {B} f) (push a b i) = Ω→ f .fst (ℓ* A B a b) (~ i) -snd (-* f) = refl - -- -Π is the same as -* -Π≡-* : ∀ {ℓ} {A : Pointed ℓ} {n m : ℕ} (f : S₊∙ (suc (n + m)) →∙ A) @@ -186,6 +169,9 @@ snd (·Π≡+* {A = A} f g i) j = lem i j -π* : ∀ {ℓ} {n m : ℕ} {A : Pointed ℓ} (f : π* n m A) → π* n m A -π* = ST.map -* +-π*^ : ∀ {ℓ} {n m : ℕ} {A : Pointed ℓ} (k : ℕ) (f : π* n m A) → π* n m A +-π*^ n = iter n -π* + 1π* : ∀ {ℓ} {n m : ℕ} {A : Pointed ℓ} → π* n m A 1π* = ∣ const∙ _ _ ∣₂ @@ -330,3 +316,81 @@ snd (π*∘∙Hom {A = A} {B = B} n m f) = → GroupEquiv (π*Gr n m A) (π*Gr n m B) fst (π*∘∙Equiv n m f) = isoToEquiv (setTruncIso (pre∘∙equiv f)) snd (π*∘∙Equiv n m f) = π*∘∙Hom n m (≃∙map f) .snd + +-- Swapping indices +π*SwapIso : ∀ {ℓ} (n m : ℕ) (A : Pointed ℓ) + → Iso (π* n m A) (π* m n A) +π*SwapIso n m A = + setTruncIso (post∘∙equiv + (isoToEquiv join-comm , push (ptSn m) (ptSn n) ⁻¹)) + + +-- This is a group iso whenever n + m > 0 +π*GrSwapIso : ∀ {ℓ} (n m : ℕ) (A : Pointed ℓ) + → (n + m > 0) + → GroupIso (π*Gr n m A) (π*Gr m n A) +fst (π*GrSwapIso n m A pos) = π*SwapIso n m A +snd (π*GrSwapIso n m A pos) = + makeIsGroupHom (elim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f g → cong ∣_∣₂ (ΣPathP ((funExt + (λ { (inl x) → refl + ; (inr x) → refl + ; (push a b i) j → main f g b a j i})) + , (sym (rUnit _) + ∙ cong (cong (fst (f +* g))) + (cong₂ _∙_ refl (∙∙lCancel _) ∙ sym (rUnit _))) + ∙ cong sym + (cong₂ _∙_ + (cong (Ω→ f .fst) (ℓ*IdL (S₊∙ n) (S₊∙ m) (ptSn m)) ∙ Ω→ f .snd) + ((cong (Ω→ g .fst) (ℓ*IdL (S₊∙ n) (S₊∙ m) (ptSn m)) ∙ Ω→ g .snd)) + ∙ rCancel _)))) + where + com : (n m : ℕ) → (n + m > 0) + → (f g : join∙ (S₊∙ n) (S₊∙ m) →∙ A) (a : _) (b : _) + → (Ω→ f .fst (ℓ* (S₊∙ n) (S₊∙ m) a b) + ∙ Ω→ g .fst (ℓ* (S₊∙ n) (S₊∙ m) a b)) + ≡ (Ω→ g .fst (ℓ* (S₊∙ n) (S₊∙ m) a b) + ∙ Ω→ f .fst (ℓ* (S₊∙ n) (S₊∙ m) a b)) + com zero zero p f g a b = ⊥.rec (snotz (+-comm 1 (fst p) ∙ snd p)) + com zero (suc m) p f g a b i = + Susp·→Ωcomm' (S₊∙ (suc m)) (S₊∙ m) + (isoToEquiv (IsoSucSphereSusp _) , IsoSucSphereSusp∙' m) + (F f) (F g) i .fst b + where + F : (f : join∙ (S₊∙ zero) (S₊∙ (suc m)) →∙ A) → Σ _ _ + F f = (λ b → Ω→ f .fst (ℓ* (S₊∙ zero) (S₊∙ (suc m)) a b)) + , cong (fst (Ω→ f)) (ℓ*IdR (S₊∙ zero) (S₊∙ (suc m)) a) ∙ Ω→ f .snd + com (suc n) m p f g a b i = + Susp·→Ωcomm' (S₊∙ (suc n)) (S₊∙ n) + ((isoToEquiv (IsoSucSphereSusp _) , IsoSucSphereSusp∙' n)) + (F f) (F g) i .fst a + where + F : (f : join∙ (S₊∙ (suc n)) (S₊∙ m) →∙ A) → Σ _ _ + F f = (λ a → Ω→ f .fst (ℓ* (S₊∙ (suc n)) (S₊∙ m) a b)) + , cong (fst (Ω→ f)) (ℓ*IdL (S₊∙ (suc n)) (S₊∙ m) b) ∙ Ω→ f .snd + + main : (f g : join∙ (S₊∙ n) (S₊∙ m) →∙ A) (a : _) (b : _) + → (Ω→ f .fst (ℓ* (S₊∙ n) (S₊∙ m) a b) + ∙ Ω→ g .fst (ℓ* (S₊∙ n) (S₊∙ m) a b)) ⁻¹ + ≡ Ω→ (f ∘∙ (join-commFun , _)) .fst (ℓ* (S₊∙ m) (S₊∙ n) b a) + ∙ Ω→ (g ∘∙ (join-commFun , _)) .fst (ℓ* (S₊∙ m) (S₊∙ n) b a) + main f g a b = + cong sym (com n m pos f g a b) + ∙ symDistr _ _ + ∙ sym (cong₂ _∙_ (main-path f) (main-path g)) + where + h : invEquiv∙ ((isoToEquiv join-comm , push (ptSn m) (ptSn n) ⁻¹)) .snd + ≡ push (ptSn n) (ptSn m) ⁻¹ + h = cong₂ _∙_ refl (∙∙lCancel _) ∙ sym (rUnit _) + + main-path : (f : join∙ (S₊∙ n) (S₊∙ m) →∙ A) → _ + main-path f = + (λ i → fst (Ω→∘∙ f (join-commFun , h i) i) ((ℓ* (S₊∙ m) (S₊∙ n) b a))) + ∙ cong (Ω→ f .fst) + (sym (PathP→compPathR∙∙ + (symP (compPathR→PathP∙∙ (join-commFun-ℓ* _ _ _ _))))) + +-π*^≡ : ∀ {ℓ} {n m : ℕ} {A : Pointed ℓ} (k : ℕ) + (f : join∙ (S₊∙ n) (S₊∙ m) →∙ A) → -π*^ k ∣ f ∣₂ ≡ ∣ -*^ k f ∣₂ +-π*^≡ zero f = refl +-π*^≡ (suc k) f = cong -π* (-π*^≡ k f) diff --git a/Cubical/Homotopy/Group/Pi4S3/BrunerieNumber.agda b/Cubical/Homotopy/Group/Pi4S3/BrunerieNumber.agda index 4d3b34727d..312346dca6 100644 --- a/Cubical/Homotopy/Group/Pi4S3/BrunerieNumber.agda +++ b/Cubical/Homotopy/Group/Pi4S3/BrunerieNumber.agda @@ -14,7 +14,8 @@ open import Cubical.Homotopy.HopfInvariant.Base open import Cubical.Homotopy.Group.Pi3S2 open import Cubical.Homotopy.Group.PinSn open import Cubical.Homotopy.BlakersMassey -open import Cubical.Homotopy.Whitehead +open import Cubical.Homotopy.WhiteheadProducts.Base +open import Cubical.Homotopy.WhiteheadProducts.Properties open import Cubical.Homotopy.Connected open import Cubical.Homotopy.Group.LES open import Cubical.Homotopy.Group.Pi4S3.S3PushoutIso2 diff --git a/Cubical/Homotopy/Group/Pi4S3/DirectProof.agda b/Cubical/Homotopy/Group/Pi4S3/DirectProof.agda index b659f92073..0528e60fdb 100644 --- a/Cubical/Homotopy/Group/Pi4S3/DirectProof.agda +++ b/Cubical/Homotopy/Group/Pi4S3/DirectProof.agda @@ -62,7 +62,7 @@ open import Cubical.Homotopy.Group.Base open import Cubical.Homotopy.Group.Pi3S2 open import Cubical.Homotopy.Group.PinSn open import Cubical.Homotopy.Hopf -open import Cubical.Homotopy.Whitehead using (joinTo⋁) +open import Cubical.Homotopy.WhiteheadProducts.Base using (joinTo⋁) open import Cubical.Homotopy.Connected open import Cubical.Homotopy.HopfInvariant.HopfMap using (hopfMap≡HopfMap') -- Only imports a simple equality of two constructions of the Hopf map. @@ -85,6 +85,7 @@ open import Cubical.HITs.Sn open import Cubical.HITs.Sn.Multiplication open import Cubical.HITs.Susp renaming (toSusp to σ) open import Cubical.HITs.Join hiding (joinS¹S¹→S³) +open import Cubical.HITs.Join.CoHSpace open import Cubical.HITs.Wedge open import Cubical.HITs.Pushout open import Cubical.HITs.SetTruncation diff --git a/Cubical/Homotopy/Group/Pi4S3/Summary.agda b/Cubical/Homotopy/Group/Pi4S3/Summary.agda index bef220f158..967ecff598 100644 --- a/Cubical/Homotopy/Group/Pi4S3/Summary.agda +++ b/Cubical/Homotopy/Group/Pi4S3/Summary.agda @@ -34,7 +34,8 @@ open import Cubical.Homotopy.HopfInvariant.Base open import Cubical.Homotopy.HopfInvariant.Homomorphism open import Cubical.Homotopy.HopfInvariant.HopfMap open import Cubical.Homotopy.HopfInvariant.Brunerie -open import Cubical.Homotopy.Whitehead +open import Cubical.Homotopy.WhiteheadProducts.Base +open import Cubical.Homotopy.WhiteheadProducts.Properties open import Cubical.Homotopy.Group.Base hiding (π) open import Cubical.Homotopy.Group.Pi3S2 open import Cubical.Homotopy.Group.Pi4S3.BrunerieNumber diff --git a/Cubical/Homotopy/Group/Properties.agda b/Cubical/Homotopy/Group/Properties.agda index 56aaaa4062..fc1aa6bc03 100644 --- a/Cubical/Homotopy/Group/Properties.agda +++ b/Cubical/Homotopy/Group/Properties.agda @@ -1,24 +1,114 @@ {-# OPTIONS --lossy-unification #-} +-- Other lemmas about homotopy groups 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.GroupoidLaws +open import Cubical.Foundations.Path open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function + +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Group.Base +open import Cubical.HITs.S1 +open import Cubical.HITs.Sn +open import Cubical.HITs.Susp open import Cubical.HITs.PropositionalTruncation open import Cubical.HITs.SetTruncation as ST open import Cubical.HITs.Truncation as TR +open import Cubical.Data.Sigma open import Cubical.Data.Nat open import Cubical.Algebra.Group.Morphisms open import Cubical.Algebra.AbGroup open import Cubical.Algebra.Group.Abelianization.Properties +private + variable + ℓ : Level + +-- Interaction of iterated inversion with other sphere inversions +-π^≡iter-Π : {k : ℕ} {A : Pointed ℓ} (n : ℕ) + → (x : π' (suc k) A) → -π^ n x ≡ ST.map (iter n -Π) x +-π^≡iter-Π zero = ST.elim (λ _ → isSetPathImplicit) λ _ → refl +-π^≡iter-Π {k = k} (suc n) = + ST.elim (λ _ → isSetPathImplicit) + λ f → cong (-π^ 1) (-π^≡iter-Π {k = k} n ∣ f ∣₂) + +-Π≡∘-S : {k : ℕ} {A : Pointed ℓ} (f : S₊∙ (suc k) →∙ A) + → -Π f ≡ (f ∘∙ (invSphere , invSpherePt)) +-Π≡∘-S {k = zero} f = + ΣPathP ((funExt (λ { base → refl + ; (loop i) → refl})) + , lUnit (snd f)) +-Π≡∘-S {k = suc k} f = + ΣPathP ((funExt (λ { north → cong (fst f) (merid (ptSn (suc k))) + ; south → refl + ; (merid a i) j + → fst f (compPath-filler (merid a) + (sym (merid (ptSn (suc k)))) (~ j) (~ i))})) + , (compPath-filler' + (cong (fst f) (sym (merid (ptSn (suc k))))) + (snd f))) + +iter-Π≡∘-S^ : {k : ℕ} {A : Pointed ℓ} (n : ℕ) (f : S₊∙ (suc k) →∙ A) + → iter n -Π f ≡ (f ∘∙ (-S^ n , -S^pt n)) +iter-Π≡∘-S^ zero f = ΣPathP (refl , lUnit (snd f)) +iter-Π≡∘-S^ {k = k} {A} (suc n) f = + cong -Π (iter-Π≡∘-S^ {k = k} {A} n f) + ∙ -Π≡∘-S (f ∘∙ (-S^ n , -S^pt n)) + ∙ ∘∙-assoc f (-S^ n , -S^pt n) (invSphere , invSpherePt) -- + ∙ cong (f ∘∙_) + (ΣPathP ((funExt (λ x → sym (invSphere-S^ n x))) + , (lem k n))) + where + fl : (k : ℕ) (n : ℕ) + → Σ[ p ∈ invSphere (invSphere (ptSn (suc k))) ≡ ptSn (suc k) ] + ((Square (cong invSphere (sym (invSpherePt {k = k})) ) + refl invSpherePt p) + × Square (cong (invSphere ∘ invSphere) (-S^pt {k = k} n)) + (cong invSphere (cong invSphere (-S^pt n) ∙ invSpherePt) + ∙ invSpherePt) + refl + p) + fl zero n .fst = refl + fl (suc k) n .fst = refl + fl zero n .snd = refl + , (cong (congS invLooper) (rUnit (cong invLooper (-S^pt n))) ∙ rUnit _) + fl (suc k) n .snd .fst i j = merid (ptSn (suc k)) (~ i ∧ ~ j) + fl (suc k) n .snd .snd i j = + hcomp (λ r → λ {(i = i0) → invSusp (invSusp (-S^pt n j)) + ; (j = i0) → invSusp (invSusp (iter n invSusp north)) + ; (j = i1) → merid (ptSn (suc k)) (~ r ∧ i)}) + (invSusp (compPath-filler (cong invSusp (-S^pt n)) + (sym (merid (ptSn (suc k)))) i j)) + + lem : (k : ℕ) (n : ℕ) + → Square (cong (-S^ n) (invSpherePt {k = k}) ∙ -S^pt n) + (-S^pt (suc n)) + (sym (invSphere-S^ n (ptSn (suc k)))) refl + lem k zero = sym (rUnit _) ∙ lUnit _ + lem k (suc n) i j = + hcomp (λ r → λ {(i = i0) → (cong (-S^ (suc n)) (invSpherePt {k = k}) + ∙ compPath-filler (cong invSphere (-S^pt n)) + invSpherePt r) j + ; (i = i1) → fl k n .snd .snd r j + ; (j = i0) → invSphere-S^ (suc n) (ptSn (suc k)) (~ i) + ; (j = i1) → fl k n .snd .fst r i + }) + (hcomp (λ r → λ {(i = i0) → cong-∙ invSphere + (cong (-S^ n) (invSpherePt {k = k})) + (-S^pt n) r j + ; (i = i1) → invSphere (doubleCompPath-filler refl + (cong invSphere (-S^pt n)) invSpherePt (~ r) j) + ; (j = i0) → invSphere-S^ (suc n) (ptSn (suc k)) (~ i) + ; (j = i1) → invSphere (invSpherePt (~ r ∨ ~ i))}) + (invSphere (lem k n i j))) + connectedFun→πEquiv : ∀ {ℓ} {A B : Pointed ℓ} (n : ℕ) (f : A →∙ B) → isConnectedFun (3 + n) (fst f) → GroupEquiv (πGr n A) (πGr n B) diff --git a/Cubical/Homotopy/HopfInvariant/Brunerie.agda b/Cubical/Homotopy/HopfInvariant/Brunerie.agda index 145356f14f..38a638f040 100644 --- a/Cubical/Homotopy/HopfInvariant/Brunerie.agda +++ b/Cubical/Homotopy/HopfInvariant/Brunerie.agda @@ -12,7 +12,7 @@ module Cubical.Homotopy.HopfInvariant.Brunerie where open import Cubical.Homotopy.HopfInvariant.Base open import Cubical.Homotopy.Group.Pi4S3.BrunerieNumber open import Cubical.Homotopy.Group.Pi4S3.S3PushoutIso -open import Cubical.Homotopy.Whitehead +open import Cubical.Homotopy.WhiteheadProducts.Base open import Cubical.ZCohomology.Base open import Cubical.ZCohomology.GroupStructure diff --git a/Cubical/Homotopy/Loopspace.agda b/Cubical/Homotopy/Loopspace.agda index d6574ab7f2..0d5753373a 100644 --- a/Cubical/Homotopy/Loopspace.agda +++ b/Cubical/Homotopy/Loopspace.agda @@ -462,6 +462,13 @@ EH-π : ∀ {ℓ} {A : Pointed ℓ} (n : ℕ) (p q : ∥ typ ((Ω^ (2 + n)) A) EH-π n = elim2 (λ x y → isOfHLevelPath 2 isSetSetTrunc _ _) λ p q → cong ∣_∣₂ (EH n p q) +-- composition of maps into loop spaces +·→Ω : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + (f g : A →∙ Ω B) + → A →∙ Ω B +fst (·→Ω f g) x = fst f x ∙ fst g x +snd (·→Ω f g) = cong₂ _∙_ (snd f) (snd g) ∙ sym (rUnit refl) + {- If A -> B -> C is a fiber sequence, and Ω B -> Ω C has a section, then Ω B is the product Ω A × Ω C. diff --git a/Cubical/Homotopy/Whitehead.agda b/Cubical/Homotopy/Whitehead.agda deleted file mode 100644 index befeb378a4..0000000000 --- a/Cubical/Homotopy/Whitehead.agda +++ /dev/null @@ -1,570 +0,0 @@ -{-# OPTIONS --lossy-unification #-} -module Cubical.Homotopy.Whitehead where - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Function -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Pointed -open import Cubical.Foundations.GroupoidLaws - -open import Cubical.Data.Nat -open import Cubical.Data.Sigma -open import Cubical.Data.Unit - -open import Cubical.HITs.Susp renaming (toSusp to σ) -open import Cubical.HITs.Pushout -open import Cubical.HITs.Sn -open import Cubical.HITs.Sn.Multiplication -open import Cubical.HITs.Join -open import Cubical.HITs.Wedge -open import Cubical.HITs.SetTruncation - -open import Cubical.Homotopy.Group.Base -open import Cubical.Homotopy.Loopspace - -open Iso -open 3x3-span - --- Whitehead product (main definition) -[_∣_]-pre : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} - → (S₊∙ (suc n) →∙ X) - → (S₊∙ (suc m) →∙ X) - → join∙ (S₊∙ n) (S₊∙ m) →∙ X -fst ([_∣_]-pre {X = X} {n = n} {m = m} f g) (inl x) = pt X -fst ([_∣_]-pre {X = X} {n = n} {m = m} f g) (inr x) = pt X -fst ([_∣_]-pre {n = n} {m = m} f g) (push a b i) = - (Ω→ g .fst (σS b) ∙ Ω→ f .fst (σS a)) i -snd ([_∣_]-pre {n = n} {m = m} f g) = refl - -[_∣_] : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} - → (S₊∙ (suc n) →∙ X) - → (S₊∙ (suc m) →∙ X) - → S₊∙ (suc (n + m)) →∙ X -[_∣_] {n = n} {m = m} f g = - [ f ∣ g ]-pre ∘∙ (sphere→Join n m , IsoSphereJoin⁻Pres∙ n m) - --- Whitehead product (as a composition) -joinTo⋁ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} - → join (typ A) (typ B) - → (Susp (typ A) , north) ⋁ (Susp (typ B) , north) -joinTo⋁ (inl x) = inr north -joinTo⋁ (inr x) = inl north -joinTo⋁ {A = A} {B = B} (push a b i) = - ((λ i → inr (σ B b i)) - ∙∙ sym (push tt) - ∙∙ λ i → inl (σ A a i)) i - -[_∣_]comp : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} - → (S₊∙ (suc n) →∙ X) - → (S₊∙ (suc m) →∙ X) - → S₊∙ (suc (n + m)) →∙ X -[_∣_]comp {n = n} {m = m} f g = - (((f ∘∙ (inv (IsoSucSphereSusp n) , IsoSucSphereSusp∙ n)) - ∨→ (g ∘∙ (inv (IsoSucSphereSusp m) , IsoSucSphereSusp∙ m)) - , cong (fst f) (IsoSucSphereSusp∙ n) ∙ snd f) - ∘∙ ((joinTo⋁ {A = S₊∙ n} {B = S₊∙ m} , sym (push tt)))) - ∘∙ (inv (IsoSphereJoin n m) , IsoSphereJoin⁻Pres∙ n m) - -[]comp≡[] : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} - → (f : (S₊∙ (suc n) →∙ X)) - → (g : (S₊∙ (suc m) →∙ X)) - → [ f ∣ g ]comp ≡ [ f ∣ g ] -[]comp≡[] {X = X} {n = n} {m} f g = - cong (_∘∙ (sphere→Join n m , IsoSphereJoin⁻Pres∙ n m)) (main n m f g) - where - ∨fun : {n m : ℕ} (f : (S₊∙ (suc n) →∙ X)) (g : (S₊∙ (suc m) →∙ X)) - → ((_ , inl north) →∙ X) - ∨fun {n = n} {m} f g = - ((f ∘∙ (inv (IsoSucSphereSusp n) , IsoSucSphereSusp∙ n)) ∨→ - (g ∘∙ (inv (IsoSucSphereSusp m) , IsoSucSphereSusp∙ m)) - , cong (fst f) (IsoSucSphereSusp∙ n) ∙ snd f) - - open import Cubical.Foundations.Path - main : (n m : ℕ) (f : (S₊∙ (suc n) →∙ X)) (g : (S₊∙ (suc m) →∙ X)) - → (∨fun f g ∘∙ (joinTo⋁ {A = S₊∙ n} {B = S₊∙ m} , sym (push tt))) - ≡ [ f ∣ g ]-pre - main n m f g = - ΣPathP ((funExt (λ { (inl x) → rp - ; (inr x) → lp - ; (push a b i) j → pushcase a b j i})) - , ((cong₂ _∙_ (symDistr _ _) refl - ∙ sym (assoc _ _ _) - ∙ cong (rp ∙_) (rCancel _) - ∙ sym (rUnit rp)) - ◁ λ i j → rp (i ∨ j))) - where - lp = cong (fst f) (IsoSucSphereSusp∙ n) ∙ snd f - rp = cong (fst g) (IsoSucSphereSusp∙ m) ∙ snd g - - help : (n : ℕ) (a : _) - → Square (cong (inv (IsoSucSphereSusp n)) (σ (S₊∙ n) a)) (σS a) - (IsoSucSphereSusp∙ n) (IsoSucSphereSusp∙ n) - help zero a = cong-∙ SuspBool→S¹ (merid a) (sym (merid (pt (S₊∙ zero)))) - ∙ sym (rUnit _) - help (suc n) a = refl - - ∙∙Distr∙ : ∀ {ℓ} {A : Type ℓ} {x y z w u : A} - {p : x ≡ y} {q : y ≡ z} {r : z ≡ w} {s : w ≡ u} - → p ∙∙ q ∙ r ∙∙ s ≡ ((p ∙ q) ∙ r ∙ s) - ∙∙Distr∙ = doubleCompPath≡compPath _ _ _ - ∙ assoc _ _ _ - ∙ cong₂ _∙_ (assoc _ _ _) refl - ∙ sym (assoc _ _ _) - - - pushcase : (a : S₊ n) (b : S₊ m) - → Square (cong (∨fun f g .fst ∘ joinTo⋁ {A = S₊∙ n} {B = S₊∙ m}) (push a b)) - (cong (fst [ f ∣ g ]-pre) (push a b)) - rp lp - pushcase a b = - (cong-∙∙ (∨fun f g .fst) _ _ _ - ∙ (λ i → cong (fst g) (PathP→compPathR∙∙ (help _ b) i) - ∙∙ symDistr lp (sym rp) i - ∙∙ cong (fst f) (PathP→compPathR∙∙ (help _ a) i)) - ∙ (λ i → cong (fst g) - (IsoSucSphereSusp∙ m - ∙∙ σS b - ∙∙ (λ j → IsoSucSphereSusp∙ m (~ j ∨ i))) - ∙∙ compPath-filler' (cong (fst g) (IsoSucSphereSusp∙ m)) (snd g) (~ i) - ∙ sym (compPath-filler' (cong (fst f) (IsoSucSphereSusp∙ n)) (snd f) (~ i)) - ∙∙ cong (fst f) - ((λ j → IsoSucSphereSusp∙ n (i ∨ j)) - ∙∙ σS a - ∙∙ sym (IsoSucSphereSusp∙ n)))) - ◁ compPathR→PathP∙∙ - ( ∙∙Distr∙ - ∙ cong₂ _∙_ (cong₂ _∙_ (cong (cong (fst g)) (sym (compPath≡compPath' _ _))) - refl) - refl - ∙ cong₂ _∙_ (cong (_∙ snd g) (cong-∙ (fst g) (IsoSucSphereSusp∙ m) (σS b)) - ∙ sym (assoc _ _ _)) - (cong (sym (snd f) ∙_) - (cong-∙ (fst f) (σS a) - (sym (IsoSucSphereSusp∙ n))) - ∙ assoc _ _ _) - ∙ sym ∙∙Distr∙ - ∙ cong₃ _∙∙_∙∙_ refl (cong₂ _∙_ refl (compPath≡compPath' _ _)) refl - ∙ λ i → compPath-filler (cong (fst g) (IsoSucSphereSusp∙ m)) (snd g) i - ∙∙ ((λ j → snd g (~ j ∧ i)) ∙∙ cong (fst g) (σS b) ∙∙ snd g) - ∙ (sym (snd f) ∙∙ cong (fst f) (σS a) ∙∙ λ j → snd f (j ∧ i)) - ∙∙ sym (compPath-filler (cong (fst f) (IsoSucSphereSusp∙ n)) (snd f) i)) - --- Homotopy group version -[_∣_]π' : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} - → π' (suc n) X → π' (suc m) X - → π' (suc (n + m)) X -[_∣_]π' = elim2 (λ _ _ → squash₂) λ f g → ∣ [ f ∣ g ] ∣₂ - --- We prove that the function joinTo⋁ used in the definition of the whitehead --- product gives an equivalence between (Susp A × Susp B) and the --- appropriate cofibre of joinTo⋁ (last two theorems in the following --- module). - -module _ (A B : Type) (a₀ : A) (b₀ : B) where - private - W = joinTo⋁ {A = (A , a₀)} {B = (B , b₀)} - - A∨B = (Susp A , north) ⋁ (Susp B , north) - - σB = σ (B , b₀) - σA = σ (A , a₀) - - cofibW = Pushout W λ _ → tt - - whitehead3x3 : 3x3-span - A00 whitehead3x3 = Susp A - A02 whitehead3x3 = B - A04 whitehead3x3 = Unit - A20 whitehead3x3 = B - A22 whitehead3x3 = A × B - A24 whitehead3x3 = A - A40 whitehead3x3 = B - A42 whitehead3x3 = B - A44 whitehead3x3 = Unit - f10 whitehead3x3 _ = south - f12 whitehead3x3 = snd - f14 whitehead3x3 _ = tt - f30 whitehead3x3 = idfun B - f32 whitehead3x3 = snd - f34 whitehead3x3 _ = tt - f01 whitehead3x3 _ = north - f21 whitehead3x3 = snd - f41 whitehead3x3 = idfun B - f03 whitehead3x3 _ = tt - f23 whitehead3x3 = fst - f43 whitehead3x3 _ = tt - H11 whitehead3x3 x = merid (fst x) - H13 whitehead3x3 _ = refl - H31 whitehead3x3 _ = refl - H33 whitehead3x3 _ = refl - - A0□→A∨B : A0□ whitehead3x3 → A∨B - A0□→A∨B (inl x) = inl x - A0□→A∨B (inr x) = inr north - A0□→A∨B (push a i) = (push tt ∙ λ i → inr (σB a (~ i))) i - - A∨B→A0□ : A∨B → A0□ whitehead3x3 - A∨B→A0□ (inl x) = inl x - A∨B→A0□ (inr north) = inl north - A∨B→A0□ (inr south) = inl north - A∨B→A0□ (inr (merid b i)) = (push b₀ ∙ sym (push b)) i - A∨B→A0□ (push a i) = inl north - - Iso-A0□-⋁ : Iso (A0□ whitehead3x3) A∨B - fun Iso-A0□-⋁ = A0□→A∨B - inv Iso-A0□-⋁ = A∨B→A0□ - rightInv Iso-A0□-⋁ (inl x) = refl - rightInv Iso-A0□-⋁ (inr north) = push tt - rightInv Iso-A0□-⋁ (inr south) = push tt ∙ λ i → inr (merid b₀ i) - rightInv Iso-A0□-⋁ (inr (merid a i)) j = lem j i - where - lem : PathP (λ i → push tt i ≡ (push tt ∙ (λ i → inr (merid b₀ i))) i) - (cong A0□→A∨B (cong A∨B→A0□ λ i → inr (merid a i))) - (λ i → inr (merid a i)) - lem = (cong-∙ A0□→A∨B (push b₀) (sym (push a)) - ∙ cong₂ _∙_ (cong (push tt ∙_) - (λ j i → inr (rCancel (merid b₀) j (~ i))) - ∙ sym (rUnit (push tt))) - (symDistr (push tt) (λ i → inr (σB a (~ i))))) - ◁ λ i j → hcomp (λ k → - λ { (i = i0) → compPath-filler' (push tt) - (compPath-filler (λ i → inr (σB a i)) - (sym (push tt)) k) k j - ; (i = i1) → inr (merid a j) - ; (j = i0) → push tt (i ∨ ~ k) - ; (j = i1) → compPath-filler' (push tt) - (λ i → inr (merid b₀ i)) k i}) - (inr (compPath-filler (merid a) - (sym (merid b₀)) (~ i) j)) - - rightInv Iso-A0□-⋁ (push a i) j = push tt (i ∧ j) - leftInv Iso-A0□-⋁ (inl x) = refl - leftInv Iso-A0□-⋁ (inr tt) = push b₀ - leftInv Iso-A0□-⋁ (push b i) j = help j i - where - help : PathP (λ i → inl north ≡ push b₀ i) - (cong A∨B→A0□ (cong A0□→A∨B (push b))) - (push b) - help = (cong-∙ A∨B→A0□ (push tt) (λ i → inr (σB b (~ i))) - ∙ (λ i → lUnit (sym (cong-∙ (A∨B→A0□ ∘ inr) - (merid b) (sym (merid b₀)) i)) (~ i)) - ∙ cong sym (cong ((push b₀ ∙ sym (push b)) ∙_) - (cong sym (rCancel (push b₀)))) - ∙ cong sym (sym (rUnit (push b₀ ∙ sym (push b))))) - ◁ λ i j → compPath-filler' (push b₀) (sym (push b)) (~ i) (~ j) - - Iso-A2□-join : Iso (A2□ whitehead3x3) (join A B) - fun Iso-A2□-join (inl x) = inr x - fun Iso-A2□-join (inr x) = inl x - fun Iso-A2□-join (push (a , b) i) = push a b (~ i) - inv Iso-A2□-join (inl x) = inr x - inv Iso-A2□-join (inr x) = inl x - inv Iso-A2□-join (push a b i) = push (a , b) (~ i) - rightInv Iso-A2□-join (inl x) = refl - rightInv Iso-A2□-join (inr x) = refl - rightInv Iso-A2□-join (push a b i) = refl - leftInv Iso-A2□-join (inl x) = refl - leftInv Iso-A2□-join (inr x) = refl - leftInv Iso-A2□-join (push a i) = refl - - isContrA4□ : isContr (A4□ whitehead3x3) - fst isContrA4□ = inr tt - snd isContrA4□ (inl x) = sym (push x) - snd isContrA4□ (inr x) = refl - snd isContrA4□ (push a i) j = push a (i ∨ ~ j) - - A4□≃Unit : A4□ whitehead3x3 ≃ Unit - A4□≃Unit = isContr→≃Unit isContrA4□ - - Iso-A□0-Susp : Iso (A□0 whitehead3x3) (Susp A) - fun Iso-A□0-Susp (inl x) = x - fun Iso-A□0-Susp (inr x) = north - fun Iso-A□0-Susp (push a i) = merid a₀ (~ i) - inv Iso-A□0-Susp x = inl x - rightInv Iso-A□0-Susp x = refl - leftInv Iso-A□0-Susp (inl x) = refl - leftInv Iso-A□0-Susp (inr x) = (λ i → inl (merid a₀ i)) ∙ push x - leftInv Iso-A□0-Susp (push a i) j = - hcomp (λ k → λ { (i = i0) → inl (merid a₀ (k ∨ j)) - ; (i = i1) → compPath-filler - (λ i₁ → inl (merid a₀ i₁)) - (push (idfun B a)) k j - ; (j = i0) → inl (merid a₀ (~ i ∧ k)) - ; (j = i1) → push a (i ∧ k)}) - (inl (merid a₀ j)) - - Iso-A□2-Susp× : Iso (A□2 whitehead3x3) (Susp A × B) - fun Iso-A□2-Susp× (inl x) = north , x - fun Iso-A□2-Susp× (inr x) = south , x - fun Iso-A□2-Susp× (push a i) = merid (fst a) i , (snd a) - inv Iso-A□2-Susp× (north , y) = inl y - inv Iso-A□2-Susp× (south , y) = inr y - inv Iso-A□2-Susp× (merid a i , y) = push (a , y) i - rightInv Iso-A□2-Susp× (north , snd₁) = refl - rightInv Iso-A□2-Susp× (south , snd₁) = refl - rightInv Iso-A□2-Susp× (merid a i , snd₁) = refl - leftInv Iso-A□2-Susp× (inl x) = refl - leftInv Iso-A□2-Susp× (inr x) = refl - leftInv Iso-A□2-Susp× (push a i) = refl - - Iso-A□4-Susp : Iso (A□4 whitehead3x3) (Susp A) - fun Iso-A□4-Susp (inl x) = north - fun Iso-A□4-Susp (inr x) = south - fun Iso-A□4-Susp (push a i) = merid a i - inv Iso-A□4-Susp north = inl tt - inv Iso-A□4-Susp south = inr tt - inv Iso-A□4-Susp (merid a i) = push a i - rightInv Iso-A□4-Susp north = refl - rightInv Iso-A□4-Susp south = refl - rightInv Iso-A□4-Susp (merid a i) = refl - leftInv Iso-A□4-Susp (inl x) = refl - leftInv Iso-A□4-Susp (inr x) = refl - leftInv Iso-A□4-Susp (push a i) = refl - - Iso-PushSusp×-Susp×Susp : - Iso (Pushout {A = Susp A × B} fst fst) (Susp A × Susp B) - Iso-PushSusp×-Susp×Susp = theIso - where - F : Pushout {A = Susp A × B} fst fst - → Susp A × Susp B - F (inl x) = x , north - F (inr x) = x , north - F (push (x , b) i) = x , σB b i - - G : Susp A × Susp B → Pushout {A = Susp A × B} fst fst - G (a , north) = inl a - G (a , south) = inr a - G (a , merid b i) = push (a , b) i - - retr : retract F G - retr (inl x) = refl - retr (inr x) = push (x , b₀) - retr (push (a , b) i) j = help j i - where - help : PathP (λ i → Path (Pushout fst fst) (inl a) (push (a , b₀) i)) - (cong G (λ i → a , σB b i)) - (push (a , b)) - help = cong-∙ (λ x → G (a , x)) (merid b) (sym (merid b₀)) - ◁ λ i j → compPath-filler - (push (a , b)) - (sym (push (a , b₀))) - (~ i) j - - theIso : Iso (Pushout fst fst) (Susp A × Susp B) - fun theIso = F - inv theIso = G - rightInv theIso (a , north) = refl - rightInv theIso (a , south) = ΣPathP (refl , merid b₀) - rightInv theIso (a , merid b i) j = - a , compPath-filler (merid b) (sym (merid b₀)) (~ j) i - leftInv theIso = retr - - Iso-A□○-PushSusp× : - Iso (A□○ whitehead3x3) (Pushout {A = Susp A × B} fst fst) - Iso-A□○-PushSusp× = - pushoutIso _ _ fst fst - (isoToEquiv Iso-A□2-Susp×) - (isoToEquiv Iso-A□0-Susp) - (isoToEquiv Iso-A□4-Susp) - (funExt (λ { (inl x) → refl - ; (inr x) → merid a₀ - ; (push a i) j → help₁ a j i})) - (funExt λ { (inl x) → refl - ; (inr x) → refl - ; (push a i) j - → fun Iso-A□4-Susp (rUnit (push (fst a)) (~ j) i)}) - where - help₁ : (a : A × B) - → PathP (λ i → north ≡ merid a₀ i) - (cong (fun Iso-A□0-Susp) - (cong (f□1 whitehead3x3) (push a))) - (merid (fst a)) - help₁ a = - (cong-∙∙ (fun Iso-A□0-Susp) - (λ i → inl (merid (fst a) i)) - (push (snd a)) - refl) - ◁ (λ i j → hcomp (λ k → λ {(i = i1) → merid (fst a) (j ∨ ~ k) - ; (j = i0) → merid (fst a) (~ k) - ; (j = i1) → merid a₀ i}) - (merid a₀ (i ∨ ~ j))) - - Iso-A□○-Susp×Susp : Iso (A□○ whitehead3x3) (Susp A × Susp B) - Iso-A□○-Susp×Susp = compIso Iso-A□○-PushSusp× Iso-PushSusp×-Susp×Susp - - Iso-A○□-cofibW : Iso (A○□ whitehead3x3) cofibW - Iso-A○□-cofibW = - pushoutIso _ _ - W (λ _ → tt) - (isoToEquiv Iso-A2□-join) (isoToEquiv Iso-A0□-⋁) - A4□≃Unit - (funExt lem) - λ _ _ → tt - where - lem : (x : A2□ whitehead3x3) - → A0□→A∨B (f1□ whitehead3x3 x) ≡ W (fun Iso-A2□-join x) - lem (inl x) = (λ i → inl (merid a₀ (~ i))) - lem (inr x) = refl - lem (push (a , b) i) j = help j i - where - help : PathP (λ i → Path (Pushout (λ _ → north) (λ _ → north)) - ((inl (merid a₀ (~ i)))) - (inr north)) - (cong A0□→A∨B (cong (f1□ whitehead3x3) (push (a , b)))) - (cong W (cong (fun Iso-A2□-join) (push (a , b)))) - help = (cong-∙∙ A0□→A∨B (λ i → inl (merid a (~ i))) (push b) refl - ∙ λ j → (λ i₂ → inl (merid a (~ i₂))) - ∙∙ compPath-filler (push tt) (λ i → inr (σB b (~ i))) (~ j) - ∙∙ λ i → inr (σB b (~ i ∧ j))) - ◁ (λ j → (λ i → inl (sym (compPath-filler - (merid a) (sym (merid a₀)) j) i)) - ∙∙ push tt - ∙∙ λ i → inr (σB b (~ i))) - - Iso₁-Susp×Susp-cofibW : Iso (Susp A × Susp B) cofibW - Iso₁-Susp×Susp-cofibW = - compIso (invIso Iso-A□○-Susp×Susp) - (compIso (3x3-Iso whitehead3x3) Iso-A○□-cofibW) - - -- Main iso - Iso-Susp×Susp-cofibJoinTo⋁ : Iso (Susp A × Susp B) cofibW - Iso-Susp×Susp-cofibJoinTo⋁ = - compIso (Σ-cong-iso-snd (λ _ → invSuspIso)) - Iso₁-Susp×Susp-cofibW - - -- The induced function A ∨ B → Susp A × Susp B satisfies - -- the following identity - Susp×Susp→cofibW≡ : Path (A∨B → Susp A × Susp B) - (Iso.inv Iso-Susp×Susp-cofibJoinTo⋁ ∘ inl) - ⋁↪ - Susp×Susp→cofibW≡ = - funExt λ { (inl x) → ΣPathP (refl , (sym (merid b₀))) - ; (inr north) → ΣPathP (refl , (sym (merid b₀))) - ; (inr south) → refl - ; (inr (merid a i)) j → lem₂ a j i - ; (push a i) j → north , (merid b₀ (~ j))} - where - f₁ = fun Iso-PushSusp×-Susp×Susp - f₂ = fun Iso-A□○-PushSusp× - f₃ = backward-l whitehead3x3 - f₄ = fun (Σ-cong-iso-snd (λ _ → invSuspIso)) - - lem : (b : B) - → cong (f₁ ∘ f₂ ∘ f₃) (push b) - ≡ (λ i → north , σB b i) - lem b = cong (cong f₁) (sym (rUnit (push (north , b)))) - - lem₂ : (a : B) - → PathP (λ i → (north , merid b₀ (~ i)) ≡ (north , south)) - (cong (f₄ ∘ f₁ ∘ f₂ ∘ f₃ ∘ A∨B→A0□ ∘ inr) (merid a)) - λ i → north , merid a i - lem₂ a = - cong (cong f₄) (cong-∙ (f₁ ∘ f₂ ∘ f₃) (push b₀) (sym (push a)) - ∙∙ cong₂ _∙_ (lem b₀ ∙ (λ j i → north , rCancel (merid b₀) j i)) - (cong sym (lem a)) - ∙∙ sym (lUnit (λ i₁ → north , σB a (~ i₁)))) - ∙ (λ i j → north , cong-∙ invSusp (merid a) (sym (merid b₀)) i (~ j) ) - ◁ λ i j → north , compPath-filler (sym (merid a)) (merid b₀) (~ i) (~ j) - --- Generalised Whitehead products -module _ {ℓ ℓ' ℓ''} {A : Pointed ℓ} - {B : Pointed ℓ'} {C : Pointed ℓ''} - (f : Susp∙ (typ A) →∙ C) (g : Susp∙ (typ B) →∙ C) where - - _·w_ : join∙ A B →∙ C - fst _·w_ (inl x) = pt C - fst _·w_ (inr x) = pt C - fst _·w_ (push a b i) = (Ω→ g .fst (σ B b) ∙ Ω→ f .fst (σ A a)) i - snd _·w_ = refl - - -- The generalised Whitehead product vanishes under suspension - isConst-Susp·w : suspFun∙ (_·w_ .fst) ≡ const∙ _ _ - isConst-Susp·w = Susp·w∙ - ∙ cong suspFun∙ (cong fst isConst-const*) - ∙ ΣPathP ((suspFunConst (pt C)) , refl) - where - const* : join∙ A B →∙ C - fst const* (inl x) = pt C - fst const* (inr x) = pt C - fst const* (push a b i) = - (Ω→ f .fst (σ A a) ∙ Ω→ g .fst (σ B b)) i - snd const* = refl - - isConst-const* : const* ≡ const∙ _ _ - fst (isConst-const* i) (inl x) = Ω→ f .fst (σ A x) i - fst (isConst-const* i) (inr x) = Ω→ g .fst (σ B x) (~ i) - fst (isConst-const* i) (push a b j) = - compPath-filler'' (Ω→ f .fst (σ A a)) (Ω→ g .fst (σ B b)) (~ i) j - snd (isConst-const* i) j = - (cong (Ω→ f .fst) (rCancel (merid (pt A))) ∙ Ω→ f .snd) j i - - Susp·w : suspFun (fst _·w_) ≡ suspFun (fst const*) - Susp·w i north = north - Susp·w i south = south - Susp·w i (merid (inl x) j) = merid (pt C) j - Susp·w i (merid (inr x) j) = merid (pt C) j - Susp·w i (merid (push a b k) j) = - hcomp (λ r → - λ {(i = i0) → fill₁ k (~ r) j - ; (i = i1) → fill₂ k (~ r) j - ; (j = i0) → north - ; (j = i1) → merid (pt C) r - ; (k = i0) → compPath-filler (merid (snd C)) (merid (pt C) ⁻¹) (~ r) j - ; (k = i1) → compPath-filler (merid (snd C)) (merid (pt C) ⁻¹) (~ r) j}) - (hcomp (λ r → - λ {(i = i0) → doubleCompPath-filler - (sym (rCancel (merid (pt C)))) - (λ k → fill₁ k i1) - (rCancel (merid (pt C))) (~ r) k j - ; (i = i1) → doubleCompPath-filler - (sym (rCancel (merid (pt C)))) - (λ k → fill₂ k i1) - (rCancel (merid (pt C))) (~ r) k j - ; (j = i0) → north - ; (j = i1) → north - ; (k = i0) → rCancel (merid (pt C)) (~ r) j - ; (k = i1) → rCancel (merid (pt C)) (~ r) j}) - (main i k j)) - where - F : Ω C .fst → (Ω^ 2) (Susp∙ (fst C)) .fst - F p = sym (rCancel (merid (pt C))) - ∙∙ cong (σ C) p - ∙∙ rCancel (merid (pt C)) - - F-hom : (p q : _) → F (p ∙ q) ≡ F p ∙ F q - F-hom p q = - cong (sym (rCancel (merid (pt C))) - ∙∙_∙∙ rCancel (merid (pt C))) - (cong-∙ (σ C) p q) - ∙ doubleCompPath≡compPath (sym (rCancel (merid (pt C)))) _ _ - ∙ cong (sym (rCancel (merid (pt C))) ∙_) - (sym (assoc _ _ _)) - ∙ assoc _ _ _ - ∙ (λ i → (sym (rCancel (merid (pt C))) - ∙ compPath-filler (cong (σ C) p) (rCancel (merid (pt C))) i) - ∙ compPath-filler' (sym (rCancel (merid (pt C)))) - (cong (σ C) q ∙ rCancel (merid (pt C))) i) - ∙ cong₂ _∙_ (sym (doubleCompPath≡compPath _ _ _)) - (sym (doubleCompPath≡compPath _ _ _)) - - main : F ((Ω→ g .fst (σ B b) ∙ Ω→ f .fst (σ A a))) - ≡ F ((Ω→ f .fst (σ A a) ∙ Ω→ g .fst (σ B b))) - main = F-hom (Ω→ g .fst (σ B b)) (Ω→ f .fst (σ A a)) - ∙ EH 0 _ _ - ∙ sym (F-hom (Ω→ f .fst (σ A a)) (Ω→ g .fst (σ B b))) - - fill₁ : (k : I) → _ - fill₁ k = compPath-filler - (merid ((Ω→ g .fst (σ B b) - ∙ Ω→ f .fst (σ A a)) k)) - (merid (pt C) ⁻¹) - - fill₂ : (k : I) → _ - fill₂ k = compPath-filler - (merid ((Ω→ f .fst (σ A a) - ∙ Ω→ g .fst (σ B b)) k)) - (merid (pt C) ⁻¹) - - Susp·w∙ : suspFun∙ (_·w_ .fst) ≡ suspFun∙ (fst const*) - Susp·w∙ = ΣPathP (Susp·w , refl) diff --git a/Cubical/Homotopy/WhiteheadProducts/Base.agda b/Cubical/Homotopy/WhiteheadProducts/Base.agda new file mode 100644 index 0000000000..9cb9befef8 --- /dev/null +++ b/Cubical/Homotopy/WhiteheadProducts/Base.agda @@ -0,0 +1,184 @@ +{-# OPTIONS --safe --lossy-unification #-} +module Cubical.Homotopy.WhiteheadProducts.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Path +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.Data.Nat +open import Cubical.Data.Sigma +open import Cubical.Data.Bool +open import Cubical.Data.Int using (ℤ ; pos) + +open import Cubical.HITs.Susp renaming (toSusp to σ) +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.Sn.Multiplication +open import Cubical.HITs.Join +open import Cubical.HITs.Wedge +open import Cubical.HITs.SetTruncation +open import Cubical.HITs.PropositionalTruncation as PT hiding (elim2) +open import Cubical.HITs.Truncation as TR hiding (elim2) +open import Cubical.HITs.S1 + +open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Group.Join +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.WhiteheadProducts.Generalised.Join.Base renaming (·wh to ·wh') + +open Iso + +-- Whitehead product (main definition) +[_∣_]-pre : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + → (S₊∙ (suc n) →∙ X) + → (S₊∙ (suc m) →∙ X) + → join∙ (S₊∙ n) (S₊∙ m) →∙ X +[_∣_]-pre {X = X} {n = n} {m = m} f g = + joinPinch∙ (S₊∙ n) (S₊∙ m) X + (λ a b → (Ω→ g .fst (σS b) ∙ Ω→ f .fst (σS a))) + +[_∣_]-pre' : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + → (S₊∙ (suc n) →∙ X) + → (S₊∙ (suc m) →∙ X) + → join∙ (S₊∙ n) (S₊∙ m) →∙ X +[_∣_]-pre' {X = X} {n = n} {m = m} f g = + ·wh' (S₊∙ n) (S₊∙ m) (f ∘∙ (inv (IsoSucSphereSusp n) , IsoSucSphereSusp∙ n)) + (g ∘∙ (inv (IsoSucSphereSusp m) , IsoSucSphereSusp∙ m)) + +[∣]-pre≡[∣]-pre' : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + (f : S₊∙ (suc n) →∙ X) (g : S₊∙ (suc m) →∙ X) + → [ f ∣ g ]-pre ≡ [ f ∣ g ]-pre' +[∣]-pre≡[∣]-pre' {n = zero} {m = zero} f g = + cong (joinPinch∙ (S₊∙ zero) (S₊∙ zero) _) + (funExt λ x → funExt λ y → cong₂ _∙_ + (cong₃ _∙∙_∙∙_ (cong sym (lUnit (g .snd))) + (cong (congS (g .fst)) + (sym (cong-∙ (inv (IsoSucSphereSusp zero)) + (merid y) (sym (merid true)) + ∙ sym (rUnit _)))) + (lUnit (g .snd))) + (cong₃ _∙∙_∙∙_ (cong sym (lUnit (f .snd))) + (cong (congS (f .fst)) + (sym (cong-∙ (inv (IsoSucSphereSusp zero)) + (merid x) (sym (merid true)) + ∙ sym (rUnit _)))) + (lUnit (f .snd)))) +[∣]-pre≡[∣]-pre' {n = zero} {m = suc m} f g = + cong (joinPinch∙ (S₊∙ zero) (S₊∙ (suc m)) _) + (funExt λ x → funExt λ y → cong₂ _∙_ + (cong₃ _∙∙_∙∙_ (cong sym (lUnit (g .snd))) + refl + (lUnit (g .snd))) + (cong₃ _∙∙_∙∙_ (cong sym (lUnit (f .snd))) + (cong (congS (f .fst)) + (sym (cong-∙ (inv (IsoSucSphereSusp zero)) + (merid x) (sym (merid true)) + ∙ sym (rUnit _)))) + (lUnit (f .snd)))) + +[∣]-pre≡[∣]-pre' {n = suc n} {m = zero} f g = + cong (joinPinch∙ (S₊∙ (suc n)) (S₊∙ zero) _) + (funExt λ x → funExt λ y → cong₂ _∙_ + (cong₃ _∙∙_∙∙_ (cong sym (lUnit (g .snd))) + (cong (congS (g .fst)) + (sym (cong-∙ (inv (IsoSucSphereSusp zero)) + (merid y) (sym (merid true)) + ∙ sym (rUnit _)))) + (lUnit (g .snd))) + (cong₃ _∙∙_∙∙_ (cong sym (lUnit (f .snd))) + refl + (lUnit (f .snd)))) +[∣]-pre≡[∣]-pre' {n = suc n} {m = suc m} f g = + cong₂ (·wh' (S₊∙ (suc n)) (S₊∙ (suc m))) (sym (∘∙-idˡ f)) (sym (∘∙-idˡ g)) + +[_∣_] : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + → (S₊∙ (suc n) →∙ X) + → (S₊∙ (suc m) →∙ X) + → S₊∙ (suc (n + m)) →∙ X +[_∣_] {n = n} {m = m} f g = + [ f ∣ g ]-pre ∘∙ (sphere→Join n m , IsoSphereJoin⁻Pres∙ n m) + +-- Homotopy group versions +[_∣_]π' : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + → π' (suc n) X → π' (suc m) X + → π' (suc (n + m)) X +[_∣_]π' = elim2 (λ _ _ → squash₂) λ f g → ∣ [ f ∣ g ] ∣₂ + +-- Join version +[_∣_]π* : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + → π' (suc n) X → π' (suc m) X + → π* n m X +[_∣_]π* = elim2 (λ _ _ → squash₂) λ f g → ∣ [ f ∣ g ]-pre ∣₂ + +-- The two versions agree +whπ*≡whπ' : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + → (f : π' (suc n) X) (g : π' (suc m) X) + → [ f ∣ g ]π' ≡ Iso.fun (Iso-π*-π' n m ) [ f ∣ g ]π* +whπ*≡whπ' {n = n} {m} = elim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f g → TR.rec (squash₂ _ _) + (λ p → cong ∣_∣₂ (ΣPathP (refl , (sym (rUnit _) + ∙ cong (cong (fst [ f ∣ g ]-pre)) p + ∙ rUnit _)))) + (help n m) + where + help : (n m : ℕ) + → hLevelTrunc 1 (IsoSphereJoin⁻Pres∙ n m + ≡ snd (≃∙map (invEquiv∙ (joinSphereEquiv∙ n m)))) + help zero zero = ∣ invEq (congEquiv F) (λ _ → pos 0) ∣ₕ + where + F : (Path (join (S₊ 0) (S₊ 0)) (inr true) (inl true)) ≃ ℤ + F = compEquiv ( (compPathlEquiv (push true true))) + (compEquiv ((Ω≃∙ (isoToEquiv (IsoSphereJoin 0 0) , refl) .fst)) + (isoToEquiv ΩS¹Isoℤ)) + help zero (suc m) = + isConnectedPath 1 + (isConnectedPath 2 + (isOfHLevelRetractFromIso 0 + (mapCompIso (IsoSphereJoin zero (suc m))) + (isConnectedSubtr 3 m + (subst (λ p → isConnected p (S₊ (2 + m))) + (+-comm 3 m) + (sphereConnected (2 + m))))) _ _) _ _ .fst + help (suc n) m = + isConnectedPath 1 + (isConnectedPath 2 + (isOfHLevelRetractFromIso 0 + (mapCompIso (IsoSphereJoin (suc n) m)) + (isConnectedSubtr 3 (n + m) + (subst (λ p → isConnected p (S₊ (2 + (n + m)))) + (+-comm 3 (n + m)) (sphereConnected (2 + (n + m)))))) _ _) _ _ .fst + +whπ'≡whπ* : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + → (f : π' (suc n) X) (g : π' (suc m) X) + → Iso.inv (Iso-π*-π' n m) [ f ∣ g ]π' ≡ [ f ∣ g ]π* +whπ'≡whπ* {n = n} {m} f g = + cong (inv (Iso-π*-π' n m)) (whπ*≡whπ' f g) + ∙ Iso.leftInv (Iso-π*-π' n m) _ + +-- Whitehead product (as a composition) +joinTo⋁ : ∀ {ℓ ℓ'} {A : Pointed ℓ} {B : Pointed ℓ'} + → join (typ A) (typ B) + → (Susp (typ A) , north) ⋁ (Susp (typ B) , north) +joinTo⋁ (inl x) = inr north +joinTo⋁ (inr x) = inl north +joinTo⋁ {A = A} {B = B} (push a b i) = + ((λ i → inr (σ B b i)) + ∙∙ sym (push tt) + ∙∙ λ i → inl (σ A a i)) i + +[_∣_]comp : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + → (S₊∙ (suc n) →∙ X) + → (S₊∙ (suc m) →∙ X) + → S₊∙ (suc (n + m)) →∙ X +[_∣_]comp {n = n} {m = m} f g = + (((f ∘∙ (inv (IsoSucSphereSusp n) , IsoSucSphereSusp∙ n)) + ∨→ (g ∘∙ (inv (IsoSucSphereSusp m) , IsoSucSphereSusp∙ m)) + , cong (fst f) (IsoSucSphereSusp∙ n) ∙ snd f) + ∘∙ ((joinTo⋁ {A = S₊∙ n} {B = S₊∙ m} , sym (push tt)))) + ∘∙ (inv (IsoSphereJoin n m) , IsoSphereJoin⁻Pres∙ n m) diff --git a/Cubical/Homotopy/WhiteheadProducts/Generalised/Join/Base.agda b/Cubical/Homotopy/WhiteheadProducts/Generalised/Join/Base.agda new file mode 100644 index 0000000000..e4d9bd45e9 --- /dev/null +++ b/Cubical/Homotopy/WhiteheadProducts/Generalised/Join/Base.agda @@ -0,0 +1,145 @@ +{-# OPTIONS --safe #-} +module Cubical.Homotopy.WhiteheadProducts.Generalised.Join.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.Data.Sigma + +open import Cubical.HITs.Susp renaming (toSusp to σ) +open import Cubical.HITs.Pushout +open import Cubical.HITs.Join +open import Cubical.HITs.Join.CoHSpace + +open import Cubical.Homotopy.Loopspace + +-- Generalised Whitehead products +module _ {ℓ ℓ' ℓ''} (A : Pointed ℓ) + (B : Pointed ℓ') {C : Pointed ℓ''} + (f : Susp∙ (typ A) →∙ C) (g : Susp∙ (typ B) →∙ C) where + + ·wh : join∙ A B →∙ C + ·wh = joinPinch∙ A B C + (λ a b → (Ω→ g .fst (σ B b) ∙ Ω→ f .fst (σ A a))) + +--- Other elementary properties and lemmas --- + -- The generalised Whitehead product vanishes under suspension + isConst-Susp·w : suspFun∙ (·wh .fst) ≡ const∙ _ _ + isConst-Susp·w = Susp·w∙ + ∙ cong suspFun∙ (cong fst isConst-const*) + ∙ ΣPathP ((suspFunConst (pt C)) , refl) + where + const* : join∙ A B →∙ C + fst const* (inl x) = pt C + fst const* (inr x) = pt C + fst const* (push a b i) = + (Ω→ f .fst (σ A a) ∙ Ω→ g .fst (σ B b)) i + snd const* = refl + + isConst-const* : const* ≡ const∙ _ _ + fst (isConst-const* i) (inl x) = Ω→ f .fst (σ A x) i + fst (isConst-const* i) (inr x) = Ω→ g .fst (σ B x) (~ i) + fst (isConst-const* i) (push a b j) = + compPath-filler'' (Ω→ f .fst (σ A a)) (Ω→ g .fst (σ B b)) (~ i) j + snd (isConst-const* i) j = + (cong (Ω→ f .fst) (rCancel (merid (pt A))) ∙ Ω→ f .snd) j i + + Susp·w : suspFun (fst ·wh) ≡ suspFun (fst const*) + Susp·w i north = north + Susp·w i south = south + Susp·w i (merid (inl x) j) = merid (pt C) j + Susp·w i (merid (inr x) j) = merid (pt C) j + Susp·w i (merid (push a b k) j) = + hcomp (λ r → + λ {(i = i0) → fill₁ k (~ r) j + ; (i = i1) → fill₂ k (~ r) j + ; (j = i0) → north + ; (j = i1) → merid (pt C) r + ; (k = i0) → compPath-filler (merid (snd C)) (merid (pt C) ⁻¹) (~ r) j + ; (k = i1) → compPath-filler (merid (snd C)) (merid (pt C) ⁻¹) (~ r) j}) + (hcomp (λ r → + λ {(i = i0) → doubleCompPath-filler + (sym (rCancel (merid (pt C)))) + (λ k → fill₁ k i1) + (rCancel (merid (pt C))) (~ r) k j + ; (i = i1) → doubleCompPath-filler + (sym (rCancel (merid (pt C)))) + (λ k → fill₂ k i1) + (rCancel (merid (pt C))) (~ r) k j + ; (j = i0) → north + ; (j = i1) → north + ; (k = i0) → rCancel (merid (pt C)) (~ r) j + ; (k = i1) → rCancel (merid (pt C)) (~ r) j}) + (main i k j)) + where + F : Ω C .fst → (Ω^ 2) (Susp∙ (fst C)) .fst + F p = sym (rCancel (merid (pt C))) + ∙∙ cong (σ C) p + ∙∙ rCancel (merid (pt C)) + + F-hom : (p q : _) → F (p ∙ q) ≡ F p ∙ F q + F-hom p q = + cong (sym (rCancel (merid (pt C))) + ∙∙_∙∙ rCancel (merid (pt C))) + (cong-∙ (σ C) p q) + ∙ doubleCompPath≡compPath (sym (rCancel (merid (pt C)))) _ _ + ∙ cong (sym (rCancel (merid (pt C))) ∙_) + (sym (assoc _ _ _)) + ∙ assoc _ _ _ + ∙ (λ i → (sym (rCancel (merid (pt C))) + ∙ compPath-filler (cong (σ C) p) (rCancel (merid (pt C))) i) + ∙ compPath-filler' (sym (rCancel (merid (pt C)))) + (cong (σ C) q ∙ rCancel (merid (pt C))) i) + ∙ cong₂ _∙_ (sym (doubleCompPath≡compPath _ _ _)) + (sym (doubleCompPath≡compPath _ _ _)) + + main : F ((Ω→ g .fst (σ B b) ∙ Ω→ f .fst (σ A a))) + ≡ F ((Ω→ f .fst (σ A a) ∙ Ω→ g .fst (σ B b))) + main = F-hom (Ω→ g .fst (σ B b)) (Ω→ f .fst (σ A a)) + ∙ EH 0 _ _ + ∙ sym (F-hom (Ω→ f .fst (σ A a)) (Ω→ g .fst (σ B b))) + + fill₁ : (k : I) → _ + fill₁ k = compPath-filler + (merid ((Ω→ g .fst (σ B b) + ∙ Ω→ f .fst (σ A a)) k)) + (merid (pt C) ⁻¹) + + fill₂ : (k : I) → _ + fill₂ k = compPath-filler + (merid ((Ω→ f .fst (σ A a) + ∙ Ω→ g .fst (σ B b)) k)) + (merid (pt C) ⁻¹) + + Susp·w∙ : suspFun∙ (·wh .fst) ≡ suspFun∙ (fst const*) + Susp·w∙ = ΣPathP (Susp·w , refl) + +-- Action on the standard loop in Ω (join A B) +cong·wh-ℓ* : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + (f : Susp∙ (typ A) →∙ C) (g : Susp∙ (typ B) →∙ C) + (a : _) (b : _) + → (cong (fst (·wh A B f g)) (ℓ* A B a b)) + ≡ ((Ω→ f .fst (σ A a) ⁻¹) + ∙∙ (Ω→ g .fst (σ B b) ∙ Ω→ f .fst (σ A a)) + ∙∙ (Ω→ g .fst (σ B b) ⁻¹)) +cong·wh-ℓ* {A = A} {B = B} f g a b = + cong-∙ (fst (·wh A B f g)) + (push (pt A) (pt B)) + (push a (pt B) ⁻¹ ∙∙ push a b ∙∙ (push (pt A) b ⁻¹)) + ∙ cong₂ _∙_ (cong₂ _∙_ gp fp + ∙ sym (rUnit refl)) + (cong-∙∙ (fst (·wh A B f g)) + (push a (pt B) ⁻¹) + (push a b) + (push (pt A) b ⁻¹) + ∙ cong₃ _∙∙_∙∙_ + (cong _⁻¹ (cong₂ _∙_ gp refl ∙ sym (lUnit _))) + refl + (cong _⁻¹ (cong₂ _∙_ refl fp ∙ sym (rUnit _)))) + ∙ sym (lUnit _) + where + fp = cong (Ω→ f .fst) (rCancel (merid (pt A))) ∙ Ω→ f .snd + gp = cong (Ω→ g .fst) (rCancel (merid (pt B))) ∙ Ω→ g .snd diff --git a/Cubical/Homotopy/WhiteheadProducts/Generalised/Join/Properties.agda b/Cubical/Homotopy/WhiteheadProducts/Generalised/Join/Properties.agda new file mode 100644 index 0000000000..1d5a43b90d --- /dev/null +++ b/Cubical/Homotopy/WhiteheadProducts/Generalised/Join/Properties.agda @@ -0,0 +1,1054 @@ +{-# OPTIONS --safe --lossy-unification #-} +module Cubical.Homotopy.WhiteheadProducts.Generalised.Join.Properties where + +{- This file contians proof of bilinearity/commutativity/jacobi and +some other basic properties of generalised Whitehead products.-} + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Path +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.Data.Unit +open import Cubical.Data.Sigma +open import Cubical.Data.Nat + +open import Cubical.HITs.Susp renaming (toSusp to σ) +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.Sn.Multiplication +open import Cubical.HITs.Join +open import Cubical.HITs.Join.CoHSpace +open import Cubical.HITs.Wedge +open import Cubical.HITs.SmashProduct + +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.WhiteheadProducts.Generalised.Join.Base + +open Iso +open 3x3-span + +-- Left bilinearity of generalised Whitehead product +module _ {ℓ ℓ' ℓ''} (A : Pointed ℓ) + (B : Pointed ℓ') {C : Pointed ℓ''} + (f g : Susp∙ (Susp (typ A)) →∙ C) + (h : Susp∙ (typ B) →∙ C) where + private + σΣA = σ (Susp∙ (typ A)) + Ω→f∙ = cong (Ω→ f .fst) (rCancel (merid north)) ∙ Ω→ f .snd + Ω→g∙ = cong (Ω→ g .fst) (rCancel (merid north)) ∙ Ω→ g .snd + + WhiteheadProdBilinₗ : ·wh (Susp∙ (typ A)) B (·Susp (Susp∙ (typ A)) f g) h + ≡ _+*_ {A = Susp∙ (typ A)} {B = B} + (·wh (Susp∙ (typ A)) B f h) + (·wh (Susp∙ (typ A)) B g h) + fst (WhiteheadProdBilinₗ i) (inl x) = + (Ω→ g .fst (σΣA x) ∙ Ω→ f .fst (σΣA x)) i + fst (WhiteheadProdBilinₗ i) (inr x) = Ω→ h .fst (σ B x) (~ i) + fst (WhiteheadProdBilinₗ i) (push a b j) = main i j + where + x = Ω→ f .fst (σΣA a) + y = Ω→ g .fst (σΣA a) + z = Ω→ h .fst (σ B b) + + fun1 fun2 fun3 : Susp∙ (typ A) →∙ Ω C + fst fun1 a = z ⁻¹ ∙ (Ω→ g .fst (σΣA a)) ⁻¹ ∙ z + snd fun1 = cong (z ⁻¹ ∙_) (cong (_∙ z) (cong sym Ω→g∙) + ∙ sym (lUnit z)) + ∙ lCancel z + fst fun2 a = ((z ∙ Ω→ f .fst (σΣA a)) ∙ Ω→ g .fst (σΣA a)) ∙ z ⁻¹ + snd fun2 = cong (_∙ z ⁻¹) + (cong₂ _∙_ (cong (z ∙_) + Ω→f∙ ∙ sym (rUnit _)) + Ω→g∙ + ∙ sym (rUnit _)) + ∙ rCancel z + fun3 = Ω→ g ∘∙ toSuspPointed (Susp∙ (typ A)) + + lem : y ⁻¹ ∙ ((z ∙ x) ∙ y) ∙ z ⁻¹ + ≡ (z ∙ x ∙ z ⁻¹) ∙ ((y ⁻¹ ∙ z) ∙ y) ∙ z ⁻¹ + lem = (sym (funExt⁻ (cong fst + (Susp·→Ωcomm A fun2 ((sym , refl) ∘∙ fun3))) a) + ∙ sym (assoc _ _ _) + ∙ sym (assoc _ _ _)) + ∙ (λ j → (z ∙ x) ∙ y ∙ z ⁻¹ ∙ (rUnit (y ⁻¹) j)) + ∙ rUnit _ + ∙ (λ i → ((z ∙ x) ∙ y ∙ z ⁻¹ ∙ y ⁻¹ + ∙ λ j → z (j ∧ i)) ∙ λ j → z (~ j ∧ i)) + ∙ cong (_∙ z ⁻¹) + (cong ((z ∙ x) ∙_) + (sym (funExt⁻ (cong fst (Susp·→Ωcomm A fun1 fun3)) a) + ∙ sym (assoc _ _ _)) + ∙ assoc _ _ _ + ∙ cong₂ _∙_ (sym (assoc z x (z ⁻¹))) + refl) + ∙ sym (assoc _ _ _) + + main : Square (cong (·wh (Susp∙ (typ A)) B + (·Susp (Susp∙ (typ A)) f g) h .fst) (push a b)) + (cong (((·wh (Susp∙ (typ A)) B f h) + +* (·wh (Susp∙ (typ A)) B g h)) .fst) (push a b)) + (y ∙ x) (z ⁻¹) + main = cong₂ _∙_ refl (sym (rUnit _) ∙ ·Suspσ f g a) + ∙ assoc z x y + ◁ doubleCompPath-filler (sym (y ∙ x)) ((z ∙ x) ∙ y) (z ⁻¹) + ▷ (doubleCompPath≡compPath _ _ _ + ∙ cong₂ _∙_ (symDistr y x) refl + ∙ sym (assoc (x ⁻¹) (y ⁻¹) _) + ∙ cong (x ⁻¹ ∙_) lem + ∙ assoc (x ⁻¹) (z ∙ x ∙ z ⁻¹) (((y ⁻¹ ∙ z) ∙ y) ∙ z ⁻¹) + ∙ cong₂ _∙_ (cong (x ⁻¹ ∙_) (assoc z x (z ⁻¹))) + (cong (_∙ z ⁻¹) (sym (assoc (y ⁻¹) z y)) + ∙ sym (assoc (y ⁻¹) (z ∙ y) _)) + ∙ sym (cong₂ _∙_ (sym (rUnit _) ∙ cong·wh-ℓ* f h a b + ∙ doubleCompPath≡compPath _ _ _) + (sym (rUnit _) ∙ cong·wh-ℓ* g h a b + ∙ doubleCompPath≡compPath _ _ _))) + + snd (WhiteheadProdBilinₗ i) j = lem j i + where + lem : Ω→ g .fst (σΣA north) ∙ Ω→ f .fst (σΣA north) + ≡ refl + lem = cong₂ _∙_ Ω→g∙ Ω→f∙ + ∙ sym (rUnit refl) + +-- Right bilinearity of generalised whitehead product +module _ {ℓ ℓ' ℓ''} (A : Pointed ℓ) + (B : Pointed ℓ') {C : Pointed ℓ''} + (f : Susp∙ (typ A) →∙ C) + (g h : Susp∙ (Susp (typ B)) →∙ C) where + private + σΣB = σ (Susp∙ (typ B)) + Ω→f∙ = cong (Ω→ f .fst) (rCancel (merid (pt A))) ∙ Ω→ f .snd + Ω→g∙ = cong (Ω→ g .fst) (rCancel (merid north)) ∙ Ω→ g .snd + Ω→h∙ = cong (Ω→ h .fst) (rCancel (merid north)) ∙ Ω→ h .snd + + WhiteheadProdBilinᵣ : ·wh A (Susp∙ (typ B)) f (·Susp (Susp∙ (typ B)) g h) + ≡ _+*_ {A = A} {B = Susp∙ (typ B)} + (·wh A (Susp∙ (typ B)) f g) + (·wh A (Susp∙ (typ B)) f h) + fst (WhiteheadProdBilinᵣ i) (inl x) = Ω→ f .fst (σ A x) i + fst (WhiteheadProdBilinᵣ i) (inr x) = + (Ω→ h .fst (σΣB x) ∙ Ω→ g .fst (σΣB x)) (~ i) + fst (WhiteheadProdBilinᵣ i) (push a b j) = main i j + where + x = Ω→ f .fst (σ A a) + y = Ω→ g .fst (σΣB b) + z = Ω→ h .fst (σΣB b) + + fun1 fun2 : Susp∙ (typ B) →∙ Ω C + fst fun1 b = x ∙ Ω→ g .fst (σΣB b) ⁻¹ ∙ x ⁻¹ + snd fun1 = cong₂ _∙_ refl + (cong₂ _∙_ (cong sym Ω→g∙) refl + ∙ sym (lUnit (Ω→ f .fst (sym (σ A a))))) + ∙ rCancel x + fun2 = Ω→ h ∘∙ (σΣB , rCancel (merid north)) + + main : Square (cong (·wh A (Susp∙ (typ B)) f + (·Susp (Susp∙ (typ B)) g h) .fst) (push a b)) + (cong (_+*_ {A = A} {B = Susp∙ (typ B)} + (·wh A (Susp∙ (typ B)) f g) + (·wh A (Susp∙ (typ B)) f h) .fst) (push a b)) + x ((z ∙ y) ⁻¹) + main = cong₂ _∙_ (·Suspσ' g h b) refl + ∙ sym (assoc y z x) + ∙ (λ _ → y ∙ (z ∙ x)) + ◁ doubleCompPath-filler (x ⁻¹) (y ∙ (z ∙ x)) ((z ∙ y) ⁻¹) + ▷ (doubleCompPath≡compPath _ _ _ + ∙ cong (x ⁻¹ ∙_) (sym (assoc _ _ _) + ∙ cong₂ _∙_ refl (cong₂ _∙_ refl (symDistr z y) + ∙ assoc _ _ _ ∙ cong₂ _∙_ (sym (assoc z x _)) refl)) + ∙ assoc _ _ _ + ∙ cong₂ _∙_ refl + ((cong₂ _∙_ refl (lUnit _ + ∙ cong (_∙ z ⁻¹) (sym (lCancel x)) + ∙ sym (assoc _ _ _)) + ∙ (assoc _ _ _)) + ∙ cong₂ _∙_ ((sym (assoc _ _ _) ∙ cong₂ _∙_ refl (sym (assoc _ _ _))) + ∙ sym (funExt⁻ (cong fst (Susp·→Ωcomm B fun1 fun2)) b)) + refl + ∙ sym (assoc _ _ _)) + ∙ assoc _ _ _ + ∙ cong₂ _∙_ (cong ((x ⁻¹ ∙ y) ∙_) (assoc _ _ _) + ∙ assoc _ _ _) refl + ∙ sym (assoc _ _ _) + ∙ sym (cong₂ _∙_ + (sym (rUnit _) + ∙ cong·wh-ℓ* f g a b + ∙ doubleCompPath≡compPath _ _ _ + ∙ (assoc _ _ _) + ∙ cong (_∙ y ⁻¹) (assoc _ _ _) + ∙ sym (assoc _ _ _)) + (sym (rUnit _) + ∙ cong·wh-ℓ* f h a b + ∙ doubleCompPath≡compPath _ _ _ + ∙ cong (x ⁻¹ ∙_) (sym (assoc _ _ _) + ∙ refl)))) + snd (WhiteheadProdBilinᵣ i) j = Ω→f∙ j i + +-- more flexible versions +WhiteheadProdBilinₗ' : ∀ {ℓ ℓ' ℓ''} (A A' : Pointed ℓ) + (e : A ≃∙ Susp∙ (typ A')) + (B : Pointed ℓ') {C : Pointed ℓ''} + (f g : Susp∙ (typ A) →∙ C) + (h : Susp∙ (typ B) →∙ C) + → ·wh A B (·Susp A f g) h + ≡ _+*_ {A = A} {B = B} (·wh A B f h) (·wh A B g h) +WhiteheadProdBilinₗ' {ℓ' = ℓ'} {ℓ''} A A' = + Equiv∙J (λ A e → (B : Pointed ℓ') {C : Pointed ℓ''} + (f g : Susp∙ (typ A) →∙ C) + (h : Susp∙ (typ B) →∙ C) + → ·wh A B (·Susp A f g) h + ≡ _+*_ {A = A} {B = B} (·wh A B f h) (·wh A B g h)) + (WhiteheadProdBilinₗ A') + +WhiteheadProdBilinᵣ' : ∀ {ℓ ℓ' ℓ''} (A : Pointed ℓ) + (B B' : Pointed ℓ') (e : B ≃∙ Susp∙ (typ B')) {C : Pointed ℓ''} + (f : Susp∙ (typ A) →∙ C) + (g h : Susp∙ (typ B) →∙ C) + → ·wh A B f (·Susp B g h) + ≡ _+*_ {A = A} {B = B} (·wh A B f g) (·wh A B f h) +WhiteheadProdBilinᵣ' {ℓ'' = ℓ''} A B B' = + Equiv∙J (λ B e → {C : Pointed ℓ''} + (f : Susp∙ (typ A) →∙ C) + (g h : Susp∙ (typ B) →∙ C) + → ·wh A B f (·Susp B g h) + ≡ _+*_ {A = A} {B = B} (·wh A B f g) (·wh A B f h)) + (WhiteheadProdBilinᵣ A B') + +-- Distributivity for suspension versions +WhiteheadProdIdL : ∀ {ℓ ℓ' ℓ''} (A : Pointed ℓ) + (B : Pointed ℓ') {C : Pointed ℓ''} + (f : Susp∙ (typ B) →∙ C) + → ·wh A B (const∙ _ _) f ≡ const∙ _ _ +fst (WhiteheadProdIdL A B {C = C} f i) (inl x) = pt C +fst (WhiteheadProdIdL A B f i) (inr x) = Ω→ f .fst (σ B x) (~ i) +fst (WhiteheadProdIdL A B f i) (push a b j) = lem i j + where + lem : Square (Ω→ f .fst (σ B b) ∙ refl ∙ refl) refl + refl (sym (Ω→ f .fst (σ B b))) + lem = (cong₂ _∙_ refl (sym (rUnit refl)) ∙ sym (rUnit _)) + ◁ λ i j → (Ω→ f .fst (σ B b) (~ i ∧ j)) +snd (WhiteheadProdIdL A B f i) = refl + +WhiteheadProdIdR : ∀ {ℓ ℓ' ℓ''} (A : Pointed ℓ) + (B : Pointed ℓ') {C : Pointed ℓ''} + (f : Susp∙ (typ A) →∙ C) + → ·wh A B f (const∙ _ _) ≡ const∙ _ _ +fst (WhiteheadProdIdR A B f i) (inl x) = Ω→ f .fst (σ A x) i +fst (WhiteheadProdIdR A B {C = C} f i) (inr x) = pt C +fst (WhiteheadProdIdR A B f i) (push a b j) = lem i j + where + lem : Square ((refl ∙ refl) ∙ Ω→ f .fst (σ A a)) refl + (Ω→ f .fst (σ A a)) refl + lem = (cong₂ _∙_ (sym (rUnit refl)) refl ∙ sym (lUnit _)) + ◁ λ i j → (Ω→ f .fst (σ A a) (i ∨ j)) +snd (WhiteheadProdIdR A B f i) j = + (cong (Ω→ f .fst) (rCancel (merid (pt A))) ∙ Ω→ f .snd) j i + +-- inversion distributes over the generalised Whitehead product +-*DistrWhitehead : ∀ {ℓ ℓ' ℓ''} (A : Pointed ℓ) + (B : Pointed ℓ') {C : Pointed ℓ''} + (f : Susp∙ (Susp (typ A)) →∙ C) + (g : Susp∙ (typ B) →∙ C) + → -* (·wh (Susp∙ (typ A)) B f g) + ≡ ·wh (Susp∙ (typ A)) B (-Susp (_ , north) f) g +-*DistrWhitehead A B f g = sym (+*IdL _) + ∙∙ cong (_+* (-* lhs)) (sym -*DistrWhiteheadLem) + ∙∙ (sym (+*Assoc _ _ _) + ∙ cong (rhs +*_) (+*InvR lhs) + ∙ +*IdR rhs) + where + lhs = ·wh (Susp∙ (typ A)) B f g + rhs = ·wh (Susp∙ (typ A)) B (-Susp (_ , north) f) g + + -*DistrWhiteheadLem : rhs +* lhs ≡ const∙ _ _ + -*DistrWhiteheadLem = + sym (WhiteheadProdBilinₗ A B _ f g) + ∙ cong₂ (·wh (Susp∙ (typ A)) B) (·SuspInvL (_ , north) f) refl -- + ∙ WhiteheadProdIdL _ _ g -- + +-- Inversion is compatible with the equivalence A * B ≃ B * A +-*Swap : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} + {C : Pointed ℓ''} (f : join∙ A B →∙ C) + → -* (f ∘∙ join-commFun∙) ≡ ((-* f) ∘∙ join-commFun∙) +fst (-*Swap {C = C} f i) (inl x) = pt C +fst (-*Swap {C = C} f i) (inr x) = pt C +fst (-*Swap {A = A} {B = B} f i) (push b a j) = main i j + where + main : (Ω→ (f ∘∙ join-commFun∙) .fst (ℓ* B A b a)) ⁻¹ + ≡ (Ω→ f .fst (ℓ* A B a b)) + main = cong sym (Ω→∘ f join-commFun∙ (ℓ* B A b a) + ∙ cong (Ω→ f .fst) + (cong₃ _∙∙_∙∙_ refl + (cong-∙ (fst join-commFun∙) _ _ + ∙ cong₂ _∙_ refl + (cong-∙∙ (fst join-commFun∙) _ _ _)) + refl + ∙ doubleCompPath≡compPath _ _ _ + ∙ assoc _ _ _ + ∙ cong (_∙ push (pt (fst A , snd A)) (pt (fst B , snd B)) ⁻¹) + (assoc _ _ _ + ∙ cong₂ _∙_ (rCancel _) refl + ∙ sym (lUnit _)))) + ∙ cong (Ω→ f .fst) (symDistr _ _) +snd (-*Swap {A = A} {B = B} f i) = + (sym (rUnit _) + ∙ cong (Ω→ f .fst) (ℓ*IdL A B (pt B)) ∙ Ω→ f .snd) (~ i) + +{- +`Anti-commutativity' of generalised whitehead products: + [f ∶ g] +(Susp A) * (Susp B) ------------------> C + | | + | | + | | + flip | | id + | | + v v +(Susp B) * (Susp A) ------------------> C + [g ∶ f] +-} +module _ {ℓ ℓ' ℓ''} (A : Pointed ℓ) + (B : Pointed ℓ') {C : Pointed ℓ''} + (f : Susp∙ (Susp (typ A)) →∙ C) + (g : Susp∙ (Susp (typ B)) →∙ C) + where + private + σΣA = σ (Susp∙ (typ A)) + σΣB = σ (Susp∙ (typ B)) + + Ω→f∙ = cong (Ω→ f .fst) (rCancel (merid north)) ∙ Ω→ f .snd + Ω→g∙ = cong (Ω→ g .fst) (rCancel (merid north)) ∙ Ω→ g .snd + Ω→-g∙ = + cong (Ω→ (-Susp (Susp∙ (typ B)) g) .fst) (rCancel (merid north)) + ∙ Ω→ (-Susp (Susp∙ (typ B)) g) .snd + + wh' : join∙ (Susp∙ (typ A)) (Susp∙ (typ B)) →∙ C + wh' = ·wh (Susp∙ (typ B)) (Susp∙ (typ A)) (-Susp _ g) f + ∘∙ (Iso.fun join-comm , push north north ⁻¹) + + -- equivalent statement (easier to prove) + anticomm : -* (·wh (Susp∙ (typ A)) (Susp∙ (typ B)) f g) + ≡ (·wh (Susp∙ (typ B)) (Susp∙ (typ A)) + (-Susp (Susp∙ (typ B)) g) f + ∘∙ join-commFun∙) + fst (anticomm i) (inl x) = Ω→ f .fst (σΣA x) i + fst (anticomm i) (inr x) = Ω→ g .fst (σΣB x) i + fst (anticomm i) (push a b i₁) = l i i₁ + where + x = Ω→ f .fst (σΣA a) + y = Ω→ g .fst (σΣB b) + + fun1 fun2 : Susp∙ (typ A) →∙ Ω C + fst fun1 a = y ∙ (Ω→ f .fst (σΣA a) ⁻¹) ∙ y ⁻¹ + snd fun1 = + cong (y ∙_) + (cong (_∙ y ⁻¹) + (cong sym ((Ω→ f ∘∙ toSuspPointed _) .snd)) + ∙ sym (lUnit _) ) + ∙ rCancel y + fun2 = Ω→ f ∘∙ toSuspPointed _ + + l : Square (cong (fst (-* ((·wh (Susp∙ (typ A)) + (Susp∙ (typ B)) f g)))) + (push a b)) + (cong (fst wh') (push a b)) + x y + l = sym (rUnit _) + ∙ cong sym (cong·wh-ℓ* f g a b) + ∙ cong₃ _∙∙_∙∙_ refl (symDistr _ _) refl + ∙ doubleCompPath≡compPath _ _ _ + ∙ assoc _ _ _ + ∙ (λ i → fst (Susp·→Ωcomm _ fun1 fun2 i) a) + ∙ cong (x ∙_) (assoc _ _ _) + ∙ sym (doubleCompPath≡compPath _ _ _) + ◁ symP (doubleCompPath-filler x (y ∙ x ⁻¹) (y ⁻¹)) + ▷ cong (_∙ x ⁻¹) (cong sym (sym + (cong-∙ (fst (-Susp (Susp (typ B) , north) g)) + (merid b) + (sym (merid north)) + ∙ cong₂ _∙_ refl Ω→g∙ ∙ sym (rUnit _)) ∙ rUnit _)) + ∙ compPath≡compPath' _ _ + snd (anticomm i) j = lem i j + where + lem : Square refl (snd wh') (Ω→ f .fst (σΣA north)) refl + lem = flipSquare Ω→f∙ + ▷ (cong sym (rUnit refl ∙ cong₂ _∙_ (sym Ω→f∙) (sym (Ω→-g∙))) + ∙ rUnit _) + + WhiteheadProdComm : ·wh (Susp∙ (typ A)) (Susp∙ (typ B)) f g + ≡ (·wh (Susp∙ (typ B)) (Susp∙ (typ A)) g f ∘∙ join-commFun∙) + WhiteheadProdComm = + preWhiteheadProdComm + ∙ cong₂ _∘∙_ (-*DistrWhitehead _ _ (-Susp (Susp∙ (typ B)) g) f + ∙ cong₂ (·wh (Susp∙ (typ B)) (Susp∙ (typ A))) + (-Susp² (Susp∙ (typ B)) g) + refl) + refl + where + preWhiteheadProdComm : ·wh (Susp∙ (typ A)) (Susp∙ (typ B)) f g + ≡ (-* (·wh (Susp∙ (typ B)) (Susp∙ (typ A)) + (-Susp (Susp∙ (typ B)) g) f) + ∘∙ join-commFun∙) + preWhiteheadProdComm = sym (-*² _) ∙ cong -* anticomm ∙ -*Swap _ + +WhiteheadProdComm' : ∀ {ℓ ℓ' ℓ''} {C : Pointed ℓ''} + (A A' : Pointed ℓ) (eA : A ≃∙ Susp∙ (typ A')) (B B' : Pointed ℓ') + + (eB : B ≃∙ Susp∙ (typ B')) + (f : _ →∙ C) (g : _ →∙ C) + → ·wh A B f g ≡ (·wh B A g f ∘∙ join-commFun∙) +WhiteheadProdComm' {C = C} A A' = + Equiv∙J (λ A _ → (B B' : Pointed _) + + (eB : B ≃∙ Susp∙ (typ B')) + (f : _ →∙ C) (g : _ →∙ C) + → ·wh A B f g ≡ (·wh B A g f ∘∙ join-commFun∙)) + λ B B' → Equiv∙J (λ B _ → (f : _ →∙ C) (g : _ →∙ C) + → ·wh (Susp∙ (typ A')) B f g + ≡ (·wh B (Susp∙ (typ A')) g f ∘∙ join-commFun∙)) + (WhiteheadProdComm _ _) + + -- (right derivator) version of the Jacobi identity. This + -- corresponds to the statement [f,[g,h]] = [[f,g],h] + [g,[f,h]] + +-- We need some 'correction functoins' to make the theorem well-typed +module _ {ℓ ℓ' ℓ'' : Level} (A : Pointed ℓ) (B : Pointed ℓ') (C : Pointed ℓ'') where + Jcorrection₁ : join∙ A (B ⋀∙ C) →∙ join∙ (A ⋀∙ B) C + fst Jcorrection₁ = + ≃∙map (invEquiv∙ (permute⋀Join≃∙ A B C)) .fst + snd Jcorrection₁ = sym (push (inl tt) (pt C)) + + Jcorrection₁⁻ : join∙ (A ⋀∙ B) C →∙ join∙ A (B ⋀∙ C) + Jcorrection₁⁻ = ≃∙map (permute⋀Join≃∙ A B C) + +Jcorrection₂ : ∀ {ℓ ℓ' ℓ'' : Level} (A : Pointed ℓ) (B : Pointed ℓ') (C : Pointed ℓ'') + → join∙ A (B ⋀∙ C) →∙ join∙ B (A ⋀∙ C) +Jcorrection₂ A B C = Jcorrection₁⁻ B A C + ∘∙ ((join→ ⋀comm→ (idfun _) , refl) + ∘∙ Jcorrection₁ A B C) + +module _ {ℓ ℓ' ℓ'' ℓ'''} (A : Pointed ℓ) + (B : Pointed ℓ') (C : Pointed ℓ'') {D : Pointed ℓ'''} + (f : Susp∙ (Susp (typ A)) →∙ D) + (g : Susp∙ (Susp (typ B)) →∙ D) + (h : Susp∙ (Susp (typ C)) →∙ D) + where + -- To state the theorem, we make some abbreviations + private + σΣA = σ (Susp∙ (typ A)) + σΣB = σ (Susp∙ (typ B)) + σΣC = σ (Susp∙ (typ C)) + + whAB = ·wh (Susp∙ (typ A)) (Susp∙ (typ B)) {D} + + whAC = ·wh (Susp∙ (typ A)) (Susp∙ (typ C)) {D} + whBC = ·wh (Susp∙ (typ B)) (Susp∙ (typ C)) {D} + + whA-BC = ·wh (Susp∙ (typ A)) ((Susp∙ (typ B)) ⋀∙ (Susp∙ (typ C))) {D} + whAB-C = ·wh ((Susp∙ (typ A)) ⋀∙ (Susp∙ (typ B))) (Susp∙ (typ C)) {D} + + whB-AC = ·wh (Susp∙ (typ B)) ((Susp∙ (typ A)) ⋀∙ (Susp∙ (typ C))) {D} + + + whB-CA = ·wh (Susp∙ (typ B)) ((Susp∙ (typ C)) ⋀∙ (Susp∙ (typ A))) {D} + + ΣB*ΣC→Σ[B⋀C] = Join→SuspSmash∙ (Susp∙ (typ B)) (Susp∙ (typ C)) + Σ[B⋀C]→ΣB*ΣC = SuspSmash→Join∙ (Susp∙ (typ B)) (Susp∙ (typ C)) + + Σ[A⋀B]→ΣA*ΣB = SuspSmash→Join∙ (Susp∙ (typ A)) (Susp∙ (typ B)) + Σ[A⋀C]→ΣA*ΣC = SuspSmash→Join∙ (Susp∙ (typ A)) (Susp∙ (typ C)) + + whB∧C = ·wh (Susp∙ (typ B)) (Susp∙ (typ C)) {D} + + Ω→f∙ = cong (Ω→ f .fst) (rCancel (merid north)) ∙ Ω→ f .snd + Ω→g∙ = cong (Ω→ g .fst) (rCancel (merid north)) ∙ Ω→ g .snd + Ω→h∙ = cong (Ω→ h .fst) (rCancel (merid north)) ∙ Ω→ h .snd + + -- We need some 'correction functions' to make the theorem well-typed + correction₁ = Jcorrection₁ (Susp∙ (typ A)) (Susp∙ (typ B)) (Susp∙ (typ C)) + + correction₁⁻ = Jcorrection₁⁻ (Susp∙ (typ A)) (Susp∙ (typ B)) (Susp∙ (typ C)) + + correction₂ = Jcorrection₂ (Susp∙ (typ A)) (Susp∙ (typ B)) (Susp∙ (typ C)) + + -- Main result + JacobiR : whA-BC f (whBC g h ∘∙ Σ[B⋀C]→ΣB*ΣC) + ≡ ((whAB-C (whAB f g ∘∙ Σ[A⋀B]→ΣA*ΣB) h ∘∙ correction₁) + +* (whB-AC g (whAC f h ∘∙ Σ[A⋀C]→ΣA*ΣC) ∘∙ correction₂)) + JacobiR = + + ΣPathP ((funExt (λ { (inl x) → lp x + ; (inr x) → rp x + ; (push a b i) j → main a b j i + })) + , flipSquare (Iso.inv ΩSuspAdjointIso f .snd)) + where + L = whA-BC f (whBC g h ∘∙ Σ[B⋀C]→ΣB*ΣC) + R = ((whAB-C (whAB f g ∘∙ Σ[A⋀B]→ΣA*ΣB) h ∘∙ correction₁) + +* (whB-AC g (whAC f h ∘∙ Σ[A⋀C]→ΣA*ΣC) ∘∙ correction₂)) + + -- The identites on point constructors. + lp : Susp (typ A) → Ω D .fst + lp = Iso.inv ΩSuspAdjointIso f .fst + + rp : (Susp∙ (typ B) ⋀ Susp∙ (typ C)) → Ω D .fst + rp = rp' ∘ ⋀→Smash + where + rpl : ∀ {ℓ} {A : Type ℓ} {x : A} (p q : x ≡ x) + → refl ≡ q + → (p ∙ q ⁻¹) ∙ p ⁻¹ ∙ q ≡ refl + rpl p = J> cong₂ _∙_ (sym (rUnit p)) (sym (rUnit _)) ∙ rCancel p + + rpr : ∀ {ℓ} {A : Type ℓ} {x : A} (q : x ≡ x) (p : x ≡ x) + → refl ≡ p + → (p ∙ q ⁻¹) ∙ p ⁻¹ ∙ q ≡ refl + rpr q = J> cong₂ _∙_ (sym (lUnit _)) (sym (lUnit _)) ∙ lCancel q + + rp' : Smash (Susp∙ (typ B)) (Susp∙ (typ C)) → Ω D .fst + rp' basel = refl + rp' baser = refl + rp' (proj b c) = ((Ω→ g .fst (σΣB b) ∙ (Ω→ h .fst (σΣC c)) ⁻¹) + ∙ ((Ω→ g .fst (σΣB b)) ⁻¹ ∙ Ω→ h .fst (σΣC c))) ⁻¹ + rp' (gluel b i) = + sym (rpl (Ω→ g .fst (σΣB b)) (Ω→ h .fst (σΣC north)) + (sym (Iso.inv ΩSuspAdjointIso h .snd)) i) + rp' (gluer c i) = + sym (rpr (Ω→ h .fst (σΣC c)) (Ω→ g .fst (σΣB north)) + (sym (Iso.inv ΩSuspAdjointIso g .snd)) i) + + apL apR : (a : Susp (typ A)) + → Susp∙ (typ B) ⋀ Susp∙ (typ C) → Ω D .fst + apL a x = lp a ⁻¹ ∙∙ cong (fst L) (push a x) ∙∙ rp x + apR a x = cong (fst R) (push a x) + + -- Some lemmas simplying the pointed functions involved + lem1 : whBC g h ∘∙ Σ[B⋀C]→ΣB*ΣC ≡ (fst (whBC g h ∘∙ Σ[B⋀C]→ΣB*ΣC) , refl) + lem1 = ΣPathP (refl , sym (rUnit _) + ∙ cong₃ _∙∙_∙∙_ + (cong sym (Iso.inv ΩSuspAdjointIso g .snd)) + (cong sym (Iso.inv ΩSuspAdjointIso h .snd)) + refl + ∙ sym (rUnit refl)) + + lem2 : whAB-C (whAB f g ∘∙ Σ[A⋀B]→ΣA*ΣB) h ∘∙ correction₁ + ≡ ((whAB-C (whAB f g ∘∙ Σ[A⋀B]→ΣA*ΣB) h ∘∙ correction₁) .fst , refl) + lem2 = ΣPathP (refl + , sym (rUnit _) + ∙ cong₃ _∙∙_∙∙_ refl + (cong sym (Iso.inv ΩSuspAdjointIso h .snd)) + refl + ∙ sym (compPath≡compPath' _ _) + ∙ cong sym (sym (rUnit _)) + ∙ cong₃ _∙∙_∙∙_ refl + (cong (cong (fst (whAB f g) ∘ fst Σ[A⋀B]→ΣA*ΣB)) + (cong sym (rCancel (merid (inl tt))))) + refl + ∙ ∙∙lCancel _) + + lem3 : whB-AC g (whAC f h ∘∙ Σ[A⋀C]→ΣA*ΣC) ∘∙ correction₂ + ≡ ((whB-AC g (whAC f h ∘∙ Σ[A⋀C]→ΣA*ΣC) ∘∙ correction₂) .fst , refl) + lem3 = ΣPathP (refl + , cong₂ _∙_ (cong (cong (fst (whB-AC g (whAC f h ∘∙ Σ[A⋀C]→ΣA*ΣC)))) lem) + refl + ∙ sym (rUnit refl)) + where + lem : snd correction₂ ≡ refl + lem = cong₂ _∙_ (cong (cong (fst (Jcorrection₁⁻ (Susp∙ (typ B)) (Susp∙ (typ A)) (Susp∙ (typ C))))) + (sym (rUnit _))) refl + ∙ rCancel _ + + + -- some more abbreviations + l : Susp (typ A) → Ω D .fst + m : Susp (typ B) → Ω D .fst + r : Susp (typ C) → Ω D .fst + l a = Ω→ f .fst (σΣA a) + m b = Ω→ g .fst (σΣB b) + r c = Ω→ h .fst (σΣC c) + + -- Goal: relate 'cong (fst L) (push a (inr (b , c))))' to 'cong + -- (fst R) (push a (inr (b , c))))' Unfolding definitions this + -- gives rise to a word problem. We could hope to automate some + -- parts of the proof in the future... + module _ (a : Susp (typ A)) (b : Susp (typ B)) (c : Susp (typ C)) where + leftId : (cong (fst L) (push a (inr (b , c)))) + ≡ ((m b ⁻¹) ∙∙ r c ∙ m b ∙∙ (r c ⁻¹)) ∙ l a + leftId = + cong₂ _∙_ ((λ j i → Ω→ (lem1 j) .fst + (σ (Susp∙ (typ B) ⋀∙ Susp∙ (typ C)) + (inr (b , c))) i) + ∙ sym (rUnit _) + ∙ cong (cong (fst (whBC g h))) + (cong-∙ (fst Σ[B⋀C]→ΣB*ΣC) + (merid (inr (b , c))) (sym (merid (inl tt)))) + ∙ cong-∙ (fst (whBC g h)) _ _ + ∙ cong₂ _∙_ (cong-∙∙ (fst (whBC g h)) _ _ _ + ∙ cong₃ _∙∙_∙∙_ + (cong sym (cong₂ _∙_ (Iso.inv ΩSuspAdjointIso h .snd) refl + ∙ sym (lUnit _))) + (λ i → Ω→ h .fst (σΣC c) ∙ Ω→ g .fst (σΣB b)) + (cong sym (cong₂ _∙_ refl (Iso.inv ΩSuspAdjointIso g .snd) + ∙ sym (rUnit _))) + ∙ refl) + (cong₂ _∙_ (Iso.inv ΩSuspAdjointIso h .snd) + (Iso.inv ΩSuspAdjointIso g .snd) ∙ sym (rUnit refl)) + ∙ sym (rUnit ((m b ⁻¹) ∙∙ r c ∙ m b ∙∙ (r c ⁻¹)))) + refl + + -- more abbreviations + ℓA-BC = ℓ* (Susp∙ (typ A)) ((Susp∙ (typ B)) ⋀∙ (Susp∙ (typ C))) + ℓAB-C = ℓ* ((Susp∙ (typ A)) ⋀∙ (Susp∙ (typ B))) (Susp∙ (typ C)) + ℓB-AC = ℓ* (Susp∙ (typ B)) ((Susp∙ (typ A)) ⋀∙ (Susp∙ (typ C))) + + correction₁-ℓ : cong (fst (correction₁)) (ℓA-BC a (inr (b , c))) + ≡ (push (inl tt) north) ⁻¹ + ∙∙ ℓAB-C (inr (a , b)) c + ∙∙ push (inl tt) north + correction₁-ℓ = cong-∙ (fst (correction₁)) _ _ + ∙ cong (sym (push (inl tt) north) ∙_) + (cong-∙∙ (fst (correction₁)) _ _ _ + ∙ doubleCompPath≡compPath _ _ _ ∙ refl) + ∙ assoc _ _ _ + ∙ cong₂ _∙_ (rCancel _) refl + ∙ sym (lUnit _) + ∙ cong₂ _∙_ (lUnit _ + ∙ (λ i → (λ j → push (inl tt) north (~ j ∨ ~ i)) + ∙ compPath-filler' (push (inl tt) north) + (push (inr (a , b)) north ⁻¹ + ∙∙ push (inr (a , b)) c + ∙∙ push (inl tt) c ⁻¹) i)) + ((λ i → push (inl tt) c + ∙∙ push (push (inr b) (~ i)) c ⁻¹ + ∙∙ push (push (inr b) (~ i)) north) + ∙ doubleCompPath≡compPath _ _ _ + ∙ assoc _ _ _ + ∙ cong₂ _∙_ (rCancel _) refl + ∙ sym (lUnit _)) + ∙ sym (assoc _ _ _) + ∙ sym (doubleCompPath≡compPath _ _ _) + + + correction₂-ℓ : cong (fst correction₂) (ℓA-BC a (inr (b , c))) + ≡ ℓB-AC b (inr (a , c)) + correction₂-ℓ = + cong-∙ (fst correction₂) _ _ + ∙ cong₂ _∙_ refl + (cong-∙∙ (fst correction₂) _ _ _ + ∙ (cong₃ _∙∙_∙∙_ (λ _ → push north (inl tt) ⁻¹) + (help a b c) + (cong sym (help north b c + ∙ cong₂ _∙_ (cong (ℓB-AC b) (sym (push (inr c))) + ∙ ℓ*IdR _ _ b) refl + ∙ sym (lUnit _))) + ∙ doubleCompPath≡compPath _ _ _ + ∙ cong (push north (inl tt) ⁻¹ ∙_) + (sym (assoc _ _ _) + ∙ (cong (ℓB-AC b (inr (a , c)) ∙_) (rCancel _) + ∙ sym (rUnit _))) + ∙ assoc _ _ _ + ∙ cong₂ _∙_ (rCancel _) refl + ∙ sym (lUnit _))) + + where + help : (a : _) (b : _) (c : _) + → cong (fst correction₂) (push a (inr (b , c))) + ≡ ℓB-AC b (inr (a , c)) ∙ push north (inl tt) + help a b c = + cong-∙∙ ((Jcorrection₁⁻ + (Susp∙ (typ B)) (Susp∙ (typ A)) (Susp∙ (typ C))) .fst + ∘ join→ ⋀comm→ (idfun (Susp (typ C)))) _ _ _ + ∙ cong₃ _∙∙_∙∙_ ((λ i → push north (push (inl a) (~ i)) + ∙∙ push b (push (inl a) (~ i)) ⁻¹ + ∙∙ push b (inl tt)) + ∙ doubleCompPath≡compPath _ _ _ + ∙ cong₂ _∙_ refl (lCancel (push b (inl tt))) + ∙ sym (rUnit _)) + refl (λ _ → push north (inl tt)) + ∙ doubleCompPath≡compPath _ _ _ ∙ assoc _ _ _ + + -- more abbreviations... + x = l a + -x = x ⁻¹ + y = m b + -y = y ⁻¹ + z = r c + -z = z ⁻¹ + + t₁ = (y ∙ -x ∙ -y ∙ x) ∙ z ∙ -x ∙ y ∙ x ∙ -y + t₂ = -y ∙ -x ∙ z ∙ x ∙ -z ∙ y + t₃ = z ∙ -x ∙ -z ∙ x + + t₃' = -x ∙ -z ∙ x + t₄ = z ∙ x ∙ -z + + fA : Susp∙ (typ A) →∙ Ω D + fst fA a = ((-y ∙ z) ∙ Ω→ f .fst (σΣA a) + ∙ -z ∙ sym (Ω→ f .fst (σΣA a))) ∙ y + snd fA = cong (λ x → ((-y ∙ z) ∙ x ∙ -z ∙ x ⁻¹) ∙ y) + (Iso.inv ΩSuspAdjointIso f .snd) + ∙ cong₂ _∙_ (cong₂ _∙_ refl (sym (lUnit _) ∙ sym (rUnit _)) + ∙ sym (assoc -y z -z) + ∙ cong (-y ∙_) (rCancel z) ∙ sym (rUnit -y)) + refl + ∙ lCancel y + + f-xyx f-xyx' f-zyz : Susp∙ (typ B) →∙ Ω D + fst f-xyx b = -x ∙ Ω→ g .fst (σΣB b) ⁻¹ ∙ x + snd f-xyx = cong₂ _∙_ refl + (cong (_∙ x) (cong sym (Iso.inv ΩSuspAdjointIso g .snd)) + ∙ sym (lUnit x)) ∙ lCancel x + fst f-xyx' b = -x ∙ Ω→ g .fst (σΣB b) ∙ x + snd f-xyx' = + cong₂ _∙_ refl + (cong (_∙ x) (Iso.inv ΩSuspAdjointIso g .snd) + ∙ sym (lUnit x)) ∙ lCancel x + fst f-zyz b = -z ∙ Ω→ g .fst (σΣB b) ⁻¹ ∙ z + snd f-zyz = cong₂ _∙_ refl + (cong (_∙ z) (cong sym (Iso.inv ΩSuspAdjointIso g .snd)) + ∙ sym (lUnit z)) + ∙ lCancel z + + f₁ f₂ fz f₃ f-yazay : Susp∙ (typ C) →∙ Ω D + fst f₁ z = (y ∙ -x ∙ -y ∙ x) ∙ Ω→ h .fst (σΣC z) ∙ -x ∙ y ∙ x ∙ -y + snd f₁ = + cong₂ _∙_ + (assoc _ _ _ ∙ assoc _ _ _) + ((cong₂ _∙_ (Iso.inv ΩSuspAdjointIso h .snd) refl + ∙ sym (lUnit _)) + ∙ sym (symDistr _ _ + ∙ cong₂ _∙_ refl + (symDistr _ _ ∙ cong₂ _∙_ refl (symDistr _ _)))) + ∙ rCancel (((y ∙ -x) ∙ -y) ∙ x) + fst f₂ z = -y ∙ -x ∙ Ω→ h .fst (σΣC z) ∙ x ∙ Ω→ h .fst (σΣC z) ⁻¹ ∙ y + snd f₂ = + cong₂ _∙_ refl + (cong₂ _∙_ refl + (cong₂ _∙_ (Iso.inv ΩSuspAdjointIso h .snd) + (cong₂ _∙_ refl + (cong₂ _∙_ (cong sym (Iso.inv ΩSuspAdjointIso h .snd)) + refl + ∙ sym (lUnit _))) + ∙ sym (lUnit (x ∙ y)))) + ∙ cong₂ _∙_ refl (assoc -x x y + ∙ cong₂ _∙_ (lCancel x) refl + ∙ sym (lUnit y)) + ∙ lCancel y + fz = (sym , refl) ∘∙ Iso.inv ΩSuspAdjointIso h + fst f₃ c = -x ∙ sym (Ω→ h .fst (σΣC c)) ∙ x + snd f₃ = + cong (-x ∙_) + (cong (_∙ x) (cong sym (Iso.inv ΩSuspAdjointIso h .snd)) + ∙ sym (lUnit x)) + ∙ lCancel x + fst f-yazay c = (-y ∙ x) ∙ Ω→ h .fst (σΣC c) ∙ -x ∙ y + snd f-yazay = + cong₂ _∙_ (sym (symDistr -x y)) + (cong₂ _∙_ (Iso.inv ΩSuspAdjointIso h .snd) refl + ∙ sym (lUnit (-x ∙ y))) + ∙ lCancel (-x ∙ y) + + + f₄ fa : Susp∙ (typ A) →∙ Ω D + fst f₄ a = z ∙ Ω→ f .fst (σΣA a) ∙ -z + snd f₄ = cong (z ∙_) (cong (_∙ -z) (Iso.inv ΩSuspAdjointIso f .snd) + ∙ sym (lUnit _)) ∙ rCancel z + fa = (sym , refl) ∘∙ Iso.inv ΩSuspAdjointIso f + + rightId₁ : cong (fst R) (push a (inr (b , c))) + ≡ (t₂ ∙ t₁) ∙ -z ∙ t₃ + rightId₁ = + cong₂ _∙_ (λ i → Ω→ (lem2 i) .fst (ℓA-BC a (inr (b , c)))) + (λ i → Ω→ (lem3 i) .fst (ℓA-BC a (inr (b , c)))) + ∙ cong₂ _∙_ + (sym (rUnit _) + ∙ cong (cong (fst (whAB-C (whAB f g ∘∙ Σ[A⋀B]→ΣA*ΣB) h))) + correction₁-ℓ + ∙ cong-∙∙ (fst (whAB-C (whAB f g ∘∙ Σ[A⋀B]→ΣA*ΣB) h)) _ _ _ + ∙ cong (λ x → x ⁻¹ + ∙∙ cong (fst (whAB-C (whAB f g ∘∙ Σ[A⋀B]→ΣA*ΣB) h)) + (ℓAB-C (inr (a , b)) c) + ∙∙ x) (cong₂ _∙_ (Iso.inv ΩSuspAdjointIso h .snd) + (Iso.inv ΩSuspAdjointIso + (whAB f g ∘∙ Σ[A⋀B]→ΣA*ΣB) .snd) + ∙ sym (rUnit refl)) + ∙ sym (rUnit _) ) + (sym (rUnit _) + ∙ cong (cong (fst (whB-AC g (whAC f h ∘∙ Σ[A⋀C]→ΣA*ΣC)))) + (correction₂-ℓ)) + ∙ cong₂ _∙_ ((λ i → cong·wh-ℓ* {A = _ , inl tt} {B = _ , north} + (lem4 i) h (inr (a , b)) c i) + ∙ cong₃ _∙∙_∙∙_ (sym (rUnit _) + ∙ cong sym (fgid' A B f g a b) + ∙ cong₃ _∙∙_∙∙_ refl + (symDistr (m b) (l a)) refl) + (cong₂ _∙_ (λ _ → r c) + (sym (rUnit _) ∙ (fgid' A B f g a b))) + (λ _ → r c ⁻¹)) + ((λ i → cong·wh-ℓ* {A = _ , north} {B = _ , inl tt} + g (lem5 i) b (inr (a , c)) i) + ∙ cong₃ _∙∙_∙∙_ (λ _ → m b ⁻¹) + (cong₂ _∙_ (sym (rUnit _) ∙ fgid' A C f h a c) + (λ _ → m b)) + (sym (rUnit _) + ∙ cong sym (fgid' A C f h a c) + ∙ cong₃ _∙∙_∙∙_ refl (symDistr (r c) (l a)) refl)) + ∙ cong₂ _∙_ (cong₃ _∙∙_∙∙_ (doubleCompPath≡compPath _ _ _ + ∙ cong₂ _∙_ refl (sym (assoc _ _ _))) + (cong (r c ∙_) (doubleCompPath≡compPath _ _ _ + ∙ cong₂ _∙_ refl (sym (assoc _ _ _)))) + refl + ∙ doubleCompPath≡compPath _ _ _ + ∙ assoc _ _ _ ∙ cong (_∙ -z) λ _ → t₁) + (doubleCompPath≡compPath _ _ _ + ∙ cong₂ _∙_ refl + (cong₂ _∙_ (cong (_∙ y) + (doubleCompPath≡compPath _ _ _) + ∙ sym (assoc _ _ y) + ∙ cong (-x ∙_) (sym (assoc _ _ y) ∙ sym (assoc _ _ _))) + ((doubleCompPath≡compPath _ _ _) + ∙ (cong (z ∙_) (sym (assoc -x -z x))))) + ∙ assoc -y _ t₃ ∙ λ _ → t₂ ∙ t₃) + ∙ sym (assoc t₁ -z (t₂ ∙ t₃)) + ∙ cong (t₁ ∙_) (assoc -z t₂ t₃ + ∙ cong (_∙ t₃) (funExt⁻ (cong fst (Susp·→Ωcomm C fz f₂)) c) + ∙ sym (assoc t₂ -z t₃)) + ∙ assoc t₁ t₂ _ + ∙ cong₂ _∙_ (funExt⁻ (cong fst (Susp·→Ωcomm C f₁ f₂)) c) refl + where + fgid' : ∀ {ℓ ℓ'} (A : Pointed ℓ) (B : Pointed ℓ') + (f : Susp∙ (Susp (typ A)) →∙ D) + (g : Susp∙ (Susp (typ B)) →∙ D) + (a : Susp (typ A)) (b : Susp (typ B)) + → cong (fst (·wh (Susp∙ (typ A)) (Susp∙ (typ B)) f g)) + (cong (fst (SuspSmash→Join∙ (Susp∙ (typ A)) (Susp∙ (typ B)))) + (σ (_ , inl tt) (inr (a , b)))) + ≡ (Ω→ f .fst (σ (Susp∙ (typ A)) a) ⁻¹ + ∙∙ Ω→ g .fst (σ (Susp∙ (typ B)) b) + ∙ Ω→ f .fst (σ (Susp∙ (typ A)) a) + ∙∙ (Ω→ g .fst (σ (Susp∙ (typ B)) b) ⁻¹)) + fgid' A B f g a b = + cong-∙ (fst (·wh (Susp∙ (typ A)) (Susp∙ (typ B)) f g) + ∘ (fst (SuspSmash→Join∙ (Susp∙ (typ A)) (Susp∙ (typ B))))) + (merid (inr (a , b))) _ + ∙ cong₂ _∙_ refl + (cong₂ _∙_ (Iso.inv ΩSuspAdjointIso g .snd) + (Iso.inv ΩSuspAdjointIso f .snd) + ∙ sym (rUnit refl)) + ∙ sym (rUnit _) + ∙ cong-∙∙ (fst (·wh (Susp∙ (typ A)) (Susp∙ (typ B)) f g)) _ _ _ + ∙ cong₃ _∙∙_∙∙_ + (cong sym (cong₂ _∙_ (Iso.inv ΩSuspAdjointIso g .snd) refl + ∙ sym (lUnit _))) + refl + (cong sym (cong₂ _∙_ refl (Iso.inv ΩSuspAdjointIso f .snd) + ∙ sym (rUnit _))) + + lem4 : whAB f g ∘∙ Σ[A⋀B]→ΣA*ΣB + ≡ ((whAB f g ∘∙ Σ[A⋀B]→ΣA*ΣB) .fst , refl) + lem4 = + ΣPathP (refl , sym (rUnit _) + ∙ cong sym (cong₂ _∙_ (Iso.inv ΩSuspAdjointIso g .snd) + (Iso.inv ΩSuspAdjointIso f .snd)) + ∙ sym (rUnit _)) + lem5 : whAC f h ∘∙ Σ[A⋀C]→ΣA*ΣC + ≡ ((whAC f h ∘∙ Σ[A⋀C]→ΣA*ΣC) .fst , refl) + lem5 = + ΣPathP (refl , sym (rUnit _) + ∙ cong sym (cong₂ _∙_ (Iso.inv ΩSuspAdjointIso h .snd) + (Iso.inv ΩSuspAdjointIso f .snd)) + ∙ sym (rUnit _)) + + rightId₂ : (t₂ ∙ t₁) ∙ -z ∙ t₃ + ≡ -y ∙ (-x ∙ t₄) ∙ y ∙ t₁ ∙ t₃' + rightId₂ = cong₂ _∙_ (cong₂ _∙_ (cong (-y ∙_) + (cong (-x ∙_) + (assoc _ _ _ ∙ assoc _ _ _ + ∙ cong₂ _∙_ (sym (assoc z x -z)) refl))) + refl) + (assoc -z z _ + ∙ cong₂ _∙_ (lCancel z) refl ∙ sym (lUnit _)) + ∙ sym (assoc _ _ _) + ∙ cong₂ _∙_ (cong (-y ∙_) (assoc _ _ _) + ∙ assoc _ _ _) + refl + ∙ sym (assoc _ _ _) + ∙ sym (assoc _ _ _) + + rightId₃ : -y ∙ (-x ∙ t₄) ∙ y ∙ t₁ ∙ t₃' + ≡ (-y ∙ z) ∙ (x ∙ -z ∙ -x) ∙ y ∙ t₃' ∙ t₁ + rightId₃ = + cong (-y ∙_) + (cong₂ _∙_ (funExt⁻ (cong fst (Susp·→Ωcomm A fa f₄)) a) + (cong (y ∙_) (funExt⁻ (cong fst (Susp·→Ωcomm C f₁ f₃)) c))) + ∙ (assoc _ _ _ + ∙ cong₂ _∙_ (cong (-y ∙_) (sym (assoc _ _ _) + ∙ cong (z ∙_) (sym (assoc _ _ _))) + ∙ assoc _ _ _) refl) + ∙ sym (assoc _ _ _) + + rightId₄ : ((-y ∙ z) ∙ (x ∙ -z ∙ -x) ∙ y ∙ t₃' ∙ t₁) + ∙ (y ∙ -z) ∙ (-y ∙ z) + ≡ (((-y ∙ z) ∙ x ∙ -z ∙ -x) ∙ y) + ∙ (-x ∙ (-y ∙ x) ∙ z ∙ -x ∙ y) ∙ -z ∙ x + rightId₄ = + cong₂ _∙_ (cong (λ t → (-y ∙ z) ∙ (x ∙ -z ∙ -x) ∙ y ∙ t₃' ∙ t) t₁≡) + (sym (assoc y -z _)) + ∙ sym (assoc _ _ _) + ∙ cong₂ _∙_ refl (sym (assoc _ _ _) + ∙ cong₂ _∙_ refl (sym (assoc _ _ _) + ∙ cong₂ _∙_ refl (sym (assoc _ _ _) + ∙ cong₂ _∙_ refl (assoc _ _ _ + ∙ cong₂ _∙_ (sym (assoc _ _ _) + ∙ cong₂ _∙_ refl (sym (assoc _ _ _) + ∙ cong (y ∙_) (sym (assoc _ _ _) + ∙ cong (z ∙_) (sym (assoc _ _ _) + ∙ cong₂ _∙_ refl (lCancel y) + ∙ sym (rUnit _))))) refl)))) + ∙ assoc _ _ _ + ∙ assoc _ y _ + ∙ cong₂ _∙_ refl id2 + where + t₁≡ : t₁ ≡ (-x ∙ -y ∙ x) ∙ y ∙ z ∙ (-x ∙ y ∙ x) ∙ -y + t₁≡ = cong₂ _∙_ + (funExt⁻ (cong fst (Susp·→Ωcomm B + (Iso.inv ΩSuspAdjointIso g) f-xyx)) b) + (cong (z ∙_) (cong (-x ∙_) + (assoc _ _ _) ∙ assoc _ _ _)) + ∙ sym (assoc _ _ _) + + id2 : t₃' ∙ ((-x ∙ -y ∙ x) ∙ y ∙ z ∙ (-x ∙ y ∙ x)) ∙ (-z ∙ -y ∙ z) + ≡ (-x ∙ (-y ∙ x) ∙ z ∙ -x ∙ y) ∙ (-z ∙ x) + id2 = cong (t₃' ∙_) + (cong₂ _∙_ (assoc _ _ _ ∙ assoc _ _ _) refl + ∙ sym (assoc _ _ _) + ∙ cong₂ _∙_ (sym (assoc _ _ _) ∙ sym (assoc _ _ _)) + (funExt⁻ (cong fst (Susp·→Ωcomm B f-xyx' f-zyz)) b) + ∙ cong₂ _∙_ (assoc _ _ _) refl + ∙ assoc _ _ _ + ∙ cong₂ _∙_ (cong₂ _∙_ refl + (assoc _ _ _ + ∙ cong (_∙ z) (sym (symDistr y z))) + ∙ assoc _ _ _ + ∙ cong (_∙ z) (sym (assoc _ _ _) + ∙ cong₂ _∙_ refl (rCancel (y ∙ z)) + ∙ sym (rUnit _))) + refl) + ∙ assoc _ _ _ + ∙ cong₂ _∙_ (assoc _ _ _ + ∙ cong (_∙ z) (assoc _ -x _ + ∙ cong₂ _∙_ (sym (assoc _ _ _) + ∙ cong (-x ∙_) (sym (assoc -z x -x) + ∙ cong (-z ∙_) (rCancel x) + ∙ sym (rUnit -z))) refl)) refl + ∙ assoc _ _ _ + ∙ assoc _ _ _ + ∙ cong (_∙ x) (sym (assoc _ _ _) + ∙ sym (assoc _ _ _) + ∙ sym (assoc _ _ _) + ∙ sym (assoc -x -z _)) + ∙ sym (assoc _ _ _) + ∙ cong (-x ∙_) (cong (_∙ x) + (funExt⁻ (cong fst (Susp·→Ωcomm C fz f-yazay)) c) + ∙ sym (assoc _ _ _)) + ∙ assoc -x _ _ + ∙ cong (_∙ (-z ∙ x)) refl + + rightId₅ : ((((-y ∙ z) ∙ x ∙ -z ∙ -x) ∙ y) ∙ x) + ∙ (-x ∙ (-y ∙ x) ∙ z ∙ -x ∙ y) ∙ -z ∙ x + ≡ (-y ∙∙ z ∙ y ∙∙ -z) ∙ x + rightId₅ = assoc _ _ _ + ∙ assoc _ _ _ + ∙ cong (_∙ x) + (cong (_∙ -z) (sym (assoc _ x _) + ∙ cong₂ _∙_ refl (assoc x -x _ + ∙ cong₂ _∙_ (rCancel x) refl + ∙ sym (lUnit _) + ∙ sym (assoc _ _ _)) + ∙ sym (assoc _ y _) + ∙ cong₂ _∙_ refl (assoc y -y _ ∙ cong₂ _∙_ (rCancel y) refl + ∙ sym (lUnit _)) + ∙ sym (assoc _ _ _) + ∙ cong₂ _∙_ refl (cong₂ _∙_ refl (assoc _ _ _ ∙ assoc _ _ _) + ∙ assoc _ _ _ + ∙ cong (_∙ y) + (cong₂ _∙_ refl (sym (symDistr x (-z ∙ -x) + ∙ cong (_∙ -x) (symDistr -z -x))) + ∙ rCancel (x ∙ -z ∙ -x)) + ∙ sym (lUnit y)) + ∙ sym (assoc _ _ _)) + ∙ sym (assoc _ _ _) + ∙ sym (doubleCompPath≡compPath _ _ _)) + + rightId : x ∙∙ cong (fst R) (push a (inr (b , c))) + ∙∙ ((y ∙ -z) ∙ (-y ∙ z)) + ≡ (-y ∙∙ z ∙ y ∙∙ -z) ∙ x + rightId = cong (x ∙∙_∙∙ (y ∙ -z) ∙ (-y ∙ z)) + (rightId₁ ∙ rightId₂ + ∙ rightId₃) + ∙ doubleCompPath≡compPath _ _ _ + ∙ cong (x ∙_) (rightId₄ ∙ refl) + ∙ assoc _ _ _ + ∙ cong₂ _∙_ (funExt⁻ (cong fst (Susp·→Ωcomm A + (Iso.inv ΩSuspAdjointIso f) fA)) a) + refl + ∙ rightId₅ + + mainId : + Square (cong (fst L) (push a (inr (b , c)))) + (cong (fst R) (push a (inr (b , c)))) + (lp a) (rp (inr (b , c))) + mainId = (leftId ∙ sym rightId) + ◁ symP (doubleCompPath-filler + x + (cong (fst R) (push a (inr (b , c)))) + ((y ∙ -z) ∙ (-y ∙ z))) + + main : (a : Susp (typ A)) + (x : Susp∙ (typ B) ⋀ Susp∙ (typ C)) + → Square (cong (fst L) (push a x)) + (cong (fst R) (push a x)) + (lp a) (rp x) + main a x = + doubleCompPath-filler (lp a ⁻¹) (cong (fst L) (push a x)) (rp x) + ▷ asFuns a x + where + asFuns : (a : Susp (typ A)) + → (x : Susp∙ (typ B) ⋀ Susp∙ (typ C)) + → apL a x ≡ apR a x + asFuns a = funExt⁻ (⋀→Homogeneous≡ (isHomogeneousPath _ _) + λ b c → sym (transport (PathP≡doubleCompPathʳ _ _ _ _) + (symP (mainId a b c)))) + +JacobiR' : + ∀ {ℓ ℓ' ℓ'' ℓ'''} {D : Pointed ℓ'''} + (A A' : Pointed ℓ) (eA : A ≃∙ Susp∙ (typ A')) + (B B' : Pointed ℓ') (eB : B ≃∙ Susp∙ (typ B')) + (C C' : Pointed ℓ'') (eC : C ≃∙ Susp∙ (typ C')) + (f : Susp∙ (typ A) →∙ D) + (g : Susp∙ (typ B) →∙ D) + (h : Susp∙ (typ C) →∙ D) + → ·wh A (B ⋀∙ C) f (·wh B C g h ∘∙ SuspSmash→Join∙ B C) + ≡ (·wh (A ⋀∙ B) C (·wh A B f g ∘∙ SuspSmash→Join∙ A B) h ∘∙ Jcorrection₁ A B C) + +* (·wh B (A ⋀∙ C) g (·wh A C f h ∘∙ SuspSmash→Join∙ A C) ∘∙ Jcorrection₂ A B C) +JacobiR' {D = D} A A' eA B B' eB C C' eC = + transport (λ i → (f : Susp∙ (typ (pA i)) →∙ D) + (g : Susp∙ (typ (pB i)) →∙ D) + (h : Susp∙ (typ (pC i)) →∙ D) + → ·wh (pA i) ((pB i) ⋀∙ (pC i)) f (·wh (pB i) (pC i) g h + ∘∙ SuspSmash→Join∙ (pB i) (pC i)) + ≡ (·wh ((pA i) ⋀∙ (pB i)) (pC i) (·wh (pA i) (pB i) f g + ∘∙ SuspSmash→Join∙ (pA i) (pB i)) h + ∘∙ Jcorrection₁ (pA i) (pB i) (pC i)) + +* (·wh (pB i) ((pA i) ⋀∙ (pC i)) g (·wh (pA i) (pC i) f h + ∘∙ SuspSmash→Join∙ (pA i) (pC i)) + ∘∙ Jcorrection₂ (pA i) (pB i) (pC i))) + (JacobiR A' B' C') + where + pA = ua∙ (fst eA) (snd eA) ⁻¹ + pB = ua∙ (fst eB) (snd eB) ⁻¹ + pC = ua∙ (fst eC) (snd eC) ⁻¹ diff --git a/Cubical/Homotopy/WhiteheadProducts/Generalised/Smash/Base.agda b/Cubical/Homotopy/WhiteheadProducts/Generalised/Smash/Base.agda new file mode 100644 index 0000000000..1f0e667abd --- /dev/null +++ b/Cubical/Homotopy/WhiteheadProducts/Generalised/Smash/Base.agda @@ -0,0 +1,123 @@ +{-# OPTIONS --safe #-} +module Cubical.Homotopy.WhiteheadProducts.Generalised.Smash.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Sigma + +open import Cubical.HITs.Susp renaming (toSusp to σ) +open import Cubical.HITs.Join +open import Cubical.HITs.Join.CoHSpace +open import Cubical.HITs.SmashProduct + +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.WhiteheadProducts.Generalised.Join.Base + +open Iso + +-- Generalised Whitehead products +module _ {ℓ ℓ' ℓ''} (A : Pointed ℓ) (B : Pointed ℓ') {C : Pointed ℓ''} + (f : Susp∙ (typ A) →∙ C) (g : Susp∙ (typ B) →∙ C) where + + -- alternative version using Suspension and smash + private + ·whΣ' : Susp∙ (Smash A B) →∙ C + fst ·whΣ' north = pt C + fst ·whΣ' south = pt C + fst ·whΣ' (merid basel i) = pt C + fst ·whΣ' (merid baser i) = pt C + fst ·whΣ' (merid (proj a b) i) = + (Ω→ f .fst (σ A a) ⁻¹ + ∙ (Ω→ g .fst (σ B b) ∙ Ω→ f .fst (σ A a)) + ∙ Ω→ g .fst (σ B b) ⁻¹) i + fst ·whΣ' (merid (gluel a i₁) i) = + (cong (λ p → (Ω→ f .fst (σ A a) ⁻¹ + ∙ (p ∙ Ω→ f .fst (σ A a)) ∙ p ⁻¹)) + (inv (ΩSuspAdjointIso {A = B}) g .snd) + ∙ cong₂ _∙_ refl (sym (rUnit _) ∙ sym (lUnit _)) + ∙ lCancel _) i₁ i + fst ·whΣ' (merid (gluer b i₁) i) = + (cong (λ p → (p ⁻¹ ∙ (Ω→ g .fst (σ B b) ∙ p) ∙ Ω→ g .fst (σ B b) ⁻¹)) + (inv (ΩSuspAdjointIso {A = A}) f .snd) + ∙ (sym (lUnit _) ∙ cong₂ _∙_ (sym (rUnit _)) refl) + ∙ lCancel _) i₁ i + snd ·whΣ' = refl + + ·whΣ : Susp∙ (A ⋀ B) →∙ C + fst ·whΣ = fst ·whΣ' ∘ suspFun ⋀→Smash + snd ·whΣ = refl + + -- This version agrees with the join version modulo the + -- equivalence Σ(A ⋀ B) ≃ A * B + ·wh≡·whΣ : ·wh A B f g ∘∙ (SuspSmash→Join∙ A B) ≡ ·whΣ + ·wh≡·whΣ = ΣPathP (funExt lem , (sym (rUnit _) + ∙ cong sym (cong₂ _∙_ + (inv ΩSuspAdjointIso g .snd) + (inv ΩSuspAdjointIso f .snd) + ∙ sym (rUnit refl)))) + where + lem : (x : _) → ·wh A B f g .fst (fun (SmashJoinIso {A = A} {B = B}) x) + ≡ ·whΣ .fst x + lem north = refl + lem south = refl + lem (merid a i) j = main j a i + where + main : Path ((A ⋀ B) → Ω C .fst) + (cong (·wh A B f g .fst ∘ SuspSmash→Join) ∘ merid) + (cong (fst ·whΣ') ∘ merid ∘ ⋀→Smash) + main = ⋀→Homogeneous≡ (isHomogeneousPath _ _) + λ x y → cong-∙∙ (·wh A B f g .fst) + (push x (pt B) ⁻¹) (push x y) (push (pt A) y ⁻¹) + ∙ doubleCompPath≡compPath _ _ _ + ∙ cong₂ _∙_ (cong sym + (cong₂ _∙_ (inv ΩSuspAdjointIso g .snd) refl + ∙ sym (lUnit _))) + (cong₂ _∙_ refl + (cong sym + (cong₂ _∙_ refl (inv ΩSuspAdjointIso f .snd) + ∙ sym (rUnit _)))) + + -- Other direction + ·whΣ≡·wh : ·wh A B f g ≡ ·whΣ ∘∙ Join→SuspSmash∙ A B + ·whΣ≡·wh = + sym (∘∙-idʳ _) + ∙ cong (·wh A B f g ∘∙_) + (ΣPathP (funExt (λ x → + sym (rightInv (SmashJoinIso {A = A} {B = B}) x)) + , ((λ i j → push (pt A) (pt B) (i ∧ ~ j)) + ▷ lUnit _))) + ∙∙ sym (∘∙-assoc (·wh A B f g) (SuspSmash→Join∙ A B) (Join→SuspSmash∙ A B)) + ∙∙ cong (_∘∙ Join→SuspSmash∙ A B) ·wh≡·whΣ + + --- PathP version + PathP-·wh-·whΣ : + PathP (λ i → isoToPath (fromSusp≅fromJoin {A = A} {B = B} {C = C}) i) + ·whΣ (·wh A B f g) + PathP-·wh-·whΣ = post∘∙equiv-ua _ _ _ ·wh≡·whΣ + where + post∘∙equiv-ua : ∀ {ℓ ℓ'} {A B : Pointed ℓ} {C : Pointed ℓ'} (e : A ≃∙ B) + (f : A →∙ C) (g : B →∙ C) → (g ∘∙ (fst (fst e) , snd e)) ≡ f + → PathP (λ i → isoToPath (post∘∙equiv {C = C} e) i) f g + post∘∙equiv-ua {B = B} {C = C} = Equiv∙J + (λ A e → (f : A →∙ C) (g : B →∙ C) → (g ∘∙ (fst (fst e) , snd e)) ≡ f + → PathP (λ i → isoToPath (post∘∙equiv {C = C} e) i) f g) + λ e f g → toPathP (cong (λ P → transport P e) + (cong ua post∘∙equivId + ∙ uaIdEquiv) + ∙ transportRefl _ + ∙ sym g + ∙ ∘∙-idˡ f) + where + post∘∙equivId : ∀ {ℓ ℓ'} {A : Pointed ℓ} {C : Pointed ℓ'} + → isoToEquiv (post∘∙equiv {C = C} (idEquiv∙ A)) ≡ idEquiv _ + post∘∙equivId = Σ≡Prop isPropIsEquiv + (funExt λ f → ΣPathP (refl + , (cong₂ _∙_ (cong (cong (fst f)) (sym (rUnit refl))) refl + ∙ sym (lUnit (snd f))))) diff --git a/Cubical/Homotopy/WhiteheadProducts/Generalised/Smash/Properties.agda b/Cubical/Homotopy/WhiteheadProducts/Generalised/Smash/Properties.agda new file mode 100644 index 0000000000..e2b2425f7e --- /dev/null +++ b/Cubical/Homotopy/WhiteheadProducts/Generalised/Smash/Properties.agda @@ -0,0 +1,365 @@ +{-# OPTIONS --safe --lossy-unification #-} + +{- This file contians proof of bilinearity/commutativity/jacobi and +some other basic properties of the smash product construction of +generalised Whitehead products.-} + +module Cubical.Homotopy.WhiteheadProducts.Generalised.Smash.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Path +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.Data.Unit +open import Cubical.Data.Sigma + +open import Cubical.HITs.Susp +open import Cubical.HITs.Pushout +open import Cubical.HITs.Join +open import Cubical.HITs.Join.CoHSpace +open import Cubical.HITs.Wedge +open import Cubical.HITs.SmashProduct +open import Cubical.HITs.SmashProduct.SymmetricMonoidal + +open import Cubical.Homotopy.WhiteheadProducts.Generalised.Smash.Base +open import Cubical.Homotopy.WhiteheadProducts.Generalised.Join.Base +open import Cubical.Homotopy.WhiteheadProducts.Generalised.Join.Properties + +open Iso + +module _ {ℓ ℓ' ℓ''} (A A' : Pointed ℓ) (e : A ≃∙ Susp∙ (typ A')) + (B : Pointed ℓ') {C : Pointed ℓ''} + (f g : Susp∙ (typ A) →∙ C) + (h : Susp∙ (typ B) →∙ C) where + -- ·whΣ version + WhiteheadProdΣBilinₗ : ·whΣ A B (·Susp A f g) h + ≡ ·Susp (A ⋀∙ B) + (·whΣ A B f h) + (·whΣ A B g h) + WhiteheadProdΣBilinₗ = + transport (λ j + → PathP-·wh-·whΣ A B + (·Susp A f g) h (~ j) + ≡ ·Susp-+*-PathP {A = A} {B} {C} (~ j) + (PathP-·wh-·whΣ A B f h (~ j)) + (PathP-·wh-·whΣ A B g h (~ j))) + (WhiteheadProdBilinₗ' A A' e B {C} f g h) + +module _ {ℓ ℓ' ℓ''} (A : Pointed ℓ) + (B B' : Pointed ℓ') (e : B ≃∙ Susp∙ (typ B')) {C : Pointed ℓ''} + (f : Susp∙ (typ A) →∙ C) + (g h : Susp∙ (typ B) →∙ C) where + -- ·whΣ version + WhiteheadProdΣBilinᵣ : + ·whΣ A B f (·Susp B g h) + ≡ ·Susp (A ⋀∙ B) (·whΣ A B f g) (·whΣ A B f h) + WhiteheadProdΣBilinᵣ = + transport (λ j + → PathP-·wh-·whΣ A B f + (·Susp B g h) (~ j) + ≡ ·Susp-+*-PathP {A = A} {B} {C} (~ j) + (PathP-·wh-·whΣ A B f g (~ j)) + (PathP-·wh-·whΣ A B f h (~ j))) + (WhiteheadProdBilinᵣ' A B B' e f g h) + +WhiteheadProdΣIdL : ∀ {ℓ ℓ' ℓ''} (A : Pointed ℓ) + (B : Pointed ℓ') {C : Pointed ℓ''} + (f : Susp∙ (typ B) →∙ C) + → ·whΣ A B (const∙ _ _) f ≡ const∙ _ _ +WhiteheadProdΣIdL A B {C = C} f = + transport (λ i → PathP-·wh-·whΣ A B (const∙ _ _) f (~ i) + ≡ ·Susp-0*-PathP (~ i)) + (WhiteheadProdIdL A B f) + +WhiteheadProdΣIdR : ∀ {ℓ ℓ' ℓ''} (A : Pointed ℓ) + (B : Pointed ℓ') {C : Pointed ℓ''} + (f : Susp∙ (typ A) →∙ C) + → ·whΣ A B f (const∙ _ _) ≡ const∙ _ _ +WhiteheadProdΣIdR A B {C = C} f = + transport (λ i → PathP-·wh-·whΣ A B f (const∙ _ _) (~ i) + ≡ ·Susp-0*-PathP (~ i)) + (WhiteheadProdIdR A B f) + +module _ {ℓ ℓ' ℓ''} (A : Pointed ℓ) + (B : Pointed ℓ') {C : Pointed ℓ''} + (f : Susp∙ (Susp (typ A)) →∙ C) + (g : Susp∙ (Susp (typ B)) →∙ C) + where + WhiteheadProdΣComm : ·whΣ (Susp∙ (typ A)) (Susp∙ (typ B)) f g + ≡ -Susp (Susp∙ (typ A) ⋀∙ Susp∙ (typ B)) + (·whΣ (Susp∙ (typ B)) (Susp∙ (typ A)) g f + ∘∙ suspFun∙ ⋀comm→) + WhiteheadProdΣComm = + sym (·wh≡·whΣ (Susp∙ (typ A)) (Susp∙ (typ B)) f g) + ∙ cong₂ _∘∙_ (WhiteheadProdComm A B f g) refl + ∙ ∘∙-assoc _ _ _ + ∙ cong (·wh (Susp∙ (typ B)) (Susp∙ (typ A)) g f ∘∙_) + (ΣPathP ((funExt (λ { north → push north north + ; south → refl + ; (merid a i) j + → ((compPath-filler' (push north north) + (cong (SuspSmash→Join + ∘ fst ((-Susp (Susp∙ (typ A) + ⋀∙ Susp∙ (typ B))) + (suspFun∙ ⋀comm→))) + (merid a))) + ▷ (funExt⁻ (sym F1≡F2) a)) (~ j) i})) + , ptLem) + ∙ sym (-Susp-∘∙ _ _ _)) + ∙ sym (-Susp-∘∙ _ _ _) + ∙ cong (-Susp (Susp∙ (typ A) ⋀∙ Susp∙ (typ B))) + (sym (∘∙-assoc _ _ _) + ∙ cong₂ _∘∙_ (·wh≡·whΣ _ _ _ _) refl) + where + ptLem : Square {A = join (Susp (typ B)) (Susp (typ A))} + (push north north ∙ sym (push north north)) + (refl ∙ sym (push north north)) + (push north north) refl + ptLem = rCancel (push north north) + ◁ (λ i j → push north north (~ j ∧ i)) + ▷ lUnit (sym (push north north)) + + F1 F2 : Susp∙ (typ A) ⋀ Susp∙ (typ B) + → ((Path (join (fst (Susp∙ (typ B))) (fst (Susp∙ (typ A)))) + (inl north) (inr north))) + F1 a i = join-commFun (SuspSmash→Join (merid a i)) + F2 a = push north north + ∙ cong (SuspSmash→Join + ∘ fst ((-Susp (Susp∙ (typ A) ⋀∙ Susp∙ (typ B))) + (suspFun∙ ⋀comm→))) + (merid a) + + F1≡F2 : F1 ≡ F2 + F1≡F2 = ⋀→Homogeneous≡ (isHomogeneousPath _ _) + λ x y → + cong-∙∙ join-commFun _ _ _ + ∙ lUnit _ + ∙ cong₂ _∙_ (sym (rCancel _)) refl + ∙ sym (assoc _ _ _) + ∙ cong₂ _∙_ refl + (sym (symDistr (cong SuspSmash→Join (merid (inr (y , x)))) + (cong SuspSmash→Join (sym (merid (inl tt))))) + ∙ cong sym (sym (cong-∙ SuspSmash→Join + (merid (inr (y , x))) + (sym (merid (inl tt))))) + ∙ cong (congS SuspSmash→Join) + (cong sym + (sym (cong-∙ (fst (suspFun∙ ⋀comm→)) + (merid (inr (x , y))) + (merid (inl tt) ⁻¹))) + ∙ rUnit _)) + +module _ {ℓ ℓ' ℓ'' ℓ'''} (A : Pointed ℓ) (B : Pointed ℓ') + (C : Pointed ℓ'') {D : Pointed ℓ'''} + (f : Susp∙ (Susp (typ A)) →∙ D) + (g : Susp∙ (Susp (typ B)) →∙ D) + (h : Susp∙ (Susp (typ C)) →∙ D) + where + -- Some abbreviations + private + whAB = ·wh (Susp∙ (typ A)) (Susp∙ (typ B)) {D} + whAC = ·wh (Susp∙ (typ A)) (Susp∙ (typ C)) {D} + + whAB-C = ·wh ((Susp∙ (typ A)) ⋀∙ (Susp∙ (typ B))) (Susp∙ (typ C)) {D} + whB-AC = ·wh (Susp∙ (typ B)) ((Susp∙ (typ A)) ⋀∙ (Susp∙ (typ C))) {D} + + Σ[A⋀B]→ΣA*ΣB = SuspSmash→Join∙ (Susp∙ (typ A)) (Susp∙ (typ B)) + Σ[A⋀C]→ΣA*ΣC = SuspSmash→Join∙ (Susp∙ (typ A)) (Susp∙ (typ C)) + + -- Main result + private + [f[gh]] = ·whΣ (Susp∙ (typ A)) (_ ⋀∙ _) + f + (·whΣ (Susp∙ (typ B)) (Susp∙ (typ C)) + g h) + + [[fg]h] = ·whΣ (_ ⋀∙ _) (Susp∙ (typ C)) + (·whΣ (Susp∙ (typ A)) (Susp∙ (typ B)) + f g) h + + [g[fh]] = ·whΣ (Susp∙ (typ B)) (_ ⋀∙ _) + g + (·whΣ (Susp∙ (typ A)) (Susp∙ (typ C)) + f h) + + JacobiΣR : [f[gh]] + ≡ ·Susp (Susp∙ (typ A) ⋀∙ (Susp∙ (typ B) ⋀∙ Susp∙ (typ C))) + ([[fg]h] ∘∙ suspFun∙ (Iso.fun SmashAssocIso)) + ([g[fh]] ∘∙ suspFun∙ (Iso.inv SmashAssocIso + ∘ (⋀comm→∙ ⋀→ idfun∙ _) + ∘ Iso.fun SmashAssocIso)) + JacobiΣR = sym (·wh≡·whΣ _ _ _ _) + ∙ cong₂ _∘∙_ + ((cong (·wh (Susp∙ (typ A)) (Susp∙ (typ B) ⋀∙ Susp∙ (typ C)) f) + (sym (·wh≡·whΣ _ _ _ _))) + ∙ JacobiR A B C f g h) + refl + ∙ fromSusp≅fromJoin⁻Pres+* _ _ + ∙ cong₂ (·Susp (Susp∙ (typ A) ⋀∙ (Susp∙ (typ B) ⋀∙ Susp∙ (typ C)))) + (∘∙-assoc _ _ _ + ∙ (cong₂ _∘∙_ lem2 lem3 + ∙ sym (∘∙-assoc _ _ _)) + ∙ cong (_∘∙ suspFun∙ (fun SmashAssocIso)) + (·wh≡·whΣ (Susp∙ (typ A) ⋀∙ Susp∙ (typ B)) + (Susp∙ (typ C)) + (·whΣ (Susp∙ (typ A)) + (Susp∙ (typ B)) f g) h)) + (∘∙-assoc _ _ _ + ∙ cong₂ _∘∙_ whB-AC≡whB-AC lem + ∙ sym (∘∙-assoc _ _ _) + ∙ cong (_∘∙ suspFun∙ compFun) + (·wh≡·whΣ (Susp∙ _) (_ ⋀∙ _) g + (·whΣ (Susp∙ (typ A)) (Susp∙ (typ C)) f h))) + where + compFun = inv SmashAssocIso + ∘ (⋀comm→∙ ⋀→ idfun∙ (Susp∙ (typ C))) + ∘ fun SmashAssocIso + + whB-AC≡whB-AC : whB-AC g (whAC f h ∘∙ Σ[A⋀C]→ΣA*ΣC) + ≡ whB-AC g (·whΣ (Susp∙ (typ A)) (Susp∙ (typ C)) f h) + whB-AC≡whB-AC = cong (whB-AC g) (·wh≡·whΣ _ _ _ _) + + Σ[B⋀[A⋀C]]→ΣB*Σ[A⋀B] = fst (SuspSmash→Join∙ (Susp∙ (typ B)) + (Susp∙ (typ A) ⋀∙ Susp∙ (typ C))) + + lem : (correction₂ A B C f g h) + ∘∙ SuspSmash→Join∙ (Susp∙ (typ A)) (Susp∙ (typ B) ⋀∙ Susp∙ (typ C)) + ≡ SuspSmash→Join∙ (Susp∙ (typ B)) (Susp∙ (typ A) ⋀∙ Susp∙ (typ C)) + ∘∙ suspFun∙ compFun + lem = ΣPathP ((funExt (λ x + → cong Σ[B⋀[A⋀C]]→ΣB*Σ[A⋀B] (cong (suspFun (inv SmashAssocIso) ∘ + Join→SuspSmash + ∘ join→ ⋀comm→ (λ c → c) + ∘ SuspSmash→Join + ∘ suspFun (fun SmashAssocIso)) + (SuspSmash→Join→SuspSmash x )) + ∙∙ cong Σ[B⋀[A⋀C]]→ΣB*Σ[A⋀B] + (cong (suspFun (inv SmashAssocIso)) + (SuspFun-Join→SuspSmash≡ ⋀comm→∙ (idfun∙ (Susp∙ (typ C))) + (SuspSmash→Join (suspFun (fun SmashAssocIso) x)))) + ∙∙ (cong Σ[B⋀[A⋀C]]→ΣB*Σ[A⋀B] + (funExt⁻ (sym (suspFunComp (inv SmashAssocIso) + (⋀comm→∙ ⋀→ idfun∙ _))) + (Join→SuspSmash (SuspSmash→Join + (suspFun (fun SmashAssocIso) x)))) + ∙∙ cong Σ[B⋀[A⋀C]]→ΣB*Σ[A⋀B] (cong (suspFun (inv SmashAssocIso + ∘ (⋀comm→∙ ⋀→ idfun∙ (Susp∙ (typ C))))) + (SuspSmash→Join→SuspSmash + (suspFun (fun SmashAssocIso) x))) + ∙∙ cong Σ[B⋀[A⋀C]]→ΣB*Σ[A⋀B] + (funExt⁻ (sym (suspFunComp (inv SmashAssocIso + ∘ (⋀comm→∙ ⋀→ idfun∙ (Susp∙ (typ C)))) + (fun SmashAssocIso))) x)))) + , ((cong₂ _∙_ refl (cong₂ _∙_ + (cong (congS (fst (Jcorrection₁⁻ + (Susp∙ (typ B)) (Susp∙ (typ A)) + (Susp∙ (typ C))))) + (sym (rUnit (sym (push (inl tt) north))))) refl + ∙ rCancel (push north (inl tt))) + ∙ sym (rUnit _)) + ◁ (flipSquare ((cong₃ _∙∙_∙∙_ refl refl (sym (rUnit _)) + ∙ ∙∙lCancel (push north (inl tt))) + ◁ (λ j i → push north (inl tt) (~ j))) + ▷ lUnit (sym (push north (inl tt)))))) + + lem2 : whAB-C (whAB f g ∘∙ Σ[A⋀B]→ΣA*ΣB) h + ≡ ·wh (Susp∙ (typ A) ⋀∙ Susp∙ (typ B)) (Susp∙ (typ C)) + (·whΣ (Susp∙ (typ A)) (Susp∙ (typ B)) f g) h + lem2 = cong₂ (·wh (Susp∙ (typ A) ⋀∙ Susp∙ (typ B)) (Susp∙ (typ C))) + (·wh≡·whΣ _ _ _ _) refl + + lem3 : (correction₁ A B C f g h) + ∘∙ SuspSmash→Join∙ (Susp∙ (typ A)) (Susp∙ (typ B) ⋀∙ Susp∙ (typ C)) + ≡ (SuspSmash→Join∙ (Susp∙ (typ A) ⋀∙ Susp∙ (typ B)) (Susp∙ (typ C)) + ∘∙ suspFun∙ (fun SmashAssocIso)) + lem3 = + ΣPathP ((funExt (λ x + → cong (SuspSmash→Join∙ (Susp∙ (typ A) ⋀∙ Susp∙ (typ B)) + (Susp∙ (typ C)) .fst) + (cong (suspFun (fun SmashAssocIso)) + (SuspSmash→Join→SuspSmash x)))) + , compPathL→PathP (cong₂ _∙_ refl (sym (rUnit _) ∙ rCancel _) + ∙ sym (rUnit _) + ∙ lUnit _) ) + + +private + JacobiΣRTy : ∀ {ℓ ℓ' ℓ'' ℓ'''} + (A : Pointed ℓ) + (B : Pointed ℓ') + (C : Pointed ℓ'') + {D : Pointed ℓ'''} → Type _ + JacobiΣRTy A B C {D} + = (f : Susp∙ (typ A) →∙ D) (g : Susp∙ (typ B) →∙ D) (h : Susp∙ (typ C) →∙ D) + → ·whΣ A (_ ⋀∙ _) f (·whΣ B C g h) + ≡ ·Susp (A ⋀∙ (B ⋀∙ C)) + (·whΣ (_ ⋀∙ _) C (·whΣ A B f g) h + ∘∙ suspFun∙ (Iso.fun SmashAssocIso)) + (·whΣ B (_ ⋀∙ _) g (·whΣ A C f h) + ∘∙ suspFun∙ (Iso.inv SmashAssocIso + ∘ (⋀comm→∙ ⋀→ idfun∙ _) + ∘ Iso.fun SmashAssocIso)) + +JacobiΣR' : ∀ {ℓ ℓ' ℓ'' ℓ'''} + (A A' : Pointed ℓ) (e : A ≃∙ Susp∙ (typ A')) + (B : Pointed ℓ') (B' : Pointed ℓ') (eB : B ≃∙ Susp∙ (typ B')) + (C : Pointed ℓ'') (C' : Pointed ℓ'') (eC : C ≃∙ Susp∙ (typ C')) + {D : Pointed ℓ'''} + (f : Susp∙ (typ A) →∙ D) (g : Susp∙ (typ B) →∙ D) (h : Susp∙ (typ C) →∙ D) + → ·whΣ A (_ ⋀∙ _) f (·whΣ B C g h) + ≡ ·Susp (A ⋀∙ (B ⋀∙ C)) + (·whΣ (_ ⋀∙ _) C (·whΣ A B f g) h + ∘∙ suspFun∙ (Iso.fun SmashAssocIso)) + (·whΣ B (_ ⋀∙ _) g (·whΣ A C f h) + ∘∙ suspFun∙ (Iso.inv SmashAssocIso + ∘ (⋀comm→∙ ⋀→ idfun∙ _) + ∘ Iso.fun SmashAssocIso)) +JacobiΣR' {ℓ = ℓ} {ℓ'} {ℓ''} {ℓ'''} A A' = + Equiv∙J (λ A eA → + (B : Pointed ℓ') (B' : Pointed ℓ') (eB : B ≃∙ Susp∙ (typ B')) + (C : Pointed ℓ'') (C' : Pointed ℓ'') (eC : C ≃∙ Susp∙ (typ C')) + {D : Pointed ℓ'''} → JacobiΣRTy A B C {D}) + λ B B' → Equiv∙J (λ B eB → + (C : Pointed ℓ'') (C' : Pointed ℓ'') (eC : C ≃∙ Susp∙ (typ C')) + {D : Pointed ℓ'''} → JacobiΣRTy (Susp∙ (typ A')) B C {D}) + λ C C' → Equiv∙J (λ C eC → {D : Pointed ℓ'''} + → JacobiΣRTy (Susp∙ (typ A')) (Susp∙ (typ B')) C {D}) + (JacobiΣR A' B' C') + +-- Other properties +whΣ-postCompR : ∀ {ℓ ℓ' ℓ''} {A : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + → {B' : Pointed ℓ'} → (e : B' ≃∙ B) + → (f : Susp∙ (typ A) →∙ C) (g : Susp∙ (typ B) →∙ C) + → (·whΣ A B' f (g ∘∙ suspFun∙ (fst (fst e)))) + ≡ (·whΣ A B f g ∘∙ suspFun∙ (idfun∙ A ⋀→ ≃∙map e)) +whΣ-postCompR {A = A} {B} {C} {B'} = + Equiv∙J (λ B' e → (f : Susp∙ (typ A) →∙ C) (g : Susp∙ (typ B) →∙ C) + → (·whΣ A B' f (g ∘∙ suspFun∙ (fst (fst e)))) + ≡ (·whΣ A B f g ∘∙ suspFun∙ (idfun∙ A ⋀→ ≃∙map e))) + λ f g → cong (·whΣ A B f) + (cong (g ∘∙_) (ΣPathP (suspFunIdFun , refl)) ∙ ∘∙-idˡ g) + ∙ (sym (∘∙-idˡ (·whΣ A B f g))) + ∙ cong₂ _∘∙_ refl + (sym (ΣPathP (suspFunIdFun , refl)) + ∙ cong suspFun∙ (sym + (cong fst ⋀→∙-idfun))) + +whΣ-postCompL : ∀ {ℓ ℓ' ℓ''} {A A' : Pointed ℓ} {B : Pointed ℓ'} {C : Pointed ℓ''} + (e : A' ≃∙ A) + (f : Susp∙ (typ A) →∙ C) (g : Susp∙ (typ B) →∙ C) + → ·whΣ A' B (f ∘∙ suspFun∙ (fst (fst e))) g + ≡ (·whΣ A B f g ∘∙ suspFun∙ (≃∙map e ⋀→ idfun∙ B)) +whΣ-postCompL {A = A} {B} {C} {B'} = + Equiv∙J (λ B e → (f : Susp∙ (typ A) →∙ B') + (g : Susp∙ (typ C) →∙ B') → + ·whΣ B C (f ∘∙ suspFun∙ (fst (fst e))) g ≡ + (·whΣ A C f g ∘∙ suspFun∙ (≃∙map e ⋀→ idfun∙ C))) + λ f g → cong₂ (·whΣ A C) (cong (f ∘∙_) suspFun∙IdFun ∙ ∘∙-idˡ f) refl + ∙ sym (∘∙-idˡ _) + ∙ cong₂ _∘∙_ refl (sym + (cong suspFun∙ (cong fst ⋀→∙-idfun) + ∙ suspFun∙IdFun)) diff --git a/Cubical/Homotopy/WhiteheadProducts/Properties.agda b/Cubical/Homotopy/WhiteheadProducts/Properties.agda new file mode 100644 index 0000000000..1c84da22cf --- /dev/null +++ b/Cubical/Homotopy/WhiteheadProducts/Properties.agda @@ -0,0 +1,1435 @@ +{-# OPTIONS --safe --lossy-unification #-} + +{- This module contains the key properties of Whitehead products: +basic algebraic properties, bilinearity, graded commutativity, Jacobi +and other minor results. It also contains a characterisation Susp A × +Susp B in terms of cofibres of a Whitehead product -} + +module Cubical.Homotopy.WhiteheadProducts.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Path +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Transport + +open import Cubical.Data.Nat +open import Cubical.Data.Sigma +open import Cubical.Data.Unit +open import Cubical.Data.FinData.Base hiding (suc) +open import Cubical.Data.Vec.Base +open import Cubical.Data.Sum as ⊎ +open import Cubical.Data.Bool hiding (_≤_) +open import Cubical.Data.Nat.IsEven +open import Cubical.Data.Empty as ⊥ + +open import Cubical.HITs.Susp renaming (toSusp to σ) +open import Cubical.HITs.Pushout +open import Cubical.HITs.Sn +open import Cubical.HITs.Sn.Multiplication +open import Cubical.HITs.Join +open import Cubical.HITs.Wedge +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.SmashProduct + +open import Cubical.Homotopy.Loopspace +open import Cubical.Homotopy.Connected +open import Cubical.Homotopy.Group.Base +open import Cubical.Homotopy.Group.Properties +open import Cubical.Homotopy.Group.Join +open import Cubical.Homotopy.WhiteheadProducts.Base +open import Cubical.Homotopy.WhiteheadProducts.Generalised.Join.Base +open import Cubical.Homotopy.WhiteheadProducts.Generalised.Join.Properties +open import Cubical.Homotopy.WhiteheadProducts.Generalised.Smash.Base +open import Cubical.Homotopy.WhiteheadProducts.Generalised.Smash.Properties + +open import Cubical.Algebra.Group +open import Cubical.Algebra.AbGroup.Base +open import Cubical.Algebra.AbGroup.Properties + +open import Cubical.Tactics.NatSolver.NatExpression renaming (∣ to ∣') +open import Cubical.Tactics.EvenOddSolver + + +open Iso +open 3x3-span +infixl 7 _·₋₁_ + +-- Equality of two slightly different versions of Whitehead products +[]comp≡[] : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + → (f : (S₊∙ (suc n) →∙ X)) + → (g : (S₊∙ (suc m) →∙ X)) + → [ f ∣ g ]comp ≡ [ f ∣ g ] +[]comp≡[] {X = X} {n = n} {m} f g = + cong (_∘∙ (sphere→Join n m , IsoSphereJoin⁻Pres∙ n m)) (main n m f g) + where + ∨fun : {n m : ℕ} (f : (S₊∙ (suc n) →∙ X)) (g : (S₊∙ (suc m) →∙ X)) + → ((_ , inl north) →∙ X) + ∨fun {n = n} {m} f g = + ((f ∘∙ (inv (IsoSucSphereSusp n) , IsoSucSphereSusp∙ n)) ∨→ + (g ∘∙ (inv (IsoSucSphereSusp m) , IsoSucSphereSusp∙ m)) + , cong (fst f) (IsoSucSphereSusp∙ n) ∙ snd f) + + main : (n m : ℕ) (f : (S₊∙ (suc n) →∙ X)) (g : (S₊∙ (suc m) →∙ X)) + → (∨fun f g ∘∙ (joinTo⋁ {A = S₊∙ n} {B = S₊∙ m} , sym (push tt))) + ≡ [ f ∣ g ]-pre + main n m f g = + ΣPathP ((funExt (λ { (inl x) → rp + ; (inr x) → lp + ; (push a b i) j → pushcase a b j i})) + , ((cong₂ _∙_ (symDistr _ _) refl + ∙ sym (assoc _ _ _) + ∙ cong (rp ∙_) (rCancel _) + ∙ sym (rUnit rp)) + ◁ λ i j → rp (i ∨ j))) + where + lp = cong (fst f) (IsoSucSphereSusp∙ n) ∙ snd f + rp = cong (fst g) (IsoSucSphereSusp∙ m) ∙ snd g + + help : (n : ℕ) (a : _) + → Square (cong (inv (IsoSucSphereSusp n)) (σ (S₊∙ n) a)) (σS a) + (IsoSucSphereSusp∙ n) (IsoSucSphereSusp∙ n) + help zero a = cong-∙ SuspBool→S¹ (merid a) (sym (merid (pt (S₊∙ zero)))) + ∙ sym (rUnit _) + help (suc n) a = refl + + ∙∙Distr∙ : ∀ {ℓ} {A : Type ℓ} {x y z w u : A} + {p : x ≡ y} {q : y ≡ z} {r : z ≡ w} {s : w ≡ u} + → p ∙∙ q ∙ r ∙∙ s ≡ ((p ∙ q) ∙ r ∙ s) + ∙∙Distr∙ = doubleCompPath≡compPath _ _ _ + ∙ assoc _ _ _ + ∙ cong₂ _∙_ (assoc _ _ _) refl + ∙ sym (assoc _ _ _) + + pushcase : (a : S₊ n) (b : S₊ m) + → Square (cong (∨fun f g .fst + ∘ joinTo⋁ {A = S₊∙ n} {B = S₊∙ m}) (push a b)) + (cong (fst [ f ∣ g ]-pre) (push a b)) + rp lp + pushcase a b = + (cong-∙∙ (∨fun f g .fst) _ _ _ + ∙ (λ i → cong (fst g) (PathP→compPathR∙∙ (help _ b) i) + ∙∙ symDistr lp (sym rp) i + ∙∙ cong (fst f) (PathP→compPathR∙∙ (help _ a) i)) + ∙ (λ i → cong (fst g) + (IsoSucSphereSusp∙ m + ∙∙ σS b + ∙∙ (λ j → IsoSucSphereSusp∙ m (~ j ∨ i))) + ∙∙ compPath-filler' (cong (fst g) + (IsoSucSphereSusp∙ m)) (snd g) (~ i) + ∙ sym (compPath-filler' (cong (fst f) + (IsoSucSphereSusp∙ n)) (snd f) (~ i)) + ∙∙ cong (fst f) + ((λ j → IsoSucSphereSusp∙ n (i ∨ j)) + ∙∙ σS a + ∙∙ sym (IsoSucSphereSusp∙ n)))) + ◁ compPathR→PathP∙∙ + ( ∙∙Distr∙ + ∙ cong₂ _∙_ (cong₂ _∙_ (cong (cong (fst g)) + (sym (compPath≡compPath' _ _))) + refl) + refl + ∙ cong₂ _∙_ (cong (_∙ snd g) + (cong-∙ (fst g) (IsoSucSphereSusp∙ m) (σS b)) + ∙ sym (assoc _ _ _)) + (cong (sym (snd f) ∙_) + (cong-∙ (fst f) (σS a) + (sym (IsoSucSphereSusp∙ n))) + ∙ assoc _ _ _) + ∙ sym ∙∙Distr∙ + ∙ cong₃ _∙∙_∙∙_ refl (cong₂ _∙_ refl (compPath≡compPath' _ _)) refl + ∙ λ i → compPath-filler (cong (fst g) (IsoSucSphereSusp∙ m)) (snd g) i + ∙∙ ((λ j → snd g (~ j ∧ i)) ∙∙ cong (fst g) (σS b) ∙∙ snd g) + ∙ (sym (snd f) ∙∙ cong (fst f) (σS a) ∙∙ λ j → snd f (j ∧ i)) + ∙∙ sym (compPath-filler (cong (fst f) (IsoSucSphereSusp∙ n)) + (snd f) i)) + +-- Standard construction of Whitehead products in terms of whΣ +[]≡·whΣ' : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + (f : S₊∙ (suc n) →∙ X) (g : S₊∙ (suc m) →∙ X) + → [ f ∣ g ] + ≡ (·whΣ (S₊∙ n) (S₊∙ m) + (f ∘∙ ≃∙map (IsoSucSphereSuspInv≃∙ n)) + (g ∘∙ ≃∙map (IsoSucSphereSuspInv≃∙ m)) + ∘∙ (suspFun∙ (inv (SphereSmashIso n m)) + ∘∙ (fun (IsoSucSphereSusp (n + m)) , IsoSucSphereSusp∙' (n + m)))) +[]≡·whΣ' {X = X} {n} {m} f g = + cong₂ _∘∙_ ((cong₂ [_∣_]-pre + (sym (∘∙-idˡ _) + ∙ cong₂ _∘∙_ refl (sym (IsoSucSphereSusp≃∙CompR n)) + ∙ sym (∘∙-assoc f _ _)) + (sym (∘∙-idˡ _) + ∙ cong₂ _∘∙_ refl (sym (IsoSucSphereSusp≃∙CompR m)) + ∙ sym (∘∙-assoc g _ _)))) refl + ∙ cong₂ _∘∙_ + (cong (joinPinch∙ (S₊∙ n) (S₊∙ m) X) (funExt (λ a → funExt λ b → + cong₂ _∙_ (Ω→lemma m b g) (Ω→lemma n a f)))) + (ΣPathP (refl , sym (lem' n m) ∙ lUnit _)) + ∙ cong₂ _∘∙_ (·whΣ≡·wh (S₊∙ n) (S₊∙ m) + (f ∘∙ ≃∙map (IsoSucSphereSuspInv≃∙ n)) + (g ∘∙ ≃∙map (IsoSucSphereSuspInv≃∙ m))) + (refl {x = ≃∙map (invEquiv∙ (joinSphereEquiv∙ n m))}) + ∙ ∘∙-assoc _ _ _ + ∙ cong₂ _∘∙_ refl + (cong (_∘∙ ≃∙map (invEquiv∙ (joinSphereEquiv∙ n m))) + (lem n m) + ∙ ∘∙-assoc _ _ _ + ∙ cong₂ _∘∙_ refl + ((cong₂ _∘∙_ (sym (∘∙-idˡ _)) refl) + ∙ rightInv (post∘∙equiv (joinSphereEquiv∙ n m)) (idfun∙ _)) + ∙ ∘∙-idˡ _) + where + Ω→lemma : (n : ℕ) (a : S₊ n) (f : S₊∙ (suc n) →∙ X) + → Ω→ ((f ∘∙ ≃∙map (IsoSucSphereSuspInv≃∙ n)) + ∘∙ ≃∙map (IsoSucSphereSusp≃∙ n)) .fst (σS a) + ≡ Ω→ (f ∘∙ ≃∙map (IsoSucSphereSuspInv≃∙ n)) .fst (σ (S₊∙ n) a) + Ω→lemma n a f = + cong (λ f → Ω→ f .fst (σS a)) + ((∘∙-assoc f _ _ + ∙ cong₂ _∘∙_ refl (IsoSucSphereSusp≃∙CompR n)) + ∙ ∘∙-idˡ f) + ∙ Ω→comp-σS n f a + + lemL : (n m : ℕ) → + suspFun (inv (SphereSmashIso n m)) + (fun (IsoSucSphereSusp (n + m)) (ptSn (suc (n + m)))) + ≡ north + lemL zero zero = refl + lemL zero (suc m) = refl + lemL (suc n) m = refl + + lemR : (n m : ℕ) → + suspFun (inv (SphereSmashIso n m)) + (fun (IsoSucSphereSusp (n + m)) (ptSn (suc (n + m)))) + ≡ south + lemR zero zero = merid (inl tt) + lemR zero (suc m) = merid (inl tt) + lemR (suc n) m = merid (inl tt) + + lemma : (n m : ℕ) (a : S₊ n) (b : S₊ m) + → Square {A = Susp (S₊∙ n ⋀ S₊∙ m)} + (merid (inr (a , b))) + (cong (suspFun (inv (SphereSmashIso n m))) + (cong (fun (IsoSucSphereSusp (n + m))) + (cong (fst (fst (joinSphereEquiv∙ n m))) (push a b)))) + (sym (lemL n m)) + (sym (lemR n m)) + lemma zero zero false false = + compPath-filler (merid (inr (false , false))) (sym (merid (inl tt))) + ▷ (cong₂ _∙_ refl (cong (sym ∘ merid) (push (inl false))) + ∙ sym (cong-∙ (suspFun (inv (SphereSmashIso zero zero))) _ _)) + lemma zero zero false true = + cong merid (sym (push (inl false))) ◁ λ i j → merid (inl tt) (~ i ∧ j) + lemma zero zero true b = + cong merid (sym (push (inr b))) ◁ λ i j → merid (inl tt) (~ i ∧ j) + lemma zero (suc m) false b = + compPath-filler (merid (inr (false , b))) (sym (merid (inl tt))) + ▷ (cong₂ _∙_ refl (cong (sym ∘ merid) (push (inl false))) + ∙ sym (cong-∙ (suspFun (inv (SphereSmashIso zero (suc m)))) _ _)) + lemma zero (suc m) true b = + cong merid (sym (push (inr b))) + ◁ compPath-filler (merid (inl tt)) (sym (merid (inl tt))) + ▷ (cong₂ _∙_ (cong merid (push (inl false))) + (cong (sym ∘ merid) (push (inl false))) + ∙ sym (cong-∙ (suspFun (inv (SphereSmashIso zero (suc m)))) _ _)) + lemma (suc n) zero a b = + compPath-filler (merid (inr (a , b))) (sym (merid (inl tt))) + ▷ (cong₂ _∙_ + (cong merid (sym (leftInv (SphereSmashIso (suc n) zero) + (inr (a , b))))) + (cong (sym ∘ merid) (sym (SphereSmashIso⁻∙ (suc n) zero))) + ∙ sym (cong-∙ (suspFun (inv (SphereSmashIso (suc n) zero))) _ _)) + lemma (suc n) (suc m) a b = + compPath-filler (merid (inr (a , b))) (sym (merid (inl tt))) + ▷ (cong₂ _∙_ + (cong merid (sym (leftInv (SphereSmashIso (suc n) (suc m)) + (inr (a , b))))) + (cong (sym ∘ merid) (sym (SphereSmashIso⁻∙ (suc n) (suc m)))) + ∙ sym (cong-∙ (suspFun (inv (SphereSmashIso (suc n) (suc m)))) _ _)) + + lem : (n m : ℕ) + → Join→SuspSmash∙ (S₊∙ n) (S₊∙ m) + ≡ (suspFun∙ (inv (SphereSmashIso n m)) + ∘∙ (fun (IsoSucSphereSusp (n + m)) , IsoSucSphereSusp∙' (n + m))) + ∘∙ ≃∙map (joinSphereEquiv∙ n m) + lem n m i .fst (inl x) = lemL n m (~ i) + lem n m i .fst (inr x) = lemR n m (~ i) + lem n m i .fst (push a b j) = lemma n m a b i j + lem zero zero i .snd j = + (sym (lUnit (refl ∙ λ _ → north)) ∙ sym (lUnit refl)) (~ i) j + lem zero (suc m) i .snd j = + (sym (lUnit (refl ∙ λ _ → north)) ∙ sym (lUnit refl)) (~ i) j + lem (suc n) m i .snd j = + (sym (lUnit (refl ∙ λ _ → north)) ∙ sym (lUnit refl)) (~ i) j + + F : _ →∙ _ + F = fun (IsoSucSphereSusp (n + m)) , IsoSucSphereSusp∙' (n + m) + + p1 : (n m : ℕ) → leftInv (joinSphereIso' (suc n) m) (inl (ptSn (suc n))) + ≡ sym (push (ptSn (suc n)) (ptSn m)) + p1 n m = + cong₂ _∙_ (cong (congS (inv (invIso SmashJoinIso))) (sym (rUnit refl))) refl + ∙ sym (lUnit _) + + p2 : (n m : ℕ) → rightInv (joinSphereIso' (suc n) m) (ptSn (suc (suc n + m))) + ≡ sym (merid (ptSn (suc n + m))) + p2 n m = cong₂ _∙_ (cong (sym ∘ merid) (IdL⌣S {n = suc n} {m = m} (ptSn (suc n)))) + (sym (rUnit refl)) + ∙ sym (rUnit _) + + compPath-filler-diag : ∀ {ℓ} {A : Type ℓ} {x y z : A} (p : x ≡ y) (q : y ≡ z) + → Path (Path A _ _) (λ i → compPath-filler p q (~ i) i) p + compPath-filler-diag p q j i = compPath-filler p q (~ i ∧ ~ j) i + + p3 : (m : ℕ) → push true (ptSn (suc m)) + ∙ leftInv (joinSphereIso' zero (suc m)) (inl true) + ≡ refl + p3 m = cong₂ _∙_ refl + (cong₂ _∙_ (cong (congS (fun SmashJoinIso)) + (sym (rUnit (λ _ → north)))) refl + ∙ sym (lUnit _)) + ∙ rCancel (push true (ptSn (suc m))) + + retEqIsoToEquiv∙ : ∀ {ℓ} {A B : Pointed ℓ} + (e : A ≃∙ B) (e∙ : invEq (fst e) (pt B) ≡ pt A) + (eq : sym (secEq (fst e) (pt B)) + ∙ cong (fst (fst e)) e∙ + ≡ sym (snd e)) + → retEq (fst e) (pt A) + ≡ cong (invEq (fst e)) (snd e) + ∙ e∙ + retEqIsoToEquiv∙ {A = A} {B} = + Equiv∙J (λ A e → (e∙ : invEq (fst e) (pt B) ≡ pt A) + (eq : sym (secEq (fst e) (pt B)) + ∙ cong (fst (fst e)) e∙ + ≡ sym (snd e)) + → retEq (fst e) (pt A) ≡ cong (invEq (fst e)) (snd e) ∙ e∙) + λ e∙ p → sym p + + lem' : (n m : ℕ) → retEq (isoToEquiv (IsoSphereJoin n m)) (inl (ptSn n)) + ≡ IsoSphereJoin⁻Pres∙ n m + lem' n m = retEqIsoToEquiv∙ + (isoToEquiv (IsoSphereJoin n m) , IsoSphereJoinPres∙ n m) + (IsoSphereJoin⁻Pres∙ n m) (IsoSphereJoinPres∙Char n m) + ∙ sym (lUnit (IsoSphereJoin⁻Pres∙ n m)) + where + lemma3 : (n m : ℕ) + → (λ i → join→Sphere≡ (suc n) m + (push (ptSn (suc n)) (ptSn m) (~ i)) (~ i)) + ≡ sym (merid (ptSn (suc (n + m)))) + lemma3 n m = flipSquare + ((λ i j → join→Sphere≡ (suc n) m + (push (ptSn (suc n)) (ptSn m) (~ i ∨ j)) (~ i)) + ▷ (cong σS (IdL⌣S {n = suc n} {m = m} (ptSn (suc n))) ∙ σS∙)) + + IsoSphereJoinPres∙Char : (n m : ℕ) + → sym (secEq (isoToEquiv (IsoSphereJoin n m)) (ptSn (suc (n + m)))) + ∙ cong (join→Sphere n m) (IsoSphereJoin⁻Pres∙ n m) + ≡ sym (IsoSphereJoinPres∙ n m) + IsoSphereJoinPres∙Char zero zero = + cong₂ _∙_ (cong sym (sym (lUnit _) + ∙ sym (lUnit (refl ∙ refl)) ∙ sym (rUnit refl))) refl + ∙ sym (rUnit refl) + IsoSphereJoinPres∙Char zero (suc n) = + cong₂ _∙_ (cong₃ _∙∙_∙∙_ + (cong sym (cong₂ _∙_ refl (sym (rUnit refl)) ∙ sym (rUnit _))) + (λ j i → compPath-filler (merid (ptSn (suc n))) + (sym (merid (ptSn (suc n)))) (i ∧ ~ j) (~ i)) + refl) refl + ∙ cong sym (sym (rUnit _) ∙ rCancel (merid (ptSn (suc n)))) + IsoSphereJoinPres∙Char (suc n) m = cong₂ _∙_ (cong₃ _∙∙_∙∙_ + (cong sym (cong₂ _∙_ refl (sym (rUnit refl)) ∙ sym (rUnit _)) + ∙ cong merid (IdL⌣S {n = suc n} {m = m} (ptSn (suc n)))) + (lemma3 n m) refl + ∙ cong sym (rCancel (merid (ptSn (suc (n + m)))))) + refl + ∙ sym (rUnit refl) + +[]≡·whΣ : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + (f : S₊∙ (2 + n) →∙ X) (g : S₊∙ (2 + m) →∙ X) + → [ f ∣ g ] + ≡ (·whΣ (S₊∙ (suc n)) (S₊∙ (suc m)) f g + ∘∙ suspFun∙ (inv (SphereSmashIso (suc n) (suc m)))) +[]≡·whΣ {X = X} {n} {m} f g = []≡·whΣ' {X = X} f g + ∙ cong₂ _∘∙_ (cong₂ (·whΣ (S₊∙ (suc n)) (S₊∙ (suc m))) + (∘∙-idˡ f) (∘∙-idˡ g)) + (∘∙-idˡ (suspFun∙ (inv (SphereSmashIso (suc n) (suc m))))) + + +-- Basic algebraic properties +[∣]π'-distrL : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + (f g : π' (suc (suc n)) X) (h : π' (suc m) X) + → [ ·π' (suc n) f g ∣ h ]π' + ≡ ·π' _ [ f ∣ h ]π' [ g ∣ h ]π' +[∣]π'-distrL {X = X} {n} {m} = + ST.elim3 (λ _ _ _ → isSetPathImplicit) + λ f g h → cong ∣_∣₂ + (lem n m (·Susp (S₊∙ (suc n)) f g) h + ∙ (cong (_∘∙ suspFun∙ (inv (SphereSmashIso (suc n) m))) + (WhiteheadProdΣBilinₗ + (S₊∙ (suc n)) (S₊∙ n) (IsoSucSphereSusp≃∙ n) (S₊∙ m) f g (h *)) + ∙ cong (_∘∙ suspFun∙ (inv (SphereSmashIso (suc n) m))) + (·Susp-suspFun (S₊∙ (suc n) ⋀∙ (S₊ m , ptSn m)) (S₊∙ (suc (n + m))) + (isoToEquiv (SphereSmashIso (suc n) m) , refl) + (·whΣ (S₊∙ (suc n)) (S₊ m , ptSn m) f (h *)) + (·whΣ (S₊∙ (suc n)) (S₊ m , ptSn m) g (h *))) + ∙ ∘∙-assoc _ _ _ + ∙ cong₂ _∘∙_ refl + (sym (suspFun∙Comp (fun (SphereSmashIso (suc n) m)) + (inv (SphereSmashIso (suc n) m))) + ∙ cong suspFun∙ (funExt (rightInv (SphereSmashIso (suc n) m))) + ∙ suspFun∙IdFun) + ∙ ∘∙-idˡ _) + ∙ cong₂ (·Susp (S₊∙ (suc (n + m)))) + (sym (lem n m f h)) + (sym (lem n m g h))) + where + _* : {m : ℕ} (h : S₊∙ (suc m) →∙ X) → _ + _* {m = m} h = h ∘∙ ≃∙map (IsoSucSphereSuspInv≃∙ m) + + module _ (n m : ℕ) (f : S₊∙ (suc (suc n)) →∙ X) (h : S₊∙ (suc m) →∙ X) where + lem : [ f ∣ h ] + ≡ (·whΣ (S₊∙ (suc n)) (S₊∙ m) f (h ∘∙ ≃∙map (IsoSucSphereSuspInv≃∙ m)) + ∘∙ (suspFun∙ (inv (SphereSmashIso (suc n) m)))) + lem = []≡·whΣ' f h + ∙ cong₂ _∘∙_ (cong₂ (·whΣ (S₊∙ (suc n)) (S₊∙ m)) (∘∙-idˡ f) refl) + (∘∙-idˡ (suspFun∙ (inv (SphereSmashIso (suc n) m)))) + +[∣]π'-distrR : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + (f : π' (suc n) X) (g h : π' (suc (suc m)) X) + → [ f ∣ ·π' (suc m) g h ]π' + ≡ ·π' _ [ f ∣ g ]π' [ f ∣ h ]π' +[∣]π'-distrR {X = X} {n} {m} = + ST.elim3 (λ _ _ _ → isSetPathImplicit) + λ f g h → cong ∣_∣₂ + ([]≡·whΣ' f (·Susp (S₊∙ (suc m)) g h) + ∙ cong₂ _∘∙_ (cong₂ (·whΣ (S₊∙ n) (S₊∙ (suc m))) refl + (∘∙-idˡ (·Susp (S₊∙ (suc m)) g h))) refl + ∙ cong₂ _∘∙_ (WhiteheadProdΣBilinᵣ (S₊∙ n) (S₊∙ (suc m)) (S₊∙ m) + (IsoSucSphereSusp≃∙ m) + (f *) g h) refl + ∙ main n m f g h + ∙ cong₂ ∙Π (cong₂ _∘∙_ + (cong₂ (·whΣ (S₊∙ n) (S₊∙ (suc m)) ) refl (sym (∘∙-idˡ g))) refl + ∙ sym ([]≡·whΣ' f g)) + (cong₂ _∘∙_ (cong₂ (·whΣ (S₊∙ n) (S₊∙ (suc m)) ) refl (sym (∘∙-idˡ h))) + refl + ∙ sym ([]≡·whΣ' f h))) + where + cor : (n m : ℕ) → _ →∙ _ + cor n m = + (suspFun∙ (inv (SphereSmashIso n (suc m))) ∘∙ + (fun (IsoSucSphereSusp (n + suc m)) , + IsoSucSphereSusp∙' (n + suc m))) + + _* : {m : ℕ} (h : S₊∙ (suc m) →∙ X) → _ + _* {m = m} h = h ∘∙ ≃∙map (IsoSucSphereSuspInv≃∙ m) + + main : (n m : ℕ) (f : S₊∙ (suc n) →∙ X) (g h : S₊∙ (suc (suc m)) →∙ X) + → (·Susp (S₊∙ n ⋀∙ S₊∙ (suc m)) + (·whΣ (S₊∙ n) (S₊∙ (suc m)) (f *) g) + (·whΣ (S₊∙ n) (S₊∙ (suc m)) (f *) h) + ∘∙ cor n m) + ≡ ∙Π (·whΣ (S₊∙ n) (S₊∙ (suc m)) + (f ∘∙ ≃∙map (IsoSucSphereSuspInv≃∙ n)) g + ∘∙ cor n m) + (·whΣ (S₊∙ n) (S₊∙ (suc m)) + (f ∘∙ ≃∙map (IsoSucSphereSuspInv≃∙ n)) h + ∘∙ cor n m) + main zero m f g h = + cong₂ _∘∙_ (·Susp-suspFun (S₊∙ zero ⋀∙ S₊∙ (suc m)) + (S₊∙ (suc m)) + (isoToEquiv (SphereSmashIso zero (suc m)) , refl) + (·whΣ (S₊∙ zero) (S₊∙ (suc m)) (f *) g) + (·whΣ (S₊∙ zero) (S₊∙ (suc m)) (f *) h)) + (∘∙-idˡ (suspFun∙ (λ y → inr (false , y)))) + ∙ ∘∙-assoc _ _ _ + ∙ cong₂ _∘∙_ refl + (ΣPathP (funExt (secEq (isoToEquiv + (congSuspIso (SphereSmashIso zero (suc m))))) + , sym (rUnit refl))) + ∙ ∘∙-idˡ _ + ∙ cong₂ (·Susp (S₊∙ (suc m))) + (cong₂ _∘∙_ refl (sym (∘∙-idˡ _))) + (cong₂ _∘∙_ refl (sym (∘∙-idˡ _))) + main (suc n) m f g h = + cong₂ _∘∙_ (·Susp-suspFun (S₊∙ (suc n) ⋀∙ S₊∙ (suc m)) + (S₊∙ (suc (n + suc m))) + (isoToEquiv (SphereSmashIso (suc n) (suc m)) , refl) + (·whΣ (S₊∙ (suc n)) (S₊∙ (suc m)) (f *) g) + (·whΣ (S₊∙ (suc n)) (S₊∙ (suc m)) (f *) h)) refl + ∙ ∘∙-assoc _ _ _ + ∙ cong₂ _∘∙_ refl (ΣPathP + ((funExt (secEq (isoToEquiv + (congSuspIso (SphereSmashIso (suc n) (suc m)))))) + , sym (rUnit _) + ∙ cong (congS (suspFun (fun (SphereSmashIso (suc n) (suc m))))) + (sym (rUnit refl)))) + ∙ ∘∙-idˡ _ + ∙ cong₂ (·Susp (S₊∙ (suc (n + suc m)))) + (cong₂ _∘∙_ refl (sym (∘∙-idˡ _))) + (cong₂ _∘∙_ refl (sym (∘∙-idˡ _))) + + +[∣]π'-idL : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + (g : π' (suc m) X) → [ 1π' (suc n) ∣ g ]π' ≡ 1π' (suc (n + m)) +[∣]π'-idL {n = n} {m} = + ST.elim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ ([]≡·whΣ' 1Π f + ∙ cong₂ _∘∙_ (cong₂ (·whΣ (S₊∙ n) (S₊∙ m)) + (ΣPathP (refl , sym (rUnit refl))) refl + ∙ WhiteheadProdΣIdL (S₊∙ n) (S₊∙ m) + (f ∘∙ ≃∙map (IsoSucSphereSuspInv≃∙ m))) refl + ∙ ΣPathP (refl , (sym (rUnit refl)))) + +[∣]π'-idR : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + (f : π' (suc n) X) → [ f ∣ 1π' (suc m) ]π' ≡ 1π' (suc (n + m)) +[∣]π'-idR {n = n} {m} = + ST.elim (λ _ → isSetPathImplicit) + λ f → cong ∣_∣₂ ([]≡·whΣ' f 1Π + ∙ cong₂ _∘∙_ (cong₂ (·whΣ (S₊∙ n) (S₊∙ m)) refl + (ΣPathP (refl , sym (rUnit refl))) + ∙ WhiteheadProdΣIdR (S₊∙ n) (S₊∙ m) + (f ∘∙ ≃∙map (IsoSucSphereSuspInv≃∙ n))) refl + ∙ ΣPathP (refl , (sym (rUnit refl)))) + +[∣]π'-invDistrL : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + (f : π' (suc (suc n)) X) (g : π' (suc m) X) + → [ -π' (suc n) f ∣ g ]π' + ≡ -π' (suc (n + m)) [ f ∣ g ]π' +[∣]π'-invDistrL {n = n} {m} f g = + sym (sym (π'-assoc (suc (n + m)) [ -π' (suc n) f ∣ g ]π' + [ f ∣ g ]π' (-π' (suc (n + m)) + [ f ∣ g ]π')) + ∙ cong₂ (·π' (suc (n + m))) refl (π'-rCancel (suc (n + m)) [ f ∣ g ]π') + ∙ π'-rUnit _ ([ -π' (suc n) f ∣ g ]π')) + ∙ cong (λ x → ·π' (suc (n + m)) x (-π' (suc (n + m)) [ f ∣ g ]π')) + (sym ([∣]π'-distrL (-π' (suc n) f) f g)) + ∙ cong₂ (·π' (suc (n + m))) + (cong [_∣ g ]π' (π'-lCancel (suc n) f) ∙ [∣]π'-idL g) refl + ∙ π'-lUnit _ (-π' (suc (n + m)) [ f ∣ g ]π') + +[∣]π'-invDistrR : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + (f : π' (suc n) X) (g : π' (suc (suc m)) X) + → [ f ∣ -π' (suc m) g ]π' + ≡ -π' (n + suc m) [ f ∣ g ]π' +[∣]π'-invDistrR {n = n} {m} f g = + sym (sym (π'-assoc (n + suc m) + [ f ∣ -π' (suc m) g ]π' + [ f ∣ g ]π' + (-π' (n + suc m) [ f ∣ g ]π')) + ∙ cong₂ (·π' (n + suc m)) refl + (π'-rCancel (n + suc m) [ f ∣ g ]π') + ∙ π'-rUnit _ ([ f ∣ -π' (suc m) g ]π')) + ∙ cong (λ x → ·π' (n + suc m) x (-π' (n + suc m) [ f ∣ g ]π')) + (sym ([∣]π'-distrR f (-π' (suc m) g) g)) + ∙ cong₂ (·π' (n + suc m)) + (cong [ f ∣_]π' (π'-lCancel (suc m) g) ∙ [∣]π'-idR f) refl + ∙ π'-lUnit _ (-π' (n + suc m) [ f ∣ g ]π') + +[∣]π'-inv^DistrL : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} (k : ℕ) + (f : π' (suc (suc n)) X) (g : π' (suc m) X) + → [ -π^ k f ∣ g ]π' ≡ -π^ k [ f ∣ g ]π' +[∣]π'-inv^DistrL {n = n} {m} zero f g = refl +[∣]π'-inv^DistrL {n = n} {m} (suc k) f g = + [∣]π'-invDistrL (-π^ k f) g + ∙ cong (-π' _) ([∣]π'-inv^DistrL k f g) + +[∣]π'-inv^DistrR : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} (k : ℕ) + (f : π' (suc n) X) (g : π' (suc (suc m)) X) + → [ f ∣ -π^ k g ]π' ≡ -π^ k [ f ∣ g ]π' +[∣]π'-inv^DistrR {n = n} {m} zero f g = refl +[∣]π'-inv^DistrR {n = n} {m} (suc k) f g = + [∣]π'-invDistrR f (-π^ k g) + ∙ cong (-π' _) ([∣]π'-inv^DistrR k f g) + +-- Commutativity of Whitehead products on π* (modulo correction iso) +[_∣_]π*-comm : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + (f : π' (suc (suc n)) X) (g : π' (suc (suc m)) X) + → [ f ∣ g ]π* ≡ fun (π*SwapIso (suc m) (suc n) X) [ g ∣ f ]π* +[_∣_]π*-comm {n = n} {m = m} = + ST.elim2 (λ _ _ → isOfHLevelPath 2 squash₂ _ _) + λ f g → cong ∣_∣₂ + (WhiteheadProdComm' + (S₊∙ (suc n)) (S₊∙ n) + (isoToEquiv (IsoSucSphereSusp n) , IsoSucSphereSusp∙' n) + (S₊∙ (suc m)) (S₊∙ m) + (isoToEquiv (IsoSucSphereSusp m) , IsoSucSphereSusp∙' m) f g + ∙ cong (·wh (S₊∙ (suc m)) (S₊∙ (suc n)) g f ∘∙_) + (ΣPathP (refl , sym (cong₂ _∙_ refl (∙∙lCancel _) ∙ sym (rUnit _))))) + +-- Graded commutativity of Whitehead products (w.r.t the following grading) +_·₋₁_ : ℕ → ℕ → ℕ +n ·₋₁ m = suc (suc n · suc m) + +[_∣_]π'-comm : ∀ {ℓ} {X : Pointed ℓ} {n m : ℕ} + (f : π' (suc (suc n)) X) (g : π' (suc (suc m)) X) + → [ f ∣ g ]π' + ≡ subst (λ k → π' (suc k) X) (+-comm (suc m) (suc n)) + (-π^ (m ·₋₁ n) [ g ∣ f ]π') +[_∣_]π'-comm {X = X} {n} {m} = + PT.rec (isPropΠ2 (λ _ _ → squash₂ _ _)) (λ main → + ST.elim2 (λ _ _ → isSetPathImplicit) + λ f g → cong ∣_∣₂ + (cong (λ f → + _∘∙_ {A = S₊∙ (suc (suc (n + suc m)))} + f (sphere→Join (suc n) (suc m) + , IsoSphereJoin⁻Pres∙ (suc n) (suc m))) + (WhiteheadProdComm' (S₊∙ (suc n)) (S₊∙ n) + (isoToEquiv (IsoSucSphereSusp n) , IsoSucSphereSusp∙' n) + (S₊∙ (suc m)) (S₊∙ m) + (isoToEquiv (IsoSucSphereSusp m) , IsoSucSphereSusp∙' m) f g) + ∙ refl) + ∙ cong ∣_∣₂ (∘∙-assoc (·wh (S₊∙ (suc m)) (S₊∙ (suc n)) g f) + join-commFun∙ + (sphere→Join (suc n) (suc m) + , (λ _ → inl (ptSn (suc n))))) + ∙ sym (fromPathP {A = λ i → π' (suc (+-comm (suc m) (suc n) i)) X} + ((-π^≡iter-Π (suc (suc m · suc n)) [ ∣ g ∣₂ ∣ ∣ f ∣₂ ]π' + ∙ cong ∣_∣₂ (iter-Π≡∘-S^ (suc (suc m · suc n)) [ g ∣ f ]) + ∙ cong ∣_∣₂ (∘∙-assoc _ _ _)) + ◁ λ i → ∣ [ g ∣ f ]-pre ∘∙ main i ∣₂))) main + + where + main' : (x : _) + → PathP (λ i → S₊ (suc (+-comm (suc m) (suc n) (~ i)))) + (-S^ (suc (suc m · suc n)) (join→Sphere (suc n) (suc m) x)) + (join→Sphere (suc m) (suc n) (join-commFun x)) + main' (inl x) i = -S^∙ (suc (suc m · suc n)) .snd i + main' (inr x) i = -S^∙ (suc (suc m · suc n)) .snd i + main' (push a b i) j = lem j i + where + lem : SquareP (λ i j → S₊ (suc (+-comm (suc m) (suc n) (~ i)))) + (cong (-S^ (suc (suc m · suc n))) (σS (a ⌣S b))) + (sym (σS (b ⌣S a))) + (λ i → -S^∙ (suc (suc m · suc n)) .snd i) + λ i → -S^∙ (suc (suc m · suc n)) .snd i + lem = cong (congS (-S^ (suc (suc m · suc n))) ∘ σS) + (comm⌣S a b) + ◁ (λ i j → cong-S^σ _ (suc (suc m · suc n)) + (transp (λ j → S₊ (+-comm (suc m) (suc n) (j ∧ ~ i))) + i (-S^ (suc (n + m · suc n)) (b ⌣S a))) (~ i) j) + ▷ (cong σS ((λ i → -S^ (suc (suc m · suc n)) + (-S^ ((suc m · suc n)) (b ⌣S a))) + ∙ cong invSphere (-S^-comp (suc m · suc n) (suc m · suc n) + (b ⌣S a) + ∙ -S^·2 (suc m · suc n) (b ⌣S a))) + ∙ σ-invSphere _ (b ⌣S a)) + + + main : ∥ PathP (λ i → S₊∙ (suc (+-comm (suc m) (suc n) i)) + →∙ join∙ (S₊∙ (suc m)) (S₊∙ (suc n))) + ((sphere→Join (suc m) (suc n) , refl) + ∘∙ -S^∙ (suc (suc m · suc n))) + (join-commFun∙ ∘∙ (sphere→Join (suc n) (suc m) , refl)) ∥₁ + main = TR.rec (isProp→isOfHLevelSuc (m + suc n) squash₁) + (λ Q → ∣ ΣPathP (fstEq , Q) ∣₁) + (isConnectedPathP _ + (isConnectedPath _ + (subst (isConnected (suc (suc (suc (m + suc n))))) + (isoToPath (invIso (joinSphereIso' (suc m) (suc n)))) + (sphereConnected (suc (suc m + suc n))) ) _ _) _ _ .fst) + where + fstEq : PathP _ _ _ + fstEq = toPathP (funExt (λ s + → ((transportRefl _ + ∙ cong (sphere→Join (suc m) (suc n)) + (sym (substCommSlice (λ n → S₊ (suc n)) (λ n → S₊ (suc n)) + (λ _ → -S^ (suc (suc m · suc n))) + (sym (+-comm (suc m) (suc n))) + (join→Sphere (suc n) (suc m) + (sphere→Join (suc n) (suc m) s)) + ∙ cong (-S^ (suc (suc m · suc n))) + (cong (subst (S₊ ∘ suc) (sym (+-comm (suc m) (suc n)))) + (Iso.rightInv (IsoSphereJoin (suc n) (suc m)) s))))) + ∙ cong (sphere→Join (suc m) (suc n)) + (fromPathP (main' (sphere→Join (suc n) (suc m) s)))) + ∙ Iso.leftInv (IsoSphereJoin (suc m) (suc n)) + (join-commFun (sphere→Join (suc n) (suc m) s)))) + +-- Time for `associativty', i.e. the Jacobi identity. +-- We need some technical lemmas: +private + assocPath : (n m l : ℕ) → suc m + (suc n + suc l) ≡ suc n + (suc m + suc l) + assocPath n m l = (+-assoc (suc m) (suc n) (suc l) + ∙ cong (_+ suc l) (+-comm (suc m) (suc n)) + ∙ +-assoc (suc n) (suc m) (suc l) ⁻¹) + +-- We will want to use the Jacobi identity already proved for whΣ. For +-- this reason, we first rephrase triple whitehead products [ f ∣ [ g +-- ∣ h ] ] and [ [ f ∣ g ] ∣ h ] in terms of whΣ. To do this, we need +-- the following correction equivalences + +tripleSmasherL≃∙ : {n m l : ℕ} + → S₊∙ (suc n + (suc m + suc l)) + ≃∙ S₊∙ (suc n) ⋀∙ (S₊∙ (suc m) ⋀∙ S₊∙ (suc l)) +tripleSmasherL≃∙ {n = n} {m} {l} = + compEquiv∙ (isoToEquiv (invIso (SphereSmashIso (suc n) (suc m + suc l))) + , SphereSmashIso⁻∙ (suc n) (suc m + suc l)) + (⋀≃ (idEquiv∙ (S₊∙ (suc n))) + ((isoToEquiv (invIso (SphereSmashIso (suc m) (suc l)))) + , (SphereSmashIso⁻∙ (suc m) (suc l))) + , refl) + +tripleSmasherR≃∙ : {n m l : ℕ} + → S₊∙ ((suc n + suc m) + suc l) + ≃∙ ((S₊∙ (suc n) ⋀∙ S₊∙ (suc m)) ⋀∙ S₊∙ (suc l)) +tripleSmasherR≃∙ {n = n} {m} {l} = + compEquiv∙ (isoToEquiv (invIso (SphereSmashIso (suc n + suc m) (suc l))) + , SphereSmashIso⁻∙ (suc n + suc m) (suc l)) + (⋀≃ ((isoToEquiv (invIso (SphereSmashIso (suc n) (suc m)))) + , (SphereSmashIso⁻∙ (suc n) (suc m))) + (idEquiv∙ (S₊∙ (suc l))) + , refl) + +-- Characterisation of triple Whitehead products +3×[]≡·whΣ-L : ∀ {ℓ} {X : Pointed ℓ} {n m l : ℕ} + (f : S₊∙ (2 + n) →∙ X) (g : S₊∙ (2 + m) →∙ X) + (h : S₊∙ (2 + l) →∙ X) + → [ f ∣ [ g ∣ h ] ] + ≡ (·whΣ (S₊∙ (suc n)) (S₊∙ (suc m) ⋀∙ S₊∙ (suc l)) f + (·whΣ (S₊∙ (suc m)) (S₊∙ (suc l)) g h) + ∘∙ suspFun∙ (fst (fst tripleSmasherL≃∙))) +3×[]≡·whΣ-L {X = X} {n} {m} {l} f g h = + cong₂ [_∣_] refl ([]≡·whΣ {X = X} g h) + ∙ []≡·whΣ f _ + ∙ cong₂ _∘∙_ + (whΣ-postCompR ((isoToEquiv (invIso (SphereSmashIso (suc m) (suc l)))) + , (SphereSmashIso⁻∙ (suc m) (suc l))) f + (·whΣ (S₊∙ (suc m)) (S₊∙ (suc l)) g h)) + refl + ∙ ∘∙-assoc _ _ _ + ∙ cong₂ _∘∙_ refl (sym (suspFun∙Comp _ _)) + +3×[]≡·whΣ-R : ∀ {ℓ} {X : Pointed ℓ} {n m l : ℕ} + (f : S₊∙ (2 + n) →∙ X) (g : S₊∙ (2 + m) →∙ X) + (h : S₊∙ (2 + l) →∙ X) + → [ [ f ∣ g ] ∣ h ] + ≡ (·whΣ (S₊∙ (suc n) ⋀∙ S₊∙ (suc m)) (S₊∙ (suc l)) + (·whΣ (S₊∙ (suc n)) (S₊∙ (suc m)) f g) h + ∘∙ suspFun∙ ((fst (fst tripleSmasherR≃∙)))) +3×[]≡·whΣ-R {X = X} {n} {m} {l} f g h = + cong₂ [_∣_] ([]≡·whΣ {X = X} f g) refl + ∙ []≡·whΣ _ _ + ∙ cong₂ _∘∙_ (whΣ-postCompL + (isoToEquiv (invIso (SphereSmashIso (suc n) (suc m))) + , (SphereSmashIso⁻∙ (suc n) (suc m))) + (·whΣ (S₊∙ (suc n)) (S₊∙ (suc m)) f g) h) refl + ∙ ∘∙-assoc _ _ _ + ∙ cong₂ _∘∙_ refl + (sym (suspFun∙Comp _ _)) + +-- Jacobi identity +[_∣_]π'Jacobi : ∀ {ℓ} {X : Pointed ℓ} {n m l : ℕ} + (f : π' (suc (suc n)) X) + (g : π' (suc (suc m)) X) + (h : π' (suc (suc l)) X) + → [ f ∣ [ g ∣ h ]π' ]π' + ≡ ·π' _ (subst (λ k → π' k X) + (cong (2 +_) (sym (+-assoc n (suc m) (suc l)))) + ([ [ f ∣ g ]π' ∣ h ]π')) + (subst (λ k → π' k X) + (cong suc (assocPath n m l)) + (-π^ (suc n · suc m) [ g ∣ [ f ∣ h ]π' ]π')) +[_∣_]π'Jacobi {X = X} {n} {m} {l} = + ST.elim3 (λ _ _ _ → isSetPathImplicit) + λ f g h → cong ∣_∣₂ + (3×[]≡·whΣ-L f g h + ∙ cong₂ _∘∙_ + (JacobiΣR' (S₊∙ (suc n)) (S₊∙ n) + (isoToEquiv (IsoSucSphereSusp n) , IsoSucSphereSusp∙' n) + (S₊∙ (suc m)) (S₊∙ m) + (isoToEquiv (IsoSucSphereSusp m) , IsoSucSphereSusp∙' m) + (S₊∙ (suc l)) (S₊∙ l) + (isoToEquiv (IsoSucSphereSusp l) , IsoSucSphereSusp∙' l) + f g h) refl + ∙ cong₂ _∘∙_ (·Susp-suspFun _ _ (invEquiv∙ tripleSmasherL≃∙) _ _) refl + ∙ ∘∙-assoc _ _ _ + ∙ cong₂ _∘∙_ refl + (sym (suspFun∙Comp _ _) + ∙ cong suspFun∙ (funExt (retEq (fst tripleSmasherL≃∙))) + ∙ suspFun∙IdFun) + ∙ ∘∙-idˡ _ + ∙ cong₂ (·Susp (S₊∙ (suc (n + suc (m + suc l))))) + (sym (sym (subst∙Id (S₊∙ ∘ (2 +_)) + (sym (+-assoc n (suc m) (suc l))) + [ [ f ∣ g ] ∣ h ]) + ∙ cong₂ _∘∙_ (3×[]≡·whΣ-R f g h) (λ _ → s1) + ∙ ∘∙-assoc _ _ _ + ∙ cong₂ _∘∙_ refl (ΣPathP + ((funExt λ { north → refl + ; south → refl + ; (merid a i) j → main a j i}) + , sym (rUnit refl)) + ∙ suspFun∙Comp _ _) + ∙ sym (∘∙-assoc _ _ _))) + (sym (cong (transport (λ i → S₊∙ (suc (assocPath n m l i)) →∙ X)) + (iter-Π≡∘-S^ deg _) + ∙ (sym (subst∙Id (S₊∙ ∘ suc) + (assocPath n m l) _) + ∙ cong₂ _∘∙_ (cong₂ _∘∙_ (3×[]≡·whΣ-L g f h) + (-S^∙suspFun deg) + ∙ ∘∙-assoc _ _ _ + ∙ cong₂ _∘∙_ refl + (sym (suspFun∙Comp (fst (fst tripleSmasherL≃∙)) + (-S^ deg)))) refl) + ∙ ∘∙-assoc _ _ _ + ∙ cong₂ _∘∙_ refl + ((cong₂ _∘∙_ refl + refl + ∙ sym (suspFun∙substLem (sym (assocPath n m l)) + (fst (fst tripleSmasherL≃∙) ∘ -S^ deg)) + ∙ final-lemma) + ∙ suspFun∙Comp _ _) + ∙ sym (∘∙-assoc _ _ _)))) + ∙ refl + ∙ cong₂ (·π' (suc (n + suc (m + suc l)))) refl + (cong (subst (λ k → π' k X) (cong suc (assocPath n m l))) + (sym (-π^≡iter-Π deg _))) + where + suspFun∙substLem : ∀ {ℓ} {X : Type ℓ} {n m : ℕ} + (p : suc n ≡ suc m) + (f : S₊ (suc m) → X) + → suspFun∙ (f ∘ subst S₊ p) + ≡ suspFun∙ f ∘∙ subst∙ (λ x → S₊∙ (suc x)) p + suspFun∙substLem p f = + cong (λ p → suspFun∙ (f ∘ subst S₊ p)) + (isSetℕ _ _ p (cong suc (cong predℕ p))) + ∙ ∘∙preSubstLem _ _ _ f + ∙ cong (λ p → suspFun∙ f ∘∙ subst∙ (λ x → S₊∙ (suc x)) p) + (isSetℕ _ _ (cong suc (cong predℕ p)) p) + where + ∘∙preSubstLem : ∀ {ℓ} {X : Type ℓ} (n m : ℕ) + (p : n ≡ m) + (f : S₊ (suc m) → X) + → suspFun∙ (f ∘ subst S₊ (cong suc p)) + ≡ (suspFun∙ f ∘∙ subst∙ (S₊∙ ∘ suc) (cong suc p)) + ∘∙preSubstLem n = J> λ f + → (cong suspFun∙ (funExt (λ x → cong f (transportRefl x))) + ∙ sym (∘∙-idˡ _)) + ∙ cong₂ _∘∙_ refl (sym (subst∙refl S₊∙)) + + deg = suc n · suc m + + meridLem1 : (a : S₊ (suc (n + suc (m + suc l)))) + → merid (fst (fst tripleSmasherL≃∙) + (-S^ deg (subst S₊ (sym (assocPath n m l)) a))) + ≡ merid (inv SmashAssocIso + ((⋀comm→∙ ⋀→ idfun∙ (S₊∙ (suc l))) + (fun SmashAssocIso (fst (fst tripleSmasherL≃∙) a)))) + meridLem1 a = + cong merid (cong (fst (fst tripleSmasherL≃∙) + ∘ -S^ deg + ∘ subst S₊ (sym (assocPath n m l))) + (sym (retEq (fst tripleSmasherL≃∙) a))) + ∙ meridLem2 (fst (fst tripleSmasherL≃∙) a) + where + meridLem2 : (a : _) + → merid (fst (fst tripleSmasherL≃∙) + (-S^ deg + (subst S₊ (sym (assocPath n m l)) + (invEq (fst tripleSmasherL≃∙) a)))) + ≡ merid (inv SmashAssocIso + ((⋀comm→∙ ⋀→ idfun∙ (S₊∙ (suc l))) + (fun SmashAssocIso a))) + meridLem2 = + ⋀→HomogeneousPathP refl (isHomogeneousPath _ _) + λ x → ⋀→HomogeneousPathP refl (isHomogeneousPath _ _) + λ y z → cong merid + (refl + ∙ cong (fst (fst tripleSmasherL≃∙) ∘ -S^ deg) + (meridLem3 x y z) + ∙ cong (fst (fst tripleSmasherL≃∙)) + (-S^² deg _) + ∙ secEq (fst tripleSmasherL≃∙) (inr (y , inr (x , z)))) + where + p = (cong suc (+-assoc n (suc m) (suc l)) + ∙ cong (_+ suc l) (sym (+-comm (suc m) (suc n)))) + ∙ cong suc (sym (+-assoc m (suc n) (suc l))) + + meridLem3 : (x : S₊ (suc n)) (y : S₊ (suc m)) (z : S₊ (suc l)) + → subst S₊ (sym (assocPath n m l)) + (x ⌣S (y ⌣S z)) + ≡ -S^ deg (y ⌣S (x ⌣S z)) + meridLem3 x y z = + cong (λ P → subst S₊ P (x ⌣S (y ⌣S z))) + (isSetℕ _ _ _ _) + ∙ fromPathP + (compPathP' {B = S₊} + (compPathP' {B = S₊} + (λ i → (assoc⌣S x y z i)) + (λ i → + (toPathP {A = λ i → S₊ (+-comm (suc m) (suc n) i)} + (sym (comm⌣S x y)) (~ i) ⌣S z)) + ▷ ((⌣Sinvₗ^ (suc m · suc n) (y ⌣S x) z + ∙ λ i → -S^ (·-comm (suc m) (suc n) i) ((y ⌣S x) ⌣S z)))) + (λ i → -S^ deg (assoc⌣S y x z (~ i)))) + + final-lemma : suspFun∙ + ((λ x → fst (fst tripleSmasherL≃∙) (-S^ deg x)) ∘ + subst S₊ (λ i → assocPath n m l (~ i))) + ≡ + suspFun∙ + ((inv SmashAssocIso ∘ + (⋀comm→∙ ⋀→ idfun∙ (S₊∙ (suc l))) ∘ fun SmashAssocIso) + ∘ invEq (invEquiv (fst tripleSmasherL≃∙))) + final-lemma = ΣPathP (( + funExt (λ { north → refl + ; south → refl + ; (merid a i) j → meridLem1 a j i})) + , refl) + + + s1 = subst∙ (S₊∙ ∘ (2 +_)) (+-assoc n (suc m) (suc l)) + s2 = subst∙ (S₊∙ ∘ suc) (sym (assocPath n m l)) + + main : (a : S₊ (suc (n + (suc m + suc l)))) + → cong (suspFun (fst (fst tripleSmasherR≃∙)) ∘ fst s1) + (merid a) + ≡ merid (fun SmashAssocIso (invEq (fst (invEquiv∙ tripleSmasherL≃∙)) a)) + main a = (λ j i → suspFun (fst (fst tripleSmasherR≃∙)) + (transp (λ i → S₊ (suc (suc + (+-assoc n (suc m) (suc l) (i ∨ j))))) j + (merid (transp (λ i → S₊ (suc + (+-assoc n (suc m) (suc l) (i ∧ j)))) + (~ j) a) i))) + ∙ cong (merid ∘ fst (fst tripleSmasherR≃∙) + ∘ subst (S₊ ∘ suc) (+-assoc n (suc m) (suc l))) + (sym (retEq (fst tripleSmasherL≃∙) a)) + ∙ lem2 (fst (fst tripleSmasherL≃∙) a) + where + lem2 : (a : _) → merid (fst (fst tripleSmasherR≃∙) + (subst (S₊ ∘ suc) (+-assoc n (suc m) (suc l)) + (invEq (fst tripleSmasherL≃∙) a))) + ≡ merid (fun SmashAssocIso a) + lem2 = ⋀→HomogeneousPathP refl (isHomogeneousPath _ _) + λ x → ⋀→HomogeneousPathP refl (isHomogeneousPath _ _) + λ y z → + cong merid (cong (fst (fst tripleSmasherR≃∙)) + (fromPathP (assoc⌣S x y z)) + ∙ secEq (fst tripleSmasherR≃∙) (inr (inr (x , y) , z))) + +-- Alternative (more standard): +jacobiPath₁ : (p q r : ℕ) + → suc (suc (q + suc (r + suc p))) ≡ + suc (suc (p + suc (q + suc r))) +jacobiPath₁ p q r = + cong suc + (+-assoc (suc q) (suc r) (suc p) + ∙ sym (+-comm (suc p) (suc q + suc r))) + +jacobiPath₂ : (p q r : ℕ) + → suc (suc (r + suc (p + suc q))) + ≡ suc (suc (p + suc (q + suc r))) +jacobiPath₂ p q r = cong suc + (+-comm (suc r) (suc p + suc q) + ∙ sym (+-assoc (suc p) (suc q) (suc r))) + +[_∣_]π'Jacobi' : ∀ {ℓ} {X : Pointed ℓ} {p q r : ℕ} + (f : π' (suc (suc p)) X) + (g : π' (suc (suc q)) X) + (h : π' (suc (suc r)) X) + → ·π' _ (-π^ (p ·₋₁ r) [ f ∣ [ g ∣ h ]π' ]π') + (·π' _ (subst (λ n → π' n X) (jacobiPath₁ p q r) + (-π^ (q ·₋₁ p) [ g ∣ [ h ∣ f ]π' ]π')) + (subst (λ n → π' n X) (jacobiPath₂ p q r) + (-π^ (r ·₋₁ q) [ h ∣ [ f ∣ g ]π' ]π'))) + ≡ 1π' _ +[_∣_]π'Jacobi' {X = X} {p} {q} {r} f g h = + cong₂ _⋄_ + (cong (-π^ (p ·₋₁ r)) + ([_∣_]π'Jacobi f g h) + ∙ AbGroupTheory.inv^Distr πX _ _ (p ·₋₁ r) + ∙ λ _ → t1 ⋄ t2) + (π'-comm _ t3 t4) + ∙ comm-4 t1 t2 t4 t3 + ∙ cong₂ _⋄_ t1+t4≡0 t2+t3≡0 + ∙ π'-rUnit _ (1π' (suc (suc (p + suc (q + suc r))))) + where + πX : AbGroup _ + πX = Group→AbGroup (π'Gr ((suc p) + (suc q + suc r)) X) + (π'-comm _) + + open AbGroupTheory πX + open GroupTheory (π'Gr ((suc p) + (suc q + suc r)) X) + open AbGroupStr (snd πX) renaming (_+_ to _⋄_ ; -_ to -πX) + + evenOddLemma : (p q r : ℕ) + → isEvenT (p ·₋₁ r + r ·₋₁ (p + suc q)) ≡ isOddT (r ·₋₁ q) + evenOddLemma p q r = ua (propBiimpl→Equiv (isPropIsEvenT _) (isPropIsOddT _) + (fst main) (snd main)) + where + _·₋₁'_ : {n : ℕ} → Expr n → Expr n → Expr n + _·₋₁'_ a b = (K 1) +' (((K 1) +' a) ·' ((K 1) +' b)) + + r·₋₁q : Expr 2 + r·₋₁q = (∣' zero) ·₋₁' (∣' one) + + p·₋₁r+r·₋₁[p+q+1] : Expr 3 + p·₋₁r+r·₋₁[p+q+1] = (∣' zero ·₋₁' ∣' two) + +' (∣' two ·₋₁' (∣' zero +' (K 1 +' ∣' one))) + + main : isEvenT (p ·₋₁ r + r ·₋₁ (p + suc q)) ↔ isOddT (r ·₋₁ q) + main = main' (evenOrOdd r) (evenOrOdd q) + where + main' : isEvenT r ⊎ isOddT r → isEvenT q ⊎ isOddT q + → isEvenT (p ·₋₁ r + r ·₋₁ (p + suc q)) ↔ isOddT (r ·₋₁ q) + main' (inl evr) (inl odq) .fst e = + ⊥.rec (¬evenAndOdd _ (e + , solveIsOdd p·₋₁r+r·₋₁[p+q+1] + ((p , ?) ∷ (q , yea odq) ∷ ((r , yea evr) ∷ [])))) + main' (inl evr) (inr x) .fst _ = + solveIsOdd r·₋₁q ((r , yea evr) ∷ (q , nay x) ∷ []) + main' (inl evr) (inl x) .snd s = + ⊥.rec (¬evenAndOdd _ + (solveIsEven r·₋₁q ((r , yea evr) ∷ (q , yea x) ∷ []) , s)) + main' (inl evr) (inr odq) .snd odrq = + solveIsEven p·₋₁r+r·₋₁[p+q+1] + ((p , ?) ∷ (q , nay odq) ∷ ((r , yea evr) ∷ [])) + main' (inr odr) _ .fst _ = solveIsOdd r·₋₁q ((r , nay odr) ∷ (q , ?) ∷ []) + main' (inr odr) evq .snd _ = + solveIsEven p·₋₁r+r·₋₁[p+q+1] ((p , ?) ∷ (q , ?) ∷ ((r , nay odr) ∷ [])) + + t1 = -π^ (p ·₋₁ r) + (subst (λ n → π' n X) + (cong (2 +_) (sym (+-assoc p (suc q) (suc r)))) + [ [ f ∣ g ]π' ∣ h ]π') + + t2 = -π^ (p ·₋₁ r) + (subst (λ k → π' k X) (cong suc (assocPath p q r)) + (-π^ (suc p · suc q) [ g ∣ [ f ∣ h ]π' ]π')) + + t3 = (subst (λ n → π' n X) (jacobiPath₁ p q r) + (-π^ (q ·₋₁ p) [ g ∣ [ h ∣ f ]π' ]π')) + + t4 = (subst (λ n → π' n X) (jacobiPath₂ p q r) + (-π^ (r ·₋₁ q) [ h ∣ [ f ∣ g ]π' ]π')) + + -π^-substComm : {n m : ℕ} (t : ℕ) (x : π' (suc n) X) (p : suc n ≡ suc m) + → -π^ t (subst (λ k → π' k X) p x) + ≡ subst (λ k → π' k X) p (-π^ t x) + -π^-substComm t x p = + cong (λ p → -π^ t (subst (λ k → π' k X) p x)) (isSetℕ _ _ p _) + ∙ sym (substCommSlice (λ k → π' (suc k) X) (λ k → π' (suc k) X) + (λ k → -π^ t) (cong predℕ p) x) + ∙ cong (λ p → subst (λ k → π' k X) p (-π^ t x)) (isSetℕ _ _ _ p) + + t2≡ : -π^ (p ·₋₁ r) + (subst (λ k → π' k X) (cong suc (assocPath p q r)) + (-π^ (suc p · suc q) [ g ∣ [ f ∣ h ]π' ]π')) + ≡ subst (λ k → π' k X) (jacobiPath₁ p q r) + (-π^ (suc p · suc q) + [ g ∣ [ h ∣ f ]π' ]π') + t2≡ = -π^-substComm (p ·₋₁ r) + (-π^ (suc p · suc q) [ g ∣ [ f ∣ h ]π' ]π') + (cong suc (assocPath p q r)) + ∙ cong (subst (λ k → π' (suc k) X) (assocPath p q r)) + ((sym (iter+ (p ·₋₁ r) (suc p · suc q) (-π' _) + [ g ∣ [ f ∣ h ]π' ]π') + ∙ cong (λ t → -π^ t [ g ∣ [ f ∣ h ]π' ]π') + (+-comm (p ·₋₁ r) (suc p · suc q) + ∙ cong ((suc p · suc q) +_) + (cong suc (·-comm (suc p) (suc r)))) + ∙ iter+ (suc p · suc q) (r ·₋₁ p) (-π' _) + [ g ∣ [ f ∣ h ]π' ]π' + ∙ cong (-π^ (suc p · suc q)) + ((cong (-π^ (r ·₋₁ p)) + ((cong [ g ∣_]π' ([_∣_]π'-comm f h) + ∙ sym (substCommSlice (λ m → π' (suc m) X) + (λ a → π' (suc (suc q + a)) X) + (λ _ f → [ g ∣ f ]π') + (+-comm (suc r) (suc p)) + (-π^ (r ·₋₁ p) [ h ∣ f ]π'))) + ∙ cong (subst (λ a → π' (suc (suc q + a)) X) + (+-comm (suc r) (suc p))) + ([∣]π'-inv^DistrR (r ·₋₁ p) g [ h ∣ f ]π')) + ∙ -π^-substComm (r ·₋₁ p) (-π^ (r ·₋₁ p) [ g ∣ [ h ∣ f ]π' ]π') + (λ i → suc (suc q + +-comm (suc r) (suc p) i)) + ∙ cong (subst (λ k → π' k X) + (λ i → suc (suc q + +-comm (suc r) (suc p) i))) + (iter+iter (-π' _) (GroupTheory.invInv (π'Gr _ _)) + (r ·₋₁ p) [ g ∣ [ h ∣ f ]π' ]π'))) + ∙ -π^-substComm (suc p · suc q) ([ g ∣ [ h ∣ f ]π' ]π') + (λ i → suc (suc q + +-comm (suc r) (suc p) i))) + ∙ refl) + ∙ compSubstℕ {A = λ n → π' n X} + (λ i → suc (suc q + +-comm (suc r) (suc p) i)) + (cong suc (assocPath p q r)) + (jacobiPath₁ p q r) + {(-π^ (suc p · suc q) [ g ∣ [ h ∣ f ]π' ]π')} + + t1≡ : t1 ≡ subst (λ n → π' n X) (jacobiPath₂ p q r) + (-π^ (p ·₋₁ r + r ·₋₁ (p + suc q)) + [ h ∣ [ f ∣ g ]π' ]π') + t1≡ = cong (-π^ (p ·₋₁ r) + ∘ subst (λ n → π' n X) + (cong (2 +_) (sym (+-assoc p (suc q) (suc r))))) + ([_∣_]π'-comm ([ f ∣ g ]π') h) + ∙ cong (-π^ (p ·₋₁ r)) + (compSubstℕ {A = λ n → π' n X} + (cong suc (+-comm (suc r) (suc (p + suc q)))) + (sym (cong suc (+-assoc (suc p) (suc q) (suc r)))) + (jacobiPath₂ p q r)) + ∙ sym (substCommSlice + (λ k → π' (suc k) X) (λ k → π' (suc k) X) + (λ k → -π^ (p ·₋₁ r)) + (cong predℕ (jacobiPath₂ p q r)) + _) + ∙ cong (subst (λ n → π' n X) (jacobiPath₂ p q r)) + (sym (iter+ (p ·₋₁ r) (r ·₋₁ (p + suc q)) (-π' _) + ([ h ∣ [ f ∣ g ]π' ]π')) + ∙ refl) + + substHomLem : (n m : ℕ) (p : n ≡ m) {x y : π' (suc n) X} + → ·π' _ (subst (λ n → π' (suc n) X) p x) (subst (λ n → π' (suc n) X) p y) + ≡ subst (λ n → π' (suc n) X) p (·π' _ x y) + substHomLem n = J> λ {x y} → + cong₂ (·π' _) (transportRefl x) (transportRefl y) + ∙ sym (transportRefl (·π' n x y)) + + t1+t4≡0 : t1 ⋄ t4 ≡ 1π' _ + t1+t4≡0 = cong₂ _⋄_ t1≡ refl + ∙ substHomLem _ _ (cong predℕ (jacobiPath₂ p q r)) + ∙ cong (subst (λ n → π' n X) (jacobiPath₂ p q r)) + ((λ _ → + ·π' _ (-π^ (p ·₋₁ r + r ·₋₁ (p + suc q)) + [ h ∣ [ f ∣ g ]π' ]π') + (-π^ (r ·₋₁ q) [ h ∣ [ f ∣ g ]π' ]π')) + ∙ GroupTheory.-iter-odd+even (π'Gr _ _) + (suc (suc (r + p · suc r + r ·₋₁ (p + suc q)))) + (r ·₋₁ q) ([ h ∣ [ f ∣ g ]π' ]π') (evenOddLemma p q r)) + ∙ λ j → transp (λ i → π' (jacobiPath₂ p q r (i ∨ j)) X) j + (1π' (jacobiPath₂ p q r j)) + + t2+t3≡0 : t2 ⋄ t3 ≡ 1π' _ + t2+t3≡0 = + cong₂ _⋄_ t2≡ refl + ∙ substHomLem _ _ (cong predℕ (jacobiPath₁ p q r)) + ∙ cong (subst (λ n → π' n X) (jacobiPath₁ p q r)) + ((λ _ → ·π' _ (-π^ (suc p · suc q) [ g ∣ [ h ∣ f ]π' ]π') + (-π^ (q ·₋₁ p) [ g ∣ [ h ∣ f ]π' ]π')) + ∙ GroupTheory.-iter-odd+even (π'Gr _ _) (suc p · suc q) + (q ·₋₁ p) + [ g ∣ [ h ∣ f ]π' ]π' + (cong isEvenT (·-comm (suc p) (suc q)))) + ∙ λ j → transp (λ i → π' (jacobiPath₁ p q r (i ∨ j)) X) j + (1π' (jacobiPath₁ p q r j)) + +-- We prove that the function joinTo⋁ used in the definition of the Whitehead +-- product gives an equivalence between (Susp A × Susp B) and the +-- appropriate cofibre of joinTo⋁ (last two theorems in the following +-- module). + +module _ (A B : Type) (a₀ : A) (b₀ : B) where + private + W = joinTo⋁ {A = (A , a₀)} {B = (B , b₀)} + + A∨B = (Susp A , north) ⋁ (Susp B , north) + + σB = σ (B , b₀) + σA = σ (A , a₀) + + cofibW = Pushout W λ _ → tt + + whitehead3x3 : 3x3-span + A00 whitehead3x3 = Susp A + A02 whitehead3x3 = B + A04 whitehead3x3 = Unit + A20 whitehead3x3 = B + A22 whitehead3x3 = A × B + A24 whitehead3x3 = A + A40 whitehead3x3 = B + A42 whitehead3x3 = B + A44 whitehead3x3 = Unit + f10 whitehead3x3 _ = south + f12 whitehead3x3 = snd + f14 whitehead3x3 _ = tt + f30 whitehead3x3 = idfun B + f32 whitehead3x3 = snd + f34 whitehead3x3 _ = tt + f01 whitehead3x3 _ = north + f21 whitehead3x3 = snd + f41 whitehead3x3 = idfun B + f03 whitehead3x3 _ = tt + f23 whitehead3x3 = fst + f43 whitehead3x3 _ = tt + H11 whitehead3x3 x = merid (fst x) + H13 whitehead3x3 _ = refl + H31 whitehead3x3 _ = refl + H33 whitehead3x3 _ = refl + + A0□→A∨B : A0□ whitehead3x3 → A∨B + A0□→A∨B (inl x) = inl x + A0□→A∨B (inr x) = inr north + A0□→A∨B (push a i) = (push tt ∙ λ i → inr (σB a (~ i))) i + + A∨B→A0□ : A∨B → A0□ whitehead3x3 + A∨B→A0□ (inl x) = inl x + A∨B→A0□ (inr north) = inl north + A∨B→A0□ (inr south) = inl north + A∨B→A0□ (inr (merid b i)) = (push b₀ ∙ sym (push b)) i + A∨B→A0□ (push a i) = inl north + + Iso-A0□-⋁ : Iso (A0□ whitehead3x3) A∨B + fun Iso-A0□-⋁ = A0□→A∨B + inv Iso-A0□-⋁ = A∨B→A0□ + rightInv Iso-A0□-⋁ (inl x) = refl + rightInv Iso-A0□-⋁ (inr north) = push tt + rightInv Iso-A0□-⋁ (inr south) = push tt ∙ λ i → inr (merid b₀ i) + rightInv Iso-A0□-⋁ (inr (merid a i)) j = lem j i + where + lem : PathP (λ i → push tt i ≡ (push tt ∙ (λ i → inr (merid b₀ i))) i) + (cong A0□→A∨B (cong A∨B→A0□ λ i → inr (merid a i))) + (λ i → inr (merid a i)) + lem = (cong-∙ A0□→A∨B (push b₀) (sym (push a)) + ∙ cong₂ _∙_ (cong (push tt ∙_) + (λ j i → inr (rCancel (merid b₀) j (~ i))) + ∙ sym (rUnit (push tt))) + (symDistr (push tt) (λ i → inr (σB a (~ i))))) + ◁ λ i j → hcomp (λ k → + λ { (i = i0) → compPath-filler' (push tt) + (compPath-filler (λ i → inr (σB a i)) + (sym (push tt)) k) k j + ; (i = i1) → inr (merid a j) + ; (j = i0) → push tt (i ∨ ~ k) + ; (j = i1) → compPath-filler' (push tt) + (λ i → inr (merid b₀ i)) k i}) + (inr (compPath-filler (merid a) + (sym (merid b₀)) (~ i) j)) + + rightInv Iso-A0□-⋁ (push a i) j = push tt (i ∧ j) + leftInv Iso-A0□-⋁ (inl x) = refl + leftInv Iso-A0□-⋁ (inr tt) = push b₀ + leftInv Iso-A0□-⋁ (push b i) j = help j i + where + help : PathP (λ i → inl north ≡ push b₀ i) + (cong A∨B→A0□ (cong A0□→A∨B (push b))) + (push b) + help = (cong-∙ A∨B→A0□ (push tt) (λ i → inr (σB b (~ i))) + ∙ (λ i → lUnit (sym (cong-∙ (A∨B→A0□ ∘ inr) + (merid b) (sym (merid b₀)) i)) (~ i)) + ∙ cong sym (cong ((push b₀ ∙ sym (push b)) ∙_) + (cong sym (rCancel (push b₀)))) + ∙ cong sym (sym (rUnit (push b₀ ∙ sym (push b))))) + ◁ λ i j → compPath-filler' (push b₀) (sym (push b)) (~ i) (~ j) + + Iso-A2□-join : Iso (A2□ whitehead3x3) (join A B) + fun Iso-A2□-join (inl x) = inr x + fun Iso-A2□-join (inr x) = inl x + fun Iso-A2□-join (push (a , b) i) = push a b (~ i) + inv Iso-A2□-join (inl x) = inr x + inv Iso-A2□-join (inr x) = inl x + inv Iso-A2□-join (push a b i) = push (a , b) (~ i) + rightInv Iso-A2□-join (inl x) = refl + rightInv Iso-A2□-join (inr x) = refl + rightInv Iso-A2□-join (push a b i) = refl + leftInv Iso-A2□-join (inl x) = refl + leftInv Iso-A2□-join (inr x) = refl + leftInv Iso-A2□-join (push a i) = refl + + isContrA4□ : isContr (A4□ whitehead3x3) + fst isContrA4□ = inr tt + snd isContrA4□ (inl x) = sym (push x) + snd isContrA4□ (inr x) = refl + snd isContrA4□ (push a i) j = push a (i ∨ ~ j) + + A4□≃Unit : A4□ whitehead3x3 ≃ Unit + A4□≃Unit = isContr→≃Unit isContrA4□ + + Iso-A□0-Susp : Iso (A□0 whitehead3x3) (Susp A) + fun Iso-A□0-Susp (inl x) = x + fun Iso-A□0-Susp (inr x) = north + fun Iso-A□0-Susp (push a i) = merid a₀ (~ i) + inv Iso-A□0-Susp x = inl x + rightInv Iso-A□0-Susp x = refl + leftInv Iso-A□0-Susp (inl x) = refl + leftInv Iso-A□0-Susp (inr x) = (λ i → inl (merid a₀ i)) ∙ push x + leftInv Iso-A□0-Susp (push a i) j = + hcomp (λ k → λ { (i = i0) → inl (merid a₀ (k ∨ j)) + ; (i = i1) → compPath-filler + (λ i₁ → inl (merid a₀ i₁)) + (push (idfun B a)) k j + ; (j = i0) → inl (merid a₀ (~ i ∧ k)) + ; (j = i1) → push a (i ∧ k)}) + (inl (merid a₀ j)) + + Iso-A□2-Susp× : Iso (A□2 whitehead3x3) (Susp A × B) + fun Iso-A□2-Susp× (inl x) = north , x + fun Iso-A□2-Susp× (inr x) = south , x + fun Iso-A□2-Susp× (push a i) = merid (fst a) i , (snd a) + inv Iso-A□2-Susp× (north , y) = inl y + inv Iso-A□2-Susp× (south , y) = inr y + inv Iso-A□2-Susp× (merid a i , y) = push (a , y) i + rightInv Iso-A□2-Susp× (north , snd₁) = refl + rightInv Iso-A□2-Susp× (south , snd₁) = refl + rightInv Iso-A□2-Susp× (merid a i , snd₁) = refl + leftInv Iso-A□2-Susp× (inl x) = refl + leftInv Iso-A□2-Susp× (inr x) = refl + leftInv Iso-A□2-Susp× (push a i) = refl + + Iso-A□4-Susp : Iso (A□4 whitehead3x3) (Susp A) + fun Iso-A□4-Susp (inl x) = north + fun Iso-A□4-Susp (inr x) = south + fun Iso-A□4-Susp (push a i) = merid a i + inv Iso-A□4-Susp north = inl tt + inv Iso-A□4-Susp south = inr tt + inv Iso-A□4-Susp (merid a i) = push a i + rightInv Iso-A□4-Susp north = refl + rightInv Iso-A□4-Susp south = refl + rightInv Iso-A□4-Susp (merid a i) = refl + leftInv Iso-A□4-Susp (inl x) = refl + leftInv Iso-A□4-Susp (inr x) = refl + leftInv Iso-A□4-Susp (push a i) = refl + + Iso-PushSusp×-Susp×Susp : + Iso (Pushout {A = Susp A × B} fst fst) (Susp A × Susp B) + Iso-PushSusp×-Susp×Susp = theIso + where + F : Pushout {A = Susp A × B} fst fst + → Susp A × Susp B + F (inl x) = x , north + F (inr x) = x , north + F (push (x , b) i) = x , σB b i + + G : Susp A × Susp B → Pushout {A = Susp A × B} fst fst + G (a , north) = inl a + G (a , south) = inr a + G (a , merid b i) = push (a , b) i + + retr : retract F G + retr (inl x) = refl + retr (inr x) = push (x , b₀) + retr (push (a , b) i) j = help j i + where + help : PathP (λ i → Path (Pushout fst fst) (inl a) (push (a , b₀) i)) + (cong G (λ i → a , σB b i)) + (push (a , b)) + help = cong-∙ (λ x → G (a , x)) (merid b) (sym (merid b₀)) + ◁ λ i j → compPath-filler + (push (a , b)) + (sym (push (a , b₀))) + (~ i) j + + theIso : Iso (Pushout fst fst) (Susp A × Susp B) + fun theIso = F + inv theIso = G + rightInv theIso (a , north) = refl + rightInv theIso (a , south) = ΣPathP (refl , merid b₀) + rightInv theIso (a , merid b i) j = + a , compPath-filler (merid b) (sym (merid b₀)) (~ j) i + leftInv theIso = retr + + Iso-A□○-PushSusp× : + Iso (A□○ whitehead3x3) (Pushout {A = Susp A × B} fst fst) + Iso-A□○-PushSusp× = + pushoutIso _ _ fst fst + (isoToEquiv Iso-A□2-Susp×) + (isoToEquiv Iso-A□0-Susp) + (isoToEquiv Iso-A□4-Susp) + (funExt (λ { (inl x) → refl + ; (inr x) → merid a₀ + ; (push a i) j → help₁ a j i})) + (funExt λ { (inl x) → refl + ; (inr x) → refl + ; (push a i) j + → fun Iso-A□4-Susp (rUnit (push (fst a)) (~ j) i)}) + where + help₁ : (a : A × B) + → PathP (λ i → north ≡ merid a₀ i) + (cong (fun Iso-A□0-Susp) + (cong (f□1 whitehead3x3) (push a))) + (merid (fst a)) + help₁ a = + (cong-∙∙ (fun Iso-A□0-Susp) + (λ i → inl (merid (fst a) i)) + (push (snd a)) + refl) + ◁ (λ i j → hcomp (λ k → λ {(i = i1) → merid (fst a) (j ∨ ~ k) + ; (j = i0) → merid (fst a) (~ k) + ; (j = i1) → merid a₀ i}) + (merid a₀ (i ∨ ~ j))) + + Iso-A□○-Susp×Susp : Iso (A□○ whitehead3x3) (Susp A × Susp B) + Iso-A□○-Susp×Susp = compIso Iso-A□○-PushSusp× Iso-PushSusp×-Susp×Susp + + Iso-A○□-cofibW : Iso (A○□ whitehead3x3) cofibW + Iso-A○□-cofibW = + pushoutIso _ _ + W (λ _ → tt) + (isoToEquiv Iso-A2□-join) (isoToEquiv Iso-A0□-⋁) + A4□≃Unit + (funExt lem) + λ _ _ → tt + where + lem : (x : A2□ whitehead3x3) + → A0□→A∨B (f1□ whitehead3x3 x) ≡ W (fun Iso-A2□-join x) + lem (inl x) = (λ i → inl (merid a₀ (~ i))) + lem (inr x) = refl + lem (push (a , b) i) j = help j i + where + help : PathP (λ i → Path (Pushout (λ _ → north) (λ _ → north)) + ((inl (merid a₀ (~ i)))) + (inr north)) + (cong A0□→A∨B (cong (f1□ whitehead3x3) (push (a , b)))) + (cong W (cong (fun Iso-A2□-join) (push (a , b)))) + help = (cong-∙∙ A0□→A∨B (λ i → inl (merid a (~ i))) (push b) refl + ∙ λ j → (λ i₂ → inl (merid a (~ i₂))) + ∙∙ compPath-filler (push tt) (λ i → inr (σB b (~ i))) (~ j) + ∙∙ λ i → inr (σB b (~ i ∧ j))) + ◁ (λ j → (λ i → inl (sym (compPath-filler + (merid a) (sym (merid a₀)) j) i)) + ∙∙ push tt + ∙∙ λ i → inr (σB b (~ i))) + + Iso₁-Susp×Susp-cofibW : Iso (Susp A × Susp B) cofibW + Iso₁-Susp×Susp-cofibW = + compIso (invIso Iso-A□○-Susp×Susp) + (compIso (3x3-Iso whitehead3x3) Iso-A○□-cofibW) + + -- Main iso + Iso-Susp×Susp-cofibJoinTo⋁ : Iso (Susp A × Susp B) cofibW + Iso-Susp×Susp-cofibJoinTo⋁ = + compIso (Σ-cong-iso-snd (λ _ → invSuspIso)) + Iso₁-Susp×Susp-cofibW + + -- The induced function A ∨ B → Susp A × Susp B satisfies + -- the following identity + Susp×Susp→cofibW≡ : Path (A∨B → Susp A × Susp B) + (Iso.inv Iso-Susp×Susp-cofibJoinTo⋁ ∘ inl) + ⋁↪ + Susp×Susp→cofibW≡ = + funExt λ { (inl x) → ΣPathP (refl , (sym (merid b₀))) + ; (inr north) → ΣPathP (refl , (sym (merid b₀))) + ; (inr south) → refl + ; (inr (merid a i)) j → lem₂ a j i + ; (push a i) j → north , (merid b₀ (~ j))} + where + f₁ = fun Iso-PushSusp×-Susp×Susp + f₂ = fun Iso-A□○-PushSusp× + f₃ = backward-l whitehead3x3 + f₄ = fun (Σ-cong-iso-snd (λ _ → invSuspIso)) + + lem : (b : B) + → cong (f₁ ∘ f₂ ∘ f₃) (push b) + ≡ (λ i → north , σB b i) + lem b = cong (cong f₁) (sym (rUnit (push (north , b)))) + + lem₂ : (a : B) + → PathP (λ i → (north , merid b₀ (~ i)) ≡ (north , south)) + (cong (f₄ ∘ f₁ ∘ f₂ ∘ f₃ ∘ A∨B→A0□ ∘ inr) (merid a)) + λ i → north , merid a i + lem₂ a = + cong (cong f₄) (cong-∙ (f₁ ∘ f₂ ∘ f₃) (push b₀) (sym (push a)) + ∙∙ cong₂ _∙_ (lem b₀ ∙ (λ j i → north , rCancel (merid b₀) j i)) + (cong sym (lem a)) + ∙∙ sym (lUnit (λ i₁ → north , σB a (~ i₁)))) + ∙ (λ i j → north , cong-∙ invSusp (merid a) (sym (merid b₀)) i (~ j)) + ◁ λ i j → north , compPath-filler (sym (merid a)) (merid b₀) (~ i) (~ j) diff --git a/Cubical/Papers/Pi4S3-JournalVersion.agda b/Cubical/Papers/Pi4S3-JournalVersion.agda index 329dab7c33..7b60a491c3 100644 --- a/Cubical/Papers/Pi4S3-JournalVersion.agda +++ b/Cubical/Papers/Pi4S3-JournalVersion.agda @@ -28,60 +28,63 @@ open import Cubical.Data.Bool as Boolean open import Cubical.Data.Unit as UnitType open import Cubical.HITs.S1 as Circle -open import Cubical.Foundations.Prelude as Prelude -open import Cubical.HITs.Susp as Suspensions -open import Cubical.HITs.Sn as Spheres +open import Cubical.Foundations.Prelude as Prelude +open import Cubical.HITs.Susp as Suspensions +open import Cubical.HITs.Sn as Spheres hiding (S) renaming (S₊ to S) -open import Cubical.HITs.Pushout as Pushouts -open import Cubical.HITs.Wedge as Wedges -open import Cubical.HITs.Join as Joins -open import Cubical.HITs.Susp as Suspension -open import Cubical.HITs.PropositionalTruncation as PT -open import Cubical.HITs.Truncation as Trunc -open import Cubical.Foundations.Univalence as Univ -open import Cubical.Homotopy.Loopspace as Loopy - -open import Cubical.Homotopy.HSpace as H-Spaces -open import Cubical.Homotopy.Group.Base as HomotopyGroups -open import Cubical.Homotopy.Group.LES as LES -open import Cubical.Homotopy.HopfInvariant.HopfMap as HopfMap -open import Cubical.Homotopy.Hopf as HopfFibration -open import Cubical.Homotopy.Connected as Connectedness +open import Cubical.HITs.Pushout as Pushouts +open import Cubical.HITs.Wedge as Wedges +open import Cubical.HITs.Join as Joins +open import Cubical.HITs.Susp as Suspension +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.Truncation as Trunc +open import Cubical.Foundations.Univalence as Univ +open import Cubical.Homotopy.Loopspace as Loopy + +open import Cubical.Homotopy.HSpace as H-Spaces +open import Cubical.Homotopy.Group.Base as HomotopyGroups +open import Cubical.Homotopy.Group.LES as LES +open import Cubical.Homotopy.HopfInvariant.HopfMap as HopfMap +open import Cubical.Homotopy.Hopf as HopfFibration +open import Cubical.Homotopy.Connected as Connectedness open S¹Hopf -open import Cubical.Homotopy.Freudenthal as Freudenthal -open import Cubical.Homotopy.Group.PinSn as Stable -open import Cubical.Homotopy.Group.Pi3S2 as π₃S² +open import Cubical.Homotopy.Freudenthal as Freudenthal +open import Cubical.Homotopy.Group.PinSn as Stable +open import Cubical.Homotopy.Group.Pi3S2 as π₃S² -- 3 -open import Cubical.Homotopy.Group.Pi4S3.S3PushoutIso as James₁ -open import Cubical.Homotopy.Group.Pi4S3.S3PushoutIso2 as James₂ -open import Cubical.HITs.S2 as Sphere -open import Cubical.Homotopy.Whitehead as Whitehead +open import Cubical.Homotopy.Group.Pi4S3.S3PushoutIso as James₁ +open import Cubical.Homotopy.Group.Pi4S3.S3PushoutIso2 as James₂ +open import Cubical.HITs.S2 as Sphere +open import Cubical.Homotopy.WhiteheadProducts.Base as Whitehead +open import Cubical.Homotopy.WhiteheadProducts.Properties as WhiteheadProps +open import Cubical.Homotopy.WhiteheadProducts.Generalised.Join.Base as WhiteheadGen open import Cubical.Homotopy.BlakersMassey module BM = BlakersMassey□ open BM -open import Cubical.Homotopy.Group.Pi4S3.BrunerieNumber as BNumber +open import Cubical.Homotopy.Group.Pi4S3.BrunerieNumber as BNumber hiding (W) -- 5 -open import Cubical.ZCohomology.Base as cohom -open import Cubical.ZCohomology.GroupStructure as cohomGr -open import Cubical.ZCohomology.Properties as cohomProps -open import Cubical.ZCohomology.RingStructure.CupProduct as cup -open import Cubical.ZCohomology.MayerVietorisUnreduced as MayerVietoris -open import Cubical.Homotopy.HopfInvariant.Base as HI -open import Cubical.Homotopy.HopfInvariant.Homomorphism as HI-hom -open import Cubical.Homotopy.HopfInvariant.Brunerie as HI-β -open import Cubical.ZCohomology.Gysin as GysinSeq -open import Cubical.Homotopy.Group.Pi4S3.Summary as π₄S³ +open import Cubical.ZCohomology.Base as cohom +open import Cubical.ZCohomology.GroupStructure as cohomGr +open import Cubical.ZCohomology.Properties as cohomProps +open import Cubical.ZCohomology.RingStructure.CupProduct as cup +open import Cubical.ZCohomology.MayerVietorisUnreduced as MayerVietoris +open import Cubical.Homotopy.HopfInvariant.Base as HI +open import Cubical.Homotopy.HopfInvariant.Homomorphism as HI-hom +open import Cubical.Homotopy.HopfInvariant.Brunerie as HI-β +open import Cubical.ZCohomology.Gysin as GysinSeq +open import Cubical.Homotopy.Group.Pi4S3.Summary as π₄S³ hiding (π) -open import Cubical.ZCohomology.RingStructure.RingLaws as cupLaws +open import Cubical.ZCohomology.RingStructure.RingLaws as cupLaws -- 6 -open import Cubical.HITs.SmashProduct.Base as SmashProd -open import Cubical.HITs.Sn.Multiplication as SMult -open import Cubical.Homotopy.Group.Join as JoinGroup -open import Cubical.Homotopy.Group.Pi4S3.DirectProof as Direct +open import Cubical.HITs.SmashProduct.Base as SmashProd +open import Cubical.HITs.Sn.Multiplication as SMult +open import Cubical.Homotopy.Group.Join as JoinGroup +open import Cubical.HITs.Join.CoHSpace as JoinCoH +open import Cubical.Homotopy.Group.Pi4S3.DirectProof as Direct ------ 2. HOMOTOPY TYPE THEORY IN Cubical Agda ------ @@ -361,7 +364,7 @@ open SMult using (comm⌣S ; assoc⌣S) open JoinGroup using (π*) -- Addition +* -open JoinGroup using (_+*_) +open JoinCoH using (_+*_) -- Proposition 6.8 open JoinGroup using (·Π≡+*) @@ -387,10 +390,10 @@ open Direct using (η₃' ; computerIsoη₃) -- Not formalised explicitly -- Definition of generalised Whitehead products ·w -open Whitehead using (_·w_) +open WhiteheadGen using (·wh) -- Proposition 6.22 (including Lemmas 19 and 20 and Proposition 6.21) -open Whitehead using (isConst-Susp·w) +open WhiteheadGen using (isConst-Susp·w) -- Theorem 6.23 -- Follows directly from above but not formalised explicitly (awaiting diff --git a/Cubical/Papers/Pi4S3.agda b/Cubical/Papers/Pi4S3.agda index 76e1fdf480..6c55512a3b 100644 --- a/Cubical/Papers/Pi4S3.agda +++ b/Cubical/Papers/Pi4S3.agda @@ -57,7 +57,7 @@ open import Cubical.Homotopy.Group.Pi3S2 as π₃S² open import Cubical.Homotopy.Group.Pi4S3.S3PushoutIso as James₁ open import Cubical.Homotopy.Group.Pi4S3.S3PushoutIso2 as James₂ open import Cubical.HITs.S2 as Sphere -open import Cubical.Homotopy.Whitehead as Whitehead +open import Cubical.Homotopy.WhiteheadProducts.Base as Whitehead open import Cubical.Homotopy.BlakersMassey module BM = BlakersMassey□ open BM diff --git a/Cubical/Tactics/EvenOddSolver.agda b/Cubical/Tactics/EvenOddSolver.agda new file mode 100644 index 0000000000..5c3edd3975 --- /dev/null +++ b/Cubical/Tactics/EvenOddSolver.agda @@ -0,0 +1,253 @@ +{-# OPTIONS --safe #-} +{- + +This module defines a rather primitive solver for the evenness of +expressions over ℕ. + +Input: Expression s over ℕ, a list of assignments v : Vec ((x : ℕ) × even? x) n +Output: Proof that ⟦ s ⟧ v is even/odd, if possible to determine. + +-} +module Cubical.Tactics.EvenOddSolver where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.FinData hiding (snotz) +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order using (zero-≤ ; isProp≤) +open import Cubical.Data.Nat.IsEven +open import Cubical.Data.Vec.Base +open import Cubical.Data.Empty as ⊥ +open import Cubical.Data.Nat.Mod +open import Cubical.Data.Fin as ΣFin hiding (Fin) +open import Cubical.Data.Sigma +open import Cubical.Data.List as L hiding (map) renaming (_++_ to _+_) + +open import Cubical.Tactics.NatSolver.NatExpression + +open Eval + +private + variable + ℓ ℓ' ℓ'' : Level + +-- Inductive type encoding whether a natural is odd or even, or +-- whether this is unknown +data even? (n : ℕ) : Type where + ? : even? n + yea : isEvenT n → even? n + nay : isOddT n → even? n + +-- First step: given a vector of naturals equipped with information +-- about whether or not they are even, return a new vector with each +-- even replaced by 0 and each odd replaced by 1. +evenOddVecLen : {n : ℕ} + → Vec (Σ[ k ∈ ℕ ] (even? k)) n → ℕ +evenOddVecLen [] = 0 +evenOddVecLen ((s , ?) ∷ xs) = suc (evenOddVecLen xs) +evenOddVecLen ((s , yea x) ∷ xs) = evenOddVecLen xs +evenOddVecLen ((s , nay x) ∷ xs) = evenOddVecLen xs + +filterEvenOddVec : {n : ℕ} + (t : Vec (Σ[ x ∈ ℕ ] (even? x)) n) + → Vec ℕ (evenOddVecLen t) +filterEvenOddVec [] = [] +filterEvenOddVec ((x , ?) ∷ t) = x ∷ filterEvenOddVec t +filterEvenOddVec ((x , yea s) ∷ t) = filterEvenOddVec t +filterEvenOddVec ((x , nay s) ∷ t) = filterEvenOddVec t + +-- Now do the same things for expressions over ℕ +filterEvenOddExpr : {n : ℕ} (s : Expr n) + (t : Vec (Σ[ x ∈ ℕ ] (even? x)) n) + → Expr (evenOddVecLen t) +filterEvenOddExpr (K k) t = K k +filterEvenOddExpr (∣ zero) ((_ , ?) ∷ t) = ∣ zero +filterEvenOddExpr (∣ (suc k)) ((_ , ?) ∷ t) = Expr↑ (filterEvenOddExpr (∣ k) t) +filterEvenOddExpr (∣ zero) ((_ , yea ev) ∷ t) = K 0 +filterEvenOddExpr (∣ (suc k)) ((_ , yea s) ∷ t) = filterEvenOddExpr (∣ k) t +filterEvenOddExpr (∣ (suc k)) ((_ , nay s) ∷ t) = filterEvenOddExpr (∣ k) t +filterEvenOddExpr (∣ zero) ((_ , nay od) ∷ t) = K 1 +filterEvenOddExpr (s +' s₁) t = filterEvenOddExpr s t +' filterEvenOddExpr s₁ t +filterEvenOddExpr (s ·' s₁) t = filterEvenOddExpr s t ·' filterEvenOddExpr s₁ t + +-- Key lemma: the evaluation of a given expression mod 2 is invariant +-- w.r.t. this kind of replacement +evalFilterEvenOdd : {n : ℕ} (s : Expr n) + (t : Vec (Σ[ x ∈ ℕ ] (even? x)) n) + → (⟦ s ⟧ (map fst t)) mod 2 + ≡ (⟦ filterEvenOddExpr s t ⟧ (filterEvenOddVec t)) mod 2 +evalFilterEvenOdd (K x) t = refl +evalFilterEvenOdd (∣ zero) ((s , ?) ∷ t) = refl +evalFilterEvenOdd (∣ zero) ((s , yea x) ∷ t) = isEvenT↔≡0 s .fst x +evalFilterEvenOdd (∣ zero) ((s , nay x) ∷ t) = isOddT↔≡1 s .fst x +evalFilterEvenOdd (∣ (suc k)) ((s , ?) ∷ t) = + evalFilterEvenOdd (∣ k) t + ∙ sym (cong (_mod 2) (Expr↑ForgetFirst + (filterEvenOddExpr (∣ k) t) s (filterEvenOddVec t))) + where + Expr↑ForgetFirst : {n : ℕ} (x : Expr n) (s : ℕ) (w : _) + → ⟦ Expr↑ x ⟧ (s ∷ w) ≡ ⟦ x ⟧ w + Expr↑ForgetFirst (K x) s w = refl + Expr↑ForgetFirst (∣ x) s w = refl + Expr↑ForgetFirst (x +' x₁) s w = + cong₂ _+_ (Expr↑ForgetFirst x s w) (Expr↑ForgetFirst x₁ s w) + Expr↑ForgetFirst (x ·' x₁) s w = + cong₂ _·_ (Expr↑ForgetFirst x s w) (Expr↑ForgetFirst x₁ s w) +evalFilterEvenOdd (∣ (suc x)) ((s , yea _) ∷ t) = evalFilterEvenOdd (∣ x) t +evalFilterEvenOdd (∣ (suc x)) ((s , nay _) ∷ t) = evalFilterEvenOdd (∣ x) t +evalFilterEvenOdd (s +' s₁) t = + mod+mod≡mod 2 (⟦ s ⟧ (map fst t)) (⟦ s₁ ⟧ (map fst t)) + ∙ cong₂ (λ x y → (x + y) mod 2) + (evalFilterEvenOdd s t) (evalFilterEvenOdd s₁ t) + ∙ sym (mod+mod≡mod 2 (⟦ filterEvenOddExpr s t ⟧ (filterEvenOddVec t)) + (⟦ filterEvenOddExpr s₁ t ⟧ (filterEvenOddVec t))) +evalFilterEvenOdd (s ·' s₁) t = + mod·mod≡mod 2 (⟦ s ⟧ (map fst t)) (⟦ s₁ ⟧ (map fst t)) + ∙ cong₂ (λ x y → (x · y) mod 2) + (evalFilterEvenOdd s t) (evalFilterEvenOdd s₁ t) + ∙ sym (mod·mod≡mod 2 (⟦ filterEvenOddExpr s t ⟧ (filterEvenOddVec t)) + (⟦ filterEvenOddExpr s₁ t ⟧ (filterEvenOddVec t))) + +-- evaluating mod 2 in different orders +evalMod2 : {n : ℕ} (s : Expr n) (t : Vec ℕ n) + → (⟦ s ⟧ t) mod 2 ≡ (⟦ s ⟧ (map (_mod 2) t)) mod 2 +evalMod2 (K x) t = refl +evalMod2 (∣ zero) (s ∷ xs) = sym (mod-idempotent s) +evalMod2 (∣ (suc x)) (s ∷ xs) = evalMod2 (∣ x) xs +evalMod2 (s +' s₁) t = + mod+mod≡mod 2 (⟦ s ⟧ t) (⟦ s₁ ⟧ t) + ∙ cong₂ (λ x y → (x + y) mod 2) (evalMod2 s t) (evalMod2 s₁ t) + ∙ sym (mod+mod≡mod 2 (⟦ s ⟧ (map (_mod 2) t)) (⟦ s₁ ⟧ (map (_mod 2) t))) +evalMod2 (s ·' s₁) t = + mod·mod≡mod 2 (⟦ s ⟧ t) (⟦ s₁ ⟧ t) + ∙ cong₂ (λ x y → (x · y) mod 2) (evalMod2 s t) (evalMod2 s₁ t) + ∙ sym (mod·mod≡mod 2 (⟦ s ⟧ (map (_mod 2) t)) (⟦ s₁ ⟧ (map (_mod 2) t))) + +-- Step 2: characterise dependent types over Vec (Fin 2) n +Fin2Lists : (n : ℕ) → List (Vec (ΣFin.Fin 2) n) +Fin2Lists zero = [ [] ] +Fin2Lists (suc n) = L.map (0 ∷_) (Fin2Lists n) + + L.map (1 ∷_) (Fin2Lists n) + +ListFinIndT : (n : ℕ) (A : Vec (ΣFin.Fin 2) n → Type) + → List (Vec (ΣFin.Fin 2) n) → Type +ListFinIndT n A [] = Unit +ListFinIndT n A (x ∷ xs) = A x × ListFinIndT n A xs + +ListFinIndT++ : {n : ℕ} (A : Vec (ΣFin.Fin 2) n → Type) + (xs ys : List (Vec (ΣFin.Fin 2) n)) + → ListFinIndT n A (xs + ys) + → ListFinIndT n A xs × ListFinIndT n A ys +ListFinIndT++ A [] ys t = tt , t +ListFinIndT++ {n = n} A (x ∷ xs) ys (t , v) .fst .fst = t +ListFinIndT++ {n = n} A (x ∷ xs) ys (t , v) .fst .snd = + fst (ListFinIndT++ A xs ys v) +ListFinIndT++ {n = n} A (x ∷ xs) ys (t , v) .snd = + snd (ListFinIndT++ A xs ys v) + +-- Key result: +ListFinInd : (n : ℕ) {A : Vec (ΣFin.Fin 2) n → Type} + → ListFinIndT n A (Fin2Lists n) + → (x : _) → A x +ListFinInd n {A = A} ind [] = ind .fst +ListFinInd (suc n) {A = A} ind (s ∷ w) with + (ListFinIndT++ A (L.map (0 ∷_) (Fin2Lists n)) + (L.map (1 ∷_) (Fin2Lists n)) ind) +... | le , ri = transport refl + (ListFinInd n {A = λ x → A (s ∷ x)} + (transport (ListFinIndT↓ n A (Fin2Lists n) s) + (lem s)) w) + where + ListFinIndT↓ : (n : ℕ) (A : Vec (ΣFin.Fin 2) (suc n) → Type) + (l : List (Vec (ΣFin.Fin 2) n)) (x : ΣFin.Fin 2) + → ListFinIndT (suc n) A (L.map (x ∷_) l) + ≡ ListFinIndT n (λ y → A (x ∷ y)) l + ListFinIndT↓ n A [] x = refl + ListFinIndT↓ n A (y ∷ l) x = cong₂ _×_ refl (ListFinIndT↓ n A l x) + + lem : (s : _) → ListFinIndT (suc n) A (L.map (_∷_ s) (Fin2Lists n)) + lem (zero , t) = + subst (λ t → ListFinIndT (suc n) A (L.map (_∷_ (zero , t)) + (Fin2Lists n))) + (isProp≤ (1 , (λ i → 2)) t) le + lem (one , t) = + subst (λ t → ListFinIndT (suc n) A (L.map (_∷_ (one , t)) + (Fin2Lists n))) + (isProp≤ (0 , (λ i → 2)) t) ri + lem (suc (suc s) , t) = + ⊥.rec (snotz (cong predℕ + (cong predℕ (+-comm (3 + s) (t .fst) ∙ snd t)))) + +private + ℕPath : (n m : ℕ) → Type + ℕPath zero zero = Unit + ℕPath zero (suc m) = ⊥ + ℕPath (suc n) zero = ⊥ + ℕPath (suc n) (suc m) = ℕPath n m + + ℕPath→≡ : {n m : ℕ} → ℕPath n m → n ≡ m + ℕPath→≡ {n = zero} {zero} p = refl + ℕPath→≡ {n = suc n} {suc m} p = cong suc (ℕPath→≡ {n = n} {m} p) + + ℕ→Fin2 : ℕ → ΣFin.Fin 2 + ℕ→Fin2 n .fst = n mod 2 + ℕ→Fin2 n .snd = mod< one n + +-- main result +module _ {n : ℕ} (s : Expr n) (x : ℕ) + (t : Vec (Σ[ k ∈ ℕ ] (even? k)) n) + {u : ListFinIndT (evenOddVecLen t) + (λ v → ℕPath (⟦ filterEvenOddExpr s t ⟧ (map fst v) mod two) (x mod two)) + (Fin2Lists (evenOddVecLen t))} where + -- The complicated type of u reduces to Unit × ... × Unit if the + -- input expression indeed is even (which Agda can fill automatically) + + solveWithConstraints : (⟦ s ⟧ (map fst t)) mod two ≡ x mod two + solveWithConstraints = + evalFilterEvenOdd s t + ∙ evalMod2 (filterEvenOddExpr s t) (filterEvenOddVec t) + ∙ ℕPath→≡ (subst (λ a → ℕPath (a mod two) (x mod two)) lem main) + where + lem : ⟦ filterEvenOddExpr s t ⟧ (map fst (map ℕ→Fin2 (filterEvenOddVec t))) + ≡ ((⟦ filterEvenOddExpr s t ⟧ (map (_mod two) (filterEvenOddVec t)))) + lem = cong (⟦ filterEvenOddExpr s t ⟧) + (sym (compMap fst ℕ→Fin2 (filterEvenOddVec t))) + + main = ListFinInd (evenOddVecLen t) + {A = λ v → ℕPath (⟦ filterEvenOddExpr s t ⟧ + (map fst v) mod two) + (x mod two)} + u (map ℕ→Fin2 (filterEvenOddVec t)) + +solveIsEven : {n : ℕ} (s : Expr n) + (t : Vec (Σ[ k ∈ ℕ ] (even? k)) n) + {u : ListFinIndT (evenOddVecLen t) + (λ v → ℕPath (⟦ filterEvenOddExpr s t ⟧ (map fst v) mod two) 0) + (Fin2Lists (evenOddVecLen t))} + → isEvenT (⟦ s ⟧ (map fst t)) +solveIsEven s t {u = u} = + isEvenT↔≡0 (⟦ s ⟧ (map fst t)) .snd + (solveWithConstraints s 0 t {u}) + +solveIsOdd : {n : ℕ} (s : Expr n) + (t : Vec (Σ[ k ∈ ℕ ] (even? k)) n) + {u : ListFinIndT (evenOddVecLen t) + (λ v → ℕPath (⟦ filterEvenOddExpr s t ⟧ (map fst v) mod two) 1) + (Fin2Lists (evenOddVecLen t))} + → isOddT (⟦ s ⟧ (map fst t)) +solveIsOdd s t {u = u} = + isOddT↔≡1 (⟦ s ⟧ (map fst t)) .snd + (solveWithConstraints s 1 t {u}) + +-- example: the expression r + (3 · p + 2 · q) is odd whenever p and r are +-- (and the status of q is unknown) +private + module _ (p q r : ℕ) (odp : isOddT p) (odr : isOddT r) where + exp : Expr 3 + exp = ∣ two +' ((K 3 ·' ∣ zero) +' (K 2 ·' ∣ one)) + + exp≡ : isEvenT (r + (3 · p + 2 · q)) + exp≡ = solveIsEven exp + ((p , nay odp) + ∷ (q , ?) + ∷ (r , nay odr) ∷ []) diff --git a/Cubical/Tactics/NatSolver/NatExpression.agda b/Cubical/Tactics/NatSolver/NatExpression.agda index b53eece18f..e83cb5017b 100644 --- a/Cubical/Tactics/NatSolver/NatExpression.agda +++ b/Cubical/Tactics/NatSolver/NatExpression.agda @@ -17,6 +17,12 @@ data Expr (n : ℕ) : Type ℓ-zero where _+'_ : Expr n → Expr n → Expr n _·'_ : Expr n → Expr n → Expr n +Expr↑ : ∀ {n} → Expr n → Expr (suc n) +Expr↑ (K x) = K x +Expr↑ (∣ x) = ∣ (suc x) +Expr↑ (e +' e₁) = Expr↑ e +' Expr↑ e₁ +Expr↑ (e ·' e₁) = Expr↑ e ·' Expr↑ e₁ + module Eval where open import Cubical.Data.Vec